Theory Kraus_Families

section ‹Kraus families›

theory Kraus_Families
  imports
    Wlog.Wlog
    Hilbert_Space_Tensor_Product.Partial_Trace
  
    Backported
    Misc_Kraus_Maps
  
  abbrevs
    "=kr" = "=kr" and "==kr" = "≡kr" and "*kr" = "*kr"
begin

unbundle cblinfun_syntax

subsection ‹Kraus families›

definition kraus_family 𝔈  bdd_above ((λF. (E,x)F. E* oCL E) ` {F. finite F  F  𝔈})  0  fst ` 𝔈
  for 𝔈 :: ((_::chilbert_space CL _::chilbert_space) × _) set

typedef (overloaded) ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family = 
  Collect kraus_family :: (('a CL 'b) × 'x) set set
  by (rule exI[of _ {}], auto simp: kraus_family_def)
setup_lifting type_definition_kraus_family

lemma kraus_familyI:
  assumes bdd_above ((λF. (E,x)F. E* oCL E) ` {F. finite F  F  𝔈})
  assumes 0  fst ` 𝔈
  shows kraus_family 𝔈
  by (meson assms kraus_family_def)

lift_definition kf_apply :: ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a,'a) trace_class  ('b,'b) trace_class is
  λ𝔈 ρ. (E𝔈. sandwich_tc (fst E) ρ) .
notation kf_apply (infixr *kr 70)

lemma kraus_family_if_finite[iff]: kraus_family 𝔈 if finite 𝔈 and 0  fst ` 𝔈
proof -
  define B where B = ((E,x)𝔈. E* oCL E)
  have ((E,x)M. E* oCL E)  B if finite M and M  𝔈 for M
    unfolding B_def
    using finite 𝔈 M  𝔈 apply (rule sum_mono2)
    by (auto intro!: positive_cblinfun_squareI)
  with that show ?thesis
    by (auto intro!: bdd_aboveI[of _ B] simp: kraus_family_def)
qed

lemma kf_apply_scaleC:
  shows kf_apply 𝔈 (c *C x) = c *C kf_apply 𝔈 x
  by (simp add: kf_apply_def cblinfun.scaleC_right case_prod_unfold sandwich_tc_scaleC_right
      flip: infsum_scaleC_right)

lemma kf_apply_abs_summable:
  assumes kraus_family 𝔈
  shows (λ(E,x). sandwich_tc E ρ) abs_summable_on 𝔈
proof -
  wlog ρ_pos: ρ  0 generalizing ρ
  proof -
    obtain ρ1 ρ2 ρ3 ρ4 where ρ_decomp: ρ = ρ1 - ρ2 + 𝗂 *C ρ3 - 𝗂 *C ρ4
      and pos: ρ1  0 ρ2  0 ρ3  0 ρ4  0
      apply atomize_elim using trace_class_decomp_4pos'[of ρ] by blast
    have norm (sandwich_tc x ρ) 
       norm (sandwich_tc x ρ1)
      + norm (sandwich_tc x ρ2)
      + norm (sandwich_tc x ρ3)
      + norm (sandwich_tc x ρ4) 
      (is _  ?S x) for x
      by (auto simp add: ρ_decomp sandwich_tc_plus sandwich_tc_minus  sandwich_tc_scaleC_right
          scaleC_add_right scaleC_diff_right norm_mult
          intro!: norm_triangle_le norm_triangle_le_diff)
    then have *: norm (sandwich_tc (fst x) ρ)  norm (?S (fst x)) for x
      by force
    show ?thesis
      unfolding case_prod_unfold
      apply (rule abs_summable_on_comparison_test[OF _ *])
      apply (intro Misc_Kraus_Maps.abs_summable_on_add abs_summable_norm abs_summable_on_scaleC_right pos)
      using hypothesis
      by (simp_all add: case_prod_unfold pos)
  qed

  have aux: trace_norm x = Re (trace x) if x  0 and trace_class x for x
    by (metis Re_complex_of_real that(1) trace_norm_pos)
  have trace_EEρ_pos: trace ((E* oCL E) oCL from_trace_class ρ)  0 for E :: 'a CL 'b
    apply (simp add: cblinfun_assoc_right trace_class_comp_right
        flip: circularity_of_trace)
    by (auto intro!: trace_pos sandwich_pos 
        simp: cblinfun_assoc_left from_trace_class_pos ρ_pos 
        simp flip: sandwich_apply)
  have trace_EEρ_lin: linear (λM. Re (trace (M oCL from_trace_class ρ))) for M
    apply (rule linear_compose[where g=Re, unfolded o_def])
    by (auto intro!: bounded_linear.linear bounded_clinear.bounded_linear
        bounded_clinear_trace_duality' bounded_linear_Re)
  have trace_EEρ_mono: mono_on (Collect ((≤) 0)) (λA. Re (trace (A oCL from_trace_class ρ))) for M
    apply (intro mono_onI Re_mono)
    apply (subst diff_ge_0_iff_ge[symmetric])
    apply (subst trace_minus[symmetric])
    by (auto intro!: trace_class_comp_right trace_comp_pos
        simp: from_trace_class_pos ρ_pos
        simp flip: cblinfun_compose_minus_left)

  from assms
  have bdd_above ((λF. ((E,x)F. E* oCL E)) ` {F. finite F  F  𝔈})
    by (simp add: kraus_family_def)
  then have bdd_above ((λF. Re (trace (((E,x)F. E* oCL E) oCL from_trace_class ρ))) ` {F. finite F  F  𝔈})
    apply (rule bdd_above_transform_mono_pos)
    by (auto intro!: sum_nonneg positive_cblinfun_squareI[OF refl] trace_EEρ_mono
        simp: case_prod_unfold)
  then have bdd_above ((λF. (E,x)F. Re (trace ((E* oCL E) oCL from_trace_class ρ))) ` {F. F  𝔈  finite F})
    apply (subst (asm) real_vector.linear_sum[where f=λM. Re (trace (M oCL from_trace_class ρ))])
    by (auto intro!: trace_EEρ_lin simp: case_prod_unfold conj_commute)
  then have (λ(E,_). Re (trace ((E* oCL E) oCL from_trace_class ρ))) summable_on 𝔈
    apply (rule nonneg_bdd_above_summable_on[rotated])
    using trace_EEρ_pos 
    by (auto simp: less_eq_complex_def)
 then have (λ(E,_). Re (trace (from_trace_class (sandwich_tc E ρ)))) summable_on 𝔈
    by (simp add: from_trace_class_sandwich_tc sandwich_apply cblinfun_assoc_right trace_class_comp_right
        flip: circularity_of_trace)
  then have (λ(E,_). trace_norm (from_trace_class (sandwich_tc E ρ))) summable_on 𝔈
    by (simp add: aux from_trace_class_pos ρ_pos  sandwich_tc_pos)
  then show (λ(E,x). sandwich_tc E ρ) abs_summable_on 𝔈
    by (simp add: norm_trace_class.rep_eq case_prod_unfold)
qed

lemma kf_apply_summable:
  shows (λ(E,x). sandwich_tc E ρ) summable_on (Rep_kraus_family 𝔈)
  apply (rule abs_summable_summable)
  apply (rule kf_apply_abs_summable)
  using Rep_kraus_family by blast


lemma kf_apply_has_sum:
  shows ((λ(E,x). sandwich_tc E ρ) has_sum kf_apply 𝔈 ρ) (Rep_kraus_family 𝔈)
  using kf_apply_summable Rep_kraus_family[of 𝔈]
  by (auto intro!: has_sum_infsum simp add: kf_apply_def kf_apply_summable case_prod_unfold)

lemma kf_apply_plus_right:
  shows kf_apply 𝔈 (x + y) = kf_apply 𝔈 x + kf_apply 𝔈 y
  using kf_apply_summable Rep_kraus_family[of 𝔈]
  by (auto intro!: infsum_add
      simp add: kf_apply_def sandwich_tc_plus scaleC_add_right case_prod_unfold)

lemma kf_apply_uminus_right:
  shows kf_apply 𝔈 (- x) = - kf_apply 𝔈 x
  using kf_apply_summable  Rep_kraus_family[of 𝔈]
  by (auto intro!: infsum_uminus 
      simp add: kf_apply_def sandwich_tc_uminus_right scaleC_minus_right case_prod_unfold)


lemma kf_apply_minus_right:
  shows kf_apply 𝔈 (x - y) = kf_apply 𝔈 x - kf_apply 𝔈 y
  by (simp only: diff_conv_add_uminus kf_apply_plus_right kf_apply_uminus_right)

lemma kf_apply_pos:
  assumes ρ  0
  shows kf_apply 𝔈 ρ  0
  by (auto intro!: infsum_nonneg_traceclass scaleC_nonneg_nonneg of_nat_0_le_iff
      sandwich_tc_pos assms simp: kf_apply_def)

lemma kf_apply_mono_right:
  assumes ρ  τ
  shows kf_apply 𝔈 ρ  kf_apply 𝔈 τ
  apply (subst diff_ge_0_iff_ge[symmetric])
  apply (subst kf_apply_minus_right[symmetric])
  apply (rule kf_apply_pos)
  using assms by (subst diff_ge_0_iff_ge)

lemma kf_apply_geq_sum:
  assumes ρ  0 and M  Rep_kraus_family 𝔈
  shows kf_apply 𝔈 ρ  ((E,_)M. sandwich_tc E ρ)
proof (cases finite M)
  case True
  have *: (λE. sandwich_tc (fst E) ρ) summable_on X if X  Rep_kraus_family 𝔈 for X
    apply (rule summable_on_subset_banach[where A=Rep_kraus_family 𝔈])
     apply (rule kf_apply_summable[unfolded case_prod_unfold])
    using assms that by blast
  show ?thesis
    apply (subst infsum_finite[symmetric])
    using assms 
    by (auto intro!: infsum_mono_neutral_traceclass * scaleC_nonneg_nonneg of_nat_0_le_iff 
        True sandwich_tc_pos
        simp: kf_apply_def case_prod_unfold)
next
  case False
  with assms show ?thesis
    by (simp add: kf_apply_pos) 
qed

lift_definition kf_domain :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family  'x set is
  λ𝔈. snd ` 𝔈.

(* lift_definition kf_remove_0 :: ‹('a::chilbert_space,'b::chilbert_space,'x) kraus_family ⇒ ('a,'b,'x) kraus_family› is
  ‹Set.filter (λ(E,x). E≠0)›
proof -
  fix 𝔈 :: ‹('a ⇒CL 'b × 'c) set›
  have *: ‹{F. finite F ∧ F ⊆ Set.filter (λp. fst p ≠ 0) 𝔈} ⊆ {F. finite F ∧ F ⊆ 𝔈}›
    by auto
  show ‹𝔈 ∈ Collect kraus_family ⟹ Set.filter (λ(E, x). E ≠ 0) 𝔈 ∈ Collect kraus_family›
    by (auto intro: bdd_above_mono2[OF _ *] simp add: kraus_family_def case_prod_unfold)
qed

lemma kf_remove_0_idem[simp]: ‹kf_remove_0 (kf_remove_0 𝔈) = kf_remove_0 𝔈›
  apply transfer'
  by auto

lemma kf_apply_remove_0[simp]:
  ‹kf_apply (kf_remove_0 𝔈) = kf_apply 𝔈›
  by (auto intro!: ext infsum_cong_neutral simp add: kf_apply_def kf_remove_0.rep_eq)
*)

lemma kf_apply_clinear[iff]: clinear (kf_apply 𝔈)
  by (auto intro!: clinearI kf_apply_plus_right kf_apply_scaleC mult.commute)


lemma kf_apply_0_right[iff]: kf_apply 𝔈 0 = 0
  by (metis ab_left_minus kf_apply_plus_right kf_apply_uminus_right)

(* lemma kraus_familyI_0:
  assumes ‹⋀E x. (E,x) ∈ 𝔈 ⟹ E = 0›
  shows ‹kraus_family 𝔈›
proof (unfold kraus_family_def, rule bdd_aboveI2)
  fix F
  assume F: ‹F ∈ {F. finite F ∧ F ⊆ 𝔈}›
  have ‹(∑(E, x)∈F. E* oCL E) = 0›
  proof (intro sum.neutral ballI)
    fix Ex assume asm: ‹Ex ∈ F›
    obtain E x where Ex: ‹Ex = (E,x)›
      by fastforce
    with asm F assms have ‹E = 0›
      by blast
    then show ‹(case Ex of (E, x) ⇒ E* oCL E) = 0›
      by (simp add: Ex)
  qed
  then show ‹(∑(E, x)∈F. E* oCL E) ≤ 0›
    by simp
qed *)

lift_definition kf_operators :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family  ('a CL 'b) set is
  image fst :: ('a CL 'b × 'x) set  ('a CL 'b) set.

subsection ‹Bound and norm›

lift_definition kf_bound :: ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a CL 'a) is
  λ𝔈. infsum_in cweak_operator_topology (λ(E,x). E* oCL E) 𝔈 .

lemma kf_bound_def':
  kf_bound 𝔈 = Rep_cblinfun_wot ((E,x)Rep_kraus_family 𝔈. compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
  unfolding kf_bound.rep_eq infsum_euclidean_eq[symmetric]
  apply transfer'
  by simp

definition kf_norm 𝔈 = norm (kf_bound 𝔈)

lemma kf_norm_sum_bdd: bdd_above ((λF. norm ((E,x)F. E* oCL E)) ` {F. F  Rep_kraus_family 𝔈  finite F})
proof -
  from Rep_kraus_family[of 𝔈]
  obtain B where B_bound: ((E,x)F. E* oCL E)  B if F  Rep_kraus_family 𝔈 and finite F for F
    by (metis (mono_tags, lifting) bdd_above.unfold imageI kraus_family_def mem_Collect_eq)
  have norm ((E,x)F. E* oCL E)  norm B if F  Rep_kraus_family 𝔈 and finite F for F
    by (metis (no_types, lifting) B_bound norm_cblinfun_mono positive_cblinfun_squareI split_def sum_nonneg that(1) that(2))
  then show bdd_above ((λF. norm ((E,x)F. E* oCL E)) ` {F. F  Rep_kraus_family 𝔈  finite F})
    by (metis (mono_tags, lifting) bdd_aboveI2 mem_Collect_eq)
qed

lemma kf_norm_geq0[iff]:
  shows kf_norm 𝔈  0
proof (cases Rep_kraus_family 𝔈  {})
  case True
  then obtain E where E  Rep_kraus_family 𝔈 by auto
  have 0  (F{F. F  Rep_kraus_family 𝔈  finite F}. norm ((E,x)F. E* oCL E))
    apply (rule cSUP_upper2[where x={E}])
    using True by (simp_all add: E  Rep_kraus_family 𝔈 kf_norm_sum_bdd)
  then show ?thesis
    by (simp add: kf_norm_def True)
next
  case False
  then show ?thesis
    by (auto simp: kf_norm_def)
qed

lemma kf_bound_has_sum:
  shows has_sum_in cweak_operator_topology (λ(E,x). E* oCL E) (Rep_kraus_family 𝔈) (kf_bound 𝔈)
proof -
  obtain B where B: finite F  F  Rep_kraus_family 𝔈  ((E,x)F. E* oCL E)  B for F
    using Rep_kraus_family[of 𝔈]
    by (auto simp: kraus_family_def case_prod_unfold bdd_above_def)
  have summable_on_in cweak_operator_topology (λ(E, x). E* oCL E) (Rep_kraus_family 𝔈)
    using B by (auto intro!: summable_wot_boundedI positive_cblinfun_squareI simp: kraus_family_def)
  then show ?thesis
    by (auto intro!: has_sum_in_infsum_in simp: kf_bound_def)
qed

lemma kraus_family_iff_summable:
  kraus_family 𝔈  summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) 𝔈  0  fst ` 𝔈
proof (intro iffI conjI)
  assume kraus_family 𝔈
  have summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) (Rep_kraus_family (Abs_kraus_family 𝔈))
    using kraus_family 𝔈 kf_bound_has_sum summable_on_in_def by blast
  with kraus_family 𝔈 show summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) 𝔈
    by (simp add: Abs_kraus_family_inverse)
  from kraus_family 𝔈 show 0  fst ` 𝔈
    using kraus_family_def by blast
next
  assume summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) 𝔈  0  fst ` 𝔈
  then show kraus_family 𝔈
    by (auto intro!: summable_wot_bdd_above[where X=𝔈] positive_cblinfun_squareI simp: kraus_family_def)
qed

lemma kraus_family_iff_summable':
  kraus_family 𝔈  (λ(E,x). Abs_cblinfun_wot (E* oCL E)) summable_on 𝔈  0  fst ` 𝔈
  apply (transfer' fixing: 𝔈)
  by (simp add: kraus_family_iff_summable)

lemma kf_bound_summable:
  shows summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) (Rep_kraus_family 𝔈)
  using kf_bound_has_sum summable_on_in_def by blast

lemma kf_bound_has_sum':
  shows ((λ(E,x). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E)) has_sum Abs_cblinfun_wot (kf_bound 𝔈)) (Rep_kraus_family 𝔈)
  using kf_bound_has_sum[of 𝔈]
  apply transfer'
  by auto

lemma kf_bound_summable':
  ((λ(E,x). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E)) summable_on Rep_kraus_family 𝔈)
  using has_sum_imp_summable kf_bound_has_sum' by blast

lemma kf_bound_is_Sup:
  shows is_Sup ((λF. (E,x)F. E* oCL E) ` {F. finite F  F  Rep_kraus_family 𝔈}) (kf_bound 𝔈)
proof -
  from Rep_kraus_family[of 𝔈]
  obtain B where finite F  F  Rep_kraus_family 𝔈  ((E,x)F. E* oCL E)  B for F
    by (metis (mono_tags, lifting) bdd_above.unfold imageI kraus_family_def mem_Collect_eq)
  then have is_Sup ((λF. (E,x)F. E* oCL E) ` {F. finite F  F  Rep_kraus_family 𝔈})
     (infsum_in cweak_operator_topology (λ(E, x). E* oCL E) (Rep_kraus_family 𝔈))
    apply (rule infsum_wot_is_Sup[OF summable_wot_boundedI[where B=B]])
    by (auto intro!: summable_wot_boundedI positive_cblinfun_squareI simp: case_prod_beta)
  then show ?thesis
    by (auto intro!: simp: kf_bound_def)
qed

lemma kf_bound_leqI:
  assumes F. finite F  F  Rep_kraus_family 𝔈  ((E,x)F. E* oCL E)  B
  shows kf_bound 𝔈  B
  using kf_bound_is_Sup[of 𝔈]
  by (simp add: assms is_Sup_def)

lemma kf_bound_pos[iff]: kf_bound 𝔈  0
  using kf_bound_is_Sup[of 𝔈]
  by (metis (no_types, lifting) empty_subsetI finite.emptyI image_iff is_Sup_def mem_Collect_eq sum.empty)


lemma not_not_singleton_kf_norm_0: 
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  assumes ¬ class.not_singleton TYPE('a)
  shows kf_norm 𝔈 = 0
  by (simp add: not_not_singleton_cblinfun_zero[OF assms] kf_norm_def)

lemma kf_norm_sum_leqI:
  assumes F. finite F  F  Rep_kraus_family 𝔈  norm ((E,x)F. E* oCL E)  B
  shows kf_norm 𝔈  B
proof -
  have bpos: B  0
    using assms[of {}] by auto
  wlog not_singleton: class.not_singleton TYPE('a) keeping bpos
    using not_not_singleton_kf_norm_0[OF negation, of 𝔈]
    by (simp add: B  0)
  have [simp]: norm (id_cblinfun :: 'a CL 'a) = 1
    apply (rule norm_cblinfun_id[internalize_sort' 'a])
     apply (rule complex_normed_vector_axioms)
    by (rule not_singleton)
  have *: selfadjoint ((E,x)F. E* oCL E) for F :: ('a CL 'b × 'c) set
    by (auto intro!: pos_imp_selfadjoint sum_nonneg intro: positive_cblinfun_squareI)
  from assms
  have F. finite F  F  Rep_kraus_family 𝔈  ((E,x)F. E* oCL E)  B *R id_cblinfun
    apply (rule less_eq_scaled_id_norm)
    by (auto intro!: * )
  then have kf_bound 𝔈  B *R id_cblinfun
    using kf_bound_leqI by blast
  then have norm (kf_bound 𝔈)  norm (B *R (id_cblinfun :: 'a CL 'a))
    apply (rule norm_cblinfun_mono[rotated])
    by simp
  then show ?thesis
    using bpos by (simp add: kf_norm_def)
qed

lemma kf_bound_geq_sum:
  assumes M  Rep_kraus_family 𝔈
  shows ((E,_)M. E* oCL E)  kf_bound 𝔈
proof (cases finite M)
  case True
  then show ?thesis
  using kf_bound_is_Sup[of 𝔈]
  apply (simp add: is_Sup_def case_prod_beta)
  using assms by blast
next
  case False
  then show ?thesis
    by simp
qed



lemma kf_norm_geq_norm_sum:
  assumes M  Rep_kraus_family 𝔈
  shows norm ((E,_)M. E* oCL E)  kf_norm 𝔈
  using kf_bound_geq_sum assms
  by (auto intro!: norm_cblinfun_mono sum_nonneg 
      intro: positive_cblinfun_squareI
      simp add: kf_norm_def case_prod_beta)

lemma kf_bound_finite: kf_bound 𝔈 = ((E,x)Rep_kraus_family 𝔈. E* oCL E) if finite (Rep_kraus_family 𝔈)
  by (auto intro!: kraus_family_if_finite simp: kf_bound_def that infsum_in_finite)

lemma kf_norm_finite: kf_norm 𝔈 = norm ((E,x)Rep_kraus_family 𝔈. E* oCL E)
  if finite (Rep_kraus_family 𝔈)
  by (simp add: kf_norm_def kf_bound_finite that)

lemma kf_apply_bounded_pos:
  assumes ρ  0
  shows norm (kf_apply 𝔈 ρ)  kf_norm 𝔈 * norm ρ
proof -
  have norm (kf_apply 𝔈 ρ) = Re (trace_tc ((E,_)Rep_kraus_family 𝔈. sandwich_tc E ρ))
    apply (subst Re_complex_of_real[symmetric])
    apply (subst norm_tc_pos)
    using ρ  0 apply (rule kf_apply_pos)
    by (simp add: kf_apply_def case_prod_unfold)
  also have  = ((E,_)Rep_kraus_family 𝔈. Re (trace_tc (sandwich_tc E ρ)))
    using kf_apply_summable[of _ 𝔈]
    by (simp_all flip: infsum_bounded_linear[of λx. Re (trace_tc x)] 
        add: case_prod_unfold bounded_linear_compose[of Re trace_tc] bounded_linear_Re
        o_def bounded_clinear.bounded_linear)
  also have   kf_norm 𝔈 * norm ρ
  proof (rule infsum_le_finite_sums)
    show (λ(E,_). Re (trace_tc (sandwich_tc E ρ))) summable_on (Rep_kraus_family 𝔈)
      unfolding case_prod_beta
      apply (rule summable_on_bounded_linear[unfolded o_def, where h=λx. Re (trace_tc x)])
      using kf_apply_summable[of _ 𝔈]
      by (simp_all flip: infsum_bounded_linear[of λx. Re (trace_tc x)] 
        add: bounded_linear_compose[of Re trace_tc] bounded_linear_Re bounded_clinear.bounded_linear 
        o_def trace_tc_scaleC assms kf_apply_def case_prod_unfold)
    fix M :: (('a CL 'b) × 'c) set assume finite M and M  Rep_kraus_family 𝔈
    have ((E,_)M. Re (trace_tc (sandwich_tc E ρ)))
        = ((E,_)M. Re (trace (E oCL from_trace_class ρ oCL E*)))
      by (simp add: trace_tc.rep_eq from_trace_class_sandwich_tc sandwich_apply scaleC_trace_class.rep_eq trace_scaleC)
    also have  = ((E,_)M. Re (trace (E* oCL E oCL from_trace_class ρ)))
      apply (subst circularity_of_trace)
      by (auto intro!: trace_class_comp_right simp: cblinfun_compose_assoc)
    also have  = Re (trace (((E,_)M. E* oCL E) oCL from_trace_class ρ))
      by (simp only: trace_class_scaleC trace_class_comp_right trace_class_from_trace_class case_prod_unfold
          flip: Re_sum trace_scaleC trace_sum cblinfun.scaleC_left cblinfun_compose_scaleC_left cblinfun_compose_sum_left)
    also have   cmod (trace (((E,_)M. E* oCL E) oCL from_trace_class ρ))
      by (rule complex_Re_le_cmod)
    also have   norm ((E,_)M. E* oCL E) * trace_norm (from_trace_class ρ)
      apply (rule cmod_trace_times)
      by simp
    also have   kf_norm 𝔈 * norm ρ
      apply (simp add: flip: norm_trace_class.rep_eq)
      apply (rule mult_right_mono)
      apply (rule kf_norm_geq_norm_sum)
      using assms M  Rep_kraus_family 𝔈 by auto
    finally show ((E,_)M. Re (trace_tc (sandwich_tc E ρ)))  kf_norm 𝔈 * norm ρ
      by -
  qed
  finally show ?thesis 
    by -
qed

lemma kf_apply_bounded:
  ― ‹We suspect the factor 4 is not needed here but don't know how to prove that.›
  norm (kf_apply 𝔈 ρ)  4 * kf_norm 𝔈 * norm ρ
proof -
  have aux: 4 * x = x + x + x + x for x :: real
    by auto
  obtain ρ1 ρ2 ρ3 ρ4 where ρ_decomp: ρ = ρ1 - ρ2 + 𝗂 *C ρ3 - 𝗂 *C ρ4
    and pos: ρ1  0 ρ2  0 ρ3  0 ρ4  0
    and norm: norm ρ1  norm ρ norm ρ2  norm ρ norm ρ3  norm ρ norm ρ4  norm ρ
    apply atomize_elim using trace_class_decomp_4pos'[of ρ] by blast
  have norm (kf_apply 𝔈 ρ) 
          norm (kf_apply 𝔈 ρ1) +
          norm (kf_apply 𝔈 ρ2) +
          norm (kf_apply 𝔈 ρ3) +
          norm (kf_apply 𝔈 ρ4)
    by (auto intro!: norm_triangle_le norm_triangle_le_diff
        simp add: ρ_decomp kf_apply_plus_right kf_apply_minus_right
        kf_apply_scaleC)
  also have   
        kf_norm 𝔈 * norm ρ1
        + kf_norm 𝔈 * norm ρ2
        + kf_norm 𝔈 * norm ρ3
        + kf_norm 𝔈 * norm ρ4
    by (auto intro!: add_mono simp add: pos kf_apply_bounded_pos)
  also have  = kf_norm 𝔈 * (norm ρ1 + norm ρ2 + norm ρ3 + norm ρ4)
    by argo
  also have   kf_norm 𝔈 * (norm ρ + norm ρ + norm ρ + norm ρ)
    by (auto intro!: mult_left_mono add_mono kf_norm_geq0 
        simp only: aux norm)
  also have  = 4 * kf_norm 𝔈 * norm ρ
    by argo
  finally show ?thesis
    by -
qed

lemma kf_apply_bounded_clinear[bounded_clinear]:
  shows bounded_clinear (kf_apply 𝔈)
  apply (rule bounded_clinearI[where K=4 * kf_norm 𝔈])
    apply (auto intro!: kf_apply_plus_right kf_apply_scaleC mult.commute)[2]
  using kf_apply_bounded
  by (simp add: mult.commute)

lemma kf_bound_from_map: ψ C kf_bound 𝔈 φ = trace_tc (𝔈 *kr tc_butterfly φ ψ)
proof -
  have has_sum_in cweak_operator_topology (λ(E,x). E* oCL E) (Rep_kraus_family 𝔈) (kf_bound 𝔈)
    by (simp add: kf_bound_has_sum)
  then have ((λ(E, x). ψ C (E* oCL E) φ) has_sum ψ C kf_bound 𝔈 φ) (Rep_kraus_family 𝔈)
    by (auto intro!: simp: has_sum_in_cweak_operator_topology_pointwise case_prod_unfold)
  moreover have ((λ(E, x). ψ C (E* oCL E) φ) has_sum trace_tc (kf_apply 𝔈 (tc_butterfly φ ψ))) (Rep_kraus_family 𝔈)
  proof -
    have *: trace_tc (sandwich_tc E (tc_butterfly φ ψ)) = ψ C (E* oCL E) φ for E :: 'a CL 'b
      by (auto intro!: simp: trace_tc.rep_eq from_trace_class_sandwich_tc
          sandwich_apply tc_butterfly.rep_eq circularity_of_trace[symmetric, of _ E]
          trace_class_comp_left cblinfun_compose_assoc trace_butterfly_comp)
    from kf_apply_has_sum Rep_kraus_family[of 𝔈]
    have ((λ(E,x). sandwich_tc E (tc_butterfly φ ψ)) has_sum kf_apply 𝔈 (tc_butterfly φ ψ)) (Rep_kraus_family 𝔈)
      by blast
    then have ((λ(E,x). trace_tc (sandwich_tc E (tc_butterfly φ ψ))) has_sum trace_tc (kf_apply 𝔈 (tc_butterfly φ ψ))) (Rep_kraus_family 𝔈)
      unfolding case_prod_unfold
      apply (rule has_sum_bounded_linear[rotated, unfolded o_def])
      by (simp add: bounded_clinear.bounded_linear)
    then
    show ?thesis
      by (simp add: * )
  qed
  ultimately show ?thesis
    using has_sum_unique by blast
qed

lemma trace_from_kf_bound: trace_tc (𝔈 *kr ρ) = trace_tc (compose_tcr (kf_bound 𝔈) ρ)
proof -
  have separating_set bounded_clinear (Collect rank1_tc)
    apply (rule separating_set_bounded_clinear_dense)
    by simp
  moreover have bounded_clinear (λa. trace_tc (𝔈 *kr a))
    by (intro bounded_linear_intros)
  moreover have bounded_clinear (λa. trace_tc (compose_tcr (kf_bound 𝔈) a))
    by (intro bounded_linear_intros)
  moreover have trace_tc (𝔈 *kr t) = trace_tc (compose_tcr (kf_bound 𝔈) t) if t  Collect rank1_tc for t
  proof -
    from that obtain x y where t: t = tc_butterfly x y
      apply (transfer' fixing: x y)
      by (auto simp: rank1_iff_butterfly)
    then have trace_tc (𝔈 *kr t) = y C kf_bound 𝔈 x
      by (simp add: kf_bound_from_map)
    also have  = trace_tc (compose_tcr (kf_bound 𝔈) (tc_butterfly x y))
      apply (transfer' fixing: x y 𝔈)
      by (simp add: trace_butterfly_comp')
    also have  = trace_tc (compose_tcr (kf_bound 𝔈) t)
      by (simp add: t)
    finally show ?thesis
      by -
  qed
  ultimately show ?thesis
    by (rule eq_from_separatingI[where P=bounded_clinear and S=Collect rank1_tc, THEN fun_cong])
qed

lemma kf_bound_selfadjoint[iff]: selfadjoint (kf_bound 𝔈)
  by (simp add: positive_selfadjointI)

lemma kf_bound_leq_kf_norm_id:
  shows kf_bound 𝔈  kf_norm 𝔈 *R id_cblinfun
  by (auto intro!: less_eq_scaled_id_norm simp: kf_norm_def)

subsection ‹Basic Kraus families›

lemma kf_emptyset[iff]: kraus_family {}
  by (auto simp: kraus_family_def)


instantiation kraus_family :: (chilbert_space, chilbert_space, type) zero begin
lift_definition zero_kraus_family :: ('a,'b,'x) kraus_family is {}
  by simp
instance..
end

lemma kf_apply_0[simp]: kf_apply 0 = 0
  by (auto simp: kf_apply_def zero_kraus_family.rep_eq)

lemma kf_bound_0[simp]: kf_bound 0 = 0
  by (metis (mono_tags, lifting) finite.intros(1) kf_bound.rep_eq kf_bound_finite sum_clauses(1) zero_kraus_family.rep_eq)

lemma kf_norm_0[simp]: kf_norm 0 = 0
  by (simp add: kf_norm_def zero_kraus_family.rep_eq)

lift_definition kf_of_op :: ('a::chilbert_space CL 'b::chilbert_space)  ('a, 'b, unit) kraus_family is
  λE::'aCL'b. if E = 0 then {} else {(E, ())}
  by (auto intro: kraus_family_if_finite)

lemma kf_of_op0[simp]: kf_of_op 0 = 0
  apply transfer'
  by simp

lemma kf_of_op_norm[simp]: kf_norm (kf_of_op E) = (norm E)2
  by (simp add: kf_of_op.rep_eq kf_norm_finite)

lemma kf_operators_in_kf_of_op[simp]: kf_operators (kf_of_op U) = (if U = 0 then {} else {U})
  apply (transfer' fixing: U)
  by simp

lemma kf_domain_of_op[simp]: kf_domain (kf_of_op A) = {()} if A  0
  apply (transfer' fixing: A)
  using that by auto

definition kf_id = kf_of_op id_cblinfun

lemma kf_domain_id[simp]: kf_domain (kf_id :: ('a::{chilbert_space,not_singleton},_,_) kraus_family) = {()}
  by (simp add: kf_id_def)

lemma kf_of_op_id[simp]: kf_of_op id_cblinfun = kf_id
  by (simp add: kf_id_def)

lemma kf_norm_id_leq1: kf_norm kf_id  1
  apply (simp add: kf_id_def del: kf_of_op_id)
  apply transfer
  by (auto intro!: power_le_one onorm_pos_le onorm_id_le)

lemma kf_norm_id_eq1[simp]: kf_norm (kf_id :: ('a :: {chilbert_space, not_singleton},'a,unit) kraus_family) = 1
  by (auto intro!: antisym kf_norm_id_leq1 simp: kf_id_def simp del: kf_of_op_id)

lemma kf_operators_in_kf_id[simp]: kf_operators kf_id = (if (id_cblinfun::'a::chilbert_spaceCL_)=0 then {} else {id_cblinfun::'aCL_})
  by (simp add: kf_id_def del: kf_of_op_id)


instantiation kraus_family :: (chilbert_space, chilbert_space, type) scaleR begin
lift_definition scaleR_kraus_family :: real  ('a::chilbert_space,'b::chilbert_space,'x) kraus_family  ('a,'b,'x) kraus_family is
  λr 𝔈. if r  0 then {} else (λ(E,x). (sqrt r *R E, x)) ` 𝔈
proof (rename_tac r 𝔈)
  fix r and 𝔈 :: ('a CL 'b × 'x) set
  assume asm: 𝔈  Collect kraus_family
  define scaled where scaled = (λ(E, y). (sqrt r *R E, y)) ` 𝔈
  show (if r  0 then {} else scaled)  Collect kraus_family
  proof (cases r > 0)
    case True
    obtain B where B: ((E,x)M. E* oCL E)  B if M  𝔈 and finite M for M
      using asm
      by (auto intro!: simp: kraus_family_def bdd_above_def)
    have ((E,x)M. E* oCL E)  r *R B if Mr𝔈: M  scaled and finite M for M
    proof -
      define M' where M' = (λ(E,x). (E /R sqrt r, x)) ` M
      then have finite M'
        using that(2) by blast
      moreover have M'  𝔈
        using Mr𝔈 True by (auto intro!: simp: M'_def scaled_def)
      ultimately have 1: ((E,x)M'. E* oCL E)  B
        using B by auto
      have 2: ((E,x)M. E* oCL E) = r *R ((E,x)M'. E* oCL E)
        apply (simp add: M'_def case_prod_unfold)
        apply (subst infsum_reindex)
        using True
        by (auto intro!: inj_onI simp: o_def infsum_scaleR_right
            simp flip: inverse_mult_distrib)
      show ?thesis
        using 1 2 True scaleR_le_cancel_left_pos by auto
    qed
    moreover have 0  fst ` scaled if r  0
      using asm that
      by (auto intro!: simp: scaled_def kraus_family_def)
    ultimately show ?thesis
      by (auto intro!: simp: kraus_family_def bdd_above_def)
  next
    case False
    then show ?thesis
      by (simp add: scaled_def)
  qed
qed
instance..
end

lemma kf_scale_apply:
  assumes r  0
  shows kf_apply (r *R 𝔈) ρ = r *R kf_apply 𝔈 ρ
proof (cases r > 0)
  case True
  then show ?thesis
    apply (simp add: scaleR_kraus_family.rep_eq kf_apply_def)
    apply (subst infsum_reindex)
    by (auto intro!: inj_onI
        simp: kf_apply_def case_prod_unfold
        o_def sandwich_tc_scaleR_left cblinfun.scaleR_left infsum_scaleR_right)
next
  case False
  with assms show ?thesis
    by (simp add: kf_apply.rep_eq scaleR_kraus_family.rep_eq)
qed

lemma kf_bound_scale[simp]: kf_bound (c *R 𝔈) = c *R kf_bound 𝔈 if c  0
  apply (rule cblinfun_cinner_eqI)
  using that
  by (simp add: kf_bound_from_map cblinfun.scaleR_left kf_scale_apply scaleR_scaleC trace_tc_scaleC)

lemma kf_norm_scale[simp]: kf_norm (c *R 𝔈) = c * kf_norm 𝔈 if c  0
  by (simp add: kf_norm_def that)

lemma kf_of_op_apply: kf_apply (kf_of_op E) ρ = sandwich_tc E ρ
  by (simp add: kf_apply_def kf_of_op.rep_eq)

lemma kf_id_apply[simp]: kf_apply kf_id ρ = ρ
  by (simp add: kf_id_def kf_of_op_apply del: kf_of_op_id)

lemma kf_scale_apply_neg:
  assumes r  0
  shows kf_apply (r *R 𝔈) = 0
  apply (transfer fixing: r)
  using assms
  by (auto intro!: ext simp: scaleR_kraus_family.rep_eq kf_apply.rep_eq)

lemma kf_apply_0_left[iff]: kf_apply 0 ρ = 0
  apply (transfer' fixing: ρ)
  by simp

lemma kf_bound_of_op[simp]: kf_bound (kf_of_op A) = A* oCL A
  by (simp add: kf_bound_def kf_of_op.rep_eq infsum_in_finite)

lemma kf_bound_id[simp]: kf_bound kf_id = id_cblinfun
  by (simp add: kf_id_def del: kf_of_op_id)

subsection ‹Filtering›

lift_definition kf_filter :: ('x  bool)  ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family is
  λ(P::'x  bool) 𝔈. Set.filter (λ(E,x). P x) 𝔈
proof (rename_tac P 𝔈, rule CollectI)
  fix P 𝔈
  assume 𝔈  Collect kraus_family
  then have kraus_family 𝔈
    by simp
  then have bdd_above (sum (λ(E, x). E* oCL E) ` {F. finite F  F  𝔈})
    by (simp add: kraus_family_def)
  then have bdd_above (sum (λ(E, x). E* oCL E) ` {F. finite F  F  Set.filter P 𝔈})
    apply (rule bdd_above_mono2)
    by auto
  moreover from kraus_family 𝔈 have 0  fst ` Set.filter P 𝔈
    by (auto simp add: kraus_family_def)
  ultimately show kraus_family (Set.filter P 𝔈)
    by (simp add: kraus_family_def)
qed

lemma kf_filter_false[simp]: kf_filter (λ_. False) 𝔈 = 0
  apply transfer by auto

lemma kf_filter_true[simp]: kf_filter (λ_. True) 𝔈 = 𝔈
  apply transfer by auto

definition kf_apply_on 𝔈 S = kf_apply (kf_filter (λx. x  S) 𝔈)
notation kf_apply_on ("_ *kr @_/ _" [71, 1000, 70] 70)

lemma kf_apply_on_pos:
  assumes ρ  0
  shows kf_apply_on 𝔈 X ρ  0
  by (simp add: kf_apply_on_def kf_apply_pos assms)

lemma kf_apply_on_bounded_clinear[bounded_clinear]:
  shows bounded_clinear (kf_apply_on 𝔈 X)
  by (simp add: kf_apply_on_def kf_apply_bounded_clinear)

lemma kf_filter_twice:
  kf_filter P (kf_filter Q 𝔈) = kf_filter (λx. P x  Q x) 𝔈
  apply (transfer' fixing: P Q)
  by auto

lemma kf_apply_on_union_has_sum:
  assumes X Y. XF  YF  XY  disjnt X Y
  shows ((λX. kf_apply_on 𝔈 X ρ) has_sum (kf_apply_on 𝔈 (F) ρ)) F
proof -
  define EE EEf where EE = Rep_kraus_family 𝔈 and EEf X = Set.filter (λ(E,x). xX) EE for X
  have inj: inj_on snd (SIGMA X:F. EEf X)
    using assms by (force intro!: simp: inj_on_def disjnt_def EEf_def)
  have snd_Sigma: snd ` (SIGMA X:F. EEf X) = EEf (F)
    apply (subgoal_tac a b x. (a, b)  EE  x  F  b  x  (a, b)  snd ` (SIGMA X:F. Set.filter (λ(E, x). x  X) EE))
     apply (auto simp add: EEf_def)[1]
    by force
  have map'_infsum: kf_apply_on 𝔈 X ρ = ((E, x)EEf X. sandwich_tc E ρ) for X
    by (simp add: kf_apply_on_def kf_apply.rep_eq EEf_def kf_filter.rep_eq EE_def case_prod_unfold)
  have has_sum: ((λ(E,x). sandwich_tc E ρ) has_sum (kf_apply_on 𝔈 X ρ)) (EEf X) for X
    using kf_apply_has_sum[of ρ kf_filter (λx. x  X) 𝔈]
    by (simp add: kf_apply_on_def kf_filter.rep_eq EEf_def EE_def)
  then have ((λ(E,x). sandwich_tc E ρ) has_sum (kf_apply_on 𝔈 (F) ρ)) (snd ` (SIGMA X:F. EEf X))
    by (simp add: snd_Sigma)
  then have ((λ(X,(E,x)). sandwich_tc E ρ) has_sum (kf_apply_on 𝔈 (F) ρ)) (SIGMA X:F. EEf X)
    apply (subst (asm) has_sum_reindex)
     apply (rule inj)
    by (simp add: o_def case_prod_unfold)
  then have ((λX. (E, x)EEf X. sandwich_tc E ρ) has_sum kf_apply_on 𝔈 (F) ρ) F
    by (rule has_sum_Sigma'_banach)
  then show ((λX. kf_apply_on 𝔈 X ρ) has_sum kf_apply_on 𝔈 (F) ρ) F
    by (auto intro: has_sum_cong[THEN iffD2, rotated] simp: map'_infsum)
qed

lemma kf_apply_on_empty[simp]: kf_apply_on E {} ρ = 0
  by (simp add: kf_apply_on_def)

lemma kf_apply_on_union_eqI:
  assumes X Y. (X,Y)F  kf_apply_on 𝔈 X ρ = kf_apply_on 𝔉 Y ρ
  assumes X Y X' Y'. (X,Y)F  (X',Y')F  (X,Y)(X',Y')  disjnt X X'
  assumes X Y X' Y'. (X,Y)F  (X',Y')F  (X,Y)(X',Y')  disjnt Y Y'
  assumes XX: XX = (fst ` F) and YY: YY = (snd ` F)
  shows kf_apply_on 𝔈 XX ρ = kf_apply_on 𝔉 YY ρ
proof -
  define F' where F' = Set.filter (λ(X,Y). X{}  Y{}) F
  have inj1: inj_on fst F'
    apply (rule inj_onI)
    using assms(2)
    unfolding F'_def
    by fastforce
  have inj2: inj_on snd F'
    apply (rule inj_onI)
    using assms(3)
    unfolding F'_def
    by fastforce
  have ((λX. kf_apply_on 𝔈 X ρ) has_sum kf_apply_on 𝔈 XX ρ) (fst ` F)
    unfolding XX
    apply (rule kf_apply_on_union_has_sum)
    using assms(2) by force
  then have ((λX. kf_apply_on 𝔈 X ρ) has_sum kf_apply_on 𝔈 XX ρ) (fst ` F')
  proof (rule has_sum_cong_neutral[OF _ _ refl, THEN iffD2, rotated -1])
    show kf_apply_on 𝔈 X ρ = 0 if X  fst ` F' - fst ` F for X
      using that by (auto simp: F'_def)
    show kf_apply_on 𝔈 X ρ = 0 if X  fst ` F - fst ` F' for X
    proof -
      from that obtain Y where (X,Y)  F and X = {}  Y = {}
        apply atomize_elim
        by (auto intro!: simp: F'_def)
      then consider (X) X = {} | (Y) Y = {}
        by auto
      then show ?thesis
      proof cases
        case X
        then show ?thesis
          by simp
      next
        case Y
        then have kf_apply_on 𝔉 Y ρ = 0
          by simp
        then show ?thesis
          using (X, Y)  F assms(1) by presburger
      qed
    qed
  qed
  then have sum1: ((λ(X,Y). kf_apply_on 𝔈 X ρ) has_sum kf_apply_on 𝔈 XX ρ) F'
    apply (subst (asm) has_sum_reindex)
    using inj1 by (auto intro!: simp: o_def case_prod_unfold)
  have ((λY. kf_apply_on 𝔉 Y ρ) has_sum kf_apply_on 𝔉 YY ρ) (snd ` F)
    unfolding YY
    apply (rule kf_apply_on_union_has_sum)
    using assms(3) by force
  then have ((λY. kf_apply_on 𝔉 Y ρ) has_sum kf_apply_on 𝔉 YY ρ) (snd ` F')
  proof (rule has_sum_cong_neutral[OF _ _ refl, THEN iffD2, rotated -1])
    show kf_apply_on 𝔉 Y ρ = 0 if Y  snd ` F' - snd ` F for Y
      using that by (auto simp: F'_def)
    show kf_apply_on 𝔉 Y ρ = 0 if Y  snd ` F - snd ` F' for Y
    proof -
      from that obtain X where (X,Y)  F and X = {}  Y = {}
        apply atomize_elim
        by (auto intro!: simp: F'_def)
      then consider (X) X = {} | (Y) Y = {}
        by auto
      then show ?thesis
      proof cases
        case Y
        then show ?thesis
          by simp
      next
        case X
        then have kf_apply_on 𝔈 X ρ = 0
          by simp
        then show ?thesis
          using (X, Y)  F assms(1) by simp
      qed
    qed
  qed
  then have ((λ(X,Y). kf_apply_on 𝔉 Y ρ) has_sum kf_apply_on 𝔉 YY ρ) F'
    apply (subst (asm) has_sum_reindex)
    using inj2 by (auto intro!: simp: o_def case_prod_unfold)
  then have sum2: ((λ(X,Y). kf_apply_on 𝔈 X ρ) has_sum kf_apply_on 𝔉 YY ρ) F'
    apply (rule has_sum_cong[THEN iffD1, rotated -1])
    using assms(1) by (auto simp: F'_def)
  from sum1 sum2 show ?thesis
    using has_sum_unique by blast
qed

lemma kf_apply_on_UNIV[simp]: kf_apply_on 𝔈 UNIV = kf_apply 𝔈
  by (auto simp: kf_apply_on_def)

lemma kf_apply_on_CARD_1[simp]: (λ𝔈. kf_apply_on 𝔈 {x::_::CARD_1}) = kf_apply
  apply (subst asm_rl[of {x} = UNIV])
  by auto

(* lemma kf_apply_eqI_from_map':
  assumes ‹⋀x. kf_apply_on 𝔈 {x} = kf_apply_on 𝔉 {x}›
  shows ‹kf_apply 𝔈 = kf_apply 𝔉›
  apply (subst kf_apply_on_UNIV[symmetric])+
  apply (rule ext)
  apply (rule kf_apply_on_union_eqI[where F=‹range (λx. ({x},{x}))› and 𝔈=𝔈 and 𝔉=𝔉])
  using assms by auto *)

lemma kf_apply_on_union_summable_on:
  assumes X Y. XF  YF  XY  disjnt X Y
  shows (λX. kf_apply_on 𝔈 X ρ) summable_on F
  using assms by (auto intro!: has_sum_imp_summable kf_apply_on_union_has_sum)

lemma kf_apply_on_union_infsum:
  assumes X Y. XF  YF  XY  disjnt X Y
  shows (XF. kf_apply_on 𝔈 X ρ) = kf_apply_on 𝔈 (F) ρ
  by (metis assms infsumI kf_apply_on_union_has_sum)


lemma kf_bound_filter:
  kf_bound (kf_filter P 𝔈)  kf_bound 𝔈
proof (unfold kf_bound.rep_eq, rule infsum_mono_neutral_wot)
  show summable_on_in cweak_operator_topology (λ(E, x). E* oCL E) (Rep_kraus_family (kf_filter P 𝔈))
    using Rep_kraus_family kraus_family_iff_summable by blast
  show summable_on_in cweak_operator_topology (λ(E, x). E* oCL E) (Rep_kraus_family 𝔈)
    using Rep_kraus_family kraus_family_iff_summable by blast
  fix Ex :: 'a CL 'b × 'c
  show (case Ex of (E, x)  E* oCL E)  (case Ex of (E, x)  E* oCL E)
    by simp
  show 0  (case Ex of (E, x)  E* oCL E)
    by (auto intro!: positive_cblinfun_squareI simp: case_prod_unfold)
  show (case Ex of (E, x)  E* oCL E)  0
    if Ex  Rep_kraus_family (kf_filter P 𝔈) - Rep_kraus_family 𝔈
    using that
    by (auto simp: kf_filter.rep_eq)
qed

lemma kf_norm_filter:
  kf_norm (kf_filter P 𝔈)  kf_norm 𝔈
  unfolding kf_norm_def
  apply (rule norm_cblinfun_mono)
  by (simp_all add: kf_bound_filter)

lemma kf_domain_filter[simp]:
  kf_domain (kf_filter P E) = Set.filter P (kf_domain E)
  apply (transfer' fixing: P)
  by force

(* lemma kf_filter_remove0:
  ‹kf_filter f (kf_remove_0 𝔈) = kf_remove_0 (kf_filter f 𝔈)›
  apply (transfer' fixing: f)
  by auto *)

lemma kf_filter_0_right[simp]: kf_filter P 0 = 0
  apply (transfer' fixing: P)
  by auto

lemma kf_apply_on_0_right[simp]: kf_apply_on 𝔈 X 0 = 0
  by (simp add: kf_apply_on_def)

lemma kf_apply_on_0_left[simp]: kf_apply_on 0 X ρ = 0
  by (simp add: kf_apply_on_def)

lemma kf_apply_on_mono3:
  assumes ρ  σ
  shows 𝔈 *kr @X ρ  𝔈 *kr @X σ
  by (simp add: assms kf_apply_mono_right kf_apply_on_def)

lemma kf_apply_on_mono2:
  assumes X  Y and ρ  0
  shows 𝔈 *kr @X ρ  𝔈 *kr @Y ρ
proof -
  wlog Y  {}
    using assms(1) negation by auto
  have [simp]: X  Y = Y
    using assms(1) by blast
  from kf_apply_on_union_infsum[where F={X, Y-X} and 𝔈=𝔈 and ρ=ρ]
  have (X{X, Y - X}. 𝔈 *kr @X ρ) = 𝔈 *kr @Y ρ
    by (auto simp: disjnt_iff sum.insert)
  then have 𝔈 *kr @X ρ + 𝔈 *kr @(Y-X) ρ = 𝔈 *kr @Y ρ 
    apply (subst (asm) sum.insert)
    using Y  {}     
    by auto
  moreover have 𝔈 *kr @(Y-X) ρ  0
    by (simp add: assms(2) kf_apply_on_pos)
  ultimately show 𝔈 *kr @X ρ  𝔈 *kr @Y ρ
    by (metis le_add_same_cancel1)
qed

lemma kf_operators_filter: kf_operators (kf_filter P 𝔈)  kf_operators 𝔈
  apply (transfer' fixing: P)
  by auto

lemma kf_equal_if_filter_equal:
  assumes x. kf_filter ((=)x) 𝔈 = kf_filter ((=)x) 𝔉
  shows 𝔈 = 𝔉
  using assms
  apply transfer'
  by fastforce

subsection ‹Equivalence›

definition kf_eq_weak 𝔈 𝔉  kf_apply 𝔈 = kf_apply 𝔉
definition kf_eq 𝔈 𝔉  (x. kf_apply_on 𝔈 {x} = kf_apply_on 𝔉 {x})

open_bundle kraus_map_syntax begin
notation kf_eq_weak (infix "=kr" 50)
notation kf_eq (infix "kr" 50)
notation kf_apply (infixr *kr 70)
notation kf_apply_on ("_ *kr @_/ _" [71, 1000, 70] 70)
end

lemma kf_eq_weak_reflI[iff]: x =kr x
  by (simp add: kf_eq_weak_def)

lemma kf_eq_weak_refl0[iff]: 0 =kr 0
  by (simp add: kf_eq_weak_def)

lemma kf_bound_cong:
  assumes 𝔈 =kr 𝔉
  shows kf_bound 𝔈 = kf_bound 𝔉
  using assms by (auto intro!: cblinfun_cinner_eqI simp: kf_eq_weak_def kf_bound_from_map)

lemma kf_norm_cong:
  assumes 𝔈 =kr 𝔉
  shows kf_norm 𝔈 = kf_norm 𝔉
  using assms
  by (simp add: kf_norm_def kf_bound_cong)


lemma kf_eq_weakI:
  assumes ρ. ρ  0  𝔈 *kr ρ = 𝔉 *kr ρ
  shows 𝔈 =kr 𝔉
  unfolding kf_eq_weak_def
  apply (rule eq_from_separatingI)
     apply (rule separating_density_ops[where B=1])
  using assms by auto

lemma kf_eqI: 
  assumes x ρ. ρ  0  𝔈 *kr @{x} ρ = 𝔉 *kr @{x} ρ
  shows 𝔈 kr 𝔉
  using kf_eq_weakI
  using assms
  by (auto simp: kf_eq_weak_def kf_eq_def kf_apply_on_def)

(* lemma kf_norm_remove_0[simp]:
  ― ‹Logically belongs in a different section but uses lemmas from this section.›
  ‹kf_norm (kf_remove_0 𝔈) = kf_norm 𝔈›
  by (auto intro!: kf_eq_weakI kf_norm_cong) *)

lemma kf_eq_weak_trans[trans]:
  F =kr G  G =kr H  F =kr H
  by (simp add: kf_eq_weak_def)

lemma kf_apply_eqI:
  assumes 𝔈 =kr 𝔉
  shows 𝔈 *kr ρ = 𝔉 *kr ρ
  using assms by (simp add: kf_eq_weak_def)

lemma kf_eq_imp_eq_weak:
  assumes 𝔈 kr 𝔉
  shows 𝔈 =kr 𝔉
  unfolding kf_eq_weak_def 
  apply (subst kf_apply_on_UNIV[symmetric])+
  apply (rule ext)
  apply (rule kf_apply_on_union_eqI[where F=range (λx. ({x},{x})) and 𝔈=𝔈 and 𝔉=𝔉])
  using assms by (auto simp: kf_eq_def)

lemma kf_filter_cong_eq[cong]:
  assumes 𝔈 = 𝔉
  assumes x. x  kf_domain 𝔈  P x = Q x
  shows kf_filter P 𝔈 = kf_filter Q 𝔉
  using assms
  apply transfer
  by (force simp: kraus_family_def)

lemma kf_filter_cong:
  assumes 𝔈 kr 𝔉
  assumes x. x  kf_domain 𝔈  P x = Q x
  shows kf_filter P 𝔈 kr kf_filter Q 𝔉
proof -
  have kf_apply (kf_filter (λxa. xa = x  P xa) 𝔈)
      = kf_apply (kf_filter (λxa. xa = x  Q xa) 𝔈) for x
  proof (cases x  kf_domain 𝔈)
    case True
    with assms have P x = Q x
      by blast
    then have (λxa. xa = x  P xa) = (λxa. xa = x  Q xa)
      by auto
    then show ?thesis
      by presburger
  next
    case False
    then have kf_filter (λxa. xa = x  P xa) 𝔈 = 0
      apply (transfer fixing: P x)
      by force
    have *: (ESet.filter (λ(E, xa). xa = x  P xa) 𝔈. sandwich_tc (fst E) ρ) = 0
      if x  snd ` Set.filter (λ(E, x). E  0) 𝔈 for 𝔈 :: ('a CL 'b × 'c) set and ρ P
      apply (rule infsum_0)
      using that
      by force
    have kf_apply (kf_filter (λxa. xa = x  P xa) 𝔈) = 0 for P
      using False
      apply (transfer fixing: x P)
      using *
      by (auto intro!: ext simp: kraus_family_def image_iff)
    then show ?thesis
      by simp
  qed
  also have kf_apply (kf_filter (λxa. xa = x  Q xa) 𝔈)
           = kf_apply (kf_filter (λxa. xa = x  Q xa) 𝔉) for x
  proof (cases Q x)
    case True
    then have (z = x  Q z)  (z = x) for z
      by auto
    with assms show ?thesis
      by (simp add: kf_eq_def kf_apply_on_def)
  next
    case False
    then have [simp]: (z = x  Q z)  False for z
      by auto
    show ?thesis
      by (simp add: kf_eq_def kf_apply_on_def)
  qed
  finally show ?thesis
    by (simp add: kf_eq_def kf_apply_on_def kf_filter_twice)
qed


lemma kf_filter_cong_weak:
  assumes 𝔈 kr 𝔉
  assumes x. x  kf_domain 𝔈  P x = Q x
  shows kf_filter P 𝔈 =kr kf_filter Q 𝔉
  by (simp add: assms kf_eq_imp_eq_weak kf_filter_cong)

lemma kf_eq_refl[iff]: 𝔈 kr 𝔈
  using kf_eq_def by blast

lemma kf_eq_trans[trans]:
  assumes 𝔈 kr 𝔉
  assumes 𝔉 kr 𝔊
  shows 𝔈 kr 𝔊
  by (metis assms(1) assms(2) kf_eq_def)

lemma kf_eq_sym[sym]:
  assumes 𝔈 kr 𝔉
  shows 𝔉 kr 𝔈
  by (metis assms kf_eq_def)

lemma kf_eq_weak_imp_eq_CARD_1:
  fixes 𝔈 𝔉 :: ('a::chilbert_space, 'b::chilbert_space, 'x::CARD_1) kraus_family
  assumes 𝔈 =kr 𝔉
  shows 𝔈 kr 𝔉
  by (metis CARD_1_UNIV assms kf_eqI kf_eq_weak_def kf_apply_on_UNIV)

lemma kf_apply_on_eqI_filter:
  assumes kf_filter (λx. x  X) 𝔈 kr kf_filter (λx. x  X) 𝔉
  shows 𝔈 *kr @X ρ = 𝔉 *kr @X ρ
proof (rule kf_apply_on_union_eqI[where F=(λx.({x},{x})) ` X])
  show (A, B)  (λx. ({x}, {x})) ` X 
       (A', B')  (λx. ({x}, {x})) ` X  (A, B)  (A', B')  disjnt A A'
    for A B A' B'
    by force
  show (A, B)  (λx. ({x}, {x})) ` X 
       (A', B')  (λx. ({x}, {x})) ` X  (A, B)  (A', B')  disjnt B B'
    for A B A' B'
    by force
  show X =  (fst ` (λx. ({x}, {x})) ` X)
    by simp
  show X =  (snd ` (λx. ({x}, {x})) ` X)
    by simp
  show 𝔈 *kr @A ρ = 𝔉 *kr @B ρ if (A, B)  (λx. ({x}, {x})) ` X for A B
  proof -
    from that obtain x where x  X and A: A = {x} and B: B = {x}
      by blast
    from x  X have *: (λx'. x = x'  x'  X) = (λx'. x' = x)
      by blast
    from assms have kf_filter ((=)x) (kf_filter (λx. x  X) 𝔈) =kr kf_filter ((=)x) (kf_filter (λx. x  X) 𝔉)
      by (simp add: kf_filter_cong_weak)
    then have kf_filter (λx'. x'=x) 𝔈 =kr kf_filter (λx'. x'=x) 𝔉
      by (simp add: kf_filter_twice * cong del: kf_filter_cong_eq)
    then have 𝔈 *kr @{x} ρ = 𝔉 *kr @{x} ρ
      by (simp add: kf_apply_on_def kf_eq_weak_def)
    then show ?thesis
      by (simp add: A B)
  qed
qed

lemma kf_apply_on_eqI:
  assumes 𝔈 kr 𝔉
  shows 𝔈 *kr @X ρ = 𝔉 *kr @X ρ
  apply (rule kf_apply_on_union_eqI[where F=(λx.({x},{x})) ` X])
  using assms by (auto simp: kf_eq_def)

lemma kf_apply_eq0I:
  assumes 𝔈 =kr 0
  shows 𝔈 *kr ρ = 0
  using assms kf_eq_weak_def by force

lemma kf_eq_weak0_imp_kf_eq0:
  assumes 𝔈 =kr 0
  shows 𝔈 kr 0
proof -
  have 𝔈 *kr @{x} ρ = 0 if ρ  0 for ρ x
  proof -
    from assms have 𝔈 *kr @UNIV ρ = 0
      by (simp add: kf_eq_weak_def)
    moreover have 𝔈 *kr @{x} ρ  𝔈 *kr @UNIV ρ
      apply (rule kf_apply_on_mono2)
      using that by auto
    moreover have 𝔈 *kr @{x} ρ  0
      using that
      by (simp add: kf_apply_on_pos)
    ultimately show ?thesis
      by (simp add: basic_trans_rules(24))
  qed
  then show ?thesis
    by (simp add: kf_eqI)
qed

lemma kf_apply_on_eq0I:
  assumes 𝔈 =kr 0
  shows 𝔈 *kr @X ρ = 0
proof -
  from assms
  have 𝔈 kr 0
    by (rule kf_eq_weak0_imp_kf_eq0)
  then have 𝔈 *kr @X ρ = 0 *kr @X ρ
    by (intro kf_apply_on_eqI_filter kf_filter_cong refl)
  then show ?thesis
    by simp
qed

(* lemma kf_remove_0_eq_weak[iff]: ‹kf_remove_0 𝔈 =kr 𝔈›
  by (simp add: kf_eq_weak_def)

lemma kf_remove_0_eq[iff]: ‹kf_remove_0 𝔈 ≡kr 𝔈›
  by (simp add: kf_eq_def kf_apply_on_def kf_filter_remove0) *)

lemma kf_filter_to_domain[simp]:
  kf_filter (λx. x  kf_domain 𝔈) 𝔈 = 𝔈
  apply transfer
  by (force simp: kraus_family_def)

lemma kf_eq_0_iff_eq_0: E =kr 0  E = 0
proof (rule iffI)
  assume asm: E =kr 0
  show E = 0
  proof (insert asm, unfold kf_eq_weak_def, transfer, rename_tac 𝔈)
    fix 𝔈 :: ('a CL 'b × 'c) set
    assume 𝔈  Collect kraus_family
    then have kraus_family 𝔈
      by simp
    have summable1: (λE. sandwich_tc (fst E) ρ) summable_on 𝔈 for ρ
      apply (rule abs_summable_summable)
      using kf_apply_abs_summable[OF kraus_family 𝔈]
      by (simp add: case_prod_unfold)
    then have summable2: (λE. sandwich_tc (fst E) ρ) summable_on 𝔈-{E} for E ρ
      apply (rule summable_on_subset_banach)
      by simp
    assume (λρ. E𝔈. sandwich_tc (fst E) ρ) = (λρ. E{}. sandwich_tc (fst E) ρ)
    then have sum0: (E𝔈. sandwich_tc (fst E) ρ) = 0 for ρ
      apply simp by meson
    have sand_Eρ_0:  sandwich_tc (fst E) ρ = 0 if E  𝔈 and ρ  0 for E ρ
    proof (rule ccontr)
      assume Eρ_neq0: sandwich_tc (fst E) ρ  0
      have Eρ_geq0: sandwich_tc (fst E) ρ  0
        by (simp add: 0  ρ sandwich_tc_pos)
      have Eρ_ge0: sandwich_tc (fst E) ρ > 0
        using Eρ_neq0 Eρ_geq0 by order
      have (E𝔈. sandwich_tc (fst E) ρ) = (E𝔈-{E}. sandwich_tc (fst E) ρ) + sandwich_tc (fst E) ρ
        apply (subst asm_rl[of 𝔈 = insert E (𝔈-{E})])
        using that apply blast
        apply (subst infsum_insert)
        by (auto intro!: summable2)
      also have   0 + sandwich_tc (fst E) ρ (is _  )
        by (simp add: 0  ρ infsum_nonneg_traceclass sandwich_tc_pos)
      also have  > 0 (is _ > )
        using Eρ_ge0
        by simp
      ultimately have (E𝔈. sandwich_tc (fst E) ρ) > 0
        by simp
      then show False
        using sum0 by simp
    qed
    then have fst E = 0 if E  𝔈 for E
      apply (rule sandwich_tc_eq0_D[where B=1])
      using that by auto
    then show 𝔈 = {}
      using kraus_family 𝔈
      by (auto simp: kraus_family_def)
  qed
next
  assume E = 0
  then show E =kr 0
    by simp
qed

lemma in_kf_domain_iff_apply_nonzero:
  x  kf_domain 𝔈  kf_apply_on 𝔈 {x}  0
proof -
  define 𝔈' where 𝔈' = Rep_kraus_family 𝔈
  have x  kf_domain 𝔈  ((E,x')Rep_kraus_family 𝔈. x'x)
    by (force simp: kf_domain.rep_eq)
  also have   ((E,x')Rep_kraus_family (kf_filter (λy. y=x) 𝔈). False)
    by (auto simp: kf_filter.rep_eq)
  also have   Rep_kraus_family (kf_filter (λy. y=x) 𝔈) = {}
    by (auto simp:)
  also have   (kf_filter (λy. y=x) 𝔈) = 0
    using Rep_kraus_family_inject zero_kraus_family.rep_eq by auto
  also have   kf_apply (kf_filter (λy. y=x) 𝔈) = 0
    apply (subst kf_eq_0_iff_eq_0[symmetric])
    by (simp add: kf_eq_weak_def)
  also have   kf_apply_on 𝔈 {x} = 0
    by (simp add: kf_apply_on_def)
  finally show ?thesis
    by auto
qed


lemma kf_domain_cong:
  assumes 𝔈 kr 𝔉
  shows kf_domain 𝔈 = kf_domain 𝔉
  apply (rule Set.set_eqI)
  using assms
  by (simp add: kf_eq_def in_kf_domain_iff_apply_nonzero)

lemma kf_eq_weak_sym[sym]:
  assumes 𝔈 =kr 𝔉
  shows 𝔉 =kr 𝔈
  by (metis assms kf_eq_weak_def)

lemma kf_eqI_from_filter_eq_weak:
  assumes x. kf_filter ((=) x) E =kr kf_filter ((=) x) F
  shows E kr F
  using assms
  apply (simp add: kf_eq_weak_def kf_eq_def kf_apply_on_def)
  apply (subst flip_eq_const)
  apply (subst flip_eq_const)
  by simp


lemma kf_eq_weak_from_separatingI:
  fixes E :: ('q::chilbert_space,'r::chilbert_space,'x) kraus_family
    and F :: ('q,'r,'y) kraus_family
  assumes separating_set (bounded_clinear :: (('q,'q) trace_class  ('r,'r) trace_class)  bool) S
  assumes ρ. ρ  S  E *kr ρ = F *kr ρ
  shows E =kr F
proof -
  have kf_apply E = kf_apply F
    by (metis assms(1) assms(2) kf_apply_bounded_clinear separating_set_def)
  then show ?thesis
    by (simp add: kf_eq_weakI)
qed

lemma kf_eq_weak_eq_trans[trans]: a =kr b  b kr c  a =kr c
  by (metis kf_eq_imp_eq_weak kf_eq_weak_def)

lemma kf_eq_eq_weak_trans[trans]: a kr b  b =kr c  a =kr c
  by (metis kf_eq_imp_eq_weak kf_eq_weak_def)

instantiation kraus_family :: (chilbert_space,chilbert_space,type) preorder begin
definition less_eq_kraus_family where 𝔈  𝔉  (x. ρ0. kf_apply_on 𝔈 {x} ρ  kf_apply_on 𝔉 {x} ρ)
definition less_kraus_family where 𝔈 < 𝔉  𝔈  𝔉  ¬ 𝔈 kr 𝔉
lemma kf_antisym: 𝔈 kr 𝔉  𝔈  𝔉  𝔉  𝔈
  for 𝔈 𝔉 :: ('a, 'b, 'c) kraus_family
  by (smt (verit, ccfv_SIG) kf_apply_on_eqI kf_eqI less_eq_kraus_family_def order.refl
      order_antisym_conv)
instance
proof (intro_classes)
  fix 𝔈 𝔉 𝔊 :: ('a, 'b, 'c) kraus_family
  show (𝔈 < 𝔉) = (𝔈  𝔉  ¬ 𝔉  𝔈)
    using kf_antisym less_kraus_family_def by auto
  show 𝔈  𝔈
    using less_eq_kraus_family_def by auto
  show 𝔈  𝔉  𝔉  𝔊  𝔈  𝔊
    by (meson basic_trans_rules(23) less_eq_kraus_family_def)
qed
end

lemma kf_apply_on_mono1: 
  assumes 𝔈  𝔉 and ρ  0 
  shows 𝔈 *kr @X ρ  𝔉 *kr @X ρ
proof -
  have [simp]: ((λx. {x}) ` X) = X for X :: 'c set
    by auto
  have ((λX. 𝔈 *kr @X ρ) has_sum 𝔈 *kr @(((λx. {x}) ` X)) ρ) ((λx. {x}) ` X)
    for 𝔈 :: ('a, 'b, 'c) kraus_family and X
    apply (rule kf_apply_on_union_has_sum)
    by auto
  then have sum: ((λX. 𝔈 *kr @X ρ) has_sum 𝔈 *kr @X ρ) ((λx. {x}) ` X)
    for 𝔈 :: ('a, 'b, 'c) kraus_family and X
    by simp
  from assms
  have leq: 𝔈 *kr @{x} ρ  𝔉 *kr @{x} ρ for x
    by (simp add: less_eq_kraus_family_def)
  show ?thesis
    using sum sum apply (rule has_sum_mono_traceclass)
    using leq by fast
qed

lemma kf_apply_mono_left: 𝔈  𝔉  ρ  0  𝔈 *kr ρ  𝔉 *kr ρ
  by (metis kf_apply_on_UNIV kf_apply_on_mono1)

lemma kf_apply_mono:
  assumes ρ  0
  assumes 𝔈  𝔉 and ρ  σ
  shows 𝔈 *kr ρ  𝔉 *kr σ
  by (meson assms(1,2,3) basic_trans_rules(23) kf_apply_mono_left kf_apply_mono_right)

lemma kf_apply_on_mono:
  assumes ρ  0
  assumes 𝔈  𝔉 and X  Y and ρ  σ
  shows 𝔈 *kr @X ρ  𝔉 *kr @Y σ
  apply (rule order.trans)
  using assms(2,1) apply (rule kf_apply_on_mono1)
  apply (rule order.trans)
  using assms(3,1) apply (rule kf_apply_on_mono2)
  using assms(4) by (rule kf_apply_on_mono3)

lemma kf_one_dim_is_id[simp]:
  fixes 𝔈 :: ('a::one_dim, 'a::one_dim, 'x) kraus_family
  shows 𝔈 =kr kf_norm 𝔈 *R kf_id
proof (rule kf_eq_weakI)
  fix t :: ('a, 'a) trace_class
  have 𝔈1pos[iff]: 𝔈 *kr 1  0
    apply (rule kf_apply_pos)
    by (metis one_cinner_one one_dim_iso_of_one one_dim_scaleC_1 tc_butterfly_pos trace_tc_butterfly
          trace_tc_one_dim_iso)

  have 𝔈t: 𝔈 *kr t = trace_tc t *C (𝔈 *kr 1) if NO_MATCH 1 tfor t
    by (metis kf_apply_scaleC one_dim_scaleC_1 trace_tc_one_dim_iso)
  have kf_bound 𝔈 = norm (𝔈 *kr 1) *R id_cblinfun
  proof (rule cblinfun_cinner_eqI)
    fix h :: 'a
    assume norm h = 1
    have h C kf_bound 𝔈 h = one_dim_iso h * cnj (one_dim_iso h) * one_dim_iso (𝔈 *kr 1)
      apply (subst kf_bound_from_map)
      by (simp add: 𝔈t cinner_scaleR_right cblinfun.scaleR_left cdot_square_norm one_dim_tc_butterfly)
    also have 1:  = one_dim_iso (𝔈 *kr 1)
      by (smt (verit, best) norm h = 1 cinner_simps(5) cnorm_eq_1 id_apply more_arith_simps(6) mult.commute
          one_dim_iso_def one_dim_iso_id one_dim_iso_is_of_complex one_dim_scaleC_1)
    also have  = trace_tc (𝔈 *kr 1)
      by simp
    also have  = norm (𝔈 *kr 1)
      apply (subst norm_tc_pos)
      by (simp_all add: 𝔈1pos)
    also have  = h C (norm (𝔈 *kr 1) *R id_cblinfun *V h)
      by (metis norm h = 1 cblinfun.scaleR_left cinner_commute cinner_scaleR_left cnorm_eq_1
          complex_cnj_complex_of_real id_cblinfun.rep_eq mult.commute mult_cancel_right1)
    finally show h C (kf_bound 𝔈 *V h) = h C (norm (𝔈 *kr 1) *R id_cblinfun *V h)
      by -
  qed
  then have kf_norm 𝔈 = cmod (one_dim_iso (𝔈 *kr 1))
    by (simp add: kf_norm_def)
  then have norm: complex_of_real (kf_norm 𝔈) = one_dim_iso (𝔈 *kr 1)
  using norm_tc_pos by fastforce

  have (one_dim_iso (𝔈 *kr t) :: complex) = one_dim_iso t * one_dim_iso (𝔈 *kr 1)
  by (metis (mono_tags, lifting) kf_apply_scaleC of_complex_one_dim_iso one_dim_iso_is_of_complex
      one_dim_iso_scaleC one_dim_scaleC_1 scaleC_one_dim_is_times)
  also have  = one_dim_iso t * complex_of_real (kf_norm 𝔈)
    by (simp add: norm)
  also have  = one_dim_iso (kf_norm 𝔈 *R t)
    by (simp add: scaleR_scaleC)
  also have  = one_dim_iso (kf_norm 𝔈 *R kf_id *kr t)
    by (simp add: kf_scale_apply)
  finally show 𝔈 *kr t = kf_norm 𝔈 *R kf_id *kr t
    using one_dim_iso_inj by blast
qed


subsection ‹Mapping and flattening›

definition kf_similar_elements :: ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a CL 'b)  ('a CL 'b × 'x) set where
  ― ‹All elements of the Kraus family that are equal up to rescaling (and belong to the same output)›
  kf_similar_elements 𝔈 E = {(F,x)  Rep_kraus_family 𝔈. (r>0. E = r *R F)}

definition kf_element_weight where
  ― ‹The total weight (norm of the square) of all similar elements›
  kf_element_weight 𝔈 E = ((F,_)kf_similar_elements 𝔈 E. norm (F* oCL F))

lemma kf_element_weight_geq0[simp]: kf_element_weight 𝔈 E  0
  by (auto intro!: infsum_nonneg simp: kf_element_weight_def)

lemma kf_similar_elements_abs_summable:
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  shows (λ(F,_). F* oCL F) abs_summable_on (kf_similar_elements 𝔈 E)
proof (cases E = 0)
  case True
  show ?thesis
    apply (rule summable_on_cong[where g=λ_. 0, THEN iffD2])
    by (auto simp: kf_similar_elements_def True)
next
  case False
  then obtain ψ where : E ψ  0
    by (metis cblinfun.zero_left cblinfun_eqI)
  define φ where φ = ((norm E)2 / (ψ C (E* *V E *V ψ))) *C ψ
  have normFF: norm (fst Fx* oCL fst Fx) = ψ C (fst Fx* *V fst Fx *V φ)
    if Fx  kf_similar_elements 𝔈 E for Fx
  proof -
    define F where F = fst Fx
    have ra. x = (ra * r) *R x if r > 0 for r and x :: 'a CL 'b
      apply (rule exI[of _ inverse r])
      using that 
      by auto
    with that obtain r where FE: F = r *R E
      apply atomize_elim
      by (auto simp: kf_similar_elements_def F_def)
    show norm (F* oCL F) = ψ C (F* *V F *V φ)
      by (simp add: False φ_def FE cblinfun.scaleR_left cblinfun.scaleR_right
          cblinfun.scaleC_left cblinfun.scaleC_right cinner_adj_right )
  qed

  have ψφ_mono: mono (λA. ψ C (A *V φ))
  proof (rule monoI)
    fix A B :: 'a CL 'a
    assume A  B
    then have B - A  0
      by auto
    then have ψ C ((B - A) *V ψ)  0
      by (simp add: cinner_pos_if_pos)
    then have ψ C ((B - A) *V φ)  0
      by (auto intro!: mult_nonneg_nonneg
          simp: φ_def cblinfun.scaleC_right divide_inverse cinner_adj_right power2_eq_square)
    then show ψ C (A *V φ)  ψ C (B *V φ)
      by (simp add: cblinfun.diff_left cinner_diff_right)
  qed

  have ψφ_linear: clinear (λA. ψ C (A *V φ))
    by (auto intro!: clinearI simp: cblinfun.add_left cinner_add_right)

  from Rep_kraus_family[of 𝔈]
  have bdd_above ((λM. (E, x)M. E* oCL E) ` {M. finite M  M  Rep_kraus_family 𝔈})
    by (simp add: kraus_family_def)
  then have bdd_above ((λM. (F, x)M. F* oCL F) ` {M. M  kf_similar_elements 𝔈 E  finite M})
    apply (rule bdd_above_mono2[OF _ _ order.refl])
    by (auto simp: kf_similar_elements_def)
  then have bdd_above ((λA. ψ C (A *V φ)) ` (λM. (F, x)M. F* oCL F) ` {M. M  kf_similar_elements 𝔈 E  finite M})
    by (rule bdd_above_image_mono[OF ψφ_mono])
  then have bdd_above ((λM. ψ C (((F, x)M. F* oCL F) *V φ)) ` {M. M  kf_similar_elements 𝔈 E  finite M})
    by (simp add: image_image)
  then have bdd_above ((λM. (F, x)M. ψ C ((F* oCL F) *V φ)) ` {M. M  kf_similar_elements 𝔈 E  finite M})
    unfolding case_prod_beta
    by (subst complex_vector.linear_sum[OF ψφ_linear, symmetric])
  then have bdd_above ((λM. (F, x)M. complex_of_real (norm (F* oCL F))) ` {M. M  kf_similar_elements 𝔈 E  finite M})
    apply (rule bdd_above_mono2[OF _ subset_refl])
    unfolding case_prod_unfold
    apply (subst sum.cong[OF refl normFF])
    by auto
  then have bdd_above ((λM. (F, x)M. norm (F* oCL F)) ` {M. M  kf_similar_elements 𝔈 E  finite M})
    by (auto simp add: bdd_above_def case_prod_unfold less_eq_complex_def)
  then have (λ(F,_). norm (F* oCL F)) summable_on (kf_similar_elements 𝔈 E)
    apply (rule nonneg_bdd_above_summable_on[rotated])
    by auto
  then show (λ(F,_). F* oCL F) abs_summable_on kf_similar_elements 𝔈 E
    by (simp add: case_prod_unfold)
qed

lemma kf_similar_elements_kf_operators:
  assumes (F,x)  kf_similar_elements 𝔈 E
  shows F  span (kf_operators 𝔈)
  using assms
  unfolding kf_similar_elements_def
  apply (transfer' fixing: E F x)
  by (metis (no_types, lifting) Product_Type.Collect_case_prodD fst_conv image_eqI span_base)

lemma kf_element_weight_neq0: kf_element_weight 𝔈 E  0 
  if (E,x)  Rep_kraus_family 𝔈 and E  0
proof -
  have 1: (E, x)  kf_similar_elements 𝔈 E
    by (auto intro!: exI[where x=1] simp: kf_similar_elements_def that)

  have kf_element_weight 𝔈 E = ((F, x)kf_similar_elements 𝔈 E. (norm F)2)
    by (simp add: kf_element_weight_def)
  moreover have   ((F, x){(E,x)}. (norm F)2) (is _  )
    apply (rule infsum_mono_neutral)
    using kf_similar_elements_abs_summable
    by (auto intro!: 1 simp: that case_prod_unfold)
  moreover have  > 0
    using that by simp
  ultimately show ?thesis
    by auto
qed


lemma kf_element_weight_0_left[simp]: kf_element_weight 0 E = 0
  by (simp add: kf_element_weight_def kf_similar_elements_def zero_kraus_family.rep_eq)

lemma kf_element_weight_0_right[simp]: kf_element_weight E 0 = 0
  by (auto intro!: infsum_0 simp add: kf_element_weight_def kf_similar_elements_def)

lemma kf_element_weight_scale:
  assumes r > 0
  shows kf_element_weight 𝔈 (r *R E) = kf_element_weight 𝔈 E
proof -
  have [simp]: (r'>0. r *R E = r' *R F)  (r'>0. E = r' *R F) for F
    apply (rule Ex_iffI[where f=λr'. r' /R r and g=λr'. r *R r'])
     apply (smt (verit, best) assms divideR_right real_scaleR_def scaleR_scaleR scaleR_simps(7)
        zero_le_scaleR_iff) 
    using assms by force
  show ?thesis
    using assms
    by (simp add: kf_similar_elements_def kf_element_weight_def)
qed

lemma kf_element_weight_kf_operators:
  assumes kf_element_weight 𝔈 E  0
  shows E  span (kf_operators 𝔈)
proof -
  from assms
  have ((F, _){(F, x). (F, x)  Rep_kraus_family 𝔈  (r>0. E = r *R F)}. norm (F* oCL F))  0
    by (simp add: kf_element_weight_def kf_similar_elements_def)
  then obtain F x r where (F, x)  Rep_kraus_family 𝔈 and E = r *R F
    by (smt (verit, ccfv_SIG) Product_Type.Collect_case_prodD infsum_0)
  then have F  kf_operators 𝔈
    by (metis fst_conv image_eqI kf_operators.rep_eq)
  with E = r *R F show ?thesis
    by (simp add: span_clauses)
qed


lemma kf_map_aux:
  fixes f :: 'x  'y and 𝔈 :: ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family
  defines B  kf_bound 𝔈
  defines filtered y  kf_filter (λx. f x = y) 𝔈
  defines flattened  {(E, y). norm (E* oCL E) = kf_element_weight (filtered y) E  E  0}
  defines good  (λ(E,y). (norm E)2 = kf_element_weight (filtered y) E  E  0)
  shows has_sum_in cweak_operator_topology (λ(E,_). E* oCL E) flattened B (is ?has_sum)
    and snd ` (SIGMA (E, y):Collect good. kf_similar_elements (filtered y) E)
      = {(F, x). (F, x)  Rep_kraus_family 𝔈  F  0} (is ?snd_sigma)
    and inj_on snd (SIGMA p:Collect good. kf_similar_elements (filtered (snd p)) (fst p)) (is ?inj_snd)
proof -
  have E_inv: kf_element_weight (filtered y) E  0 if good (E,y) for E y
    using that by (auto simp:  good_def)

  show snd_sigma: ?snd_sigma
  proof (intro Set.set_eqI iffI)
    fix Fx
    assume asm: Fx  snd ` (SIGMA (E, y):Collect good. kf_similar_elements (filtered y) E)
    obtain F x where Fx_def: Fx = (F,x) by fastforce
    with asm obtain E y where Fx_rel_E: (F, x)  kf_similar_elements (filtered y) E and good (E,y)
      by auto
    then have (F, x)  Rep_kraus_family 𝔈
      by (simp add: kf_similar_elements_def filtered_def kf_filter.rep_eq)
    from Fx_rel_E obtain r where E = r *R F
      by (smt (verit) kf_similar_elements_def mem_Collect_eq prod.sel(1) split_def)
    with good (E,y) have F  0
      by (simp add: good_def)
    with (F, x)  Rep_kraus_family 𝔈 show Fx  {(F, x). (F, x)  Rep_kraus_family 𝔈  F  0}
      by (simp add: Fx_def)
  next
    fix Fx
    assume asm: Fx  {(F, x). (F, x)  Rep_kraus_family 𝔈  F  0}
    obtain F x where Fx_def: Fx = (F,x) by fastforce
    from asm have Fx_𝔈: (F, x)  Rep_kraus_family 𝔈 and [simp]: F  0
      by (auto simp: Fx_def)
    have weight_fx_F_not0: kf_element_weight (filtered (f x)) F  0
      using Fx_𝔈 by (simp_all add: filtered_def kf_filter.rep_eq kf_element_weight_neq0)
    then have weight_fx_F_pos: kf_element_weight (filtered (f x)) F > 0
      using kf_element_weight_geq0 
      by (metis less_eq_real_def)
    define E where E = (sqrt (kf_element_weight (filtered (f x)) F) / norm F) *R F
    have [simp]: E  0
      by (auto intro!: weight_fx_F_not0 simp: E_def)
    have E_F_same: kf_element_weight (filtered (f x)) E = kf_element_weight (filtered (f x)) F
      by (simp add: E_def kf_element_weight_scale weight_fx_F_pos)
    have good (E, f x)
      apply (simp add: good_def E_F_same)
      by (simp add: E_def)
    have 1: sqrt (kf_element_weight (filtered (f x)) F) / norm F > 0
      by (auto intro!: divide_pos_pos weight_fx_F_pos)
    then have (F, x)  kf_similar_elements (filtered (f x)) E
      by (auto intro!: (F, x)  Rep_kraus_family 𝔈 simp: kf_similar_elements_def E_def F  0
         filtered_def kf_filter.rep_eq)
    with good (E,f x)
    show Fx  snd ` (SIGMA (E, y):Collect good. kf_similar_elements (filtered y) E)
      by (force intro!: image_eqI[where x=((E,()),F,x)] simp: Fx_def filtered_def)
  qed

  show inj_snd: ?inj_snd
  proof (rule inj_onI)
    fix EFx EFx' :: ('a CL 'b × 'y) × 'a CL 'b × 'x
    assume EFx_in: EFx  (SIGMA p:Collect good. kf_similar_elements (filtered (snd p)) (fst p))
      and EFx'_in: EFx'  (SIGMA p:Collect good. kf_similar_elements (filtered (snd p)) (fst p))
    assume snd_eq: snd EFx = snd EFx'
    obtain E F x y where [simp]: EFx = ((E,y),F,x)
      by (metis (full_types) old.unit.exhaust surj_pair)
    obtain E' F' x' y' where [simp]: EFx' = ((E',y'),(F',x')) 
      by (metis (full_types) old.unit.exhaust surj_pair)
    from snd_eq have [simp]: F' = F and [simp]: x' = x
      by auto
    from EFx_in have good (E,y) and F_rel_E: (F, x)  kf_similar_elements (filtered y) E
      by auto
    from EFx'_in have good (E',y') and F_rel_E': (F, x)  kf_similar_elements (filtered y') E'
      by auto
    from good (E,y) have E  0
      by (simp add: good_def)
    from good (E',y') have E'  0
      by (simp add: good_def)
    from F_rel_E obtain r where ErF: E = r *R F and r > 0
      by (auto intro!: simp: kf_similar_elements_def)
    from F_rel_E' obtain r' where E'rF: E' = r' *R F and r' > 0
      by (auto intro!: simp: kf_similar_elements_def)

    from EFx_in have y = f x
      by (auto intro!: simp: filtered_def kf_similar_elements_def kf_filter.rep_eq)
    moreover from EFx'_in have y' = f x'
      by (auto intro!: simp: filtered_def kf_similar_elements_def kf_filter.rep_eq)
    ultimately have [simp]: y = y'
      by simp

    define r'' where r'' = r' / r
    with E'rF ErF E  0
    have E'_E: E' = r'' *R E
      by auto
    with r' > 0 r > 0 E'  0
    have [simp]: r'' > 0
      by (fastforce simp: r''_def)
    from E'_E have kf_element_weight (filtered y') E' = kf_element_weight (filtered y) E
      by (simp add: kf_element_weight_scale)
    with good (E,y) good (E',y') have (norm E')2 = (norm E)2
      by (auto intro!: simp: good_def)
    with E' = r'' *R E
    have E' = E
      using 0 < r'' by force
    then show EFx = EFx'
      by simp
  qed

  show ?has_sum
  proof (unfold has_sum_in_cweak_operator_topology_pointwise, intro allI)
    fix ψ φ :: 'a
    define B' where B' = ψ C B φ
    define normal where normal E y = E /R sqrt (kf_element_weight (filtered y) E) for E y
    have has_sum_in cweak_operator_topology (λ(F,x). F* oCL F) (Rep_kraus_family 𝔈) B
      using B_def kf_bound_has_sum by blast
    then have ((λ(F,x). ψ C (F* oCL F) φ) has_sum B') (Rep_kraus_family 𝔈)
      by (simp add: B'_def has_sum_in_cweak_operator_topology_pointwise case_prod_unfold)
    then have ((λ(F,x). ψ C (F* oCL F) φ) has_sum B') {(F,x)Rep_kraus_family 𝔈. F  0}
      apply (rule has_sum_cong_neutral[THEN iffD2, rotated -1])
      by (auto simp: zero_cblinfun_wot_def)
    then have ((λ(F,x). ψ C (F* oCL F) φ) has_sum B')
           (snd ` (SIGMA (E,y):Collect good. kf_similar_elements (filtered y) E))
      by (simp add: snd_sigma)
    then have ((λ((E,x), (F,y)). ψ C (F* oCL F) φ) has_sum B')
       (SIGMA (E,y):Collect good. kf_similar_elements (filtered y) E)
      apply (subst (asm) has_sum_reindex)
      by (auto intro!: inj_on_def inj_snd simp: o_def case_prod_unfold)
    then have sum1: ((λ((E,y), (F,x)). (norm F)2 *R (ψ C (normal E y* oCL normal E y) φ)) has_sum B')
        (SIGMA (E,y):Collect good. kf_similar_elements (filtered y) E)
      apply (rule has_sum_cong[THEN iffD2, rotated])
      apply (subgoal_tac b aa ba r.
       (r * norm aa)2 = kf_element_weight (filtered b) (complex_of_real r *C aa) 
       aa  0 
       0 < r 
       (complex_of_real (norm aa))2 *
       (inverse (complex_of_real (sqrt (kf_element_weight (filtered b) (complex_of_real r *C aa)))) *
        (inverse (complex_of_real (sqrt (kf_element_weight (filtered b) (complex_of_real r *C aa)))) *
         (complex_of_real r * complex_of_real r))) 
       1 
       is_orthogonal ψ (aa* *V aa *V φ))
      apply (auto intro!: simp: good_def normal_def sandwich_tc_scaleR_left power_inverse real_sqrt_pow2 E_inv
          kf_similar_elements_def kf_element_weight_scale 
          cblinfun.scaleC_left cblinfun.scaleC_right cinner_scaleC_right scaleR_scaleC)[1]
      by (smt (verit) Extra_Ordered_Fields.mult_sign_intros(5) Extra_Ordered_Fields.sign_simps(5) inverse_eq_iff_eq left_inverse more_arith_simps(11) of_real_eq_0_iff of_real_mult power2_eq_square power_inverse real_inv_sqrt_pow2 zero_less_norm_iff)
    then have ((λ(E,y). (F, x)kf_similar_elements (filtered y) E.
                (norm F)2 *R (ψ C (normal E y* oCL normal E y) φ)) has_sum B') (Collect good)
      by (auto intro!: has_sum_Sigma'_banach simp add: case_prod_unfold)
    then have ((λ(E,y). ((F, x)kf_similar_elements (filtered y) E.
                              (norm F)2) *R (ψ C (normal E y* oCL normal E y) φ))
                        has_sum B') (Collect good)
      apply (rule has_sum_cong[THEN iffD1, rotated])
      apply simp
      by (smt (verit) Complex_Vector_Spaces.infsum_scaleR_left cblinfun.scaleR_left infsum_cong split_def)
    then have ((λ(E,y). kf_element_weight (filtered y) E *R 
                              (ψ C (normal E y* oCL normal E y) φ)) has_sum B') (Collect good)
      by (simp add: kf_element_weight_def)
    then have ((λ(E,_). ψ C (E* oCL E) φ) has_sum B') (Collect good)
      apply (rule has_sum_cong[THEN iffD1, rotated])
      by (auto  intro!: field_class.field_inverse
          simp add: normal_def sandwich_tc_scaleR_left power_inverse real_sqrt_pow2 E_inv
          cblinfun.scaleR_left scaleR_scaleC
          simp flip: inverse_mult_distrib semigroup_mult.mult_assoc of_real_mult
          split!: prod.split)
    then have ((λ(E,_). ψ C (E* oCL E) φ) has_sum B') flattened
      by (simp add: flattened_def good_def)
    then show ((λx. ψ C ((case x of (E, uu_)  E* oCL E) *V φ)) has_sum B') flattened
      by (simp add: case_prod_unfold)
  qed
qed


lift_definition kf_map :: ('x  'y)  ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a, 'b, 'y) kraus_family is
  λf 𝔈. {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E  E  0}
proof (rename_tac f 𝔈)
  fix f :: 'x  'y and 𝔈 :: ('a, 'b, 'x) kraus_family
  define filtered flattened B 
    where filtered y = kf_filter (λx. f x = y) 𝔈
      and flattened = {(E, y). norm (E* oCL E) = kf_element_weight (filtered y) E  E  0}
      and B = kf_bound 𝔈 
    for y

  from kf_map_aux[of f 𝔈]
  have bound_has_sum: has_sum_in cweak_operator_topology (λ(E,_). E* oCL E) flattened B
    by (simp_all add: filtered_def flattened_def B_def)

  have nonzero: 0  fst ` flattened
    by (auto intro!: simp: flattened_def)

  from bound_has_sum nonzero show flattened  Collect kraus_family
    by (auto simp: kraus_family_iff_summable summable_on_in_def)
qed

lemma
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  shows kf_apply_map[simp]: kf_apply (kf_map f 𝔈) = kf_apply 𝔈
    and kf_map_bound: kf_bound (kf_map f 𝔈) = kf_bound 𝔈
    and kf_map_norm[simp]: kf_norm (kf_map f 𝔈) = kf_norm 𝔈
proof (rule ext)
  fix ρ :: ('a, 'a) trace_class
  define filtered good flattened B normal σ
    where filtered y = kf_filter (λx. f x = y) 𝔈
      and good = (λ(E,y). (norm E)2 = kf_element_weight (filtered y) E  E  0)
      and flattened = {(E, y). norm (E* oCL E) = kf_element_weight (filtered y) E  E  0}
      and B = kf_bound 𝔈 
      and normal E y = E /R sqrt (kf_element_weight (filtered y) E)
      and σ = 𝔈 *kr ρ
    for E y
  have E_inv: kf_element_weight (filtered y) E  0 if good (E,y) for E y
    using that by (auto simp:  good_def)

  from kf_map_aux[of f 𝔈]
  have snd_sigma: snd ` (SIGMA (E, y):Collect good. kf_similar_elements (filtered y) E)
      = {(F, x). (F, x)  Rep_kraus_family 𝔈  F  0}
    and inj_snd: inj_on snd (SIGMA p:Collect good. kf_similar_elements (filtered (snd p)) (fst p))
    and bound_has_sum: has_sum_in cweak_operator_topology (λ(E,_). E* oCL E) flattened B
    by (simp_all add: good_def filtered_def flattened_def B_def)

  show kf_apply (kf_map f 𝔈) ρ = σ
  proof -
    have (λ(F,x). sandwich_tc F ρ) summable_on Rep_kraus_family 𝔈
      using kf_apply_summable by (simp add: case_prod_unfold)
    then have ((λ(F,x). sandwich_tc F ρ) has_sum σ) (Rep_kraus_family 𝔈)
      by (simp add: σ_def kf_apply_def split_def)
    then have ((λ(F,x). sandwich_tc F ρ) has_sum σ) {(F,x)Rep_kraus_family 𝔈. F  0}
      apply (rule has_sum_cong_neutral[THEN iffD2, rotated -1])
      by auto
    then have ((λ(F,x). sandwich_tc F ρ) has_sum σ)
           (snd ` (SIGMA (E,y):Collect good. kf_similar_elements (filtered y) E))
      by (simp add: snd_sigma)
    then have ((λ((E,x), (F,y)). sandwich_tc F ρ) has_sum σ)
       (SIGMA (E,y):Collect good. kf_similar_elements (filtered y) E)
      apply (subst (asm) has_sum_reindex)
      by (auto intro!: inj_on_def inj_snd simp: o_def case_prod_unfold)
    then have sum1: ((λ((E,y), (F,_)). (norm F)2 *R sandwich_tc (normal E y) ρ) has_sum σ)
                (SIGMA (E,y):Collect good. kf_similar_elements (filtered y) E)
      apply (rule has_sum_cong[THEN iffD2, rotated])
      apply (subgoal_tac b a r. (r * norm a)2 = kf_element_weight (filtered b) a 
       a  0  0 < r  ((norm a)2 * inverse (kf_element_weight (filtered b) a) * r2) *R sandwich_tc a ρ = sandwich_tc a ρ)
       apply (auto intro!: real_vector.scale_one simp: good_def normal_def sandwich_tc_scaleR_left power_inverse real_sqrt_pow2 E_inv
          kf_similar_elements_def kf_element_weight_scale)[1]
      by (metis (no_types, opaque_lifting) Extra_Ordered_Fields.sign_simps(5) linorder_not_less more_arith_simps(11) mult_eq_0_iff norm_le_zero_iff order.refl power2_eq_square right_inverse scale_one)
    then have ((λ(E,y). (F, x)kf_similar_elements (filtered y) E.
                (norm F)2 *R sandwich_tc (normal E y) ρ) has_sum σ) (Collect good)
      by (auto intro!: has_sum_Sigma'_banach simp add: case_prod_unfold)
    then have ((λ(E,y). ((F, x)kf_similar_elements (filtered y) E. (norm F)2) *R sandwich_tc (normal E y) ρ)
                        has_sum σ) (Collect good)
      apply (rule has_sum_cong[THEN iffD1, rotated])
      apply simp
      by (smt (verit) Complex_Vector_Spaces.infsum_scaleR_left cblinfun.scaleR_left infsum_cong split_def)
    then have ((λ(E,y). kf_element_weight (filtered y) E *R sandwich_tc (normal E y) ρ) has_sum σ) (Collect good)
      by (simp add: kf_element_weight_def)
    then have ((λ(E,_). sandwich_tc E ρ) has_sum σ) (Collect good)
      apply (rule has_sum_cong[THEN iffD1, rotated])
      by (auto intro!: simp: normal_def sandwich_tc_scaleR_left power_inverse real_sqrt_pow2 E_inv)
    then have ((λ(E,_). sandwich_tc E ρ) has_sum σ) flattened
      by (simp add: flattened_def good_def)
    then show kf_map f 𝔈 *kr ρ = σ
      by (simp add: kf_apply.rep_eq kf_map.rep_eq flattened_def
          case_prod_unfold infsumI filtered_def)
  qed

  from bound_has_sum show bound: kf_bound (kf_map f 𝔈) = B
    apply (simp add: kf_bound_def flattened_def kf_map.rep_eq B_def filtered_def)
    using has_sum_in_infsum_in has_sum_in_unique hausdorff_cweak_operator_topology summable_on_in_def
    by blast

  from bound show kf_norm (kf_map f 𝔈) = kf_norm 𝔈
    by (simp add: kf_norm_def B_def)
qed

abbreviation kf_flatten  kf_map (λ_. ())

text ‹Like constkf_map, but with a much simpler definition.
      However, only makes sense for injective functions.›
lift_definition kf_map_inj :: ('x  'y)  ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a, 'b, 'y) kraus_family is
  λf 𝔈. (λ(E,x). (E, f x)) ` 𝔈
proof (rule CollectI, rule kraus_familyI, rename_tac f 𝔈)
  fix f :: 'x  'y and 𝔈 :: ('a CL 'b × 'x) set
  assume 𝔈  Collect kraus_family
  then obtain B where B: ((E, x)F. E* oCL E)  B if F  {F. finite F  F  𝔈} for F
    by (auto simp: kraus_family_def bdd_above_def)
  show bdd_above ((λF. (E, x)F. E* oCL E) ` {F. finite F  F  (λ(E, x). (E, f x)) ` 𝔈})
  proof (rule bdd_aboveI2)
    fix F assume F  {F. finite F  F  (λ(E, x). (E, f x)) ` 𝔈}
    then obtain F' where finite F' and F'  𝔈 and F_F': F = (λ(E, x). (E, f x)) ` F'
      and inj: inj_on (λ(E, x). (E, f x)) F'
      by (metis (no_types, lifting) finite_imageD mem_Collect_eq subset_image_inj)
    have ((E, x)F'. E* oCL E)  B
      by (auto intro!: B finite F' F'  𝔈)
    moreover have ((E, x)F. E* oCL E)  ((E, x)F'. E* oCL E)
      apply (simp add: F_F' inj sum.reindex)
      by (simp add: case_prod_beta)
    ultimately show ((E, x)F. E* oCL E)  B
      by simp
  qed
  from 𝔈  Collect kraus_family
  show 0  fst ` (λ(E,x). (E, f x)) ` 𝔈
    by (force simp: kraus_family_def)
qed                            

lemma kf_element_weight_map_inj:
  assumes inj_on f (kf_domain 𝔈)
  shows kf_element_weight (kf_map_inj f 𝔈) E = kf_element_weight 𝔈 E
proof -
  wlog E  0
    using negation by simp
  have inj2: inj_on (λ(E, x). (E, f x)) {(F, x). (F, x)  Rep_kraus_family 𝔈  (r>0. E = r *R F)}
  proof (rule inj_onI)
    fix Fy Gx :: 'c CL 'd × 'a
    obtain F y G x where [simp]: Fy = (F,y) Gx = (G,x)
      by (auto simp: prod_eq_iff)
    assume Fy  {(F, y). (F, y)  Rep_kraus_family 𝔈  (r>0. E = r *R F)} and Gx  {(G, x). (G, x)  Rep_kraus_family 𝔈  (r>0. E = r *R G)}
    then have Fy_𝔈: (F, y)  Rep_kraus_family 𝔈 and ErF: r>0. E = r *R F and Gx_𝔈: (G, x)  Rep_kraus_family 𝔈 and ErG: r>0. E = r *R G
      by auto
    from ErF E  0 have F  0
      by auto
    with Fy_𝔈 have y  kf_domain 𝔈
      by (force simp add: kf_domain.rep_eq)
    from ErG E  0 have G  0
      by auto
    with Gx_𝔈 have x  kf_domain 𝔈
      by (force simp add: kf_domain.rep_eq)
    assume (case Fy of (F, y)  (F, f y)) = (case Gx of (G, x)  (G, f x))
    then have [simp]: F = G and f y = f x
      by auto
    with assms x  kf_domain 𝔈 y  kf_domain 𝔈
    have y = x
      by (simp add: inj_onD)
    then show Fy = Gx
      by simp
  qed

  have kf_element_weight (kf_map_inj f 𝔈) E
     = ((F, _){(F, x). (F, x)  (λ(E,x). (E, f x)) ` Rep_kraus_family 𝔈  (r>0. E = r *R F)}. (norm F)2)
    by (simp add: kf_element_weight_def assms kf_similar_elements_def kf_map_inj.rep_eq)
  also have  = ((F, _)(λ(E,x). (E, f x)) ` {(F, x). (F, x)  Rep_kraus_family 𝔈  (r>0. E = r *R F)}. (norm F)2)
    apply (rule arg_cong2[where f=infsum])
    by auto
  also have  = ((F, _){(F, x). (F, x)  Rep_kraus_family 𝔈  (r>0. E = r *R F)}. (norm F)2)
    apply (subst infsum_reindex)
     apply (rule inj2)
    using assms by (simp add: o_def case_prod_unfold)
  also have  = kf_element_weight 𝔈 E
    by (simp add: kf_element_weight_def assms kf_similar_elements_def)
  finally show ?thesis
    by -
qed

lemma kf_eq_weak_kf_map_left: kf_map f F =kr G if F =kr G
  using that by (simp add: kf_eq_weak_def kf_apply_map)

lemma kf_eq_weak_kf_map_right: F =kr kf_map f G if F =kr G
  using that by (simp add: kf_eq_weak_def kf_apply_map)

lemma kf_filter_map:
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  shows kf_filter P (kf_map f 𝔈) = kf_map f (kf_filter (λx. P (f x)) 𝔈)
proof -
  have (E,x)  Set.filter (λ(E, y). P y) {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E  E  0}
    (E,x)  {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. f x = y) (kf_filter (λx. P (f x)) 𝔈)) E  E  0}
    for E x and 𝔈 :: ('a, 'b, 'x) kraus_family
  proof (cases P x)
    case True
    then show ?thesis
      apply (simp add: kf_filter_twice)
      by (metis (mono_tags, lifting) kf_filter_cong_eq)
  next
    case False
    then have [simp]: (λz. f z = x  P (f z)) = (λ_. False)
      by auto
    from False show ?thesis
      apply (simp add: kf_filter_twice)
      by (smt (verit, ccfv_SIG) kf_element_weight_0_left kf_filter_cong_eq kf_filter_false norm_eq_zero zero_eq_power2)
  qed
  then show ?thesis
    apply (transfer' fixing: P f)
    by blast
qed

lemma kf_filter_map_inj:
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  shows kf_filter P (kf_map_inj f 𝔈) = kf_map_inj f (kf_filter (λx. P (f x)) 𝔈)
  apply (transfer' fixing: P f)
  by (force simp: case_prod_beta image_iff)

lemma kf_map_kf_map_inj_comp:
  assumes inj_on f (kf_domain 𝔈)
  shows kf_map g (kf_map_inj f 𝔈) = kf_map (g o f) 𝔈
proof (transfer' fixing: f g 𝔈)
  from assms
  have inj_on f (kf_domain (kf_filter (λx. g (f x) = y) 𝔈)) for y
    apply (rule inj_on_subset) by force
  then show {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. g x = y) (kf_map_inj f 𝔈)) E  E  0} =
             {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. (g  f) x = y) 𝔈) E  E  0}
    by (simp add: kf_filter_map_inj kf_element_weight_map_inj assms)
qed

(* lemma kf_similar_elements_remove_0[simp]: ‹kf_similar_elements (kf_remove_0 𝔈) E = kf_similar_elements 𝔈 E› if ‹E ≠ 0›
  using that by (auto simp: kf_similar_elements_def kf_remove_0.rep_eq)

lemma kf_element_weight_remove_0[simp]: ‹kf_element_weight (kf_remove_0 𝔈) E = kf_element_weight 𝔈 E›
  apply (cases ‹E = 0›)
   apply (simp add: kf_element_weight_0_right)
  by (simp add: kf_element_weight_def) *)

lemma kf_element_weight_eqweak0:
  assumes 𝔈 =kr 0
  shows kf_element_weight 𝔈 E = 0
  apply (subst kf_eq_0_iff_eq_0[THEN iffD1])
  using assms by auto

lemma kf_map_inj_kf_map_comp:
  assumes inj_on g (f ` kf_domain 𝔈)
  shows kf_map_inj g (kf_map f 𝔈) = kf_map (g o f) 𝔈
proof (transfer' fixing: f g 𝔈, rule Set.set_eqI)
  fix Ex :: 'd CL 'e × 'b
  obtain E x where [simp]: Ex = (E,x)
    by (auto simp: prod_eq_iff)
  have Ex  (λ(E, x). (E, g x)) ` {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E  E  0} 
    (y. x = g y  (norm E)2 = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E)  E  0
    by auto
  also have   (norm E)2 = kf_element_weight (kf_filter (λz. g (f z) = x) 𝔈) E  E  0
  proof (rule iffI)
    assume asm: (y. x = g y  (norm E)2 = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E)  E  0
    then obtain y where xy: x = g y and weight: (norm E)2 = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E
      by auto
    from asm have E  0
      by simp
    with weight have ¬ kf_filter (λx. f x = y) 𝔈 =kr 0
      using kf_element_weight_eqweak0 by fastforce
    then have ¬ kf_filter (λx. f x = y  x  kf_domain 𝔈) 𝔈 =kr 0
      by (smt (verit, del_insts) kf_eq_0_iff_eq_0 kf_filter_cong_eq)
    then have (λx. f x = y  x  kf_domain 𝔈)  (λz. False)
      using kf_filter_false[of 𝔈]
      by (metis kf_eq_weak_refl0)
    then have yf𝔈: y  f ` kf_domain 𝔈
      by fast
    have kf_element_weight ((kf_filter (λx. f x = y) 𝔈)) E = kf_element_weight ((kf_filter (λz. g (f z) = x) 𝔈)) E
      apply (rule arg_cong2[where f=kf_element_weight, OF _ refl])
      apply (rule kf_filter_cong_eq[OF refl])
      using yf𝔈 assms xy
      by (meson image_eqI inj_onD)
    then
    have kf_element_weight (kf_filter (λx. f x = y) 𝔈) E = kf_element_weight (kf_filter (λz. g (f z) = x) 𝔈) E
      by simp
    with weight E  0
    show (norm E)2 = kf_element_weight (kf_filter (λz. g (f z) = x) 𝔈) E  E  0
      by simp
  next
    assume asm: (norm E)2 = kf_element_weight (kf_filter (λz. g (f z) = x) 𝔈) E  E  0
    then have E  0
      by simp
    from asm have ¬ kf_filter (λz. g (f z) = x) 𝔈 =kr 0
      using kf_element_weight_eqweak0 by fastforce
    then have ¬ kf_filter (λz. g (f z) = x  z  kf_domain 𝔈) 𝔈 =kr 0
      by (smt (verit, ccfv_threshold) kf_eq_0_iff_eq_0 kf_filter_cong_eq)
    then have (λz. g (f z) = x  z  kf_domain 𝔈)  (λz. False)
      using kf_filter_false[of 𝔈]
      by (metis kf_eq_weak_refl0)
    then obtain z where z  kf_domain 𝔈 and gfz: g (f z) = x
      using kf_filter_false[of 𝔈]
      by auto
    have kf_element_weight ((kf_filter (λx. f x = f z) 𝔈)) E = kf_element_weight ((kf_filter (λz. g (f z) = x) 𝔈)) E
      apply (rule arg_cong2[where f=kf_element_weight, OF _ refl])
      apply (rule kf_filter_cong_eq[OF refl])
      using assms gfz z  kf_domain 𝔈
      by (metis image_eqI inj_onD)
    with asm E  0
    show (y. x = g y  (norm E)2 = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E)  E  0
      by (auto intro!: exI[of _ f z] simp flip: gfz)
  qed
  also have   Ex  {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. (g  f) x = y) 𝔈) E  E  0}
    by simp
  finally show Ex  (λ(E, x). (E, g x)) ` {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. f x = y) 𝔈) E  E  0} 
         Ex  {(E, y). norm (E* oCL E) = kf_element_weight (kf_filter (λx. (g  f) x = y) 𝔈) E  E  0}
    by -
qed

lemma kf_apply_map_inj[simp]:
  assumes inj_on f (kf_domain 𝔈)
  shows kf_map_inj f 𝔈 *kr ρ = 𝔈 *kr ρ
proof -
  define EE where EE = Set.filter (λ(E,x). E0) (Rep_kraus_family 𝔈)
  have kf_map_inj f 𝔈 *kr ρ = (E(λ(E, x). (E, f x)) ` Rep_kraus_family 𝔈. sandwich_tc (fst E) ρ)
    by (simp add: kf_apply.rep_eq kf_map_inj.rep_eq)
  also have  = (E(λ(E, x). (E, f x)) ` EE. sandwich_tc (fst E) ρ)
    apply (rule infsum_cong_neutral)
    by (force simp: EE_def)+
  also have  = infsum ((λE. sandwich_tc (fst E) ρ)  (λ(E, x). (E, f x))) EE
    apply (rule infsum_reindex)
    apply (subgoal_tac aa ba b. xRep_kraus_family 𝔈. yRep_kraus_family 𝔈. f (snd x) = f (snd y)  snd x = snd y 
       (aa, ba)  Rep_kraus_family 𝔈  f b = f ba  (aa, b)  Rep_kraus_family 𝔈  aa  0  b = ba)
    using assms
    by (auto intro!: simp: inj_on_def kf_domain.rep_eq EE_def)
  also have  = ((E,x)EE. sandwich_tc E ρ)
    by (simp add: o_def case_prod_unfold)
  also have  = ((E,x)Rep_kraus_family 𝔈. sandwich_tc E ρ)
    apply (rule infsum_cong_neutral)
    by (auto simp: EE_def)
  also have  = 𝔈 *kr ρ
    by (metis (no_types, lifting) infsum_cong kf_apply.rep_eq prod.case_eq_if)
  finally show kf_map_inj f 𝔈 *kr ρ = 𝔈 *kr ρ
    by -
qed

lemma kf_map_inj_eq_kf_map:
  assumes inj_on f (kf_domain 𝔈)
  shows kf_map_inj f 𝔈 kr kf_map f 𝔈
proof (rule kf_eqI)
  fix x ρ
  define 𝔈fx where 𝔈fx = kf_filter (λz. f z = x) 𝔈
  from assms have inj_𝔈fx: inj_on f (kf_domain 𝔈fx)
    by (simp add: inj_on_def kf_domain.rep_eq 𝔈fx_def kf_filter.rep_eq)
  have kf_map_inj f 𝔈 *kr @{x} ρ = kf_filter (λz. z=x) (kf_map_inj f 𝔈) *kr ρ
    by (simp add: kf_apply_on_def)
  also have  = kf_map_inj f 𝔈fx *kr ρ
    apply (rule arg_cong[where f=λt. kf_apply t ρ])
    unfolding 𝔈fx_def
    apply (transfer' fixing: f)
    by force
  also have  = 𝔈fx *kr ρ
    using inj_𝔈fx by (rule kf_apply_map_inj)
  also have  = kf_map f (kf_filter (λz. f z = x) 𝔈) *kr ρ
    by (simp add: 𝔈fx_def)
  also have  = kf_apply (kf_filter (λxa. xa = x) (kf_map f 𝔈)) ρ
    apply (subst kf_filter_map)
    by simp
  also have  = kf_map f 𝔈 *kr @{x} ρ
    by (simp add: kf_apply_on_def)
  finally show kf_map_inj f 𝔈 *kr @{x} ρ = kf_map f 𝔈 *kr @{x} ρ
    by -
qed

lemma kf_map_inj_id[simp]: kf_map_inj id 𝔈 = 𝔈
  apply transfer' by simp

lemma kf_map_id: kf_map id 𝔈 kr 𝔈
  by (metis inj_on_id kf_eq_sym kf_map_inj_eq_kf_map kf_map_inj_id)

lemma kf_map_inj_bound[simp]:
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  assumes inj_on f (kf_domain 𝔈)
  shows kf_bound (kf_map_inj f 𝔈) = kf_bound 𝔈
  by (metis assms kf_eq_imp_eq_weak kf_map_inj_eq_kf_map kf_bound_cong kf_map_bound)

lemma kf_map_inj_norm[simp]:
  fixes 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'x) kraus_family
  assumes inj_on f (kf_domain 𝔈)
  shows kf_norm (kf_map_inj f 𝔈) = kf_norm 𝔈
  using assms kf_eq_imp_eq_weak kf_map_inj_eq_kf_map kf_norm_cong by fastforce

lemma kf_map_cong_weak:
  assumes 𝔈 =kr 𝔉
  shows kf_map f 𝔈 =kr kf_map g 𝔉
  by (metis assms kf_eq_weak_def kf_apply_map)

lemma kf_flatten_cong_weak:
  assumes 𝔈 =kr 𝔉
  shows kf_flatten 𝔈 =kr kf_flatten 𝔉
  using assms by (rule kf_map_cong_weak)

lemma kf_flatten_cong:
  assumes 𝔈 =kr 𝔉
  shows kf_flatten 𝔈 kr kf_flatten 𝔉
  by (simp add: assms kf_eq_weak_imp_eq_CARD_1 kf_flatten_cong_weak)

lemma kf_map_twice:
  kf_map f (kf_map g 𝔈) kr kf_map (f  g) 𝔈
  apply (rule kf_eqI)
  by (simp add: kf_filter_map kf_apply_on_def)

lemma kf_map_cong:
  assumes x. x  kf_domain 𝔈  f x = g x
  assumes 𝔈 kr 𝔉
  shows kf_map f 𝔈 kr kf_map g 𝔉
proof -
  have kf_filter (λy. f y = x) 𝔈 =kr kf_filter (λy. g y = x) 𝔉 for x
    apply (rule kf_filter_cong_weak)
    using assms by auto
  then have 𝔈 *kr @(f-`{x}) ρ = 𝔉 *kr @(g-`{x}) ρ for x ρ
    by (auto intro!: kf_apply_eqI simp add: kf_apply_on_def)
  then show ?thesis
    apply (rule_tac kf_eqI)
    by (simp add: kf_apply_on_def kf_filter_map)
qed

lemma kf_map_inj_cong_eq:
  assumes x. x  kf_domain 𝔈  f x = g x
  assumes 𝔈 = 𝔉
  shows kf_map_inj f 𝔈 = kf_map_inj g 𝔉
  using assms
  apply transfer'
  by force

(* lemma kf_map_remove_0[simp]:
  ‹kf_map f (kf_remove_0 𝔈) = kf_map f 𝔈›
  apply (transfer' fixing: f)
  by (simp add: kf_map.rep_eq kf_filter_remove0)

lemma kf_domain_remove_0[simp]: ‹kf_domain (kf_remove_0 E) = kf_domain E›
  apply transfer'
  by auto

lemma kf_remove_0_map[simp]:
  ‹kf_remove_0 (kf_map f 𝔈) = kf_map f 𝔈›
  apply (transfer' fixing: f)
  by auto *)



lemma kf_domain_map[simp]:
  kf_domain (kf_map f 𝔈) = f ` kf_domain 𝔈
proof (rule Set.set_eqI, rule iffI)
  fix x assume x  kf_domain (kf_map f 𝔈)
  then obtain a where (norm a)2 = kf_element_weight (kf_filter (λxa. f xa = x) 𝔈) a and a  0
    by (auto intro!: simp: kf_domain.rep_eq kf_map.rep_eq)
  then have kf_element_weight (kf_filter (λxa. f xa = x) 𝔈) a  0
    by force
  then have ((E, _)kf_similar_elements (kf_filter (λxa. f xa = x) 𝔈) a. (norm E)2)  0
    by (simp add: kf_element_weight_def)
  from this[unfolded not_def, rule_format, OF infsum_0]
  obtain E x' where rel_ops: (E,x')  kf_similar_elements (kf_filter (λxa. f xa = x) 𝔈) a
    and (norm E)2  0
    by fast
  then have E  0
    by force
  with rel_ops obtain E' where E'  0 and (E',x')  Rep_kraus_family (kf_filter (λxa. f xa = x) 𝔈)
    apply atomize_elim
    by (auto simp: kf_similar_elements_def)
  then have (E',x')  Rep_kraus_family 𝔈 and f x' = x
    by (auto simp: kf_filter.rep_eq)
  with E'  0 have x'  kf_domain 𝔈
    by (force simp: kf_domain.rep_eq)
  with f x' = x
  show x  f ` kf_domain 𝔈
    by fast
next
  fix x assume x  f ` kf_domain 𝔈
  then obtain y where x = f y and y  kf_domain 𝔈
    by blast
  then obtain E where E  0 and (E,y)  Rep_kraus_family 𝔈
    using Rep_kraus_family by (force simp: kf_domain.rep_eq kraus_family_def)
  then have Ey: (E,y)  Rep_kraus_family (kf_filter (λz. f z=x) 𝔈)
    by (simp add: kf_filter.rep_eq x = f y)
  then have kf_bound (kf_filter (λz. f z=x) 𝔈)  0
  proof -
    define B where B = kf_bound (kf_filter (λz. f z=x) 𝔈)
    have has_sum_in cweak_operator_topology (λ(E, x). E* oCL E) {(E,y)} (E* oCL E)
      apply (subst asm_rl[of E* oCL E = ((E,x){(E,y)}. E* oCL E)], simp)
      apply (rule has_sum_in_finite)
      by auto
    moreover have has_sum_in cweak_operator_topology (λ(E, x). E* oCL E) (Rep_kraus_family (kf_filter (λz. f z = x) 𝔈)) B
      using kf_bound_has_sum B_def by blast
    ultimately have B  E* oCL E
      apply (rule has_sum_mono_neutral_wot)
      using Ey positive_cblinfun_squareI by auto
    then show B  0
      by (meson E  0 basic_trans_rules(24) op_square_nondegenerate positive_cblinfun_squareI)
  qed
  then have kf_bound (kf_map f (kf_filter (λz. f z=x) 𝔈))  0
    by (simp add: kf_map_bound)
  then have kf_bound (kf_filter (λz. z=x) (kf_map f 𝔈))  0
    by (simp add: kf_filter_map)
  from this[unfolded not_def kf_bound.rep_eq, rule_format, OF infsum_in_0]
  obtain E' x' where (E',x')  Rep_kraus_family (kf_filter (λz. z=x) (kf_map f 𝔈))
    and E'  0
    by fastforce
  then have (E',x')  Rep_kraus_family (kf_map f 𝔈) and x' = x
    by (auto simp: kf_filter.rep_eq)
  with E'  0 show x  kf_domain (kf_map f 𝔈)
    by (auto simp: kf_domain.rep_eq image_iff Bex_def)
qed


lemma kf_apply_on_map[simp]:
  (kf_map f E) *kr @X ρ = E *kr @(f -` X) ρ
  by (auto intro!: simp: kf_apply_on_def kf_filter_map)

lemma kf_apply_on_map_inj[simp]:
  assumes inj_on f ((f -` X)  kf_domain E)
  shows  kf_map_inj f E *kr @X ρ = E *kr @(f -` X) ρ
proof -
  from assms
  have inj_on f (Set.filter (λx. f x  X) (kf_domain E))
    by (smt (verit, del_insts) IntI Set.member_filter inj_onD inj_onI vimage_eq)
  then show ?thesis
    by (auto intro!: simp: kf_apply_on_def kf_filter_map_inj)
qed

lemma kf_map0[simp]: kf_map f 0 = 0
  apply transfer'
  by auto

lemma kf_map_inj_kr_eq_weak:
  assumes inj_on f (kf_domain 𝔈)
  shows kf_map_inj f 𝔈 =kr 𝔈
  by (simp add: assms kf_eq_weakI)

lemma kf_map_inj_0[simp]: kf_map_inj f 0 = 0
  apply (transfer' fixing: f)
  by simp

lemma kf_domain_map_inj[simp]: kf_domain (kf_map_inj f 𝔈) = f ` kf_domain 𝔈
  apply transfer'
  by force

lemma kf_operators_kf_map:
  kf_operators (kf_map f 𝔈)  span (kf_operators 𝔈)
proof (rule subsetI)
  fix E
  assume E  kf_operators (kf_map f 𝔈)
  then obtain b where (norm E)2 = kf_element_weight (kf_filter (λx. f x = b) 𝔈) E  E  0 
    by (auto simp add: kf_operators.rep_eq kf_map.rep_eq)
  then have kf_element_weight (kf_filter (λx. f x = b) 𝔈) E  0
    by force
  then have E  span (kf_operators (kf_filter (λx. f x = b) 𝔈))
    by (rule kf_element_weight_kf_operators)
  then show E  span (kf_operators 𝔈)
    using kf_operators_filter[of (λx. f x = b) 𝔈]
    by (meson basic_trans_rules(31) span_mono)
qed

lemma kf_operators_kf_map_inj[simp]: kf_operators (kf_map_inj f 𝔈) = kf_operators 𝔈
  apply transfer' by force


subsection ‹Addition›

lift_definition kf_plus :: ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family  ('a,'b,'y) kraus_family  ('a,'b,'x+'y) kraus_family is
  λ𝔈 𝔉. (λ(E,x). (E, Inl x)) ` 𝔈  (λ(F,y). (F, Inr y)) ` 𝔉
proof (rename_tac 𝔈 𝔉)
  fix 𝔈 :: ('a CL 'b × 'x) set and 𝔉 :: ('a CL 'b × 'y) set
  assume 𝔈  Collect kraus_family and 𝔉  Collect kraus_family
  then have kraus_family 𝔈 and kraus_family 𝔉
    by auto
  then have kraus_family ((λ(E, x). (E, Inl x)) ` 𝔈  (λ(F, y). (F, Inr y)) ` 𝔉)
    by (force intro!: summable_on_Un_disjoint
      summable_on_reindex[THEN iffD2] inj_onI
      simp: kraus_family_iff_summable' o_def case_prod_unfold conj_commute)
  then show (λ(E, x). (E, Inl x)) ` 𝔈  (λ(F, y). (F, Inr y)) ` 𝔉  Collect kraus_family
    by simp
qed

instantiation kraus_family :: (chilbert_space, chilbert_space, type) plus begin
definition plus_kraus_family where 𝔈 + 𝔉 = kf_map (λxy. case xy of Inl x  x | Inr y  y) (kf_plus 𝔈 𝔉)
instance..
end

lemma kf_plus_apply:
  fixes 𝔈 :: ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family
    and 𝔉 :: ('a, 'b, 'y) kraus_family
  shows kf_apply (kf_plus 𝔈 𝔉) ρ = kf_apply 𝔈 ρ + kf_apply 𝔉 ρ
proof -
  have kf_apply (kf_plus 𝔈 𝔉) ρ = 
    (EF(λ(E,x). (E, Inl x)) ` Rep_kraus_family 𝔈  (λ(F,y). (F, Inr y)) ` Rep_kraus_family 𝔉. sandwich_tc (fst EF) ρ)
    by (simp add: kf_plus.rep_eq kf_apply_def case_prod_unfold)
  also have  = (EF(λ(E,x). (E, Inl x :: 'x+'y)) ` Rep_kraus_family 𝔈. sandwich_tc (fst EF) ρ)
                + (EF(λ(F,y). (F, Inr y :: 'x+'y)) ` Rep_kraus_family 𝔉. sandwich_tc (fst EF) ρ)
    apply (subst infsum_Un_disjoint)
    using kf_apply_summable
    by (auto intro!: summable_on_reindex[THEN iffD2] inj_onI
        simp: o_def case_prod_unfold kf_apply_summable)
  also have  = (ERep_kraus_family 𝔈. sandwich_tc (fst E) ρ) + (FRep_kraus_family 𝔉. sandwich_tc (fst F) ρ)
    apply (subst infsum_reindex)
     apply (auto intro!: inj_onI)[1]
    apply (subst infsum_reindex)
     apply (auto intro!: inj_onI)[1]
    by (simp add:  o_def case_prod_unfold)
  also have  = kf_apply 𝔈 ρ + kf_apply 𝔉 ρ
    by (simp add: kf_apply_def)
  finally show ?thesis
    by -
qed

lemma kf_plus_apply': (𝔈 + 𝔉) *kr ρ = 𝔈 *kr ρ + 𝔉 *kr ρ
  by (simp add: kf_plus_apply plus_kraus_family_def)

lemma kf_plus_0_left[simp]: kf_plus 0 𝔈 = kf_map_inj Inr 𝔈
  apply transfer' by auto

lemma kf_plus_0_right[simp]: kf_plus 𝔈 0 = kf_map_inj Inl 𝔈
  apply transfer' by auto

lemma kf_plus_0_left'[simp]: 0 + 𝔈 kr 𝔈
proof -
  define merge where merge xy = (case xy of Inl x  x | Inr y  y) for xy :: 'c + 'c
  have 0 + 𝔈 = kf_map merge (kf_map_inj Inr 𝔈)
    by (simp add: plus_kraus_family_def merge_def[abs_def])
  also have  kr kf_map merge (kf_map Inr 𝔈)
    by (auto intro!: kf_map_cong kf_map_inj_eq_kf_map)
  also have  kr kf_map (merge o Inr) 𝔈
  by (simp add: kf_map_twice)
  also have  = kf_map id 𝔈
    apply (rule arg_cong2[where f=kf_map])
    by (auto simp: merge_def)
  also have  kr 𝔈
    by (simp add: kf_map_id)
  finally show ?thesis
    by -
qed

lemma kf_plus_0_right': 𝔈 + 0 kr 𝔈
proof -
  define merge where merge xy = (case xy of Inl x  x | Inr y  y) for xy :: 'c + 'c
  have 𝔈 + 0 = kf_map merge (kf_map_inj Inl 𝔈)
    by (simp add: plus_kraus_family_def merge_def[abs_def])
  also have  kr kf_map merge (kf_map Inl 𝔈)
    by (auto intro!: kf_map_cong kf_map_inj_eq_kf_map)
  also have  kr kf_map (merge o Inl) 𝔈
  by (simp add: kf_map_twice)
  also have  = kf_map id 𝔈
    apply (rule arg_cong2[where f=kf_map])
    by (auto simp: merge_def)
  also have  kr 𝔈
    by (simp add: kf_map_id)
  finally show ?thesis
    by -
qed

lemma kf_plus_bound: kf_bound (kf_plus 𝔈 𝔉) = kf_bound 𝔈 + kf_bound 𝔉
proof -
  define l r where l = (λ(E::'aCL'b, x)  (E, Inl x :: 'c+'d))
    and r = (λ(F::'aCL'b, y)  (F, Inr y :: 'c+'d))
  have Abs_cblinfun_wot (kf_bound (kf_plus 𝔈 𝔉)) 
    = ((E, x)l ` Rep_kraus_family 𝔈  r ` Rep_kraus_family 𝔉. compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
    by (simp add: kf_bound_def' kf_plus.rep_eq Rep_cblinfun_wot_inverse flip: l_def r_def)
  also have  = ((E, x)l ` Rep_kraus_family 𝔈. compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
                + ((E, x)r ` Rep_kraus_family 𝔉. compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
    apply (rule infsum_Un_disjoint)
      apply (metis (no_types, lifting) ext Un_empty_right l_def image_empty kf_bound_summable' kf_plus.rep_eq
        zero_kraus_family.rep_eq) 
     apply (metis (no_types, lifting) ext r_def empty_subsetI image_empty kf_bound_summable' kf_plus.rep_eq
        sup.absorb_iff2 zero_kraus_family.rep_eq) 
    by (auto intro!: simp: l_def r_def)
  also have  = Abs_cblinfun_wot (kf_bound (kf_map_inj Inl 𝔈 :: (_,_,'c+'d) kraus_family)) + Abs_cblinfun_wot (kf_bound (kf_map_inj Inr 𝔉 :: (_,_,'c+'d) kraus_family))
    by (simp add: kf_bound_def' Rep_cblinfun_wot_inverse l_def r_def kf_map_inj.rep_eq case_prod_unfold)
  also have  = Abs_cblinfun_wot (kf_bound 𝔈 + kf_bound 𝔉)
    by (simp add: kf_map_inj_bound plus_cblinfun_wot.abs_eq)
  finally show ?thesis
    by (metis (no_types, lifting) Rep_cblinfun_wot_inverse kf_bound_def' plus_cblinfun_wot.rep_eq)
qed

lemma kf_plus_bound': kf_bound (𝔈 + 𝔉) = kf_bound 𝔈 + kf_bound 𝔉
  by (simp add: kf_map_bound kf_plus_bound plus_kraus_family_def)

lemma kf_norm_triangle: kf_norm (kf_plus 𝔈 𝔉)  kf_norm 𝔈 + kf_norm 𝔉
  by (simp add: kf_norm_def kf_plus_bound norm_triangle_ineq)

lemma kf_norm_triangle': kf_norm (𝔈 + 𝔉)  kf_norm 𝔈 + kf_norm 𝔉
  by (simp add: kf_norm_def kf_plus_bound' norm_triangle_ineq)

lemma kf_plus_map_both:
  kf_plus (kf_map f 𝔈) (kf_map g 𝔉) = kf_map (map_sum f g) (kf_plus 𝔈 𝔉)
proof -
  have 1: kf_filter (λx. map_sum f g x = Inl y) (kf_plus 𝔈 𝔉) = 
          kf_map_inj Inl (kf_filter (λx. f x = y) 𝔈) for y
    apply (transfer' fixing: f g y)
    by force
  have 2: kf_filter (λx. map_sum f g x = Inr y) (kf_plus 𝔈 𝔉) = 
          kf_map_inj Inr (kf_filter (λx. g x = y) 𝔉) for y
    apply (transfer' fixing: f g y)
    by force
  show ?thesis
    apply (transfer' fixing: f g 𝔈 𝔉)
    apply (rule Set.set_eqI)
    subgoal for x
      apply (cases snd x)
      by (auto intro!: simp: 1 2 kf_element_weight_map_inj split!: sum.split)
    by -
qed


subsection ‹Composition›

lemma kf_comp_dependent_raw_norm_aux:
  fixes 𝔈 :: 'a  ('e::chilbert_space, 'f::chilbert_space, 'g) kraus_family
    and 𝔉 :: ('b::chilbert_space, 'e, 'a) kraus_family
  assumes B: x. x  kf_domain 𝔉  kf_norm (𝔈 x)  B
  assumes [simp]: B  0
  assumes finite C
  assumes C_subset: C  (λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))
  shows ((E,x)C. E* oCL E)  (B * kf_norm 𝔉) *R id_cblinfun
proof -
  define BF :: 'b CL 'b where BF = kf_norm 𝔉 *R id_cblinfun
  then have kf_bound 𝔉  BF
    by (simp add: kf_bound_leq_kf_norm_id)
  then have BF: ((F, y)M. (F* oCL F))  BF if M  Rep_kraus_family 𝔉 and finite M for M
    using dual_order.trans kf_bound_geq_sum that(1) by blast
  define BE :: 'e CL 'e where BE = B *R id_cblinfun
  define 𝔉x𝔈 where 𝔉x𝔈 = (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))
  have BE: ((E, x)M. (E* oCL E))  BE if y  kf_domain 𝔉 and M  Rep_kraus_family (𝔈 y) and finite M for M y
  proof -
    from B that(1,2)
    have norm ((E, x)M. E* oCL E)  B
      by (smt (verit) kf_norm_geq_norm_sum that)
    then show ?thesis
      by (auto intro!: less_eq_scaled_id_norm pos_selfadjoint sum_nonneg intro: positive_cblinfun_squareI simp: BE_def)
  qed

  define A where A = ((E,x)C. E* oCL E)
  define CE CF where CE y = (λ(_,(F,E,y,x)). (E,x)) ` Set.filter (λ(_,(F,E,y',x)). y'=y) C 
    and CF = (λ(_,(F,E,y,x)). (F,y)) ` C for y
  with finite C have [simp]: finite (CE y) finite CF for y
    by auto
  have C_C1C2: C  (λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):CF. CE y)
  proof (rule subsetI)
    fix c assume c  C
    then obtain EF E F x y where c_def: c = (EF,(F,E,y,x))
      by (metis surj_pair)
    from c  C have EF_def: EF = E oCL F
      using C_subset by (auto intro!: simp: c_def)
    from c  C have 1: (F,y)  CF
      apply (simp add: CF_def c_def)
      by force
    from c  C have 2: (E,x)  CE y
      apply (simp add: CE_def c_def)
      by force
    from 1 2 show c  (λ((F, y), E, x). (E oCL F, (F,E,y,x))) ` (SIGMA (F, y):CF. CE y)
      apply (simp add: c_def EF_def)
      by force
  qed

  have CE_sub_𝔈: CE y  Rep_kraus_family (𝔈 y) and CF  Rep_kraus_family 𝔉 for y
    using C_subset by (auto simp add: CE_def CF_def 𝔉x𝔈_def case_prod_unfold)
  have CE_BE: ((E, x)CE y. (E* oCL E))  BE if y  kf_domain 𝔉 for y
    using BE[where y=y] CE_sub_𝔈[of y] that
    by auto

  have A  ((E,x)  (λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):CF. CE y). E* oCL E)
    using C_C1C2 by (auto intro!: finite_imageI sum_mono2 positive_cblinfun_squareI simp: A_def simp flip: adj_cblinfun_compose)[1]
  also have  = (((F,y), (E,x))(SIGMA (F,y):CF. CE y). (F* oCL (E* oCL E) oCL F))
    apply (subst sum.reindex)
    by (auto intro!: inj_onI simp: case_prod_unfold cblinfun_compose_assoc)
  also have  = ((F, y)CF. sandwich (F*) ((E, x)CE y. (E* oCL E)))
    apply (subst sum.Sigma[symmetric])
    by (auto intro!: simp: case_prod_unfold sandwich_apply cblinfun_compose_sum_right cblinfun_compose_sum_left simp flip: )
  also have   ((F, y)CF. sandwich (F*) BE)
  proof (rule sum_mono)
    fix i :: 'b CL 'e × 'a assume i  CF
    obtain F y where i: i = (F,y)
      by force
    have 1: sandwich (F*) *V ((E,x)CE y. E* oCL E)  sandwich (F*) *V BE if y  kf_domain 𝔉
      apply (rule sandwich_mono)
      using that CE_BE by simp
    have F = 0 if y  kf_domain 𝔉
        using C_subset CF_def CF  Rep_kraus_family 𝔉 i  CF that i
        by (smt (verit, ccfv_SIG) Set.basic_monos(7) Set.member_filter case_prodI image_iff kf_domain.rep_eq prod.sel(2))
    then have 2: sandwich (F*) *V ((E,x)CE y. E* oCL E)  sandwich (F*) *V BE if y  kf_domain 𝔉
      using that by simp
    from 1 2 show (case i of (F, y)  sandwich (F*) *V ((E, x)CE y. E* oCL E))
          (case i of (F, y)  sandwich (F*) *V BE)
      by (auto simp: case_prod_unfold i)
  qed
  also have  = B *R ((F, y)CF. F* oCL F)
    by (simp add: scaleR_sum_right case_prod_unfold sandwich_apply BE_def)
  also have   B *R BF
    using BF by (simp add: CF  Rep_kraus_family 𝔉 scaleR_left_mono case_prod_unfold)
  also have B *R BF = (B * kf_norm 𝔉) *R id_cblinfun
    by (simp add: BF_def)
  finally show A  (B * kf_norm 𝔉) *R id_cblinfun
    by -
qed

lift_definition kf_comp_dependent_raw :: ('x  ('b::chilbert_space,'c::chilbert_space,'y) kraus_family)  ('a::chilbert_space,'b,'x) kraus_family
                     ('a, 'c, ('a CL 'b) × ('b CL 'c) × 'x × 'y) kraus_family is
  λ𝔈 𝔉. if bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉) then
    Set.filter (λ(EF,_). EF0) ((λ((F,y), (E::'bCL'c,x::'y)). (E oCL F, (F,E,y,x))) ` (SIGMA (F::'aCL'b,y::'x):Rep_kraus_family 𝔉. (Rep_kraus_family (𝔈 y))))
    else {}
proof (rename_tac 𝔈 𝔉)
  fix 𝔈 :: 'x  ('b, 'c, 'y) kraus_family and 𝔉 :: ('a, 'b, 'x) kraus_family
  show (if bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
        then Set.filter (λ(EF,_). EF0) ((λ((F, y), E, x). (E oCL F, (F, E, y, x))) ` (SIGMA (F, y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y)))
        else {})
        Collect kraus_family
  proof (cases bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉))
    case True
    obtain B where 𝔈_uniform: y  kf_domain 𝔉  kf_norm (𝔈 y)  B and B  0 for y
    proof atomize_elim
      from True
      obtain B0 where y  kf_domain 𝔉  kf_norm (𝔈 y)  B0 for y
        by (auto simp: bdd_above_def)
      then show B. (y. y  kf_domain 𝔉  kf_norm (𝔈 y)  B)  0  B
        apply (rule_tac exI[of _ max 0 B0])
        by force
    qed
    define 𝔉x𝔈 where 𝔉x𝔈 = (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))
    have bdd_above ((λM. (E,x)M. E* oCL E) `
            {M. M  Set.filter (λ(EF, _). EF  0) ((λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` 𝔉x𝔈)  finite M})
    proof (rule bdd_aboveI, rename_tac A)
      fix A :: 'a CL 'a
      assume A  (λM. (E, x)M. E* oCL E) ` {M. M  Set.filter (λ(EF, _). EF  0) ((λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` 𝔉x𝔈)  finite M}
      then obtain C where A_def: A = ((E,x)C. E* oCL E)
        and C𝔉𝔈: C  Set.filter (λ(EF,_). EF  0) ((λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` 𝔉x𝔈)
        and [simp]: finite C
        by auto
      from kf_comp_dependent_raw_norm_aux[OF 𝔈_uniform B  0 finite C]
      show A  (B * kf_norm 𝔉) *R id_cblinfun
        using C𝔉𝔈
        by (force intro!: simp: A_def 𝔉x𝔈_def)
    qed
    then have kraus_family (Set.filter (λ(EF,_). EF0) ((λ((F, y), E, x). (E oCL F, (F,E,y,x))) ` (SIGMA (F, y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))))
      by (auto intro!: kraus_familyI simp: conj_commute 𝔉x𝔈_def)
    then show ?thesis
      using True by simp
  next
    case False
    then show ?thesis
      by (auto simp: kraus_family_def)
  qed
qed

lemma kf_comp_dependent_raw_norm_leq:
  fixes 𝔈 :: 'a  ('b::chilbert_space, 'c::chilbert_space, 'd) kraus_family
    and 𝔉 :: ('e::chilbert_space, 'b, 'a) kraus_family
  assumes x. x  kf_domain 𝔉  kf_norm (𝔈 x)  B
  assumes B  0
  shows kf_norm (kf_comp_dependent_raw 𝔈 𝔉)  B * kf_norm 𝔉
proof -
  wlog not_singleton: class.not_singleton TYPE('e)
    using not_not_singleton_kf_norm_0[OF negation, of 𝔉]
    using not_not_singleton_kf_norm_0[OF negation, of kf_comp_dependent_raw 𝔈 𝔉]
    by simp
  show ?thesis
  proof (rule kf_norm_sum_leqI)
    fix F assume finite F and F_subset: F  Rep_kraus_family (kf_comp_dependent_raw 𝔈 𝔉)
    have [simp]: norm (id_cblinfun :: 'e CL 'e) = 1
      apply (rule norm_cblinfun_id[internalize_sort' 'a])
      apply (rule complex_normed_vector_axioms)
      by (rule not_singleton)
    from assms have bdd: bdd_above ((λx. kf_norm (𝔈 x)) ` kf_domain 𝔉)
      by fast
    have ((E, x)F. E* oCL E)  (B * kf_norm 𝔉) *R id_cblinfun
      using assms finite F apply (rule kf_comp_dependent_raw_norm_aux)
      using F_subset by (auto simp: kf_comp_dependent_raw.rep_eq bdd)
    then have norm ((E, x)F. E* oCL E)  norm ((B * kf_norm 𝔉) *R (id_cblinfun :: 'e CL 'e))
      apply (rule norm_cblinfun_mono[rotated])
      using positive_cblinfun_squareI 
      by (auto intro!: sum_nonneg)
    then show norm ((E, x)F. E* oCL E)  B * kf_norm 𝔉
      using B  0 by auto
  qed
qed

hide_fact kf_comp_dependent_raw_norm_aux

definition kf_comp_dependent 𝔈 𝔉 = kf_map (λ(F,E,y,x). (y,x)) (kf_comp_dependent_raw 𝔈 𝔉)

definition kf_comp 𝔈 𝔉 = kf_comp_dependent (λ_. 𝔈) 𝔉

lemma kf_comp_dependent_norm_leq:
  assumes x. x  kf_domain 𝔉  kf_norm (𝔈 x)  B
  assumes B  0
  shows kf_norm (kf_comp_dependent 𝔈 𝔉)  B * kf_norm 𝔉
  using assms by (auto intro!: kf_comp_dependent_raw_norm_leq simp: kf_comp_dependent_def)

lemma kf_comp_norm_leq:
  shows kf_norm (kf_comp 𝔈 𝔉)  kf_norm 𝔈 * kf_norm 𝔉
  by (auto intro!: kf_comp_dependent_norm_leq simp: kf_comp_def)

lemma kf_comp_dependent_raw_apply:
  fixes 𝔈 :: 'y  ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family
    and 𝔉 :: ('c::chilbert_space, 'a, 'y) kraus_family
  assumes bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  shows kf_comp_dependent_raw 𝔈 𝔉 *kr ρ
            = ((F,y)Rep_kraus_family 𝔉. 𝔈 y *kr sandwich_tc F ρ)
proof -
  have sum2: (λ(F, x). sandwich_tc F ρ) summable_on (Rep_kraus_family 𝔉)
    using kf_apply_summable[of ρ 𝔉] (* kraus𝔉 *) by (simp add: case_prod_unfold)
  have (λE. sandwich_tc (fst E) ρ) summable_on 
          (Set.filter (λ(E,x). E  0) ((λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))))
    using kf_apply_summable[of _ kf_comp_dependent_raw 𝔈 𝔉] assms
    by (simp add: kf_comp_dependent_raw.rep_eq case_prod_unfold)
  then have (λE. sandwich_tc (fst E) ρ) summable_on 
          (λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))
    apply (rule summable_on_cong_neutral[THEN iffD1, rotated -1])
    by force+
  then have sum1: (λ((F,y), (E,x)). sandwich_tc (E oCL F) ρ) summable_on (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y))
    apply (subst (asm) summable_on_reindex)
    by (auto intro!: inj_onI simp: o_def case_prod_unfold)
  have kf_comp_dependent_raw 𝔈 𝔉 *kr ρ
          = (E(Set.filter (λ(E,x). E  0) ((λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y)))). sandwich_tc (fst E) ρ)
    using assms by (simp add: kf_apply_def kf_comp_dependent_raw.rep_eq case_prod_unfold)
  also have  = (E(λ((F,y), (E,x)). (E oCL F, (F,E,y,x))) ` (SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y)). sandwich_tc (fst E) ρ)
    apply (rule infsum_cong_neutral)
    by force+
  also have  = (((F,y), (E,x))(SIGMA (F,y):Rep_kraus_family 𝔉. Rep_kraus_family (𝔈 y)). sandwich_tc (E oCL F) ρ)
    apply (subst infsum_reindex)
    by (auto intro!: inj_onI simp: o_def case_prod_unfold)
  also have  = ((F,y)Rep_kraus_family 𝔉. (E,x)Rep_kraus_family (𝔈 y). sandwich_tc (E oCL F) ρ)
    apply (subst infsum_Sigma'_banach[symmetric])
    using sum1 by (auto simp: case_prod_unfold)
  also have  = ((F,y)Rep_kraus_family 𝔉. (E,x)Rep_kraus_family (𝔈 y). sandwich_tc E (sandwich_tc F ρ))
    by (simp add: sandwich_tc_compose)
  also have  = ((F,y)Rep_kraus_family 𝔉. kf_apply (𝔈 y) (sandwich_tc F ρ))
    by (simp add: kf_apply_def case_prod_unfold)
  finally show ?thesis
    by -
qed

lemma kf_comp_dependent_apply:
  fixes 𝔈 :: 'y  ('a::chilbert_space, 'b::chilbert_space, 'x) kraus_family
    and 𝔉 :: ('c::chilbert_space, 'a, 'y) kraus_family
  assumes bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  shows kf_comp_dependent 𝔈 𝔉 *kr ρ
      = ((F,y)Rep_kraus_family 𝔉. 𝔈 y *kr sandwich_tc F ρ)
  using assms by (simp add: kf_comp_dependent_def kf_apply_map
      kf_comp_dependent_raw_apply)

lemma kf_comp_apply:
  shows kf_apply (kf_comp 𝔈 𝔉) = kf_apply 𝔈  kf_apply 𝔉
proof (rule ext, rename_tac ρ)
  fix ρ :: ('a, 'a) trace_class

  have sumF: (λ(F, y). sandwich_tc F ρ) summable_on Rep_kraus_family 𝔉
    by (rule kf_apply_summable)
  have kf_comp 𝔈 𝔉 *kr ρ = ((F,y)Rep_kraus_family 𝔉. 𝔈 *kr sandwich_tc F ρ)
    by (auto intro!: kf_comp_dependent_apply simp: kf_comp_def)
  also have  = kf_apply 𝔈 ((F,y)Rep_kraus_family 𝔉. sandwich_tc F ρ)
    apply (subst infsum_bounded_linear[symmetric, where h=kf_apply 𝔈])
    using sumF by (auto intro!: bounded_clinear.bounded_linear kf_apply_bounded_clinear
        simp: o_def case_prod_unfold)
  also have  = (kf_apply 𝔈  kf_apply 𝔉) ρ
    by (simp add: o_def kf_apply_def case_prod_unfold)
  finally show kf_apply (kf_comp 𝔈 𝔉) ρ = (kf_apply 𝔈  kf_apply 𝔉) ρ
    by -
qed

lemma kf_comp_cong_weak: kf_comp F G =kr kf_comp F' G'
  if F =kr F' and G =kr G'
  by (metis kf_eq_weak_def kf_comp_apply that)

lemma kf_comp_dependent_raw_assoc: 
  fixes 𝔈 :: 'f  ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: 'g  ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  defines reorder :: 'a CL 'c × 'c CL 'd × ('a CL 'b × 'b CL 'c × 'g × 'f) × 'e
                    'a CL 'b × 'b CL 'd × 'g × 'b CL 'c × 'c CL 'd × 'f × 'e 
             λ(FG::'a CL 'c, E::'c CL 'd, (G::'a CL 'b, F::'b CL 'c, g::'g, f::'f), e::'e). 
              (G, E oCL F, g, F, E, f, e)
  assumes bdd_above (range (kf_norm o 𝔈))
  assumes bdd_above (range (kf_norm o 𝔉))
  shows kf_comp_dependent_raw (λg::'g. kf_comp_dependent_raw 𝔈 (𝔉 g)) 𝔊
        = kf_map_inj reorder (kf_comp_dependent_raw (λ(_,_,_,f). 𝔈 f) (kf_comp_dependent_raw 𝔉 𝔊))
    (is ?lhs = ?rhs)
proof (rule Rep_kraus_family_inject[THEN iffD1])
  from assms have bdd_E: bdd_above ((kf_norm o 𝔈) ` X) for X
    by (simp add: bdd_above_mono2)
  from assms have bdd_F: bdd_above ((kf_norm o 𝔉) ` X) for X
    by (simp add: bdd_above_mono2)
  have bdd1: bdd_above ((λx. kf_norm (kf_comp_dependent_raw 𝔈 (𝔉 x))) ` X) for X
  proof -
    from bdd_F[where X=UNIV] obtain BF where BF: kf_norm (𝔉 x)  BF for x
      by (auto simp: bdd_above_def)
    moreover from bdd_E[where X=UNIV] obtain BE where BE: kf_norm (𝔈 x)  BE for x
      by (auto simp: bdd_above_def)
    ultimately have kf_norm (kf_comp_dependent_raw (λx. 𝔈 x) (𝔉 x))  BE * BF for x
      by (smt (verit, best) kf_comp_dependent_raw_norm_leq kf_norm_geq0 landau_omega.R_mult_left_mono)
    then show ?thesis
      by (auto intro!: bdd_aboveI)
  qed
  have bdd2: bdd_above ((kf_norm  (λ(_::'a CL 'b, _::'b CL 'c, _::'g, y::'f). 𝔈 y)) ` X) for X
    using assms(2) by (auto simp: bdd_above_def)
  define EE FF GG where EE f = Rep_kraus_family (𝔈 f) and FF g = Rep_kraus_family (𝔉 g) and GG = Rep_kraus_family 𝔊 for f g
  have Rep_kraus_family ?lhs
        = (Set.filter (λ(E,x). E  0) ((λ((F,y), (E, x)). (E oCL F, F, E, y, x)) `
          (SIGMA (F, y):GG. Rep_kraus_family (kf_comp_dependent_raw 𝔈 (𝔉 y)))))
    apply (subst kf_comp_dependent_raw.rep_eq)
    using bdd1 by (simp add:  GG_def)
  also have  = (Set.filter (λ(E,x). E  0) ((λ((G, g), (EF, x)). (EF oCL G, G, EF, g, x)) `
    (SIGMA (G, g):GG. Set.filter (λ(E,x). E  0) ((λ((F, f), (E, e)). (E oCL F, F, E, f, e)) ` (SIGMA (F, f):FF g. EE f)))))
    unfolding EE_def FF_def
    apply (subst kf_comp_dependent_raw.rep_eq)
    using assms bdd_E by (simp add: case_prod_beta)
  also have  = (Set.filter (λ(E,x). E  0) ((λ((G, g), (EF, x)). (EF oCL G, G, EF, g, x)) `
    (SIGMA (G, g):GG. (λ((F, f), (E, e)). (E oCL F, F, E, f, e)) ` (SIGMA (F, f):FF g. EE f))))
    by force
  also have  = (Set.filter (λ(E,x). E  0) ((λ((F, y), (E, x)). (E oCL F, reorder (F, E, y, x))) ` 
         (SIGMA (FG, _, _, _, y):(λ((G, g), (F, f)). (F oCL G, G, F, g, f)) ` (SIGMA (G, g):GG. FF g). EE y)))
    by (force simp: reorder_def image_iff case_prod_unfold cblinfun_compose_assoc)
  also have  = (Set.filter (λ(E,x). E  0) ((λ((F, y), (E, x)). (E oCL F, reorder (F, E, y, x))) ` 
         (SIGMA (FG, _, _, _, y):Set.filter (λ(E,x). E  0) ((λ((G, g), (F, f)). (F oCL G, G, F, g, f)) ` (SIGMA (G, g):GG. FF g)). EE y)))
    by force
  also have  = (Set.filter (λ(E,x). E  0) ((λ((F, y), (E, x)). (E oCL F, reorder (F, E, y, x))) `
    (SIGMA (FG, (_, _, _, f)):Rep_kraus_family (kf_comp_dependent_raw 𝔉 𝔊). EE f)))
    apply (rule arg_cong[where f=Set.filter _])
    apply (subst kf_comp_dependent_raw.rep_eq)
    using assms bdd_F
    by (simp add: flip: FF_def GG_def)
  also have  = (Set.filter (λ(E,x). E  0) ((λ(E,z). (E, reorder z)) ` (λ((F, y), (E, x)). (E oCL F, F, E, y, x)) `
    (SIGMA (FG, (_, _, _, f)):Rep_kraus_family (kf_comp_dependent_raw 𝔉 𝔊). EE f)))
    by (simp add: image_image case_prod_beta)
  also have  = (λ(E,z). (E, reorder z)) ` (Set.filter (λ(E,x). E  0) ((λ((F, y), (E, x)). (E oCL F, F, E, y, x)) `
    (SIGMA (FG, (_, _, _, f)):Rep_kraus_family (kf_comp_dependent_raw 𝔉 𝔊). EE f)))
    by force+
  also have  = (λ(E,x). (E,reorder x)) ` Rep_kraus_family
     (kf_comp_dependent_raw (λ(_, _, _, y). 𝔈 y) (kf_comp_dependent_raw 𝔉 𝔊))
    apply (subst (2) kf_comp_dependent_raw.rep_eq)
    using bdd2 by (simp add: case_prod_unfold EE_def)
  also have  = Rep_kraus_family ?rhs
    by (simp add: kf_map_inj.rep_eq case_prod_beta)
  finally show Rep_kraus_family ?lhs = Rep_kraus_family ?rhs
    by -
qed

lemma kf_filter_comp_dependent:
  fixes 𝔉 :: 'e  ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'e) kraus_family
  assumes bdd_above ((kf_norm o 𝔉) ` kf_domain 𝔈)
  shows kf_filter (λ(e,f). F e f  E e) (kf_comp_dependent 𝔉 𝔈)
      = kf_comp_dependent (λe. kf_filter (F e) (𝔉 e)) (kf_filter E 𝔈)
proof -
  from assms
  have bdd2: bdd_above ((λe. kf_norm (kf_filter (F e) (𝔉 e))) ` kf_domain 𝔈)
    apply (rule bdd_above_mono2)
    by (auto intro!: kf_norm_filter)
  then have bdd3: bdd_above ((λx. kf_norm (kf_filter (F x) (𝔉 x))) `
                              kf_domain (kf_filter E 𝔈))
    apply (rule bdd_above_mono2)
    by auto
  show ?thesis
    unfolding kf_comp_dependent_def kf_filter_map
    apply (rule arg_cong[where f=kf_map _])
    using assms bdd2 bdd3 apply (transfer' fixing: F E)
    by (auto intro!: simp: kf_filter.rep_eq case_prod_unfold image_iff Bex_def)
qed

lemma kf_comp_assoc_weak: 
  fixes 𝔈 :: ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  shows kf_comp (kf_comp 𝔈 𝔉) 𝔊 =kr kf_comp 𝔈 (kf_comp 𝔉 𝔊)
  apply (rule kf_eq_weakI)
  by (simp add: kf_comp_apply)

lemma kf_comp_dependent_raw_cong_left:
  assumes bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  assumes bdd_above ((kf_norm o 𝔈') ` kf_domain 𝔉)
  assumes x. x  snd ` Rep_kraus_family 𝔉  𝔈 x = 𝔈' x
  shows kf_comp_dependent_raw 𝔈 𝔉 = kf_comp_dependent_raw 𝔈' 𝔉
proof -
  show ?thesis
  apply (rule Rep_kraus_family_inject[THEN iffD1])
    using assms
  by (force simp: kf_comp_dependent_def kf_comp_dependent_raw.rep_eq 
      image_iff case_prod_beta Bex_def) 
qed

(* lemma kf_remove_0_comp_dependent_raw:
  ‹kf_remove_0 (kf_comp_dependent_raw 𝔈 𝔉) =
      kf_remove_0 (kf_comp_dependent_raw (λx. kf_remove_0 (𝔈 x)) (kf_remove_0 𝔉))›
  apply transfer'
  by (auto intro!: simp: image_iff case_prod_beta kf_remove_0.rep_eq Bex_def) *)

lemma kf_comp_dependent_cong_left:
  assumes bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  assumes bdd_above ((kf_norm o 𝔈') ` kf_domain 𝔉)
  assumes x. x  kf_domain 𝔉  𝔈 x = 𝔈' x
  shows kf_comp_dependent 𝔈 𝔉 = kf_comp_dependent 𝔈' 𝔉
proof -
  have kf_comp_dependent 𝔈 𝔉 = kf_map (λ(F, E, y). y) (kf_comp_dependent_raw 𝔈 𝔉)
    by (simp add: kf_comp_dependent_def id_def)
  also have  = kf_map (λ(F, E, y). y) ((kf_comp_dependent_raw (λx. (𝔈' x)) (𝔉)))
    apply (rule arg_cong[where f=λt. kf_map _ t])
    apply (rule kf_comp_dependent_raw_cong_left[OF assms])
    using assms by (auto intro!: simp: kf_domain.rep_eq)
  also have  = kf_comp_dependent 𝔈' 𝔉
    by (simp add: kf_comp_dependent_def id_def)
  finally show ?thesis
    by -
qed

lemma kf_domain_comp_dependent_raw_subset:
  kf_domain (kf_comp_dependent_raw 𝔈 𝔉)  UNIV × UNIV × (SIGMA x:kf_domain 𝔉. kf_domain (𝔈 x))
  by (auto intro!: simp: kf_comp_dependent_raw.rep_eq kf_domain.rep_eq image_iff Bex_def)

lemma kf_domain_comp_dependent_subset:
  kf_domain (kf_comp_dependent 𝔈 𝔉)  (SIGMA x:kf_domain 𝔉. kf_domain (𝔈 x))
  apply (simp add: kf_comp_dependent_def kf_domain_map id_def)
  by (auto intro!: simp: kf_comp_dependent_raw.rep_eq kf_domain.rep_eq image_iff Bex_def)

lemma kf_domain_comp_subset: kf_domain (kf_comp 𝔈 𝔉)  kf_domain 𝔉 × kf_domain 𝔈
  by (metis Sigma_cong kf_comp_def kf_domain_comp_dependent_subset)

lemma kf_apply_comp_dependent_cong:
  fixes 𝔈 :: 'f  ('b::chilbert_space,'c::chilbert_space,'e1) kraus_family
    and 𝔈' :: 'f  ('b::chilbert_space,'c::chilbert_space,'e2) kraus_family
    and 𝔉 𝔉' :: ('a::chilbert_space,'b::chilbert_space,'f) kraus_family
  assumes bdd: bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  assumes bdd': bdd_above ((kf_norm o 𝔈') ` kf_domain 𝔉')
  assumes f  kf_domain 𝔉  kf_apply_on (𝔈 f) E = kf_apply_on (𝔈' f) E'
  assumes kf_apply_on 𝔉 {f} = kf_apply_on 𝔉' {f}
  shows kf_apply_on (kf_comp_dependent 𝔈 𝔉) ({f}×E) = kf_apply_on (kf_comp_dependent 𝔈' 𝔉') ({f}×E')
proof (rule ext)
  fix ρ :: ('a, 'a) trace_class

  have rewrite_comp: kf_apply_on (kf_comp_dependent 𝔈 𝔉) ({f}×E) = 
        kf_apply (kf_comp (kf_filter (λx. xE) (𝔈 f))
                                           (kf_filter (λx. x=f) 𝔉))
    if bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉) 
    for E and 𝔈 :: 'f  ('b::chilbert_space,'c::chilbert_space,'e) kraus_family 
      and 𝔉 :: ('a::chilbert_space,'b::chilbert_space,'f) kraus_family
  proof -
    have bdd_filter: bdd_above ((kf_norm  (λf. kf_filter (λx. xE) (𝔈 f))) ` kf_domain 𝔉)
      apply (rule bdd_above_mono2[OF _ subset_refl, rotated])
      using kf_norm_filter apply blast
      using that
      by (metis (mono_tags, lifting) bdd_above_mono2 comp_apply kf_norm_filter order.refl)
    have aux: (λ(x,y). yE  x=f) = (λx. x{f}×E)
      by auto
    have *: kf_filter (λx. x{f}×E) (kf_comp_dependent 𝔈 𝔉)
        = kf_comp_dependent (λf. kf_filter (λx. xE) (𝔈 f))
               (kf_filter (λx. x=f) 𝔉)
      using kf_filter_comp_dependent[where 𝔈=𝔉 and 𝔉=𝔈 and F=λ_ x. xE and E=λx. x=f,
            OF that]
      unfolding aux by auto
    have kf_apply_on (kf_comp_dependent 𝔈 𝔉) ({f}×E)
      = kf_apply (kf_comp_dependent (λf. kf_filter (λx. xE) (𝔈 f))
                                                     (kf_filter (λx. x=f) 𝔉))
      by (auto intro!: simp: kf_apply_on_def * )
    also have  = kf_apply
     (kf_comp_dependent (λ_. kf_filter (λx. xE) (𝔈 f))
       (kf_filter (λx. x=f) 𝔉))
      apply (rule arg_cong[where f="λz. kf_apply z"])
      apply (rule kf_comp_dependent_cong_left)
      using bdd_filter by auto
    also have  = kf_apply (kf_comp (kf_filter (λx. xE) (𝔈 f))
       (kf_filter (λx. x=f) 𝔉))
      by (simp add: kf_comp_def)
    finally show ?thesis
      by -
  qed

  have rew_𝔈: kf_apply (kf_filter (λx. xE) (𝔈 f))
      = kf_apply (kf_filter (λx. xE') (𝔈' f))
    if f  kf_domain 𝔉
    using assms(3)[OF that]
    by (simp add: kf_apply_on_def)
  have rew_𝔉: kf_apply (kf_filter (λf'. f' = f) 𝔉')
     = kf_apply (kf_filter (λf'. f' = f) 𝔉)
    using assms(4)
    by (simp add: kf_apply_on_def)
  have 𝔉_0: kf_apply (kf_filter (λf'. f' = f) 𝔉) = 0
    if f  kf_domain 𝔉
  proof -
(*     have *: ‹(x = f ∧ x ∈ kf_domain 𝔉) ⟷ False› for x
      using that by simp *)

    have kf_filter (λf'. f' = f) 𝔉 kr
             kf_filter (λf'. f' = f) (kf_filter (λx. x  kf_domain 𝔉) 𝔉)
      by (auto intro!: kf_filter_cong intro: kf_eq_sym simp: kf_filter_to_domain)
    also have  kr (kf_filter (λ_. False) 𝔉)
      using that apply (simp add: kf_filter_twice del: kf_filter_false)
      by (smt (verit) kf_eq_def kf_filter_cong_eq)
    also have  kr 0
      by (simp add: kf_filter_false)
    finally show ?thesis
      by (metis kf_apply_0 kf_eq_imp_eq_weak kf_eq_weak_def)
  qed

  show kf_comp_dependent 𝔈 𝔉 *kr @({f} × E) ρ =
    kf_comp_dependent 𝔈' 𝔉' *kr @({f} × E') ρ
    apply (cases f  kf_domain 𝔉)
    by (auto intro!: ext simp add: rewrite_comp[OF bdd] rewrite_comp[OF bdd']
        kf_comp_apply rew_𝔈 rew_𝔉 𝔉_0)
qed



lemma kf_comp_dependent_cong_weak:
  fixes 𝔈 :: 'f  ('b::chilbert_space,'c::chilbert_space,'e1) kraus_family
    and 𝔈' :: 'f  ('b::chilbert_space,'c::chilbert_space,'e2) kraus_family
    and 𝔉 𝔉' :: ('a::chilbert_space,'b::chilbert_space,'f) kraus_family
  assumes bdd: bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  assumes eq: x. x  kf_domain 𝔉  𝔈 x =kr 𝔈' x
  assumes 𝔉 kr 𝔉'
  shows kf_comp_dependent 𝔈 𝔉 =kr kf_comp_dependent 𝔈' 𝔉'
proof -
  have kf_apply_on (kf_comp_dependent 𝔈 𝔉) ({f}×UNIV) = kf_apply_on (kf_comp_dependent 𝔈' 𝔉') ({f}×UNIV) for f
  proof -
    note bdd
    moreover have bdd_above ((kf_norm  𝔈') ` kf_domain 𝔉')
      by (metis (no_types, lifting) assms(1) assms(2) assms(3) comp_apply image_cong kf_norm_cong kf_domain_cong)
    moreover have kf_apply_on (𝔈 x) UNIV = kf_apply_on (𝔈' x) UNIV if x  kf_domain 𝔉 for x
      using assms(2) kf_eq_weak_def that
      by (metis kf_apply_on_UNIV)
    moreover have kf_apply_on 𝔉 {f} = kf_apply_on 𝔉' {f}
      by (meson assms(3) kf_eq_def)
    ultimately show ?thesis
      by (rule kf_apply_comp_dependent_cong)
  qed
  then have kf_apply_on (kf_comp_dependent 𝔈 𝔉) (f. {f}×UNIV) = kf_apply_on (kf_comp_dependent 𝔈' 𝔉') (f. {f}×UNIV)
    apply (rule_tac ext)
    apply (rule kf_apply_on_union_eqI[where F=range (λf. ({f}×UNIV,{f}×UNIV))])
    by auto
  moreover have (f. {f} × UNIV) = UNIV
    by fast
  ultimately show ?thesis
    by (metis kf_eq_weak_def kf_apply_on_UNIV)
qed


lemma kf_comp_dependent_assoc: 
  fixes 𝔈 :: 'g  'f  ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: 'g  ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  assumes bdd_E: bdd_above ((kf_norm o case_prod 𝔈) ` (SIGMA x:kf_domain 𝔊. kf_domain (𝔉 x)))
  assumes bdd_F: bdd_above ((kf_norm o 𝔉) ` kf_domain 𝔊)
  shows (kf_comp_dependent (λg. kf_comp_dependent (𝔈 g) (𝔉 g)) 𝔊) kr
  kf_map (λ((g,f),e). (g,f,e)) (kf_comp_dependent (λ(g,f). 𝔈 g f) (kf_comp_dependent 𝔉 𝔊))
    (is ?lhs kr ?rhs)
proof (rule kf_eqI)
  fix gfe :: 'g × 'f × 'e and ρ
  obtain g f e where gfe_def: gfe = (g,f,e)
    apply atomize_elim
    apply (rule exI[of _ fst gfe])
    apply (rule exI[of _ fst (snd gfe)])
    apply (rule exI[of _ snd (snd gfe)])
    by simp
  have aux: (λx. (fst (fst x), snd (fst x), snd x) = gfe) = (λx. x=((g,f),e))
    by (auto simp: gfe_def)
  have bdd1: bdd_above ((λx. kf_norm (kf_filter (λx. x = f) (𝔉 x))) ` kf_domain 𝔊)
    using kf_norm_filter bdd_F
    by (metis (mono_tags, lifting) bdd_above_mono2 o_apply order.refl)
  from bdd_E have bdd2: bdd_above ((kf_norm  𝔈 g) ` kf_domain (𝔉 g)) if g  kf_domain 𝔊 for g
    apply (rule bdd_above_mono)
    using that
    by (force simp: image_iff)
  from bdd_E have bdd3: bdd_above ((λx. kf_norm (kf_filter (λx. x = e) (case_prod 𝔈 x))) ` (SIGMA x:kf_domain 𝔊. kf_domain (𝔉 x)))
    apply (rule bdd_above_mono2)
    by (auto simp: kf_norm_filter)
  then have bdd4: bdd_above ((λx. kf_norm (kf_filter (λx. x = e) (𝔈 (fst x) (snd x)))) ` X)
    if X  (SIGMA x:kf_domain 𝔊. kf_domain (𝔉 x)) for X
    apply (rule bdd_above_mono2)
    using that by (auto simp: bdd_above_def)
  have bdd5: bdd_above ((kf_norm  (λg. kf_comp_dependent (𝔈 g) (𝔉 g))) ` X) 
    if X𝔊: X  kf_domain 𝔊 for X
  proof -
    from bdd_E
    obtain BE' where BE': g  kf_domain 𝔊  x  kf_domain (𝔉 g)  kf_norm (𝔈 g x)  BE' for g x
      by (auto simp: bdd_above_def)
    define BE where BE = max BE' 0
    from BE' have BE: g  kf_domain 𝔊  x  kf_domain (𝔉 g)  kf_norm (𝔈 g x)  BE and BE  0 for g x
      by (force simp: BE_def)+
    then have BE  0
      by (smt (z3) kf_norm_geq0)
    from bdd_F obtain BF where BF: kf_norm (𝔉 x)  BF if x  kf_domain 𝔊 for x
      by (auto simp: bdd_above_def)
    have kf_norm (kf_comp_dependent (𝔈 g) (𝔉 g))
                   BE * kf_norm (𝔉 g) if g  kf_domain 𝔊 for g
      apply (rule kf_comp_dependent_norm_leq[OF BE BE  0])
      using that by auto
    then have kf_norm (kf_comp_dependent (𝔈 g) (𝔉 g))  BE * BF if g  kf_domain 𝔊 for g
      apply (rule order_trans)
      using BF BE  0 that
      by (auto intro!: mult_left_mono)
    with X𝔊 show ?thesis
      by (auto intro!: bdd_aboveI)
  qed

  have ?lhs *kr @{gfe} ρ
       = kf_comp_dependent 
            (λg. kf_filter (λx. x=(f,e)) (kf_comp_dependent (𝔈 g) (𝔉 g)))
            (kf_filter (λx. x=g) 𝔊) *kr ρ
    unfolding kf_apply_on_def
    apply (subst kf_filter_comp_dependent[symmetric])
     apply (rule bdd5, rule order_refl)
    apply (subst asm_rl[of (λ(x, y). y = (f, e)  x = g) = (λx. x  {gfe})])
    by (auto simp: gfe_def)
  also have  = kf_comp_dependent
            (λg. kf_comp_dependent
                       (λf. kf_filter (λx. x=e) (𝔈 g f)) (kf_filter (λx. x=f) (𝔉 g)))
            (kf_filter (λx. x=g) 𝔊) *kr ρ
    apply (rule kf_apply_eqI)
    apply (rule kf_comp_dependent_cong_weak[OF _ _ kf_eq_refl])
     apply fastforce
    apply (subst kf_filter_comp_dependent[symmetric])
    using bdd2 apply fastforce
    apply (subst asm_rl[of (λ(ea, fa). fa = e  ea = f) = (λx. x = (f, e))])
    by auto
  also have  = kf_comp_dependent
            (λ_. kf_comp_dependent
                       (λf. kf_filter (λx. x=e) (𝔈 g f)) (kf_filter (λx. x=f) (𝔉 g)))
            (kf_filter (λx. x=g) 𝔊) *kr ρ
    apply (rule arg_cong[where f=λt. kf_apply t ρ])
    apply (rule kf_comp_dependent_cong_left)
    by (auto intro!: simp: kf_domain.rep_eq kf_filter.rep_eq)
  also have  = kf_comp_dependent
            (λ_. kf_comp_dependent
                       (λ_. kf_filter (λx. x=e) (𝔈 g f)) (kf_filter (λx. x=f) (𝔉 g)))
            (kf_filter (λx. x=g) 𝔊) *kr ρ
    apply (rule arg_cong[where f=λt. kf_comp_dependent t _ *kr ρ])
    apply (rule ext)
    apply (rule kf_comp_dependent_cong_left)
    by (auto intro!: simp: kf_domain.rep_eq kf_filter.rep_eq)
  also have  = kf_comp
            (kf_comp
                       (kf_filter (λx. x=e) (𝔈 g f)) (kf_filter (λx. x=f) (𝔉 g)))
            (kf_filter (λx. x=g) 𝔊) *kr ρ
    by (simp add: kf_comp_def)
  also have  = kf_comp (kf_filter (λx. x=e) (𝔈 g f))
                          (kf_comp (kf_filter (λx. x=f) (𝔉 g)) (kf_filter (λx. x=g) 𝔊)) *kr ρ
    by (simp add: kf_comp_assoc_weak[unfolded kf_eq_weak_def])
  also have  = kf_comp_dependent (λ_. kf_filter (λx. x=e) (𝔈 g f))
            (kf_comp_dependent (λ_. kf_filter (λx. x=f) (𝔉 g))
                                         (kf_filter (λx. x=g) 𝔊)) *kr ρ
    by (simp add: kf_comp_def)
  also have  = kf_comp_dependent (λ(g,f). kf_filter (λx. x=e) (𝔈 g f))
            (kf_comp_dependent (λ_. kf_filter (λx. x=f) (𝔉 g))
                                         (kf_filter (λx. x=g) 𝔊)) *kr ρ
    apply (rule arg_cong[where f=λt. t *kr ρ])
    apply (rule kf_comp_dependent_cong_left)
    apply force
    using kf_domain_comp_dependent_subset apply (fastforce intro!: bdd4 simp: o_def case_prod_unfold)
    using kf_domain_comp_dependent_subset[of (λ_. kf_filter (λx. x = f) (𝔉 g)) kf_filter (λx. x = g) 𝔊]
    by (auto intro!: simp: kf_filter.rep_eq case_prod_unfold)
  also have  = kf_comp_dependent (λ(g,f). kf_filter (λx. x=e) (𝔈 g f))
            (kf_comp_dependent (λg. kf_filter (λx. x=f) (𝔉 g))
                                         (kf_filter (λx. x=g) 𝔊)) *kr ρ
    apply (rule arg_cong[where f=λt. kf_comp_dependent _ t *kr _])
    apply (rule kf_comp_dependent_cong_left)
    by (auto intro!: simp: kf_domain.rep_eq kf_filter.rep_eq)
  also have  = kf_comp_dependent (λ(g,f). kf_filter (λx. x=e) (𝔈 g f))
            (kf_filter (λx. x=(g,f)) (kf_comp_dependent 𝔉 𝔊)) *kr ρ
    apply (subst kf_filter_comp_dependent[symmetric])
    using bdd_F apply (simp add: bdd_above_mono2)
    apply (subst asm_rl[of (λ(e, fa). fa = f  e = g) = (λx. x = (g, f))])
    by auto
  also have  = kf_filter (λx. x=((g,f),e))
       (kf_comp_dependent (λ(g,f). 𝔈 g f) (kf_comp_dependent 𝔉 𝔊)) *kr ρ
    unfolding case_prod_beta
    apply (subst kf_filter_comp_dependent[symmetric])
    using bdd_E kf_domain_comp_dependent_subset[of 𝔉 𝔊] apply (fastforce simp: bdd_above_def)
    apply (subst asm_rl[of (λ(ea, fa). fa = e  ea = (g, f)) = (λx. x = ((g, f), e))])
    by auto
  also have  = kf_filter (λx. x=gfe)
        (kf_map (λ((g, f), e). (g, f, e))
       (kf_comp_dependent (λ(g,f). 𝔈 g f) (kf_comp_dependent 𝔉 𝔊))) *kr ρ
    apply (simp add: kf_filter_map case_prod_beta aux)
    apply (subst aux)
    by simp
  also have  = ?rhs *kr @{gfe} ρ
    by (simp add: kf_apply_on_def)
  finally show ?lhs *kr @{gfe} ρ = ?rhs *kr @{gfe} ρ
    by -
qed

lemma kf_comp_dependent_assoc_weak:
  fixes 𝔈 :: 'g  'f  ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: 'g  ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  assumes bdd_E: bdd_above ((kf_norm o case_prod 𝔈) ` (SIGMA x:kf_domain 𝔊. kf_domain (𝔉 x)))
  assumes bdd_F: bdd_above ((kf_norm o 𝔉) ` kf_domain 𝔊)
  shows kf_comp_dependent (λg. kf_comp_dependent (𝔈 g) (𝔉 g)) 𝔊 =kr
         kf_comp_dependent (λ(g,f). 𝔈 g f) (kf_comp_dependent 𝔉 𝔊)
  using kf_comp_dependent_assoc[OF assms, THEN kf_eq_imp_eq_weak]
  by (metis (no_types, lifting) kf_apply_map kf_eq_weak_def)

lemma kf_comp_dependent_comp_assoc_weak:
  fixes 𝔈 :: ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: 'g  ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  assumes bdd_above ((kf_norm o 𝔉) ` kf_domain 𝔊)
  shows kf_comp_dependent (λg. kf_comp 𝔈 (𝔉 g)) 𝔊 =kr
         kf_comp 𝔈 (kf_comp_dependent 𝔉 𝔊)
  using kf_comp_dependent_assoc_weak[where 𝔈=λ_ _. 𝔈 and 𝔊=𝔊, OF _ assms]
  by (fastforce simp: case_prod_unfold kf_comp_def)

lemma kf_comp_comp_dependent_assoc_weak:
  fixes 𝔈 :: 'f  ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  assumes bdd_E: bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  shows kf_comp (kf_comp_dependent 𝔈 𝔉) 𝔊 =kr
         kf_comp_dependent (λ(_,f). 𝔈 f) (kf_comp 𝔉 𝔊)
proof -
  from bdd_E have bdd_above ((kf_norm  (λ(g, y). 𝔈 y)) ` (kf_domain 𝔊 × kf_domain 𝔉))
    by (auto intro!: simp: bdd_above_def)
  then have kf_comp_dependent (λ_. kf_comp_dependent 𝔈 𝔉) 𝔊 =kr kf_comp_dependent (λ(_, f). 𝔈 f) (kf_comp_dependent (λ_. 𝔉) 𝔊)
    apply (rule kf_comp_dependent_assoc_weak)
    by auto
  then show ?thesis
    by (simp add: case_prod_unfold kf_comp_def)
qed

lemma kf_comp_assoc:
  fixes 𝔈 :: ('c::chilbert_space,'d::chilbert_space,'e) kraus_family
    and 𝔉 :: ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔊 :: ('a::chilbert_space,'b::chilbert_space,'g) kraus_family
  shows kf_comp (kf_comp 𝔈 𝔉) 𝔊 kr
   kf_map (λ((g,f),e). (g,f,e)) (kf_comp 𝔈 (kf_comp 𝔉 𝔊))
  apply (simp add: kf_comp_def)
  apply (rule kf_eq_trans)
   apply (rule kf_comp_dependent_assoc)
  by (auto simp: case_prod_unfold)

lemma kf_comp_dependent_cong:
  fixes 𝔈 𝔈' :: 'f  ('b::chilbert_space,'c::chilbert_space,'e) kraus_family
    and 𝔉 𝔉' :: ('a::chilbert_space,'b::chilbert_space,'f) kraus_family
  assumes bdd: bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  assumes x. x  kf_domain 𝔉  𝔈 x kr 𝔈' x
  assumes 𝔉 kr 𝔉'
  shows kf_comp_dependent 𝔈 𝔉 kr kf_comp_dependent 𝔈' 𝔉'
proof (rule kf_eqI)
  fix ρ :: ('a, 'a) trace_class
  fix x :: 'f × 'e

  note bdd
  moreover have bdd': bdd_above ((kf_norm  𝔈') ` kf_domain 𝔉')
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) image_cong kf_eq_imp_eq_weak kf_norm_cong kf_domain_cong o_apply)
  moreover have kf_apply_on (𝔈 xa) {snd x} = kf_apply_on (𝔈' xa) {snd x} if xa  kf_domain 𝔉 for xa
    by (meson assms(2) kf_eq_def that)
  moreover have kf_apply_on 𝔉 {fst x} = kf_apply_on 𝔉' {fst x}
    by (meson assms(3) kf_eq_def)
  ultimately have kf_apply_on (kf_comp_dependent 𝔈 𝔉) ({fst x} × {snd x}) = kf_apply_on (kf_comp_dependent 𝔈' 𝔉') ({fst x} × {snd x})
    by (rule kf_apply_comp_dependent_cong)
  then show kf_comp_dependent 𝔈 𝔉 *kr @{x} ρ = kf_comp_dependent 𝔈' 𝔉' *kr @{x} ρ
    by simp
qed


lemma kf_comp_cong:
  fixes 𝔈 𝔈' :: ('b::chilbert_space,'c::chilbert_space,'e) kraus_family
    and 𝔉 𝔉' :: ('a::chilbert_space,'b::chilbert_space,'f) kraus_family
  assumes 𝔈 kr 𝔈'
  assumes 𝔉 kr 𝔉'
  shows kf_comp 𝔈 𝔉 kr kf_comp 𝔈' 𝔉'
  by (auto intro!: kf_comp_dependent_cong assms
      simp add: kf_comp_def)

lemma kf_bound_comp_dependent_raw_of_op:
  shows kf_bound (kf_comp_dependent_raw 𝔈 (kf_of_op U))
       = sandwich (U*) (kf_bound (𝔈 ()))
proof -
  write compose_wot (infixl "oW" 55)
  define EE where EE = Rep_kraus_family (𝔈 ())

  have sum1: summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) EE
    by (simp add: EE_def kf_bound_summable)
  then have sum2: summable_on_in cweak_operator_topology (λ(E,x). U* oCL (E* oCL E)) EE
    by (simp add: case_prod_unfold summable_on_in_wot_compose_left)

  have kf_bound (kf_comp_dependent_raw 𝔈 (kf_of_op U)) = 
        infsum_in cweak_operator_topology (λ(E, x). E* oCL E)
          (Set.filter (λ(EF,_). EF  0) (((λ((F, y), (E, x)). ((E oCL F, F, E, y, x))) ` (SIGMA (F, y):if U=0 then {} else {(U, ())}. EE))))
    by (simp add: kf_bound.rep_eq kf_comp_dependent_raw.rep_eq kf_of_op.rep_eq EE_def)
  also have  = infsum_in cweak_operator_topology (λ(E, x). E* oCL E)
          ((λ((F, y), (E, x)). ((E oCL F, F, E, y, x))) ` (SIGMA (F, y):{(U, ())}. EE))
    apply (cases U=0; rule infsum_in_cong_neutral)
    by force+
  also have   = infsum_in cweak_operator_topology (λ(E, x). E* oCL E)
                   ((λ(E,x). (E oCL U, U, E, (), x)) ` EE)
    apply (rule arg_cong[where f=infsum_in _ _])
    by force
  also have  = infsum_in cweak_operator_topology (λ(E, x). (E oCL U)* oCL (E oCL U)) EE
    apply (subst infsum_in_reindex)
    by (auto intro!: inj_onI simp: o_def case_prod_unfold infsum_in_reindex)
  also have  = infsum_in cweak_operator_topology (λ(E, x). U* oCL (E* oCL E) oCL U) EE
    by (metis (no_types, lifting) adj_cblinfun_compose cblinfun_assoc_left(1))
  also have  = U* oCL infsum_in cweak_operator_topology (λ(E,x). E* oCL E) EE oCL U
    using sum1 sum2 by (simp add: case_prod_unfold infsum_in_wot_compose_right infsum_in_wot_compose_left)
  also have  = sandwich (U*) (kf_bound (𝔈 ()))
    by (simp add: EE_def kf_bound.rep_eq sandwich_apply)
  finally show ?thesis
    by -
qed


lemma kf_bound_comp_dependent_of_op:
  shows kf_bound (kf_comp_dependent 𝔈 (kf_of_op U)) = sandwich (U*) (kf_bound (𝔈 ()))
  by (simp add: kf_comp_dependent_def kf_map_bound kf_bound_comp_dependent_raw_of_op)

lemma kf_bound_comp_of_op:
  shows kf_bound (kf_comp 𝔈 (kf_of_op U)) = sandwich (U*) (kf_bound 𝔈)
  by (simp add: kf_bound_comp_dependent_of_op kf_comp_def)

lemma kf_norm_comp_dependent_of_op_coiso:
  assumes isometry (U*)
  shows kf_norm (kf_comp_dependent 𝔈 (kf_of_op U)) = kf_norm (𝔈 ())
  using assms
  by (simp add: kf_bound_comp_dependent_of_op kf_norm_def sandwich_apply
      norm_isometry_compose norm_isometry_compose')


lemma kf_norm_comp_of_op_coiso:
  assumes isometry (U*)
  shows kf_norm (kf_comp 𝔈 (kf_of_op U)) = kf_norm 𝔈
  using assms
  by (simp add: kf_bound_comp_of_op kf_norm_def sandwich_apply
      norm_isometry_compose norm_isometry_compose')

lemma kf_bound_comp_dependend_raw_iso:
  assumes isometry U
  shows kf_bound (kf_comp_dependent_raw (λ_. kf_of_op U) 𝔈)
       = kf_bound 𝔈
proof -
  write compose_wot (infixl "oW" 55)
  define EE where EE = Rep_kraus_family 𝔈

  have bdd_above ((λx. (norm U)2) ` kf_domain 𝔈)
    by auto
  then have kf_bound (kf_comp_dependent_raw (λ_. kf_of_op U) 𝔈) = 
        infsum_in cweak_operator_topology (λ(E, x). E* oCL E)
           (Set.filter (λ(EF, _). EF  0) ((λ((F, y), (E, x)). (E oCL F, F, E, y, ())) ` (SIGMA (F, y):EE. if U=0 then {} else {(U, ())})))
    by (simp add: kf_bound.rep_eq kf_comp_dependent_raw.rep_eq kf_of_op.rep_eq EE_def)
  also have  = infsum_in cweak_operator_topology (λ(E, x). E* oCL E)
           ((λ((F, y), (E, x)). (E oCL F, F, E, y, ())) ` (SIGMA (F, y):EE. {(U, ())}))
    apply (cases U = 0; rule infsum_in_cong_neutral)
    by force+
  also have   = infsum_in cweak_operator_topology (λ(E, x). E* oCL E)
                   ((λ(E,x). (U oCL E, E, U, x, ())) ` EE)
    apply (rule arg_cong[where f=infsum_in _ _])
    by force
  also have  = infsum_in cweak_operator_topology (λ(E, x). (U oCL E)* oCL (U oCL E)) EE
    apply (subst infsum_in_reindex)
    by (auto intro!: inj_onI simp: o_def case_prod_unfold infsum_in_reindex)
  also have  = infsum_in cweak_operator_topology (λ(E, x). E* oCL (U* oCL U) oCL E) EE
    by (metis (no_types, lifting) adj_cblinfun_compose cblinfun_assoc_left(1))
  also have  = infsum_in cweak_operator_topology (λ(E,x). E* oCL E) EE
    using assms by simp
  also have  = kf_bound 𝔈
    by (simp add: EE_def kf_bound.rep_eq sandwich_apply)
  finally show ?thesis
    by -
qed

lemma kf_bound_comp_dependent_iso:
  assumes isometry U
  shows kf_bound (kf_comp_dependent (λ_. kf_of_op U) 𝔈) = kf_bound 𝔈
  using assms by (simp add: kf_comp_dependent_def kf_map_bound kf_bound_comp_dependend_raw_iso)

lemma kf_bound_comp_iso:
  assumes isometry U
  shows kf_bound (kf_comp (kf_of_op U) 𝔈) = kf_bound 𝔈
  using assms by (simp add: kf_bound_comp_dependent_iso kf_comp_def)

lemma kf_norm_comp_dependent_iso:
  assumes isometry U
  shows kf_norm (kf_comp_dependent (λ_. kf_of_op U) 𝔈) = kf_norm 𝔈
  using assms
  by (simp add: kf_bound_comp_dependent_iso kf_norm_def sandwich_apply
      norm_isometry_compose norm_isometry_compose')

lemma kf_norm_comp_iso:
  assumes isometry U
  shows kf_norm (kf_comp (kf_of_op U) 𝔈) = kf_norm 𝔈
  using assms
  by (simp add: kf_bound_comp_iso kf_norm_def sandwich_apply
      norm_isometry_compose norm_isometry_compose')


lemma kf_comp_dependent_raw_0_right[simp]: kf_comp_dependent_raw 𝔈 0 = 0
  apply transfer'
  by (auto intro!: simp: zero_kraus_family.rep_eq)

lemma kf_comp_dependent_raw_0_left[simp]: kf_comp_dependent_raw 0 𝔈 = 0
  apply transfer'
  by (auto intro!: simp: zero_kraus_family.rep_eq)

lemma kf_comp_dependent_0_left[simp]: kf_comp_dependent (λ_. 0) E = 0
proof -
  have bdd_above ((kf_norm  0) ` kf_domain E)
    by auto
  then have kf_comp_dependent 0 E =kr 0
    by (auto intro!: ext simp: kf_eq_weak_def kf_comp_dependent_apply split_def)
  then have (kf_comp_dependent 0 E) = 0
    using kf_eq_0_iff_eq_0 by auto
  then show kf_comp_dependent (λ_. 0) E = 0
    by (simp add: kf_comp_dependent_def zero_fun_def)
qed

lemma kf_comp_dependent_0_right[simp]: kf_comp_dependent E 0 = 0
  by (auto intro!: ext simp add: kf_eq_weak_def kf_comp_dependent_def)

lemma kf_comp_0_left[simp]: kf_comp 0 E = 0
  using kf_comp_dependent_0_left[of E]
  by (simp add: kf_comp_def zero_fun_def)

lemma kf_comp_0_right[simp]: kf_comp E 0 = 0
  using kf_comp_dependent_0_right[of λ_. E]
  by (simp add: kf_comp_def)

lemma kf_filter_comp:
  fixes 𝔉 :: ('b::chilbert_space,'c::chilbert_space,'f) kraus_family
    and 𝔈 :: ('a::chilbert_space,'b::chilbert_space,'e) kraus_family
  shows kf_filter (λ(e,f). F f  E e) (kf_comp 𝔉 𝔈)
      = kf_comp (kf_filter F 𝔉) (kf_filter E 𝔈)
  unfolding kf_comp_def
  apply (rule kf_filter_comp_dependent)
  by auto

lemma kf_comp_dependent_invalid:
  assumes ¬ bdd_above ((kf_norm o 𝔈) ` kf_domain 𝔉)
  shows kf_comp_dependent 𝔈 𝔉 = 0
  by (metis (no_types, lifting) Rep_kraus_family_inject assms kf_comp_dependent_def kf_comp_dependent_raw.rep_eq kf_map0 zero_kraus_family.rep_eq)

lemma kf_comp_dependent_map_left:
  kf_comp_dependent (λx. kf_map (f x) (E x)) F
     kr kf_map (λ(x,y). (x, f x y)) (kf_comp_dependent E F)
proof (cases bdd_above ((kf_norm  E) ` kf_domain F))
  case True
  show ?thesis
  proof (rule kf_eqI)
    fix xy :: 'c × 'd and ρ
    obtain x y where xy: xy = (x, y)
      by force
    define F' where F' x = kf_filter (λx'. x' = x) F for x
    define E'f where E'f y e = kf_filter (λx. f e x = y) (E e) for e y
    have bdd2: bdd_above ((λx. kf_norm (E'f y x)) ` kf_domain (F' x))
      apply (simp add: E'f_def F'_def)
      by fastforce
    have kf_comp_dependent (λx. kf_map (f x) (E x)) F *kr @{xy} ρ
        = kf_filter (λ(x',y'). y'=y  x'=x) (kf_comp_dependent (λx. kf_map (f x) (E x)) F) *kr ρ
      (is ?lhs = _)
      apply (simp add: kf_apply_on_def xy case_prod_unfold)
      by (metis fst_conv prod.collapse snd_conv)
    also have  = kf_apply (kf_comp_dependent (λe. kf_filter (λy'. y' = y)
         (kf_map (f e) (E e))) (F' x)) ρ
      using True by (simp add: kf_filter_comp_dependent F'_def)
    also have  = kf_apply (kf_comp_dependent
                 (λe. kf_map (f e) (E'f y e)) (F' x)) ρ
      by (simp add: kf_filter_map E'f_def)
    also have  = kf_apply (kf_comp_dependent (E'f y) (F' x)) ρ
      apply (rule kf_apply_eqI)
      apply (rule kf_comp_dependent_cong_weak)
      by (simp_all add: bdd2 kf_eq_weak_def)
    also have  = kf_apply 
       (kf_filter (λ(x',y'). f x' y' = y  x' = x) (kf_comp_dependent E F)) ρ
      apply (subst kf_filter_comp_dependent)
      using True by (simp_all add: o_def F'_def E'f_def[abs_def])
    also have  = kf_apply (kf_map (λ(x,y). (x, f x y))
       (kf_filter (λ(x',y'). f x' y' = y  x' = x) (kf_comp_dependent E F))) ρ
      by simp
    also have 
      = kf_apply (kf_filter (λ(x',y'). y'=y  x'=x)
       (kf_map (λ(x,y). (x, f x y)) (kf_comp_dependent E F))) ρ
      by (simp add: kf_filter_map case_prod_unfold)
    also have  = kf_map (λ(x,y). (x, f x y)) (kf_comp_dependent E F) *kr @{xy} ρ
      apply (simp add: kf_apply_on_def xy case_prod_unfold)
      by (metis fst_conv prod.collapse snd_conv)
    finally show ?lhs = 
      by -
  qed
next
  case False
  then show ?thesis
    by (simp add: kf_comp_dependent_invalid)
qed

(* lemma kf_comp_dependent_map_left_weak:
  ‹kf_comp_dependent (λx. kf_map (f x) (E x)) F
     =kr kf_comp_dependent E F›
by - *)

lemma kf_comp_dependent_map_right:
  kf_comp_dependent E (kf_map f F)
     kr kf_map (λ(x,y). (f x, y)) (kf_comp_dependent (λx. E (f x)) F)
proof (cases bdd_above ((kf_norm  E) ` kf_domain (kf_map f F)))
  case True
  show ?thesis
  proof (rule kf_eqI)
    fix xy :: 'c × 'd and ρ
    obtain x y where xy: xy = (x, y)
      by force
    define F'f where F'f x = kf_filter (λxa. f xa = x) F for x
    define E' where E' y e = kf_filter (λy'. y' = y) (E e) for e y
    have bdd2: bdd_above ((kf_norm  E' y) ` kf_domain (kf_map f (F'f x)))
      apply (simp add: E'_def F'f_def)
      by fastforce
    have bdd3: bdd_above ((kf_norm  (λx. E (f x))) ` kf_domain F)
      by (metis (no_types, lifting) ext True comp_apply image_comp kf_domain_map)
    have bdd4: bdd_above ((kf_norm  (λ_. E' y x)) ` kf_domain (F'f x))
      by fastforce                                   
    have kf_comp_dependent E (kf_map f F) *kr @{xy} ρ
        = kf_filter (λ(x',y'). y'=y  x'=x) (kf_comp_dependent E (kf_map f F)) *kr ρ
      (is ?lhs = _)
      apply (simp add: kf_apply_on_def xy case_prod_unfold)
      by (metis fst_conv prod.collapse snd_conv)
    also have  = kf_apply (kf_comp_dependent (E' y) (kf_filter (λx'. x' = x) (kf_map f F))) ρ
      using True by (simp add: kf_filter_comp_dependent F'f_def E'_def[abs_def])
    also have  = kf_apply (kf_comp_dependent
                 (E' y) (kf_map f (F'f x))) ρ
      by (simp add: kf_filter_map F'f_def)
    also have  = kf_apply (kf_comp (E' y x) (kf_map f (F'f x))) ρ
      unfolding kf_comp_def
      apply (rule kf_apply_eqI)
      using bdd2 apply (rule kf_comp_dependent_cong_weak)
      by (auto simp: F'f_def)
    also have  = kf_apply (E' y x) (kf_apply (F'f x) ρ)
      by (simp add: kf_comp_apply)
    also have  = kf_apply (kf_comp (E' y x) (F'f x)) ρ
      by (simp add: kf_comp_apply)
    also have  = kf_apply (kf_comp_dependent (λe. kf_filter (λy'. y' = y) (E (f e))) (F'f x)) ρ
      unfolding kf_comp_def
      apply (rule kf_apply_eqI)
      using bdd4 apply (rule kf_comp_dependent_cong_weak)
      by (auto intro!: simp: F'f_def E'_def)
    also have  = kf_apply (kf_filter (λ(x',y'). y' = y  f x' = x) (kf_comp_dependent (λx. E (f x)) F)) ρ
      using bdd3 by (simp add: kf_filter_comp_dependent F'f_def[abs_def] E'_def[abs_def])
    also have  = kf_apply (kf_filter (λ(x',y'). y'=y  x'=x)
              (kf_map (λ(x, y). (f x, y)) (kf_comp_dependent (λx. E (f x)) F))) ρ
      by (simp add: kf_filter_map case_prod_unfold)
    also have  = kf_map (λ(x, y). (f x, y)) (kf_comp_dependent (λx. E (f x)) F) *kr @{xy} ρ
      apply (simp add: kf_apply_on_def xy case_prod_unfold)
      by (metis fst_conv prod.collapse snd_conv)

    finally show ?lhs = 
      by -
  qed
next
  case False
  have not_bdd2: ¬ bdd_above ((kf_norm  (λx. E (f x))) ` kf_domain F)
    by (metis (no_types, lifting) False comp_apply image_comp image_cong kf_domain_map)
  show ?thesis
    using False not_bdd2
    by (simp add: kf_comp_dependent_invalid)
qed

lemma kf_comp_dependent_raw_map_inj_right:
  kf_comp_dependent_raw E (kf_map_inj f F)
     = kf_map_inj (λ(E,F,x,y). (E, F, f x, y)) (kf_comp_dependent_raw (λx. E (f x)) F)
proof -
  have (λ((F, y), (E, x)). (E oCL F, F, E, y, x)) ` (SIGMA (F, y):Rep_kraus_family (kf_map_inj f F). Rep_kraus_family (E y)) =
           (λ((F, y),(E,x)). (E oCL F, F, E, f y, x)) ` (SIGMA (F, y):Rep_kraus_family F. Rep_kraus_family (E (f y)))
    by (auto intro!: image_eqI simp: Sigma_image_left kf_map_inj.rep_eq)
  then show ?thesis
    apply (transfer' fixing: f)
    by (simp add: image_image case_prod_unfold filter_image)
qed

lemma kf_comp_dependent_map_inj_right:
  assumes inj_on f (kf_domain F)
  shows kf_comp_dependent E (kf_map_inj f F)
     = kf_map_inj (λ(x,y). (f x, y)) (kf_comp_dependent (λx. E (f x)) F)
proof -
  have dom: kf_domain (kf_comp_dependent_raw (λx. E (f x)) F)  UNIV × UNIV × kf_domain F × UNIV
  proof -
    have kf_domain (kf_comp_dependent_raw (λx. E (f x)) F)  UNIV × UNIV × (SIGMA x:kf_domain F. kf_domain (E (f x)))
      by (rule kf_domain_comp_dependent_raw_subset)
    also have   UNIV × UNIV × kf_domain F × UNIV
      by auto
    finally show ?thesis
      by -
  qed
  have inj2: inj_on (λ(E, F, x, y). (E, F, f x, y)) (kf_domain (kf_comp_dependent_raw (λx. E (f x)) F))
  proof -
    have inj_on (λ(E, F, x, y). (E, F, f x, y)) (UNIV × UNIV × kf_domain F × UNIV)
      using assms by (auto simp: inj_on_def)
    with dom show ?thesis
      by (rule subset_inj_on[rotated])
  qed
  have inj3: inj_on (λ(x, y). (f x, y)) ((λ(F, E, x). x) ` kf_domain (kf_comp_dependent_raw (λx. E (f x)) F))
  proof -
    from dom have ((λ(F, E, x). x) ` kf_domain (kf_comp_dependent_raw (λx. E (f x)) F))  kf_domain F × UNIV
      by auto
    moreover have inj_on (λ(x, y). (f x, y)) (kf_domain F × UNIV)
      using assms by (auto simp: o_def inj_on_def)
    ultimately show ?thesis
      by (rule subset_inj_on[rotated])
  qed
  have kf_comp_dependent E (kf_map_inj f F) = kf_map (λ(F, E, y). y) (kf_comp_dependent_raw E (kf_map_inj f F))
    by (simp add: kf_comp_dependent_def id_def)
  also have  = kf_map (λ(F,E,y). y) (kf_map_inj (λ(E,F,x,y). (E, F, f x, y)) (kf_comp_dependent_raw (λx. E (f x)) F))
    by (simp add: kf_comp_dependent_raw_map_inj_right)
  also have  = kf_map (λ(E,F,x,y). (f x, y)) (kf_comp_dependent_raw (λx. E (f x)) F)
    apply (subst kf_map_kf_map_inj_comp)
     apply (rule inj2)
    using assms by (simp add: o_def case_prod_unfold)
  also have  = kf_map_inj (λ(x, y). (f x, y)) (kf_map (λ(F, E, x). x) (kf_comp_dependent_raw (λx. E (f x)) F))
    apply (subst kf_map_inj_kf_map_comp)
     apply (rule inj3)
    by (simp add: o_def case_prod_unfold)
  also have  = kf_map_inj (λ(x,y). (f x, y)) (kf_comp_dependent (λx. E (f x)) F)
    by (simp add: kf_comp_dependent_def id_def)
  finally show ?thesis
    by -
qed

lemma kf_comp_dependent_map_right_weak:
  kf_comp_dependent E (kf_map f F)
     =kr kf_comp_dependent (λx. E (f x)) F
  by (smt (verit) kf_apply_eqI kf_apply_map kf_comp_dependent_cong kf_comp_dependent_invalid kf_comp_dependent_map_right kf_eq_def
      kf_eq_imp_eq_weak kf_eq_weakI)

lemma kf_comp_map_left:
  kf_comp (kf_map f E) F kr kf_map (λ(x,y). (x, f y)) (kf_comp E F)
  by (simp add: kf_comp_def kf_comp_dependent_map_left)

lemma kf_comp_map_right:
  kf_comp E (kf_map f F) kr kf_map (λ(x,y). (f x, y)) (kf_comp E F)
  using kf_comp_dependent_map_right[where E=λ_. E and f=f and F=F]
  by (simp add: kf_comp_def)

lemma kf_comp_map_both:
  kf_comp (kf_map e E) (kf_map f F) kr kf_map (λ(x,y). (f x, e y)) (kf_comp E F)
  apply (rule kf_comp_map_left[THEN kf_eq_trans])
  apply (rule kf_map_cong[THEN kf_eq_trans, OF refl])
  apply (rule kf_comp_map_right)
  apply (rule kf_map_twice[THEN kf_eq_trans])
  by (simp add: o_def case_prod_unfold)

lemma kf_apply_commute:
  assumes kf_operators 𝔉  commutant (kf_operators 𝔈)
  shows kf_apply 𝔉 o kf_apply 𝔈 = kf_apply 𝔈 o kf_apply 𝔉
proof (rule eq_from_separatingI[OF separating_density_ops[where B=1], rotated 3])
  show 0 < (1 :: real)
    by simp
  show clinear ((*kr) 𝔉  (*kr) 𝔈)
    by (simp add: clinear_compose)
  show clinear ((*kr) 𝔈  (*kr) 𝔉)
    using clinear_compose by blast
  fix t :: ('a, 'a) trace_class
  assume t  {t. 0  t  norm t  1}
  then have t  0
    by simp
  from assms
  have E  kf_operators 𝔈. F  kf_operators 𝔉. E oCL F = F oCL E
    unfolding commutant_def by auto
  then show ((*kr) 𝔉  (*kr) 𝔈) t = ((*kr) 𝔈  (*kr) 𝔉) t
  proof (transfer fixing: t, tactic FILTER (fn st => Thm.nprems_of st = 1) all_tac)
      (* The FILTER tactic forces the transfer method to produce subgoals where the transferring worked *)
    fix 𝔈 :: ('a CL 'a × 'c) set and 𝔉 :: ('a CL 'a × 'b) set
    assume 𝔉  Collect kraus_family and 𝔈  Collect kraus_family
    then have [iff]: kraus_family 𝔉  kraus_family 𝔈
      by auto
    assume comm: Efst ` 𝔈. Ffst ` 𝔉. E oCL F = F oCL E
    have sum1: (λE. sandwich_tc (fst E) t) summable_on 𝔈
      apply (rule abs_summable_summable)
      apply (rule kf_apply_abs_summable[unfolded case_prod_unfold])
      by simp
    have sum2: (λy. sandwich_tc a (sandwich_tc (fst y) t)) summable_on 𝔈 for a
      apply (rule summable_on_bounded_linear[where h=sandwich_tc _])
      by (simp_all add: bounded_clinear.bounded_linear bounded_clinear_sandwich_tc sum1)
    have sum3: (λF. sandwich_tc (fst F) t) summable_on 𝔉 for t
      apply (rule abs_summable_summable)
      apply (rule kf_apply_abs_summable[unfolded case_prod_unfold])
      by simp

    have ((λρ. F𝔉. sandwich_tc (fst F) ρ)  (λρ. E𝔈. sandwich_tc (fst E) ρ)) t
      = (F𝔉. E𝔈. sandwich_tc (fst F) (sandwich_tc (fst E) t))  (is ?lhs = _)
      apply (subst infsum_bounded_linear[where h=sandwich_tc _])
      by (simp_all add: bounded_clinear.bounded_linear bounded_clinear_sandwich_tc sum1)
    also have  = (E𝔈. F𝔉. sandwich_tc (fst F) (sandwich_tc (fst E) t))
      apply (rule infsum_swap_positive_tc)
      using t  0 by (simp_all add: sum2 sum3 sandwich_tc_pos)
    also have  = (E𝔈. F𝔉. sandwich_tc (fst E) (sandwich_tc (fst F) t))
      apply (intro infsum_cong)
      apply (subst sandwich_tc_compose[THEN fun_cong, unfolded o_def, symmetric])+
      using comm
      by simp
    also have  = ((λρ. F𝔈. sandwich_tc (fst F) ρ)  (λρ. E𝔉. sandwich_tc (fst E) ρ)) t
      apply (subst infsum_bounded_linear[where h=sandwich_tc _])
      by (simp_all add: bounded_clinear.bounded_linear bounded_clinear_sandwich_tc sum3)
    finally show ?lhs = 
      by -
  qed
qed

lemma kf_comp_commute_weak:
  assumes kf_operators 𝔉  commutant (kf_operators 𝔈)
  shows kf_comp 𝔉 𝔈 =kr kf_comp 𝔈 𝔉
  apply (rule kf_eq_weakI)
  apply (simp add: kf_comp_apply)
  using  kf_apply_commute[OF assms, unfolded o_def]
  by meson

lemma kf_comp_commute:
  assumes kf_operators 𝔉  commutant (kf_operators 𝔈)
  shows kf_comp 𝔉 𝔈 kr kf_map prod.swap (kf_comp 𝔈 𝔉)
proof (rule kf_eqI_from_filter_eq_weak)
  fix xy :: 'c × 'b
  obtain x y where xy: xy = (x,y)
    by (simp add: prod_eq_iff)
  have *: kf_operators (kf_filter ((=) y) 𝔉)  commutant (kf_operators (kf_filter ((=) x) 𝔈))
    using kf_operators_filter commutant_antimono[OF kf_operators_filter] assms
    by fastforce
  have kf_filter ((=) xy) (kf_comp 𝔉 𝔈) = kf_comp (kf_filter ((=) y) 𝔉) (kf_filter ((=) x) 𝔈)
    apply (simp add: xy prod_eq_iff[abs_def] case_prod_unfold flip: kf_filter_comp)
    by meson
  also have  =kr kf_comp (kf_filter ((=) x) 𝔈) (kf_filter ((=) y) 𝔉)
    apply (rule kf_comp_commute_weak)
    by (simp add: * )
  also have  = kf_filter ((=) (prod.swap xy)) (kf_comp 𝔈 𝔉)
    apply (simp add: xy prod_eq_iff[abs_def] case_prod_unfold flip: kf_filter_comp)
    by meson
  also have  =kr kf_filter ((=) xy) (kf_map prod.swap (kf_comp 𝔈 𝔉))
    apply (simp add: kf_filter_map)
    by (metis (mono_tags, lifting) kf_eq_refl kf_eq_weak_kf_map_right kf_filter_cong_weak swap_swap)
  finally show kf_filter ((=) xy) (kf_comp 𝔉 𝔈) =kr kf_filter ((=) xy) (kf_map prod.swap (kf_comp 𝔈 𝔉))
    by -
qed

lemma kf_comp_apply_on_singleton:
  kf_comp 𝔈 𝔉 *kr @{x} ρ = 𝔈 *kr @{snd x} (𝔉 *kr @{fst x} ρ)
proof -
  have kf_comp 𝔈 𝔉 *kr @{x} ρ = kf_filter (λ(x1,x2). x2 = snd x  x1 = fst x) (kf_comp 𝔈 𝔉) *kr ρ
    by (simp add: kf_apply_on_def prod_eq_iff case_prod_unfold conj_commute)
  also have  = kf_comp (kf_filter (λx2. x2 = snd x) 𝔈) (kf_filter (λx1. x1 = fst x) 𝔉) *kr ρ
    by (simp add: kf_filter_comp)
  also have  = 𝔈 *kr @{snd x} (𝔉 *kr @{fst x} ρ)
    by (simp add: kf_apply_on_def kf_comp_apply)
  finally show ?thesis
    by -
qed

lemma kf_comp_dependent_apply_on_singleton:
  assumes bdd_above ((kf_norm  𝔈) ` kf_domain 𝔉)
  shows kf_comp_dependent 𝔈 𝔉 *kr @{x} ρ = 𝔈 (fst x) *kr @{snd x} (𝔉 *kr @{fst x} ρ)
proof -
  have kf_comp_dependent 𝔈 𝔉 *kr @{x} ρ = kf_comp_dependent 𝔈 𝔉 *kr @({fst x}×{snd x}) ρ
    by fastforce
  also have  = kf_comp_dependent (λ_. 𝔈 (fst x)) 𝔉 *kr @({fst x}×{snd x}) ρ
    apply (rule kf_apply_comp_dependent_cong[THEN fun_cong])
    using assms by auto
  also have  = kf_comp (𝔈 (fst x)) 𝔉 *kr @({x}) ρ
    by (simp add: kf_comp_def)
  also have  = 𝔈 (fst x) *kr @{snd x} (𝔉 *kr @{fst x} ρ)
    by (simp add: kf_comp_apply_on_singleton)
  finally show ?thesis
    by -
qed

lemma kf_comp_id_left: kf_comp kf_id 𝔈 kr kf_map (λx. (x,())) 𝔈
proof (rule kf_eqI_from_filter_eq_weak)
  fix xy :: 'c × unit
  define x where x = fst xy
  then have xy: xy = (x,())
    by (auto intro!: prod.expand)
  have kf_filter ((=) xy) (kf_comp kf_id 𝔈) = kf_comp (kf_filter (λ_. True) kf_id) (kf_filter ((=)x) 𝔈)
    by (auto intro!: arg_cong2[where f=kf_filter] simp add: xy simp flip: kf_filter_comp simp del: kf_filter_true)
  also have  =kr kf_filter ((=)x) 𝔈
    by (simp add: kf_comp_apply kf_eq_weakI)
  also have  =kr kf_map (λx. (x,())) (kf_filter ((=)x) 𝔈)
    by (simp add: kf_eq_weak_def)
  also have  = kf_filter ((=) xy) (kf_map (λx. (x,())) 𝔈)
    by (simp add: kf_filter_map xy)
  finally show kf_filter ((=) xy) (kf_comp kf_id 𝔈) =kr 
    by -
qed

lemma kf_comp_id_right: kf_comp 𝔈 kf_id kr kf_map (λx. ((),x)) 𝔈
proof (rule kf_eqI_from_filter_eq_weak)
  fix xy :: unit × 'c
  define y where y = snd xy
  then have xy: xy = ((),y)
    by (auto intro!: prod.expand simp: y_def)
  have kf_filter ((=) xy) (kf_comp 𝔈 kf_id) = kf_comp (kf_filter ((=)y) 𝔈) (kf_filter (λ_. True) kf_id)
    by (auto intro!: arg_cong2[where f=kf_filter] simp add: xy simp flip: kf_filter_comp simp del: kf_filter_true)
  also have  =kr kf_filter ((=)y) 𝔈
    by (simp add: kf_comp_apply kf_eq_weakI)
  also have  =kr kf_map (Pair ()) (kf_filter ((=)y) 𝔈)
    by (simp add: kf_eq_weak_def)
  also have  = kf_filter ((=) xy) (kf_map (λx. ((),x)) 𝔈)
    by (simp add: kf_filter_map xy)
  finally show kf_filter ((=) xy) (kf_comp 𝔈 kf_id) =kr 
    by -
qed

lemma kf_comp_dependent_raw_kf_plus_left:
  fixes 𝔇 :: 'f  ('b::chilbert_space, 'c::chilbert_space, 'd) kraus_family
  fixes 𝔈 :: 'f  ('b::chilbert_space, 'c::chilbert_space, 'e) kraus_family
  fixes 𝔉 :: ('a::chilbert_space, 'b, 'f) kraus_family
  assumes bdd_above ((λx. kf_norm (𝔇 x)) ` kf_domain 𝔉)
  assumes bdd_above ((λx. kf_norm (𝔈 x)) ` kf_domain 𝔉)
  shows kf_comp_dependent_raw (λx. kf_plus (𝔇 x) (𝔈 x)) 𝔉 =
    kf_map_inj (λx. case x of Inl (F,D,f,d)  (F, D, f, Inl d) | Inr (F,E,f,e)  (F, E, f, Inr e))
               (kf_plus (kf_comp_dependent_raw 𝔇 𝔉) (kf_comp_dependent_raw 𝔈 𝔉))
proof (rule Rep_kraus_family_inject[THEN iffD1], rule Set.set_eqI, rename_tac tuple)
  fix tuple :: ('a CL 'c) × ('a CL 'b) × ('b CL 'c) × 'f × ('d + 'e)
  have [simp]: bdd_above ((λx. kf_norm (kf_plus (𝔇 x) (𝔈 x))) ` kf_domain 𝔉)
    apply (rule bdd_above_mono2[rotated])
      apply (rule order_refl)
     apply (rule kf_norm_triangle)
    using assms by (rule bdd_above_plus)
  obtain EF F E f y where tuple: tuple = (EF,F,E,f,y)
    by (auto simp: prod_eq_iff)
  have tuple  Rep_kraus_family (kf_comp_dependent_raw (λx. kf_plus (𝔇 x) (𝔈 x)) 𝔉)
     EF = E oCL F  EF  0  (F,f)  Rep_kraus_family 𝔉  (E,y)  Rep_kraus_family (kf_plus (𝔇 f) (𝔈 f))
    by (auto intro!: image_eqI simp add: tuple kf_comp_dependent_raw.rep_eq)
  also have   EF = E oCL F  EF  0  (F,f)  Rep_kraus_family 𝔉  
                      (case y of Inl d  (E,d)  Rep_kraus_family (𝔇 f)
                               | Inr e  (E,e)  Rep_kraus_family (𝔈 f))
    apply (cases y)
    by (auto simp: kf_plus.rep_eq)
  also have   (case y of Inl d  (EF, F, E, f, d)  Rep_kraus_family (kf_comp_dependent_raw 𝔇 𝔉)
                             | Inr e  (EF, F, E, f, e)  Rep_kraus_family (kf_comp_dependent_raw 𝔈 𝔉))
    apply (cases y)
    by (force intro!: simp: kf_comp_dependent_raw.rep_eq assms)+
  also have   (case y of Inl d  (EF, Inl (F, E, f, d))  Rep_kraus_family (kf_plus (kf_comp_dependent_raw 𝔇 𝔉) (kf_comp_dependent_raw 𝔈 𝔉))
                             | Inr e  (EF, Inr (F, E, f, e))  Rep_kraus_family (kf_plus (kf_comp_dependent_raw 𝔇 𝔉) (kf_comp_dependent_raw 𝔈 𝔉)))
    apply (cases y)
    by (auto intro!: simp: kf_plus.rep_eq)
  also have  
       (tuple  Rep_kraus_family (kf_map_inj
              (λx. case x of Inl (F, D, f, d)  (F, D, f, Inl d) | Inr (F, E, f, e)  (F, E, f, Inr e))
              (kf_plus (kf_comp_dependent_raw 𝔇 𝔉) (kf_comp_dependent_raw 𝔈 𝔉))))
    apply (cases y)
    by (force simp: tuple kf_map_inj.rep_eq split!: sum.split_asm prod.split)+
  finally 
  show (tuple  Rep_kraus_family (kf_comp_dependent_raw (λx. kf_plus (𝔇 x) (𝔈 x)) 𝔉)) 
       (tuple  Rep_kraus_family (kf_map_inj
              (λx. case x of Inl (F, D, f, d)  (F, D, f, Inl d) | Inr (F, E, f, e)  (F, E, f, Inr e))
              (kf_plus (kf_comp_dependent_raw 𝔇 𝔉) (kf_comp_dependent_raw 𝔈 𝔉))))
    by -
qed


lemma kf_comp_dependent_kf_plus_left:
  assumes bdd_above ((λx. kf_norm (𝔇 x)) ` kf_domain 𝔉)
  assumes bdd_above ((λx. kf_norm (𝔈 x)) ` kf_domain 𝔉)
  shows kf_comp_dependent (λx. kf_plus (𝔇 x) (𝔈 x)) 𝔉 =
    kf_map_inj (λx. case x of Inl (f,d)  (f, Inl d) | Inr (f,e)  (f, Inr e)) (kf_plus (kf_comp_dependent 𝔇 𝔉) (kf_comp_dependent 𝔈 𝔉))
  apply (simp add: kf_comp_dependent_def kf_comp_dependent_raw_kf_plus_left[OF assms] kf_plus_map_both)
  apply (subst kf_map_kf_map_inj_comp)
   apply (auto simp add: inj_on_def split!: sum.split_asm)[1]
  apply (subst kf_map_inj_kf_map_comp)
   apply (auto simp add: inj_on_def split!: sum.split_asm)[1]
  apply (rule arg_cong2[where f=kf_map])
  by (auto intro!: ext simp add: o_def split!: sum.split)

lemma kf_map_inj_twice:
  shows kf_map_inj f (kf_map_inj g 𝔈) = kf_map_inj (f o g) 𝔈
  apply (transfer' fixing: f g)
  by (simp add: image_image case_prod_unfold)


lemma kf_comp_dependent_kf_plus_left':
  assumes bdd_above ((λx. kf_norm (𝔇 x)) ` kf_domain 𝔉)
  assumes bdd_above ((λx. kf_norm (𝔈 x)) ` kf_domain 𝔉)
  shows kf_plus (kf_comp_dependent 𝔇 𝔉) (kf_comp_dependent 𝔈 𝔉) = 
    kf_map_inj (λ(f,de). case de of Inl d  Inl (f,d) | Inr e  Inr (f,e)) (kf_comp_dependent (λx. kf_plus (𝔇 x) (𝔈 x)) 𝔉)
  apply (subst kf_comp_dependent_kf_plus_left[OF assms])
  apply (subst kf_map_inj_twice)
  apply (rewrite at kf_map_inj  _ to id DEADID.rel_mono_strong)
  by (auto intro!: ext simp: split!: sum.split)


lemma kf_comp_dependent_plus_left:
  assumes bdd_above ((λx. kf_norm (𝔇 x)) ` kf_domain 𝔉)
  assumes bdd_above ((λx. kf_norm (𝔈 x)) ` kf_domain 𝔉)
  shows kf_comp_dependent (λx. 𝔇 x + 𝔈 x) 𝔉 kr kf_comp_dependent 𝔇 𝔉 + kf_comp_dependent 𝔈 𝔉
proof -
  have kf_comp_dependent (λx. 𝔇 x + 𝔈 x) 𝔉 kr 
      kf_map (λ(x, y). (x, case y of Inl x  x | Inr y  y)) (kf_comp_dependent (λx. kf_plus (𝔇 x) (𝔈 x)) 𝔉)
    unfolding plus_kraus_family_def
    by (rule kf_comp_dependent_map_left)
  also have  = kf_map (λ(x, y). (x, case y of Inl x  x | Inr y  y))
     (kf_map_inj (case_sum (λ(f, d). (f, Inl d)) (λ(f, e). (f, Inr e)))
       (kf_plus (kf_comp_dependent 𝔇 𝔉) (kf_comp_dependent 𝔈 𝔉)))
    by (simp add: kf_comp_dependent_kf_plus_left[OF assms])
  also have  = kf_map (λy. case y of Inl x  x | Inr y  y)
     (kf_plus (kf_comp_dependent 𝔇 𝔉) (kf_comp_dependent 𝔈 𝔉))
    apply (subst kf_map_kf_map_inj_comp)
     apply (auto intro!: inj_onI split!: sum.split_asm)[1]
    apply (rule arg_cong2[where f=kf_map, OF _ refl])
    by (auto intro!: ext simp add: o_def split!: sum.split)
  also have  = kf_comp_dependent 𝔇 𝔉 + kf_comp_dependent 𝔈 𝔉
    by (simp add: plus_kraus_family_def)
  finally show ?thesis
    by -
qed

subsection ‹Infinite sums›

lift_definition kf_infsum :: ('a  ('b::chilbert_space,'c::chilbert_space,'x) kraus_family)  'a set  ('b,'c,'a×'x) kraus_family is
  λ𝔈 A. if summable_on_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A
         then (λ(a,(E,x)). (E,(a,x))) ` Sigma A (λa. Rep_kraus_family (𝔈 a)) else {}
proof (rule CollectI, rename_tac 𝔈 A)
  fix 𝔈 :: 'a  ('b, 'c, 'x) kraus_family and A
  define 𝔈' where 𝔈' a = Rep_kraus_family (𝔈 a) for a
  show kraus_family (if summable_on_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A 
                      then (λ(a,(E,x)). (E,(a,x))) ` Sigma A 𝔈'
                      else {})
  proof (cases summable_on_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A)
    case True
    have kraus_family ((λ(a, E, x). (E, a, x)) ` Sigma A 𝔈')
    proof (intro kraus_familyI bdd_aboveI)
      fix C assume C  (λF. (E, x)F. E* oCL E) ` {F. finite F  F  (λ(a, E, x). (E, a, x)) ` Sigma A 𝔈'}
      then obtain F where finite F and F_subset: F  (λ(a, E, x). (E, a, x)) ` Sigma A 𝔈'
        and C_def: C = ((E, x)F. E* oCL E)
        by blast
      define B F' where B = infsum_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A
        and F' = (λ(E,a,x). (a,E,x)) ` F

      have [iff]: finite F'
        using finite F by (auto intro!: simp: F'_def)
      have F'_subset: F'  Sigma A 𝔈'
        using F_subset by (auto simp: F'_def)
      have Sigma_decomp: (SIGMA a:(λx. fst x) ` F'. snd ` Set.filter (λ(a',Ex). a'=a) F') = F'
        by force

      have C = ((a, (E, x))F'. E* oCL E)
        unfolding F'_def
        apply (subst sum.reindex)
        by (auto intro!: inj_onI simp: C_def case_prod_unfold)
      also have  = (afst ` F'. (E, x)snd ` Set.filter (λ(a',Ex). a'=a) F'. E* oCL E)
        apply (subst sum.Sigma)
        by (auto intro!: finite_imageI simp: Sigma_decomp)
      also have  = (afst ` F'. infsum_in cweak_operator_topology (λ(E, x). E* oCL E) (snd ` Set.filter (λ(a',Ex). a'=a) F'))
        apply (rule sum.cong[OF refl])
        apply (rule infsum_in_finite[symmetric])
        by auto
      also have   (afst ` F'. infsum_in cweak_operator_topology (λ(E, x). E* oCL E) (𝔈' a))
      proof (rule sum_mono, rule infsum_mono_neutral_wot)
        fix a assume a  fst ` F'
        show summable_on_in cweak_operator_topology (λ(E, x). E* oCL E) (snd ` Set.filter (λ(a', Ex). a' = a) F')
          by (auto intro!: summable_on_in_finite)
        show summable_on_in cweak_operator_topology (λ(E, x). E* oCL E) (𝔈' a)
          using 𝔈'_def[abs_def] kf_bound_has_sum summable_on_in_def by blast
        show (case Ex of (E, x)  E* oCL E)  (case Ex of (E, x)  E* oCL E) for Ex
          by blast
        have snd ` Set.filter (λ(a', Ex). a' = a) F'  𝔈' a
          using F'_subset by auto
        then show (case Ex of (E, x)  E* oCL E)  0
          if Ex  snd ` Set.filter (λ(a', Ex). a' = a) F' - 𝔈' a for Ex
          using that by blast
        show 0  (case Ex of (E, x)  E* oCL E) for Ex
          by (auto intro!: positive_cblinfun_squareI simp: case_prod_unfold)
      qed
      also have  = (afst ` F'. kf_bound (𝔈 a))
        unfolding 𝔈'_def
        apply (transfer' fixing: F')
        by simp
      also have  = infsum_in cweak_operator_topology (λa. kf_bound (𝔈 a)) (fst ` F')
        apply (rule infsum_in_finite[symmetric])
        by auto
      also have   infsum_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A
      proof (rule infsum_mono_neutral_wot)
        show summable_on_in cweak_operator_topology (λa. kf_bound (𝔈 a)) (fst ` F')
          by (auto intro!: summable_on_in_finite)
        show summable_on_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A
          using True by blast
        show kf_bound (𝔈 a)  kf_bound (𝔈 a) for a
          by blast
        show kf_bound (𝔈 a)  0 if a  fst ` F' - A for a
          using F'_subset that by auto
        show 0  kf_bound (𝔈 a) for a
          by simp
      qed
      also have  = B
        using B_def by blast
      finally show C  B
        by -
    next
      show 0  fst ` (λ(a, E, x). (E, a, x)) ` Sigma A 𝔈'
        using Rep_kraus_family by (force simp: 𝔈'_def kraus_family_def)
    qed
    with True show ?thesis 
      by simp
  next
    case False
    then show ?thesis
      by (auto intro!: bdd_aboveI simp add: kraus_family_def)
  qed
qed

definition kf_summable :: ('a  ('b::chilbert_space,'c::chilbert_space,'x) kraus_family)  'a set  bool where
  kf_summable 𝔈 A  summable_on_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A

lemma kf_summable_from_abs_summable:
  assumes sum: (λa. kf_norm (𝔈 a)) summable_on A
  shows kf_summable 𝔈 A
  using assms
  by (simp add: summable_imp_wot_summable abs_summable_summable kf_summable_def kf_norm_def)

lemma kf_infsum_apply:
  assumes kf_summable 𝔈 A
  shows kf_infsum 𝔈 A *kr ρ = (aA. 𝔈 a *kr ρ)
proof -
  from kf_apply_summable[of ρ kf_infsum 𝔈 A]
  have summable: (λ(a, E, x). sandwich_tc E ρ) summable_on (SIGMA a:A. Rep_kraus_family (𝔈 a))
    using assms
    by (simp add: kf_summable_def kf_infsum.rep_eq case_prod_unfold summable_on_reindex inj_on_def prod.expand o_def)
  have kf_infsum 𝔈 A *kr ρ = ((E,ax)(λ(a, E, x)  (E, a, x)) ` (SIGMA a:A. Rep_kraus_family (𝔈 a)). sandwich_tc E ρ)
    using assms unfolding kf_summable_def
    apply (transfer' fixing: 𝔈)
    by (simp add: case_prod_unfold)
  also have  = ((a, E, x)  (SIGMA a:A. Rep_kraus_family (𝔈 a)). sandwich_tc E ρ)
    apply (subst infsum_reindex)
    by (auto intro!: inj_onI simp: o_def case_prod_unfold)
  also have  = (aA. (E, b)Rep_kraus_family (𝔈 a). sandwich_tc E ρ)
    apply (subst infsum_Sigma'_banach[symmetric])
    using summable by auto
  also have  = (aA. 𝔈 a *kr ρ)
    by (metis (no_types, lifting) infsum_cong kf_apply.rep_eq split_def)
  finally show ?thesis
    by -
qed

lemma kf_infsum_apply_summable:
  assumes kf_summable 𝔈 A
  shows (λa. 𝔈 a *kr ρ) summable_on A
proof -
  from kf_apply_summable[of ρ kf_infsum 𝔈 A]
  have summable: (λ(a, E, x). sandwich_tc E ρ) summable_on (SIGMA a:A. Rep_kraus_family (𝔈 a))
    using assms
    by (simp add: kf_summable_def kf_infsum.rep_eq case_prod_unfold summable_on_reindex inj_on_def prod.expand o_def)
  from summable_on_Sigma_banach[OF this]
  have (λa. (E, x)Rep_kraus_family (𝔈 a). sandwich_tc E ρ) summable_on A
    by simp
  then show ?thesis
    by (metis (mono_tags, lifting) infsum_cong kf_apply.rep_eq split_def summable_on_cong)
qed

lemma kf_bound_infsum:
  fixes f :: 'a  ('b::chilbert_space,'c::chilbert_space,'x) kraus_family
  assumes kf_summable f A
  shows kf_bound (kf_infsum f A) = infsum_in cweak_operator_topology (λa. kf_bound (f a)) A
proof -
  have pos: 0  compose_wot (adj_wot a) a for a :: ('b,'c) cblinfun_wot
    apply transfer'
    using positive_cblinfun_squareI by blast
  have sum3: (λx. (E,_)Rep_kraus_family (f x). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E)) summable_on A
  proof -
    define b where b x = kf_bound (f x) for x
    have (λx. Abs_cblinfun_wot (b x)) summable_on A
      using assms unfolding kf_summable_def
      apply (subst (asm) b_def[symmetric])
      apply (transfer' fixing: b A)
      by simp
    then show ?thesis
      by (simp add: Rep_cblinfun_wot_inverse kf_bound_def' b_def)
  qed
  have sum2: (λ(E, _). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E)) summable_on
         Rep_kraus_family (f x) if x  A for x
    by (rule kf_bound_summable')
  have sum1: (λ(_, E, _). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E)) summable_on
    (SIGMA a:A. Rep_kraus_family (f a))
    apply (rule summable_on_Sigma_wotI)
    using sum3 sum2
    by (auto intro!: pos simp: case_prod_unfold)

  have Abs_cblinfun_wot (kf_bound (kf_infsum f A)) =
        ((E, x)Rep_kraus_family (kf_infsum f A).
               compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
    by (simp add: kf_bound_def' Rep_cblinfun_wot_inverse)
  also have  = ((E, x)(λ(a, E, xa). (E, a, xa)) ` (SIGMA a:A. Rep_kraus_family (f a)).
       compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
    using assms by (simp add: kf_infsum.rep_eq kf_summable_def)
  also have  = ((a, E, x)(SIGMA a:A. Rep_kraus_family (f a)). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
    apply (subst infsum_reindex)
    by (auto intro!: inj_onI simp: o_def case_prod_unfold)
  also have  = (aA. (E, x)Rep_kraus_family (f a). compose_wot (adj_wot (Abs_cblinfun_wot E)) (Abs_cblinfun_wot E))
    apply (subst infsum_Sigma_topological_monoid)
    using sum1 sum2 by auto
  also have  = (aA. Abs_cblinfun_wot (kf_bound (f a)))
    by (simp add: kf_bound_def' Rep_cblinfun_wot_inverse)
  also have  = Abs_cblinfun_wot (infsum_in cweak_operator_topology (λa. kf_bound (f a)) A)
    apply (simp only: infsum_euclidean_eq[symmetric])
    apply (transfer' fixing: f A)
    by simp
  finally show ?thesis
    apply (rule Abs_cblinfun_wot_inject[THEN iffD1, rotated -1])
    by simp_all
qed

lemma kf_norm_infsum:
  assumes sum: (λa. kf_norm (𝔈 a)) summable_on A
  shows kf_norm (kf_infsum 𝔈 A)  (aA. kf_norm (𝔈 a))
proof -
  have kf_norm (kf_infsum 𝔈 A) = norm (infsum_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A)
    unfolding kf_norm_def
    apply (subst kf_bound_infsum)
    by (simp_all add: kf_summable_from_abs_summable assms)
  also have   (aA. norm (kf_bound (𝔈 a)))
    apply (rule triangle_ineq_wot)
    using sum by (simp add: kf_norm_def)
  also have   (aA. kf_norm (𝔈 a))
    by (smt (verit, del_insts) infsum_cong kf_norm_def)
  finally show ?thesis
    by -
qed

lemma kf_filter_infsum:
  assumes kf_summable 𝔈 A
  shows kf_filter P (kf_infsum 𝔈 A) 
           = kf_infsum (λa. kf_filter (λx. P (a,x)) (𝔈 a)) {aA. x. P (a,x)}
    (is ?lhs = ?rhs)
proof -
  have summ: summable_on_in cweak_operator_topology (λa. kf_bound (kf_filter (λx. P (a, x)) (𝔈 a)))
      {a  A. x. P (a, x)}
  proof (rule summable_wot_boundedI)
    fix F assume finite F and F_subset: F  {a  A. x. P (a, x)}
    have (aF. kf_bound (kf_filter (λx. P (a, x)) (𝔈 a)))
         (aF. kf_bound (𝔈 a))
      by (meson kf_bound_filter sum_mono)
    also have  = infsum_in cweak_operator_topology (λa. kf_bound (𝔈 a)) F
      apply (rule infsum_in_finite[symmetric])
      by (auto intro!: finite F)
    also have   infsum_in cweak_operator_topology (λa. kf_bound (𝔈 a)) A
      apply (rule infsum_mono_neutral_wot)
      using F_subset assms
      by (auto intro!: finite F intro: summable_on_in_finite simp: kf_summable_def)
    finally show (aF. kf_bound (kf_filter (λx. P (a, x)) (𝔈 a)))  
      by -
  next
    show 0  kf_bound (kf_filter (λy. P (x, y)) (𝔈 x)) for x
      by simp
  qed
  have Rep_kraus_family ?lhs = Rep_kraus_family ?rhs
    using assms by (force simp add: kf_filter.rep_eq kf_infsum.rep_eq assms summ kf_summable_def)
  then show ?lhs = ?rhs
    by (simp add: Rep_kraus_family_inject)
qed

lemma kf_infsum_empty[simp]: kf_infsum 𝔈 {} = 0
  apply transfer' by simp

lemma kf_infsum_singleton[simp]: kf_infsum 𝔈 {a} = kf_map_inj (λx. (a,x)) (𝔈 a)
  apply (rule Rep_kraus_family_inject[THEN iffD1])
  by (force simp add: kf_infsum.rep_eq summable_on_in_finite kf_map_inj.rep_eq)

lemma kf_infsum_invalid: 
  assumes ¬ kf_summable 𝔈 A
  shows kf_infsum 𝔈 A = 0
  using assms
  apply transfer'
  unfolding kf_summable_def
  by simp


lemma kf_infsum_cong:
  fixes 𝔈 𝔉 :: 'a  ('b::chilbert_space, 'c::chilbert_space, 'x) kraus_family
  assumes a. a  A  𝔈 a kr 𝔉 a
  shows kf_infsum 𝔈 A kr kf_infsum 𝔉 A
proof (cases kf_summable 𝔈 A)
  case True
  then have True': kf_summable 𝔉 A
    unfolding kf_summable_def
    apply (rule summable_on_in_cong[THEN iffD1, rotated])
    by (simp add: kf_bound_cong assms kf_eq_imp_eq_weak)
  show ?thesis
  proof (rule kf_eqI)
    fix ax :: 'a × 'x and ρ
    obtain a x where ax_def: ax = (a,x)
      by fastforce
    have *: {a'. a' = a  a'  A} = (if aA then {a} else {})
      by auto
    have kf_infsum 𝔈 A *kr @{ax} ρ = (if aA then 𝔈 a *kr @{x} ρ else 0)
      by (simp add: ax_def kf_apply_on_def True kf_filter_infsum * 
          kf_apply_map_inj inj_on_def)
    also from assms have  = (if aA then 𝔉 a *kr @{x} ρ else 0)
      by (auto intro!: kf_apply_on_eqI)
    also have  = kf_infsum 𝔉 A *kr @{ax} ρ
      by (simp add: ax_def kf_apply_on_def True' kf_filter_infsum * 
          kf_apply_map_inj inj_on_def)
    finally show kf_infsum 𝔈 A *kr @{ax} ρ = kf_infsum 𝔉 A *kr @{ax} ρ
      by -
  qed
next
  case False
  then have False': ¬ kf_summable 𝔉 A
    unfolding kf_summable_def
    apply (subst (asm) assms[THEN kf_eq_imp_eq_weak, 
          THEN kf_bound_cong, THEN summable_on_in_cong])
    by auto
  show ?thesis
    by (simp add: kf_infsum_invalid False False')
qed

subsection ‹Trace-preserving maps›

definition kf_trace_preserving 𝔈  (ρ. trace_tc (𝔈 *kr ρ) = trace_tc ρ)

definition kf_trace_reducing 𝔈  (ρ0. trace_tc (𝔈 *kr ρ)  trace_tc ρ)

lemma kf_trace_reducing_iff_norm_leq1: kf_trace_reducing 𝔈  kf_norm 𝔈  1
proof (unfold kf_trace_reducing_def, intro iffI allI impI)
  assume assm: kf_norm 𝔈  1
  fix ρ :: ('a, 'a) trace_class
  assume ρ  0
  have trace_tc (𝔈 *kr ρ) = norm (𝔈 *kr ρ)
    by (simp add: 0  ρ kf_apply_pos norm_tc_pos)
  also have   kf_norm 𝔈 * norm ρ
    using 0  ρ complex_of_real_mono kf_apply_bounded_pos by blast
  also have   norm ρ
    by (metis assm complex_of_real_mono kf_norm_geq0 mult_left_le_one_le norm_ge_zero)
  also have  = trace_tc ρ
    by (simp add: 0  ρ norm_tc_pos)
  finally show trace_tc (𝔈 *kr ρ)  trace_tc ρ
    by -
next
  assume assm[rule_format]: ρ0. trace_tc (𝔈 *kr ρ)  trace_tc ρ
  have kf_bound 𝔈  id_cblinfun
  proof (rule cblinfun_leI)
    fix x
    have x C kf_bound 𝔈 x = trace_tc (𝔈 *kr tc_butterfly x x)
      by (simp add: kf_bound_from_map)
    also have   trace_tc (tc_butterfly x x)
      apply (rule assm)
      by simp
    also have  = x C id_cblinfun x
      by (simp add: tc_butterfly.rep_eq trace_butterfly trace_tc.rep_eq)
  finally show x C kf_bound 𝔈 x  x C id_cblinfun x
      by -
  qed
  then show kf_norm 𝔈  1
    by (smt (verit, best) kf_norm_def kf_bound_pos norm_cblinfun_id_le norm_cblinfun_mono)
qed

lemma kf_trace_preserving_iff_bound_id: kf_trace_preserving 𝔈  kf_bound 𝔈 = id_cblinfun
proof (unfold kf_trace_preserving_def, intro iffI allI)
  assume asm[rule_format]: ρ. trace_tc (𝔈 *kr ρ) = trace_tc ρ
  have ψ C kf_bound 𝔈 ψ = ψ C id_cblinfun ψ for ψ
  proof -
    have ψ C kf_bound 𝔈 ψ = trace_tc (𝔈 *kr tc_butterfly ψ ψ)
      by (simp add: kf_bound_from_map)
    also have  = trace_tc (tc_butterfly ψ ψ)
      by (simp add: asm)
    also have  = ψ C id_cblinfun ψ
      by (simp add: tc_butterfly.rep_eq trace_butterfly trace_tc.rep_eq)
    finally show ?thesis
      by -
  qed
  then show kf_bound 𝔈 = id_cblinfun
    using cblinfun_cinner_eq0I[where a=kf_bound 𝔈 - id_cblinfun]
    by (simp add: cblinfun.real.diff_left cinner_diff_right)
next
  assume asm: kf_bound 𝔈 = id_cblinfun
  fix ρ
  have trace_tc (𝔈 *kr ρ) = trace_tc (compose_tcr (kf_bound 𝔈) ρ)
    by (rule trace_from_kf_bound)
  also have  = trace_tc ρ
    using asm by fastforce
  finally show trace_tc (𝔈 *kr ρ) = trace_tc ρ
    by -
qed

lemma kf_trace_norm_preserving: kf_norm 𝔈  1 if kf_trace_preserving 𝔈
  apply (rule kf_trace_reducing_iff_norm_leq1[THEN iffD1])
  using that
  by (simp add: kf_trace_preserving_def kf_trace_reducing_def)

lemma kf_trace_norm_preserving_eq: 
  fixes 𝔈 :: ('a::{chilbert_space,not_singleton}, 'b::chilbert_space, 'c) kraus_family
  assumes kf_trace_preserving 𝔈
  shows kf_norm 𝔈 = 1
  using assms by (simp add: kf_trace_preserving_iff_bound_id kf_norm_def)

lemma kf_trace_preserving_map[simp]: kf_trace_preserving (kf_map f 𝔈)  kf_trace_preserving 𝔈
  by (simp add: kf_map_bound kf_trace_preserving_iff_bound_id)

lemma kf_trace_reducing_map[simp]: kf_trace_reducing (kf_map f 𝔈)  kf_trace_reducing 𝔈
  by (simp add: kf_trace_reducing_iff_norm_leq1)

lemma kf_trace_preserving_id[iff]: kf_trace_preserving kf_id
  by (simp add: kf_trace_preserving_iff_bound_id)

lemma kf_trace_reducing_id[iff]: kf_trace_reducing kf_id
  by (simp add: kf_norm_id_leq1 kf_trace_reducing_iff_norm_leq1)

subsection ‹Sampling›

lift_definition kf_sample :: ('x  real)  ('a::chilbert_space, 'a, 'x) kraus_family is
  λp. if (x. p x  0)  p summable_on UNIV then Set.filter (λ(E,_). E0) (range (λx. (sqrt (p x) *R id_cblinfun, x))) else {}
proof -
  fix p :: 'x  real
  show ?thesis p
  proof (cases (x. p x  0)  p summable_on UNIV)
    case True
    then have p abs_summable_on UNIV
      by simp
    from abs_summable_iff_bdd_above[THEN iffD1, OF this]
    obtain B where F_B: finite F  (xF. p x)  B for F
      apply atomize_elim
      using True by (auto simp: bdd_above_def)
    have ((E,x)F. E* oCL E)  B *R id_cblinfun
      if finite F and F  Set.filter (λ(E,_). E0) (range (λx. (sqrt (p x) *R id_cblinfun, x)))
      for F :: ('aCL'a × 'x) set
    proof -
      from that
      obtain F' where finite F' and F_def: F = (λx. (sqrt (p x) *R id_cblinfun, x)) ` F'
        using finite_subset_filter_image
        by meson
      then have ((E,x)F. E* oCL E) = (xF'. (sqrt (p x) *R id_cblinfun)* oCL (sqrt (p x) *R id_cblinfun))
        by (simp add: sum.reindex inj_on_def)
      also have  = (xF'. p x *R id_cblinfun)
        using True by simp
      also have  = (xF'. p x) *R id_cblinfun
        by (metis scaleR_left.sum)
      also have   B *R id_cblinfun
        using F. finite F  (xF. p x)  B finite F' positive_id_cblinfun scaleR_right_mono by blast
      finally show ?thesis
        by -
    qed
    then have kraus_family (Set.filter (λ(E,_). E0) (range (λx. (sqrt (p x) *R (id_cblinfun ::'aCL_), x))))
      by (force intro!: bdd_aboveI[where M=B *R id_cblinfun] simp: kraus_family_def case_prod_unfold)
    then show ?thesis
      using True by simp
  next
    case False
    then show ?thesis
      by auto
  qed
qed

lemma kf_sample_norm:
  fixes p :: 'x  real
  assumes x. p x  0
  assumes p summable_on UNIV
  shows kf_norm (kf_sample p :: ('a::{chilbert_space,not_singleton},'a,'x) kraus_family)
             = (x. p x)
proof -
  define B :: 'aCL'a where B = kf_bound (kf_sample p)
  obtain ψ :: 'a where norm ψ = 1
    using ex_norm1_not_singleton by blast

  have ψ C (B *V ψ) = ψ C ((x. p x) *R id_cblinfun *V ψ) 
    if norm ψ = 1 for ψ
  proof -
    have has_sum_in cweak_operator_topology (λ(E, x). E* oCL E) (Rep_kraus_family (kf_sample p)) B
      using B_def kf_bound_has_sum by blast
    then have has_sum_in cweak_operator_topology (λ(E, x). E* oCL E) (Set.filter (λ(E,_). E0) (range (λx. (sqrt (p x) *R id_cblinfun, x)))) B
      by (simp add: kf_sample.rep_eq assms)
    then have has_sum_in cweak_operator_topology (λ(E, x). E* oCL E) (range (λx. (sqrt (p x) *R id_cblinfun, x))) B
      apply (rule has_sum_in_cong_neutral[THEN iffD1, rotated -1])
      by auto
    then have has_sum_in cweak_operator_topology (λx. norm (p x) *R id_cblinfun) UNIV B
      by (simp add: has_sum_in_reindex inj_on_def o_def)
    then have has_sum_in cweak_operator_topology (λx. p x *R id_cblinfun) UNIV B
      apply (rule has_sum_in_cong[THEN iffD1, rotated])
      by (simp add: assms(1))
    then have has_sum_in euclidean (λx. ψ C (p x *R id_cblinfun) ψ) UNIV (ψ C B ψ)
      apply (rule has_sum_in_comm_additive[rotated 3, OF cweak_operator_topology_cinner_continuous, unfolded o_def])
      by (simp_all add: Modules.additive_def cblinfun.add_left cinner_simps)
    then have ((λx. of_real (p x)) has_sum (ψ C B ψ)) UNIV
      apply (simp add: scaleR_scaleC has_sum_euclidean_iff)
      using norm ψ = 1 cnorm_eq_1 by force
    then have ψ C B ψ = (x. of_real (p x))
      by (simp add: infsumI)
    also have  = of_real (x. p x)
      by (metis infsum_of_real)
    also have  = ψ C ((x. p x) *R id_cblinfun) ψ
      using norm ψ = 1 by (simp add: scaleR_scaleC cnorm_eq_1)
    finally show ?thesis
      by -
  qed
  then have B = (x. p x) *R id_cblinfun
    by (rule cblinfun_cinner_eqI)
  then have norm B = norm (x. p x)
    by simp
  also have  = (x. p x)
    by (simp add: abs_of_nonneg assms infsum_nonneg)
  finally show ?thesis
    by (simp add: kf_norm_def B_def)
qed

subsection ‹Trace›

lift_definition kf_trace :: 'a set  ('a::chilbert_space, 'b::one_dim, 'a) kraus_family is
  λB. if is_onb B then (λx. (vector_to_cblinfun x*, x)) ` B else {}
proof (rename_tac B)
  fix B :: 'a set
  define family :: ('a CL 'b × 'a) set where family = (λx. (vector_to_cblinfun x*, x)) ` B
  have kraus_family family if is_onb B
  proof -
    have ((E, x)F. E* oCL E)  id_cblinfun if finite F and FB: F  family for F :: ('a CL 'b × 'a) set
    proof -
      obtain G where finite G and G  B and FG: F = (λx. (vector_to_cblinfun x*, x)) ` G
        apply atomize_elim
        using finite F and FB
        apply (simp add: family_def)
        by (meson finite_subset_image)
      from G  B have [simp]: is_ortho_set G
        by (meson is_onb B is_onb_def is_ortho_set_antimono)
      from G  B have [simp]: e  G  norm e = 1 for e
        by (meson Set.basic_monos(7) is_onb B is_onb_def)
      have [simp]: inj_on (λx. (vector_to_cblinfun x*, x)) G
        by (meson inj_onI prod.inject)
      have ((E, x)F. E* oCL E) = (xG. selfbutter x)
        by (simp add: FG sum.reindex flip: butterfly_def_one_dim)
      also have (xG. selfbutter x)  id_cblinfun
        apply (rule sum_butterfly_leq_id)
        by auto
      finally show ?thesis
        by -
    qed
    moreover have 0  fst ` family
      apply (simp add: family_def image_image)
      using is_onb B apply (simp add: is_onb_def)
      by (smt (verit) adj_0 double_adj imageE norm_vector_to_cblinfun norm_zero)
    ultimately show ?thesis
      by (auto intro!: bdd_aboveI[where M=id_cblinfun] kraus_familyI)
  qed
  then
  show (if is_onb B then family else {})  Collect kraus_family
    by auto
qed

lemma kf_trace_is_trace: 
  assumes is_onb B
  shows kf_trace B *kr ρ = one_dim_iso (trace_tc ρ)
proof -
  define ρ' where ρ' = from_trace_class ρ
  have kf_apply (kf_trace B) ρ = (xB. sandwich_tc (vector_to_cblinfun x*) ρ)
    apply (simp add: kf_apply.rep_eq kf_trace.rep_eq assms)
    apply (subst infsum_reindex)
     apply (meson inj_onI prod.simps(1))
    by (simp add: o_def)
  also have  = (xB. one_dim_iso (x C (ρ' x)))
    apply (intro infsum_cong from_trace_class_inject[THEN iffD1])
    apply (subst from_trace_class_sandwich_tc)
    by (simp add: sandwich_apply flip: ρ'_def)
  also have  = one_dim_iso (xB. (x C (ρ' x)))
    by (metis (mono_tags, lifting) ρ'_def infsum_cblinfun_apply infsum_cong assms one_cblinfun.rep_eq trace_class_from_trace_class trace_exists)
  also have  = one_dim_iso (trace ρ')
    by (metis ρ'_def trace_class_from_trace_class trace_alt_def[OF assms])
  also have  = one_dim_iso (trace_tc ρ)
    by (simp add: ρ'_def trace_tc.rep_eq)
  finally show ?thesis
    by -
qed

lemma kf_eq_weak_kf_trace:
  assumes is_onb A and is_onb B
  shows kf_trace A =kr kf_trace B
  by (auto simp: kf_eq_weak_def kf_trace_is_trace assms)

lemma trace_is_kf_trace:
  assumes is_onb B
  shows trace_tc t = one_dim_iso (kf_trace B *kr t)
  by (simp add: kf_trace_is_trace assms)

lemma kf_trace_bound[simp]:
  assumes is_onb B
  shows kf_bound (kf_trace B) = id_cblinfun
  using assms
  by (auto intro!: cblinfun_cinner_eqI simp: kf_bound_from_map kf_trace_is_trace trace_tc_butterfly)

lemma kf_trace_norm_eq1[simp]:
  fixes B :: 'a::{chilbert_space, not_singleton} set
  assumes is_onb B
  shows kf_norm (kf_trace B) = 1
  using assms by (simp add: kf_trace_bound kf_norm_def)

lemma kf_trace_norm_leq1[simp]:
  fixes B :: 'a::chilbert_space set
  assumes is_onb B
  shows kf_norm (kf_trace B)  1
  by (simp add: assms kf_norm_def norm_cblinfun_id_le)

subsection ‹Constant maps›


lift_definition kf_constant_onedim :: ('b,'b) trace_class  ('a::one_dim, 'b::chilbert_space, unit) kraus_family is
  λt::('b,'b) trace_class. if t  0 then
    Set.filter (λ(E,_). E0) ((λv. (vector_to_cblinfun v,())) ` spectral_dec_vecs_tc t)
  else {}
proof (rule CollectI, rename_tac t)
  fix t :: ('b,'b) trace_class
  show kraus_family (if t  0 then Set.filter (λ(E,_). E0) ((λv. (vector_to_cblinfun v :: 'aCL'b,())) ` spectral_dec_vecs_tc t) else {})
  proof (cases t  0)
    case True
    have kraus_family (Set.filter (λ(E,_). E0) ((λv. (vector_to_cblinfun v :: 'aCL'b,())) ` spectral_dec_vecs_tc t))
    proof (intro kraus_familyI bdd_aboveI, rename_tac E)
      fix E :: 'a CL 'a
      assume E  (λF. (E, x)F. E* oCL E) ` {F. finite F  F  Set.filter (λ(E,_). E0) ((λv. (vector_to_cblinfun v, ())) ` spectral_dec_vecs_tc t)}
      then obtain F where E_def: E = ((E, x)F. E* oCL E) and finite F and F  Set.filter (λ(E,_). E0) ((λv. (vector_to_cblinfun v, ())) ` spectral_dec_vecs_tc t)
        by blast
      then obtain F' where F_def: F = (λv. (vector_to_cblinfun v, ())) ` F' and finite F' and F'_subset: F'  spectral_dec_vecs_tc t
        by (meson finite_subset_filter_image)
      have inj: inj_on (λv. (vector_to_cblinfun v :: 'aCL'b, ())) F'
      proof (rule inj_onI, rule ccontr)
        fix x y
        assume x  F' and y  F'
        assume eq: (vector_to_cblinfun x :: 'aCL'b, ()) = (vector_to_cblinfun y, ())
        assume x  y
        have ortho: is_ortho_set (spectral_dec_vecs (from_trace_class t))
          using True
          by (auto intro!: spectral_dec_vecs_ortho trace_class_compact pos_selfadjoint
              simp: selfadjoint_tc.rep_eq from_trace_class_pos)
        with x  y F'_subset x  F' y  F'
        have x C y = 0
          by (auto simp: spectral_dec_vecs_tc.rep_eq is_ortho_set_def)
        then have (vector_to_cblinfun x :: 'aCL'b)* oCL (vector_to_cblinfun y :: 'aCL'b) = 0
          by simp
        with eq have (vector_to_cblinfun x :: 'aCL'b)= 0
          by force
        then have norm x = 0
          by (smt (verit, del_insts) norm_vector_to_cblinfun norm_zero)
        with ortho F'_subset x  F' show False
          by (auto simp: spectral_dec_vecs_tc.rep_eq is_ortho_set_def)
      qed
      have E = ((E, x)F. E* oCL E)
        by (simp add: E_def)
      also have  = (vF'. (vector_to_cblinfun v :: 'aCL'b)* oCL vector_to_cblinfun v)
        unfolding F_def
        apply (subst sum.reindex)
        by (auto intro!: inj)
      also have  = (vF'. ((norm (vector_to_cblinfun v :: 'aCL'b))2) *R 1)
        by (auto intro!: sum.cong simp: power2_norm_eq_cinner scaleR_scaleC)
      also have  = (vF'. (norm (vector_to_cblinfun v :: 'aCL'b))2) *R 1
        by (metis scaleR_left.sum)
      also have  = (vF'. (norm (vector_to_cblinfun v :: 'aCL'b))2) *R 1
        using finite F' by force
      also have   (vspectral_dec_vecs_tc t. (norm (vector_to_cblinfun v :: 'aCL'b))2) *R 1
        apply (intro scaleR_right_mono infsum_mono_neutral)
        using F'_subset
        by (auto intro!: one_dim_cblinfun_one_pos spectral_dec_vec_tc_norm_summable True
            simp: finite F' )
      finally show E  (vspectral_dec_vecs_tc t. (norm (vector_to_cblinfun v :: 'aCL'b))2) *R 1
        by -
    next
      show 0  fst ` Set.filter (λ(E, _). E  0) ((λv. (vector_to_cblinfun v, ())) ` spectral_dec_vecs_tc t)
        by auto
    qed
    with True show ?thesis
      by simp
  next
    case False
    then show ?thesis
      by simp
  qed
qed

definition kf_constant :: ('b,'b) trace_class  ('a::chilbert_space, 'b::chilbert_space, unit) kraus_family where
  kf_constant ρ = kf_flatten (kf_comp (kf_constant_onedim ρ :: (complex,_,_) kraus_family) (kf_trace some_chilbert_basis))

lemma kf_constant_onedim_invalid: ¬ ρ  0  kf_constant_onedim ρ = 0
  apply transfer'
  by simp

lemma kf_constant_invalid: ¬ ρ  0  kf_constant ρ = 0
  by (simp add: kf_constant_def kf_constant_onedim_invalid)

lemma kf_constant_onedim_apply: 
  assumes ρ  0
  shows kf_apply (kf_constant_onedim ρ) σ = one_dim_iso σ *C ρ
proof -
  have kf_apply (kf_constant_onedim ρ) σ
         = ((E,x)Set.filter (λ(E,_). E0) ((λv. (vector_to_cblinfun v, ())) ` spectral_dec_vecs_tc ρ). sandwich_tc E σ)
    by (simp add: assms kf_apply.rep_eq kf_constant_onedim.rep_eq case_prod_unfold)
  also have  = ((E,x)(λv. (vector_to_cblinfun v, ())) ` spectral_dec_vecs_tc ρ. sandwich_tc E σ)
    apply (rule infsum_cong_neutral)
    by auto
  also have  = (v  spectral_dec_vecs_tc ρ. sandwich_tc (vector_to_cblinfun v) σ)
    apply (subst infsum_reindex)
    using vector_to_cblinfun_inj[of UNIV]
    by (auto simp: o_def inj_on_def)
  also have  = (v  spectral_dec_vecs_tc ρ. one_dim_iso σ *C tc_butterfly v v)
    apply (rule infsum_cong)
    apply (subst one_dim_scaleC_1[symmetric])
    apply (rule from_trace_class_inject[THEN iffD1])
    apply (simp only: sandwich_tc_def compose_tcl.rep_eq compose_tcr.rep_eq scaleC_trace_class.rep_eq
        tc_butterfly.rep_eq cblinfun_compose_scaleC_right cblinfun_compose_scaleC_left)
    by (simp flip: butterfly_def_one_dim)
  also have  = one_dim_iso σ *C (v  spectral_dec_vecs_tc ρ. tc_butterfly v v)
    using infsum_scaleC_right by fastforce
  also have  = one_dim_iso σ *C ρ
    by (simp add: assms butterfly_spectral_dec_vec_tc_has_sum infsumI)
  finally show ?thesis
    by -
qed

lemma kf_constant_apply: 
  assumes ρ  0
  shows kf_apply (kf_constant ρ) σ = trace_tc σ *C ρ
  using assms by (simp add: kf_constant_def kf_comp_apply kf_trace_is_trace kf_constant_onedim_apply)

lemma kf_bound_constant_onedim[simp]: 
  fixes ρ :: ('a::chilbert_space, 'a) trace_class
  assumes ρ  0
  shows kf_bound (kf_constant_onedim ρ) = norm ρ *R id_cblinfun
proof (rule cblinfun_cinner_eqI)
  fix ψ :: 'b assume norm ψ = 1
  have ψ C kf_bound (kf_constant_onedim ρ) ψ = trace_tc (kf_apply (kf_constant_onedim ρ) (tc_butterfly ψ ψ))
    by (rule kf_bound_from_map)
  also have  = trace_tc (trace_tc (tc_butterfly ψ ψ) *C ρ)
    by (simp add: kf_constant_onedim_apply assms)
  also have  = trace_tc ρ
    by (metis norm ψ = 1 cinner_complex_def complex_cnj_one complex_vector.vector_space_assms(4) norm_mult norm_one norm_tc_butterfly norm_tc_pos of_real_hom.hom_one one_cinner_one tc_butterfly_pos)
  also have  = ψ C (trace_tc ρ *C id_cblinfun) ψ
    by (metis norm ψ = 1 cblinfun.scaleC_left cinner_scaleC_right cnorm_eq_1 id_apply id_cblinfun.rep_eq of_complex_def one_dim_iso_id one_dim_iso_is_of_complex scaleC_conv_of_complex)
  also have  = ψ C (norm ρ *R id_cblinfun) ψ
    by (simp add: assms norm_tc_pos scaleR_scaleC)
  finally show ψ C kf_bound (kf_constant_onedim ρ) ψ = ψ C (norm ρ *R id_cblinfun) ψ
    by -
qed

lemma kf_bound_constant[simp]: 
  fixes ρ :: ('a::chilbert_space, 'a) trace_class
  assumes ρ  0
  shows kf_bound (kf_constant ρ) = norm ρ *R id_cblinfun
  apply (rule cblinfun_cinner_eqI)
  using assms
  by (simp add: kf_bound_from_map kf_constant_apply trace_tc_butterfly norm_tc_pos scaleR_scaleC trace_tc_scaleC)
  
lemma kf_norm_constant_onedim[simp]:
  assumes ρ  0
  shows kf_norm (kf_constant_onedim ρ) = norm ρ
  using assms
  by (simp add: kf_bound_constant kf_norm_def)

lemma kf_norm_constant:
  assumes ρ  0
  shows kf_norm (kf_constant ρ :: ('a::{chilbert_space,not_singleton},'b::chilbert_space,_) kraus_family) = norm ρ
  using assms by (simp add: kf_norm_def norm_cblinfun_id)

lemma kf_norm_constant_leq:
  shows kf_norm (kf_constant ρ)  norm ρ
  apply (cases ρ  0)
   apply (simp add: kf_norm_def)
   apply (metis Groups.mult_ac(2) mult_cancel_right1 mult_left_mono norm_cblinfun_id_le norm_ge_zero)
  by (simp add: kf_constant_invalid)

lemma kf_comp_constant_right:
  assumes [iff]: t  0
  shows kf_map fst (kf_comp E (kf_constant t)) kr kf_constant (E *kr t)
proof (rule kf_eqI)
  fix ρ :: ('b, 'b) trace_class assume [iff]: ρ  0
  have [simp]: fst -` {()} = UNIV
    by auto
  have [simp]: {()} = UNIV
    by auto
  fix x
  have kf_map fst (kf_comp E (kf_constant t)) *kr @{x} ρ = kf_comp E (kf_constant t) *kr ρ
    by (simp add: kf_apply_on_map)
  also have  = E *kr trace_tc ρ *C t
    by (simp add: kf_comp_apply kf_constant_apply)
  also have  = kf_constant (E *kr t) *kr @{x} ρ
    by (simp add: kf_constant_apply kf_apply_pos kf_apply_scaleC)
  finally show kf_map fst (kf_comp E (kf_constant t)) *kr @{x} ρ = kf_constant (E *kr t) *kr @{x} ρ
    by -
qed

lemma kf_comp_constant_right_weak:
  assumes [iff]: t  0
  shows kf_comp E (kf_constant t) =kr kf_constant (E *kr t)
  by (metis assms kf_apply_map kf_comp_constant_right kf_eq_imp_eq_weak kf_eq_weak_def)


subsection ‹Tensor products›

lemma kf_tensor_raw_bound_aux:
  fixes 𝔈 :: ('a ell2 CL 'b ell2 × 'x) set and 𝔉 :: ('c ell2 CL 'd ell2 × 'y) set
  assumes S. finite S  S  𝔈  ((E, x)S. E* oCL E)  B
  assumes S. finite S  S  𝔉  ((E, x)S. E* oCL E)  C
  assumes finite U
  assumes U  ((λ((E, x), F, y). (E o F, E, F, x, y)) ` (𝔈 × 𝔉))
  shows ((G, z)U. G* oCL G)  B o C
proof -
  from assms(1)[where S={}] have [simp]: B  0
    by simp
  define f :: (('a ell2 CL 'b ell2 × 'x) × ('c ell2 CL 'd ell2 × 'y))  _
    where f = (λ((E,x), (F,y)). (E o F, (E,F,x,y)))
  from assms
  obtain V where V_subset: V  𝔈 × 𝔉 and [simp]: finite V and U = f ` V
    apply (simp flip: f_def)
    by (meson finite_subset_image)
  define W where W = fst ` V × snd ` V
  have inj_on f W
    by (auto intro!: inj_onI simp: f_def)
  from finite V have [simp]: finite W
    using W_def by blast
  have W  V
    by (auto intro!: image_eqI simp: W_def)
  have ((G, z)U. G* oCL G)  ((G, z)f ` W. G* oCL G)
    using U = f ` V V  W
    by (auto intro!: sum_mono2 positive_cblinfun_squareI)
  also have  = (((E,x),(F,y))W. (E o F)* oCL (E o F))
    apply (subst sum.reindex)
    using inj_on f W
    by (auto simp: case_prod_unfold f_def)
  also have  = (((E,x),(F,y))W. (E* oCL E) o (F* oCL F))
    by (simp add: comp_tensor_op tensor_op_adjoint)
  also have  = ((E,x)fst`V. E* oCL E) o ((F,y)snd`V. F* oCL F)
    unfolding W_def
    apply (subst sum.Sigma[symmetric])
      apply simp
     apply simp
    apply (simp add: case_prod_beta tensor_op_cbilinear.sum_left) 
    by (simp add:  tensor_op_cbilinear.sum_right)
  also have   B o C
    using V_subset
    by (auto intro!: tensor_op_mono assms sum_nonneg intro: positive_cblinfun_squareI)
  finally show ?thesis
    by-
qed


lift_definition kf_tensor_raw :: ('a ell2, 'b ell2, 'x) kraus_family  ('c ell2, 'd ell2, 'y) kraus_family  
          (('a×'c) ell2, ('b×'d) ell2, (('a ell2CL'b ell2)×('c ell2CL'd ell2)×'x×'y)) kraus_family is
  λ𝔈 𝔉. (λ((E,x), (F,y)). (E o F, (E,F,x,y))) ` (𝔈×𝔉)
proof (rename_tac 𝔈 𝔉, intro CollectI)
  fix 𝔈 :: ('a ell2 CL 'b ell2 × 'x) set and 𝔉 :: ('c ell2 CL 'd ell2 × 'y) set
  assume 𝔈  Collect kraus_family and 𝔉  Collect kraus_family
  then have kraus_family 𝔈 and kraus_family 𝔉
    by auto
  define tensor where tensor = ((λ((E, x), F, y). (E o F, E, F, x, y)) ` (𝔈 × 𝔉))
  show kraus_family tensor
  proof (intro kraus_familyI)
    from kraus_family 𝔈
    obtain B where B: ((E, x)S. E* oCL E)  B if finite S and S  𝔈 for S
      apply atomize_elim
      by (auto simp: kraus_family_def bdd_above_def)
    from B[where S={}] have [simp]: B  0
      by simp
    from kraus_family 𝔉
    obtain C where C: ((F, x)T. F* oCL F)  C if finite T and T  𝔉 for T
      apply atomize_elim
      by (auto simp: kraus_family_def bdd_above_def)
    have ((G, z)U. G* oCL G)  B o C if finite U and U  tensor for U
      using that by (auto intro!: kf_tensor_raw_bound_aux B C simp: tensor_def)
    then show bdd_above ((λF. (E, x)F. E* oCL E) ` {F. finite F  F  tensor})
      by fast
    show 0  fst ` tensor
    proof (rule notI)
      assume 0  fst ` tensor
      then obtain E F x y where EF0: E o F = 0 and Ex: (E,x)  𝔈 and Fy: (F,y)  𝔉
        by (auto intro!: simp: tensor_def)
      from kraus_family 𝔈 Ex
      have E  0
        by (force simp: kraus_family_def)
      from kraus_family 𝔉 Fy
      have F  0
        by (force simp: kraus_family_def)
      from E  0 F  0 have E o F  0
        using tensor_op_nonzero by blast
      with EF0 show False
        by simp
    qed
  qed
qed

lemma kf_apply_tensor_raw_as_infsum:
  kf_tensor_raw 𝔈 𝔉 *kr ρ = (((E,x),(F,y))Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉. sandwich_tc (E o F) ρ)
proof -
  have inj: inj_on (λ((E, x), F, y). (E o F, E, F, x, y)) (Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉)
    by (auto intro!: inj_onI)
  show kf_apply (kf_tensor_raw 𝔈 𝔉) ρ
      = (((E,x),(F,y))Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉. sandwich_tc (E o F) ρ)
    apply (simp add: kf_apply.rep_eq kf_tensor_raw.rep_eq infsum_reindex inj o_def)
    by (simp add: case_prod_unfold)
qed

lemma kf_apply_tensor_raw:
  shows kf_tensor_raw 𝔈 𝔉 *kr tc_tensor ρ σ = tc_tensor (𝔈 *kr ρ) (𝔉 *kr σ)
proof -
  have inj: inj_on (λ((E, x), F, y). (E o F, E, F, x, y)) (Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉)
    by (auto intro!: inj_onI)
  have [simp]: bounded_linear (λx. tc_tensor x (kf_apply 𝔉 σ))
    by (intro bounded_linear_intros)
  have [simp]: bounded_linear (tc_tensor (sandwich_tc E ρ)) for E
    by (intro bounded_linear_intros)
  have sum2: (λ(E, x). sandwich_tc E ρ) summable_on Rep_kraus_family 𝔈
    using kf_apply_summable by blast
  have sum3: (λ(F, y). sandwich_tc F σ) summable_on Rep_kraus_family 𝔉
    using kf_apply_summable by blast

  from kf_apply_summable[of _ kf_tensor_raw 𝔈 𝔉]
  have sum1: (λ((E, x), F, y). sandwich_tc (E o F) (tc_tensor ρ σ)) summable_on Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉
    apply (simp add: kf_apply.rep_eq kf_tensor_raw.rep_eq summable_on_reindex inj o_def)
    by (simp add: case_prod_unfold)

  have kf_apply (kf_tensor_raw 𝔈 𝔉) (tc_tensor ρ σ)
      = (((E,x),(F,y))Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉. sandwich_tc (E o F) (tc_tensor ρ σ))
    by (rule kf_apply_tensor_raw_as_infsum)
  also have  = ((E,x)Rep_kraus_family 𝔈. (F,y)Rep_kraus_family 𝔉. sandwich_tc (E o F) (tc_tensor ρ σ))
    apply (subst infsum_Sigma_banach[symmetric])
    using sum1 by (auto intro!: simp: case_prod_unfold)
  also have  = ((E,x)Rep_kraus_family 𝔈. (F,y)Rep_kraus_family 𝔉. tc_tensor (sandwich_tc E ρ) (sandwich_tc F σ))
    by (simp add: sandwich_tc_tensor)
  also have  = ((E,x)Rep_kraus_family 𝔈. tc_tensor (sandwich_tc E ρ) ((F,y)Rep_kraus_family 𝔉. (sandwich_tc F σ)))
    apply (subst infsum_bounded_linear[where h=tc_tensor (sandwich_tc _ ρ), symmetric])
      apply (auto intro!: sum3)[2]
    by (simp add: o_def case_prod_unfold)
  also have  = ((E,x)Rep_kraus_family 𝔈. tc_tensor (sandwich_tc E ρ) (kf_apply 𝔉 σ))
    by (simp add: kf_apply_def case_prod_unfold)
  also have  = tc_tensor ((E,x)Rep_kraus_family 𝔈. sandwich_tc E ρ) (kf_apply 𝔉 σ)
    apply (subst infsum_bounded_linear[where h=λx. tc_tensor x (kf_apply 𝔉 σ), symmetric])
      apply (auto intro!: sum2)[2]
    by (simp add: o_def case_prod_unfold)
  also have  = tc_tensor (kf_apply 𝔈 ρ) (kf_apply 𝔉 σ)
    by (simp add: kf_apply_def case_prod_unfold)
  finally show ?thesis
    by -
qed

hide_fact kf_tensor_raw_bound_aux

definition kf_tensor 𝔈 𝔉 = kf_map (λ(E, F, x, y). (x,y)) (kf_tensor_raw 𝔈 𝔉)

lemma kf_apply_tensor:
  kf_tensor 𝔈 𝔉 *kr tc_tensor ρ σ = tc_tensor (𝔈 *kr ρ) (𝔉 *kr σ)
  by (auto intro!: simp: kf_tensor_def kf_apply_map kf_apply_tensor_raw)

lemma kf_apply_tensor_as_infsum:
  kf_tensor 𝔈 𝔉 *kr ρ = (((E,x),(F,y))Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉. sandwich_tc (E o F) ρ)
  by (simp add: kf_tensor_def kf_apply_tensor_raw_as_infsum)


lemma kf_bound_tensor_raw:
  kf_bound (kf_tensor_raw 𝔈 𝔉) = kf_bound 𝔈 o kf_bound 𝔉
proof (rule cblinfun_cinner_tensor_eqI)
  fix ψ φ

  from kf_bound_summable[of kf_tensor_raw 𝔈 𝔉]
  have sum1: summable_on_in cweak_operator_topology (λ((E,x),(F,y)). (E o F)* oCL (E o F))
     (Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉)
    unfolding kf_tensor_raw.rep_eq
    apply (subst (asm) summable_on_in_reindex)
    by (auto simp add: kf_tensor_raw.rep_eq case_prod_unfold inj_on_def o_def)
  have sum4: summable_on_in cweak_operator_topology (λ(E,x). E* oCL E) (Rep_kraus_family 𝔈)
    using kf_bound_summable by blast
  have sum5: summable_on_in cweak_operator_topology (λ(F,y). F* oCL F) (Rep_kraus_family 𝔉)
    using kf_bound_summable by blast
  have sum2: (λ(E, x). ψ C ((E* oCL E) *V ψ)) abs_summable_on Rep_kraus_family 𝔈
    using kf_bound_summable by (auto intro!: summable_on_in_cweak_operator_topology_pointwise 
        simp add: case_prod_unfold simp flip: summable_on_iff_abs_summable_on_complex
        simp del: cblinfun_apply_cblinfun_compose)
  have sum3: (λ(F,y). φ C ((F* oCL F) *V φ)) abs_summable_on Rep_kraus_family 𝔉
    using kf_bound_summable by (auto intro!: summable_on_in_cweak_operator_topology_pointwise
        simp add: case_prod_unfold simp flip: summable_on_iff_abs_summable_on_complex
        simp del: cblinfun_apply_cblinfun_compose)

  have (ψ s φ) C (kf_bound (kf_tensor_raw 𝔈 𝔉) *V ψ s φ)
      = (ψ s φ) C
        (infsum_in cweak_operator_topology ((λ(E, x). E* oCL E)  (λ((E, x), F, y). (E o F, E, F, x, y)))
          (Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉) *V
         ψ s φ)
    unfolding kf_bound.rep_eq kf_tensor_raw.rep_eq
    apply (subst infsum_in_reindex)
    by (auto simp add: inj_on_def case_prod_unfold)
  also have  = (ψ s φ) C
        (infsum_in cweak_operator_topology (λ((E,x),(F,y)). (E o F)* oCL (E o F))
            (Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉) *V
         ψ s φ)
    by (simp add: o_def case_prod_unfold)
  also have  = (((E,x),(F,y))  Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉.
                         (ψ s φ) C ((E o F)* oCL (E o F)) (ψ s φ))
    apply (subst infsum_in_cweak_operator_topology_pointwise)
    using sum1 by (auto intro!: simp: case_prod_unfold)
  also have  = (((E,x),(F,y))  Rep_kraus_family 𝔈 × Rep_kraus_family 𝔉.
                     (ψ C (E* oCL E) ψ) * (φ C (F* oCL F) φ))
    apply (rule infsum_cong)
    by (simp_all add: tensor_op_adjoint tensor_op_ell2)
  also have  = ((E,x)  Rep_kraus_family 𝔈. ψ C (E* oCL E) ψ)
                      * ((F,y)  Rep_kraus_family 𝔉. φ C (F* oCL F) φ)
    apply (subst infsum_product')
    using sum2 sum3 by (simp_all add: case_prod_unfold)
  also have  = (ψ C kf_bound 𝔈 ψ) * (φ C kf_bound 𝔉 φ)
    unfolding kf_bound.rep_eq case_prod_unfold
    apply (subst infsum_in_cweak_operator_topology_pointwise[symmetric])
    using sum4 apply (simp add: case_prod_unfold)
    apply (subst infsum_in_cweak_operator_topology_pointwise[symmetric])
    using sum5 apply (simp add: case_prod_unfold)
    by (rule refl)
  also have  = (ψ s φ) C ((kf_bound 𝔈 o kf_bound 𝔉) *V ψ s φ)
    by (simp add: tensor_op_ell2)
  finally show (ψ s φ) C (kf_bound (kf_tensor_raw 𝔈 𝔉) *V ψ s φ) =
       (ψ s φ) C ((kf_bound 𝔈 o kf_bound 𝔉) *V ψ s φ)
    by -
qed


lemma kf_bound_tensor:
  kf_bound (kf_tensor 𝔈 𝔉) = kf_bound 𝔈 o kf_bound 𝔉
  by (simp add: kf_tensor_def kf_map_bound kf_bound_tensor_raw) 

lemma kf_norm_tensor:
  kf_norm (kf_tensor 𝔈 𝔉) = kf_norm 𝔈 * kf_norm 𝔉
  by (auto intro!: norm_cblinfun_mono
      simp add: kf_norm_def kf_bound_tensor
      simp flip: tensor_op_norm)

lemma kf_tensor_cong_weak:
  assumes 𝔈 =kr 𝔈'
  assumes 𝔉 =kr 𝔉'
  shows kf_tensor 𝔈 𝔉 =kr kf_tensor 𝔈' 𝔉'
proof (rule kf_eq_weakI)
  show kf_apply (kf_tensor 𝔈 𝔉) ρ = kf_apply (kf_tensor 𝔈' 𝔉') ρ for ρ
  proof (rule eq_from_separatingI2x[where x=ρ, OF separating_set_bounded_clinear_tc_tensor])
    show bounded_clinear (kf_apply (kf_tensor 𝔈 𝔉))
      by (simp add: kf_apply_bounded_clinear)
    show bounded_clinear (kf_apply (kf_tensor 𝔈' 𝔉'))
      by (simp add: kf_apply_bounded_clinear)
    have 𝔈𝔈': kf_apply 𝔈 ρ = kf_apply 𝔈' ρ for ρ
      by (metis assms(1) kf_eq_weak_def)
    have 𝔉𝔉': kf_apply 𝔉 ρ = kf_apply 𝔉' ρ for ρ
      by (metis assms(2) kf_eq_weak_def)
    fix ρ :: ('a ell2, 'a ell2) trace_class and σ :: ('e ell2, 'e ell2) trace_class
    show kf_apply (kf_tensor 𝔈 𝔉) (tc_tensor ρ σ) = kf_apply (kf_tensor 𝔈' 𝔉') (tc_tensor ρ σ)
      by (auto intro!: simp: kf_apply_tensor 𝔈𝔈' 𝔉𝔉'
          simp flip: tensor_ell2_ket tensor_tc_butterfly)
  qed
qed

lemma kf_filter_tensor:
  kf_filter (λ(x,y). P x  Q y) (kf_tensor 𝔈 𝔉) = kf_tensor (kf_filter P 𝔈) (kf_filter Q 𝔉)
  apply (simp add: kf_tensor_def kf_filter_map)
  apply (rule arg_cong[where f=kf_map _])
  apply transfer
  by (force simp add: image_iff case_prod_unfold)

lemma kf_filter_tensor_singleton:
  kf_filter ((=) x) (kf_tensor 𝔈 𝔉) = kf_tensor (kf_filter ((=) (fst x)) 𝔈) (kf_filter ((=) (snd x)) 𝔉)
  by (simp add: kf_filter_tensor[symmetric] case_prod_unfold prod_eq_iff)

lemma kf_tensor_cong:
  fixes 𝔈 𝔈' :: ('a ell2, 'b ell2, 'x) kraus_family
    and 𝔉 𝔉' :: ('c ell2, 'd ell2, 'y) kraus_family
  assumes 𝔈 kr 𝔈'
  assumes 𝔉 kr 𝔉'
  shows kf_tensor 𝔈 𝔉 kr kf_tensor 𝔈' 𝔉'
proof (rule kf_eqI)
  fix xy :: 'x × 'y and ρ
  obtain x y where [simp]: xy = (x,y)
    by fastforce
  have aux1: (λxy'. xy' = (x, y)) = (λ(x', y'). x' = x  y' = y)
    by auto
  have kf_apply_on (kf_tensor 𝔈 𝔉) {xy}
         = kf_apply (kf_tensor (kf_filter (λz. z = x) 𝔈) (kf_filter (λz. z = y) 𝔉))
    apply (simp add: kf_apply_on_def aux1 kf_filter_tensor)
    apply (subst aux1)
    by (simp add: kf_apply_on_def aux1 kf_filter_tensor)
  also have  = kf_apply (kf_tensor (kf_filter (λz. z = x) 𝔈') (kf_filter (λz. z = y) 𝔉'))
    apply (rule ext)
    apply (rule kf_apply_eqI)
    apply (rule kf_tensor_cong_weak)
    by (auto intro!: kf_filter_cong_weak assms)
  also have  = kf_apply_on (kf_tensor 𝔈' 𝔉') {xy}
    apply (simp add: kf_apply_on_def aux1 kf_filter_tensor)
    apply (subst aux1)
    by (simp add: kf_apply_on_def aux1 kf_filter_tensor)
  finally show kf_tensor 𝔈 𝔉 *kr @{xy} ρ = kf_tensor 𝔈' 𝔉' *kr @{xy} ρ
    by simp
qed


lemma kf_tensor_compose_distrib_weak:
  shows kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 )
     =kr kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 )
  by (auto intro!: eq_from_separatingI2[OF separating_set_bounded_clinear_tc_tensor]
      kf_apply_bounded_clinear comp_bounded_clinear
      simp: kf_eq_weak_def kf_apply_tensor kf_comp_apply)

lemma kf_tensor_compose_distrib:
  shows kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 )
     kr kf_map (λ((e,g),(f,h)). ((e,f),(g,h))) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 ))
proof (rule kf_eqI_from_filter_eq_weak)
  fix efgh :: ('e × 'f) × ('g × 'h)
  obtain e f g h where efgh: efgh = ((e,f),(g,h))
    apply atomize_elim
    apply (cases efgh)
    by auto
  have kf_filter ((=) efgh) (kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 ))
         =kr kf_filter (λ(x,y). (e,f)=x  (g,h)=y) (kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 ))
    by (smt (verit, best) case_prod_unfold kf_eq_def kf_filter_cong_weak split_pairs2 efgh)
  also have  = kf_tensor (kf_filter ((=) (e,f)) (kf_comp 𝔈 𝔉)) (kf_filter ((=) (g,h)) (kf_comp 𝔊 ))
    by (simp add: kf_filter_tensor)
  also have  = kf_tensor (kf_filter (λ(e',f'). f=f'  e=e') (kf_comp 𝔈 𝔉)) (kf_filter (λ(g',h'). h=h'  g=g') (kf_comp 𝔊 ))
    apply (intro arg_cong2[where f=kf_tensor] arg_cong2[where f=kf_filter])
    by auto
  also have  = kf_tensor (kf_comp (kf_filter ((=) f) 𝔈) (kf_filter ((=) e) 𝔉)) (kf_comp (kf_filter ((=) h) 𝔊) (kf_filter ((=) g) ))
    by (simp add: kf_filter_comp)
  also have  =kr kf_comp (kf_tensor (kf_filter ((=) f) 𝔈) (kf_filter ((=) h) 𝔊)) (kf_tensor (kf_filter ((=) e) 𝔉) (kf_filter ((=) g) ))
    by (rule kf_tensor_compose_distrib_weak)
  also have  =kr kf_filter (λ((e',g'), (f',h')). (f = f'  h = h')  (e = e'  g = g')) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 ))
    by (simp add: case_prod_unfold flip: kf_filter_tensor kf_filter_comp)
  also have  =kr kf_map (λ((e,g),(f,h)). ((e,f),(g,h))) (kf_filter (λ((e',g'), (f',h')). (f = f'  h = h')  (e = e'  g = g')) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 )))
    by (simp add: kf_eq_weak_def)
  also have  =kr kf_filter (λ((e',f'), (g',h')). (f = f'  h = h')  (e = e'  g = g')) (kf_map (λ((e,g),(f,h)). ((e,f),(g,h))) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 )))
    by (simp add: kf_filter_map case_prod_unfold)
  also have  = kf_filter ((=) efgh) (kf_map (λ((e,g),(f,h)). ((e,f),(g,h))) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 )))
    apply (rule arg_cong2[where f=kf_filter])
    by (auto simp: efgh)
  finally show kf_filter ((=) efgh) (kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 )) =kr 
    by -
qed

lemma kf_tensor_compose_distrib':
  shows kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 )
     kr kf_map (λ((e,f),(g,h)). ((e,g),(f,h))) (kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 ))
proof -
  have aux: ((λ((e,f),(g,h)). ((e,g),(f,h))) o (λ((e,g),(f,h)). ((e,f),(g,h)))) = id
    by auto
  have kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 ) kr kf_map ((λ((e,f),(g,h)). ((e,g),(f,h))) o (λ((e,g),(f,h)). ((e,f),(g,h)))) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 ))
    by (simp add: aux kf_eq_sym kf_map_id)
  also have  kr kf_map (λ((e,f),(g,h)). ((e,g),(f,h))) (kf_map (λ((e,g),(f,h)). ((e,f),(g,h))) (kf_comp (kf_tensor 𝔈 𝔊) (kf_tensor 𝔉 )))
    using kf_eq_sym kf_map_twice by blast
  also have  kr kf_map (λ((e,f),(g,h)). ((e,g),(f,h))) (kf_tensor (kf_comp 𝔈 𝔉) (kf_comp 𝔊 ))
    using kf_eq_sym kf_map_cong[OF refl] kf_tensor_compose_distrib by blast
  finally show ?thesis
    by -
qed


definition kf_tensor_right :: ('extra ell2, 'extra ell2) trace_class  ('qu ell2, ('qu×'extra) ell2, unit) kraus_family where
  ― ‹termkf_tensor_right ρ maps termσ to termσ o ρ
  kf_tensor_right ρ = kf_map_inj (λ_. ()) (kf_comp (kf_tensor kf_id (kf_constant_onedim ρ)) (kf_of_op (tensor_ell2_right (ket ()))))
definition kf_tensor_left :: ('extra ell2, 'extra ell2) trace_class  ('qu ell2, ('extra×'qu) ell2, unit) kraus_family where
  ― ‹termkf_tensor_right ρ maps termσ to termρ o σ
  kf_tensor_left ρ = kf_map_inj (λ_. ()) (kf_comp (kf_tensor (kf_constant_onedim ρ) kf_id) (kf_of_op (tensor_ell2_left (ket ()))))

lemma kf_apply_tensor_right[simp]:
  assumes ρ  0
  shows kf_tensor_right ρ *kr σ = tc_tensor σ ρ
proof -
  have *: sandwich_tc (tensor_ell2_right (ket ())) σ = tc_tensor σ (tc_butterfly (ket()) (ket()))
    apply transfer'
    using sandwich_tensor_ell2_right' by blast
  show ?thesis
    by (simp add: kf_tensor_right_def kf_apply_map_inj inj_on_def kf_comp_apply
        kf_of_op_apply * kf_apply_tensor kf_constant_onedim_apply assms trace_tc_butterfly
        flip: trace_tc_one_dim_iso)
qed
lemma kf_apply_tensor_left[simp]:
  assumes ρ  0
  shows kf_tensor_left ρ *kr σ = tc_tensor ρ σ
proof -
  have *: sandwich_tc (tensor_ell2_left (ket ())) σ = tc_tensor (tc_butterfly (ket()) (ket())) σ
    apply transfer'
    using sandwich_tensor_ell2_left' by blast
  show ?thesis
    by (simp add: kf_tensor_left_def kf_apply_map_inj inj_on_def kf_comp_apply
        kf_of_op_apply * kf_apply_tensor kf_constant_onedim_apply assms trace_tc_butterfly
        flip: trace_tc_one_dim_iso)
qed

lemma kf_bound_tensor_right[simp]:
  assumes ρ  0
  shows kf_bound (kf_tensor_right ρ) = norm ρ *C id_cblinfun
proof (rule cblinfun_cinner_eqI)
  fix ψ :: 'b ell2
  have ψ C kf_bound (kf_tensor_right ρ) ψ = trace_tc (kf_tensor_right ρ *kr tc_butterfly ψ ψ)
    by (simp add: kf_bound_from_map)
  also have  = trace_tc (tc_tensor (tc_butterfly ψ ψ) ρ)
    using assms by (simp add: kf_apply_tensor_right)
  also have  = trace_tc (tc_butterfly ψ ψ) * trace_tc ρ
    by (metis (no_types, lifting) assms norm_tc_pos norm_tc_tensor of_real_mult tc_butterfly_pos tc_tensor_pos)
  also have  = norm ρ * trace_tc (tc_butterfly ψ ψ)
    by (simp add: assms norm_tc_pos)
  also have  = ψ C (norm ρ *C id_cblinfun) ψ
    by (simp add: trace_tc_butterfly)
  finally show ψ C (kf_bound (kf_tensor_right ρ)) ψ = ψ C (norm ρ *C id_cblinfun) ψ
    by -
qed
lemma kf_bound_tensor_left[simp]:
  assumes ρ  0
  shows kf_bound (kf_tensor_left ρ) = norm ρ *C id_cblinfun
proof (rule cblinfun_cinner_eqI)
  fix ψ :: 'b ell2
  have ψ C kf_bound (kf_tensor_left ρ) ψ = trace_tc (kf_tensor_left ρ *kr tc_butterfly ψ ψ)
    by (simp add: kf_bound_from_map)
  also have  = trace_tc (tc_tensor ρ (tc_butterfly ψ ψ))
    using assms by (simp add: kf_apply_tensor_left)
  also have  = trace_tc ρ * trace_tc (tc_butterfly ψ ψ)
    by (metis (no_types, lifting) assms norm_tc_pos norm_tc_tensor of_real_mult tc_butterfly_pos tc_tensor_pos)
  also have  = norm ρ * trace_tc (tc_butterfly ψ ψ)
    by (simp add: assms norm_tc_pos)
  also have  = ψ C (norm ρ *C id_cblinfun) ψ
    by (simp add: trace_tc_butterfly)
  finally show ψ C (kf_bound (kf_tensor_left ρ)) ψ = ψ C (norm ρ *C id_cblinfun) ψ
    by -
qed


lemma kf_norm_tensor_right[simp]:
  assumes ρ  0
  shows kf_norm (kf_tensor_right ρ) = norm ρ
  using assms by (simp add: kf_norm_def)
lemma kf_norm_tensor_left[simp]:
  assumes ρ  0
  shows kf_norm (kf_tensor_left ρ) = norm ρ
  using assms by (simp add: kf_norm_def)


lemma kf_trace_preserving_tensor:
  assumes kf_trace_preserving 𝔈 and kf_trace_preserving 𝔉
  shows kf_trace_preserving (kf_tensor 𝔈 𝔉)
  by (metis assms(1,2) kf_bound_tensor kf_trace_preserving_iff_bound_id tensor_id)

lemma kf_trace_reducing_tensor:
  assumes kf_trace_reducing 𝔈 and kf_trace_reducing 𝔉
  shows kf_trace_reducing (kf_tensor 𝔈 𝔉)
  using assms 
  by (auto intro!: mult_le_one simp: kf_trace_reducing_iff_norm_leq1 kf_norm_tensor kf_norm_geq0)

lemma kf_tensor_map_left:
  kf_tensor (kf_map f 𝔈) 𝔉 kr kf_map (apfst f) (kf_tensor 𝔈 𝔉)
proof (rule kf_eqI_from_filter_eq_weak)
  fix xy :: 'e × 'f
  obtain x y where xy: xy = (x,y)
    apply atomize_elim
    by (auto intro!: prod.exhaust)
  have kf_filter ((=) xy) (kf_tensor (kf_map f 𝔈) 𝔉) = kf_tensor (kf_filter ((=)x) (kf_map f 𝔈)) (kf_filter ((=)y) 𝔉)
    by (auto intro!: arg_cong2[where f=kf_filter] simp add: xy simp flip: kf_filter_tensor)
  also have  = kf_tensor (kf_map f (kf_filter (λx'. x = f x') 𝔈)) (kf_filter ((=) y) 𝔉)
    by (simp add: kf_filter_map)
  also have  =kr kf_tensor (kf_filter (λx'. x = f x') 𝔈) (kf_filter ((=) y) 𝔉)
    apply (rule kf_tensor_cong_weak)
    by (simp_all add: kf_eq_weak_def)
  also have  =kr kf_filter (λ(x',y'). x = f x'  y = y') (kf_tensor 𝔈 𝔉)
    by (auto intro!: arg_cong2[where f=kf_filter] simp add: xy simp flip: kf_filter_tensor)
  also have  =kr kf_map (apfst f) (kf_filter (λ(x',y'). x = f x'  y = y') (kf_tensor 𝔈 𝔉))
    by (simp add: kf_eq_weak_def)
  also have  = kf_filter ((=) xy) (kf_map (apfst f) (kf_tensor 𝔈 𝔉))
    by (auto intro!: arg_cong2[where f=kf_map] arg_cong2[where f=kf_filter] simp add: xy kf_filter_map simp flip: )
  finally show kf_filter ((=) xy) (kf_tensor (kf_map f 𝔈) 𝔉) =kr 
    by -
qed

lemma kf_tensor_map_right:
  kf_tensor 𝔈 (kf_map f 𝔉) kr kf_map (apsnd f) (kf_tensor 𝔈 𝔉)
proof (rule kf_eqI_from_filter_eq_weak)
  fix xy :: 'e × 'f
  obtain x y where xy: xy = (x,y)
    apply atomize_elim
    by (auto intro!: prod.exhaust)
  have kf_filter ((=) xy) (kf_tensor 𝔈 (kf_map f 𝔉)) = kf_tensor (kf_filter ((=)x) 𝔈) (kf_filter ((=)y) (kf_map f 𝔉))
    by (auto intro!: arg_cong2[where f=kf_filter] simp add: xy simp flip: kf_filter_tensor)
  also have  = kf_tensor (kf_filter ((=)x) 𝔈) (kf_map f (kf_filter (λy'. y = f y') 𝔉))
    by (simp add: kf_filter_map)
  also have  =kr kf_tensor (kf_filter ((=)x) 𝔈) (kf_filter (λy'. y = f y') 𝔉)
    apply (rule kf_tensor_cong_weak)
    by (simp_all add: kf_eq_weak_def)
  also have  =kr kf_filter (λ(x',y'). x = x'  y = f y') (kf_tensor 𝔈 𝔉)
    by (auto intro!: arg_cong2[where f=kf_filter] simp add: xy simp flip: kf_filter_tensor)
  also have  =kr kf_map (apsnd f) (kf_filter (λ(x',y'). x = x'  y = f y') (kf_tensor 𝔈 𝔉))
    by (simp add: kf_eq_weak_def)
  also have  = kf_filter ((=) xy) (kf_map (apsnd f) (kf_tensor 𝔈 𝔉))
    by (auto intro!: arg_cong2[where f=kf_map] arg_cong2[where f=kf_filter] simp add: xy kf_filter_map simp flip: )
  finally show kf_filter ((=) xy) (kf_tensor 𝔈 (kf_map f 𝔉)) =kr 
    by -
qed

lemma kf_tensor_map_both:
  kf_tensor (kf_map f 𝔈) (kf_map g 𝔉) kr kf_map (map_prod f g) (kf_tensor 𝔈 𝔉)
  apply (rule kf_tensor_map_left[THEN kf_eq_trans])
  apply (rule kf_map_cong[THEN kf_eq_trans, OF refl])
   apply (rule kf_tensor_map_right)
  apply (rule kf_map_twice[THEN kf_eq_trans])
  by (simp add: o_def map_prod_def case_prod_unfold)

lemma kf_tensor_raw_map_inj_both:
  kf_tensor_raw (kf_map_inj f 𝔈) (kf_map_inj g 𝔉) = kf_map_inj (λ(E,F,x,y). (E,F,f x,g y)) (kf_tensor_raw 𝔈 𝔉)
  apply (transfer' fixing: f g)
  by force

lemma kf_domain_tensor_raw_subset:
  kf_domain (kf_tensor_raw 𝔈 𝔉)  kf_operators 𝔈 × kf_operators 𝔉 × kf_domain 𝔈 × kf_domain 𝔉
  apply transfer'
  by force

lemma kf_tensor_map_inj_both:
  assumes inj_on f (kf_domain 𝔈)
  assumes inj_on g (kf_domain 𝔉)
  shows kf_tensor (kf_map_inj f 𝔈) (kf_map_inj g 𝔉) = kf_map_inj (map_prod f g) (kf_tensor 𝔈 𝔉)
proof -
  from assms
  have inj1: inj_on (λ(E, F, x, y). (E, F, f x, g y)) (kf_domain (kf_tensor_raw 𝔈 𝔉))
    using kf_domain_tensor_raw_subset[of 𝔈 𝔉]
    apply (simp add: inj_on_def ball_def split!: prod.split)
    by blast+
  from assms
  have inj2: inj_on (map_prod f g) ((λ(E, F, x). x) ` kf_domain (kf_tensor_raw 𝔈 𝔉))
    using kf_domain_tensor_raw_subset[of 𝔈 𝔉]
    apply (simp add: inj_on_def ball_def split!: prod.split)
    by blast+
  have kf_tensor (kf_map_inj f 𝔈) (kf_map_inj g 𝔉) = kf_map (λ(E, F, y). y) (kf_map_inj (λ(E, F, x, y). (E, F, f x, g y)) (kf_tensor_raw 𝔈 𝔉))
    by (simp add: kf_tensor_def kf_tensor_raw_map_inj_both id_def)
  also have  = kf_map (λ(E, F, x, y). (f x, g y)) (kf_tensor_raw 𝔈 𝔉)
    apply (subst kf_map_kf_map_inj_comp)
     apply (rule inj1)
    by (simp add: case_prod_unfold o_def)
  also have  = kf_map_inj (map_prod f g) (kf_map (λ(E, F, x). x) (kf_tensor_raw 𝔈 𝔉))
    apply (subst kf_map_inj_kf_map_comp)
    apply (rule inj2)
    by (simp add: o_def case_prod_unfold map_prod_def)
  also have  = kf_map_inj (map_prod f g) (kf_tensor 𝔈 𝔉)
    by (simp add: kf_tensor_def kf_tensor_raw_map_inj_both id_def)
  finally show ?thesis
    by -
qed

lemma kf_operators_tensor_raw:
  shows kf_operators (kf_tensor_raw 𝔈 𝔉) = {E o F | E F. E  kf_operators 𝔈  F  kf_operators 𝔉}
  apply (simp add: kf_operators.rep_eq kf_tensor_raw.rep_eq)
  by (force simp: case_prod_unfold)

lemma kf_operators_tensor:
  shows kf_operators (kf_tensor 𝔈 𝔉)  span {E o F | E F. E  kf_operators 𝔈  F  kf_operators 𝔉}
proof -
  have kf_operators (kf_tensor 𝔈 𝔉)  span (kf_operators (kf_tensor_raw 𝔈 𝔉))
    by (simp add: kf_operators_kf_map kf_tensor_def)
  also have  = span {E o F | E F. E  kf_operators 𝔈  F  kf_operators 𝔉}
    by (metis kf_operators_tensor_raw)
  finally show ?thesis
    by -
qed

lemma kf_domain_tensor: kf_domain (kf_tensor 𝔈 𝔉) = kf_domain 𝔈 × kf_domain 𝔉
proof (intro Set.set_eqI iffI)
  fix xy assume xy_dom: xy  kf_domain (kf_tensor 𝔈 𝔉)
  obtain x y where xy: xy = (x,y)
    by (auto simp: prod_eq_iff)
  from xy_dom obtain E F where (E,F,x,y)  kf_domain (kf_tensor_raw 𝔈 𝔉)
    by (force simp add: kf_tensor_def xy)
  then obtain EF where EFEFxy: (EF,E,F,x,y)  Rep_kraus_family (kf_tensor_raw 𝔈 𝔉)
    by (auto simp: kf_domain_def)
  then have EF = E o F
    by (force simp: kf_tensor_raw.rep_eq case_prod_unfold)
  from EFEFxy have EF  0
    using Rep_kraus_family
    by (force simp: kraus_family_def)
  with  EF = E o F have E  0 and F  0
    by fastforce+
  from EFEFxy have (E,x)  Rep_kraus_family 𝔈
    apply (transfer' fixing: E x)
    by auto
  with E  0 have x  kf_domain 𝔈
    apply (transfer' fixing: E x)
    by force
  from EFEFxy have (F,y)  Rep_kraus_family 𝔉
    apply (transfer' fixing: F y)
    by auto
  with F  0 have y  kf_domain 𝔉
    apply (transfer' fixing: F y)
    by force
  from  x  kf_domain 𝔈  y  kf_domain 𝔉
  show xy  kf_domain 𝔈 × kf_domain 𝔉
    by (simp add: xy)
next
  fix xy assume xy_dom: xy  kf_domain 𝔈 × kf_domain 𝔉
  then obtain x y where xy: xy = (x,y) and xE: x  kf_domain 𝔈 and yF: y  kf_domain 𝔉
    by (auto simp: prod_eq_iff)
  from xE obtain E where Ex: (E,x)  Rep_kraus_family 𝔈
    by (auto simp: kf_domain_def)
  from yF obtain F where Fy: (F,y)  Rep_kraus_family 𝔉
    by (auto simp: kf_domain_def)
  from Ex Fy have (E o F, E, F, x, y)  Rep_kraus_family (kf_tensor_raw 𝔈 𝔉)
    by (force simp: kf_tensor_raw.rep_eq case_prod_unfold)
  moreover 
  have E  0 and F  0
    using Ex Fy Rep_kraus_family
    by (force simp: kraus_family_def)+
  then have E o F  0
    by (simp add: tensor_op_nonzero)
  ultimately have (E, F, x, y)  kf_domain (kf_tensor_raw 𝔈 𝔉)
    by (force simp: kf_domain_def)
  then show xy  kf_domain (kf_tensor 𝔈 𝔉)
    by (force simp: kf_tensor_def xy case_prod_unfold)
qed

lemma kf_tensor_raw_0_left[simp]: kf_tensor_raw 0 𝔈 = 0
  apply transfer'
  by simp

lemma kf_tensor_raw_0_right[simp]: kf_tensor_raw 𝔈 0 = 0
  apply transfer'
  by simp

lemma kf_tensor_0_left[simp]: kf_tensor 0 𝔈 = 0
  by (simp add: kf_tensor_def)

lemma kf_tensor_0_right[simp]: kf_tensor 𝔈 0 = 0
  by (simp add: kf_tensor_def)

lemma kf_tensor_of_op:
  kf_tensor (kf_of_op A) (kf_of_op B) = kf_map (λ(). ((),())) (kf_of_op (A o B))
proof -
  wlog Aneq0: A  0
    using negation
    by simp
  wlog Bneq0: B  0 keeping Aneq0
    using negation
    by simp
  have [simp]: (λ(). ((), ())) = Pair ()
    by auto
  have kf_tensor_raw (kf_of_op A) (kf_of_op B) = kf_map_inj (λ(). (A, B, (), ())) (kf_of_op (A o B))
    apply (transfer' fixing: A B)
    by (simp add: case_unit_Unity tensor_op_nonzero)
  then show ?thesis
    by (simp add: kf_tensor_def kf_map_kf_map_inj_comp inj_on_def o_def case_unit_Unity)
qed


subsection ‹Partial trace›

definition kf_partial_trace_right :: (('a×'b) ell2, 'a ell2, 'b) kraus_family where
  kf_partial_trace_right = kf_map (λ((_,b),_). inv ket b)
  (kf_comp (kf_of_op (tensor_ell2_right (ket ())*))
   (kf_tensor kf_id (kf_trace (range ket))))

definition kf_partial_trace_left :: (('a×'b) ell2, 'b ell2, 'a) kraus_family where
  kf_partial_trace_left = kf_map_inj snd (kf_comp kf_partial_trace_right (kf_of_op swap_ell2))

lemma partial_trace_is_kf_partial_trace: 
  fixes t :: (('a × 'b) ell2, ('a × 'b) ell2) trace_class
  shows partial_trace t = kf_partial_trace_right *kr t
proof -
  have partial_trace t = kf_apply (kf_of_op (tensor_ell2_right (ket ())*))
    (kf_apply (kf_tensor kf_id (kf_trace (range ket))) t)
  proof (rule eq_from_separatingI2x[where x=t, OF separating_set_bounded_clinear_tc_tensor])
    show bounded_clinear partial_trace
      by simp
    show bounded_clinear
     (λt. kf_apply (kf_of_op (tensor_ell2_right (ket ())*))
           (kf_apply (kf_tensor kf_id (kf_trace (range ket))) t))
      by (intro bounded_linear_intros)
    fix ρ :: ('a ell2, 'a ell2) trace_class and σ :: ('b ell2, 'b ell2) trace_class
    have trace (from_trace_class σ) *C from_trace_class ρ =
      tensor_ell2_right (ket ())* oCL from_trace_class ρ o of_complex (trace (from_trace_class σ)) oCL tensor_ell2_right (ket ())
      by (auto intro!: cblinfun_eqI simp: tensor_op_ell2 ket_CARD_1_is_1)
    then show partial_trace (tc_tensor ρ σ) =
           kf_apply (kf_of_op (tensor_ell2_right (ket ())*))
            (kf_apply (kf_tensor kf_id (kf_trace (range ket))) (tc_tensor ρ σ))
      by (auto intro!: from_trace_class_inject[THEN iffD1]
          simp: partial_trace_tensor kf_apply_tensor kf_trace_is_trace kf_of_op_apply
          from_trace_class_sandwich_tc sandwich_apply trace_tc.rep_eq tc_tensor.rep_eq scaleC_trace_class.rep_eq)
  qed
  then show ?thesis
    by (simp add: kf_partial_trace_right_def kf_comp_apply)
qed

lemma partial_trace_ignores_kraus_family:
  assumes kf_trace_preserving 𝔈
  shows partial_trace (kf_tensor 𝔉 𝔈 *kr ρ) = 𝔉 *kr partial_trace ρ
proof (rule eq_from_separatingI2x[where x=ρ, OF separating_set_bounded_clinear_tc_tensor])
  show bounded_clinear (λa. partial_trace (kf_tensor 𝔉 𝔈 *kr a))
    by (intro bounded_linear_intros)
  show bounded_clinear (λa. 𝔉 *kr partial_trace a)
    by (intro bounded_linear_intros)
  fix ρ :: ('e ell2, 'e ell2) trace_class and σ :: ('a ell2, 'a ell2) trace_class
  from assms
  show partial_trace (kf_tensor 𝔉 𝔈 *kr tc_tensor ρ σ) =
           𝔉 *kr partial_trace (tc_tensor ρ σ)
    by (simp add: kf_apply_tensor partial_trace_tensor kf_trace_preserving_def kf_apply_scaleC)
qed

lemma kf_partial_trace_bound[simp]:
  shows kf_bound kf_partial_trace_right = id_cblinfun
  by (simp add: kf_partial_trace_right_def kf_map_bound
      unitary_tensor_ell2_right_CARD_1 kf_bound_comp_iso kf_bound_tensor
      kf_trace_bound)

lemma kf_partial_trace_norm[simp]:
  shows kf_norm kf_partial_trace_right = 1
  by (simp add: kf_norm_def)

lemma kf_partial_trace_right_apply_singleton:
    kf_partial_trace_right *kr @{x} ρ = sandwich_tc (tensor_ell2_right (ket x)*) ρ
proof -
  have kf_partial_trace_right *kr @{x} (tc_tensor (tc_butterfly (ket a) (ket b)) (tc_butterfly (ket c) (ket d))) =
       sandwich_tc (tensor_ell2_right (ket x)*) (tc_tensor (tc_butterfly (ket a) (ket b)) (tc_butterfly (ket c) (ket d))) for a b :: 'a and c d :: 'b
  proof -
    have aux1: (λxa. (case xa of (x, xa)  (case x of (uu_, b)  λ_. inv ket b) xa)  {x}) = (λ(e,f). True  inv ket (snd e) = x)
      by auto
    have aux2: (λe. inv ket (snd e) = x) = (λ(a,b). True  inv ket b = x)
      by auto
    have kf_partial_trace_right *kr @{x} (tc_tensor (tc_butterfly (ket a) (ket b)) (tc_butterfly (ket c) (ket d))) =
       sandwich_tc (tensor_ell2_right (ket ())*)
        (tc_tensor (tc_butterfly (ket a) (ket b)) (kf_apply (kf_filter (λb. inv ket b = x) (kf_trace (range ket))) (tc_butterfly (ket c) (ket d))))
      by (auto simp only: kf_apply_on_def kf_partial_trace_right_def
          kf_filter_map aux1 kf_filter_comp kf_of_op_apply
          kf_filter_true kf_filter_tensor aux2 kf_apply_map
          kf_comp_apply o_def kf_apply_tensor kf_id_apply)
    also have  = sandwich_tc (tensor_ell2_right (ket ())*) (tc_tensor (tc_butterfly (ket a) (ket b)) (of_bool (x=c  x=d) *R tc_butterfly (ket ()) (ket ())))
    proof (rule arg_cong[where f=λx. sandwich_tc _ (tc_tensor _ x)])
      have kf_apply (kf_filter (λb. inv ket b = x) (kf_trace (range ket))) (tc_butterfly (ket c) (ket d))
         = sandwich_tc (vector_to_cblinfun (ket x)*) (tc_butterfly (ket c) (ket d))
        apply (transfer' fixing: x)
        apply (subst infsum_single[where i=((vector_to_cblinfun (ket x))*, ket x)])
        by auto
      also have  = of_bool (x=c  x=d) *R tc_butterfly (ket ()) (ket ())
        by (auto simp add: sandwich_tc_butterfly ket_CARD_1_is_1 cinner_ket)
      finally show kf_apply (kf_filter (λb. inv ket b = x) (kf_trace (range ket))) (tc_butterfly (ket c) (ket d))
           = of_bool (x=c  x=d) *R tc_butterfly (ket ()) (ket ())
        by -
    qed
    also have  = sandwich_tc (tensor_ell2_right (ket x)*) (tc_tensor (tc_butterfly (ket a) (ket b)) (tc_butterfly (ket c) (ket d)))
      by (auto simp: tensor_tc_butterfly sandwich_tc_butterfly)
    finally show ?thesis
      by -
  qed
  then show ?thesis
    apply (rule_tac eq_from_separatingI2x[where x=ρ])
       apply (rule separating_set_bounded_clinear_tc_tensor_nested)
        apply (rule separating_set_tc_butterfly_nested)
         apply (rule separating_set_ket)
        apply (rule separating_set_ket)
       apply (rule separating_set_tc_butterfly_nested)
        apply (rule separating_set_ket)
       apply (rule separating_set_ket)
    by (auto intro!: kf_apply_on_bounded_clinear bounded_clinear_sandwich_tc separating_set_tc_butterfly_nested simp: )
qed

lemma kf_partial_trace_left_apply_singleton:
  kf_partial_trace_left *kr @{x} ρ = sandwich_tc (tensor_ell2_left (ket x)*) ρ
proof -
  have aux1: (λxa. snd xa = x) = (λ(e,f). f=x  True)
    by auto
  have aux2: (λxa. xa  {x}) = (λxa. xa = x)
    by auto
  have inj_snd: inj_on (snd :: unit×'b  'b) X for X
    by (auto intro!: inj_onI)
  have aux3: tensor_ell2_right (ket x)* *V ket (b, a) = of_bool (x=a) *R ket b for x a :: 'x and b :: 'y
    by (smt (verit) cinner_ket_same of_bool_eq(1) of_bool_eq(2) of_real_1 of_real_hom.hom_0_iff orthogonal_ket scaleR_scaleC tensor_ell2_ket tensor_ell2_right_adj_apply)
  have aux4: tensor_ell2_left (ket x)* *V ket (a, b) = of_bool (x=a) *R ket b for x a :: 'x and b :: 'y
    by (smt (verit, del_insts) cinner_ket_same of_bool_eq(1) of_bool_eq(2) of_real_1 of_real_hom.hom_0_iff orthogonal_ket scaleR_scaleC tensor_ell2_ket tensor_ell2_left_adj_apply)
  have aux5: tensor_ell2_right (ket x)* oCL swap_ell2 = tensor_ell2_left (ket x)*
    apply (rule equal_ket)
    by (auto intro!: simp: aux3 aux4)
  have kf_partial_trace_left *kr @{x} ρ
     = kf_partial_trace_right *kr @{x} (sandwich_tc swap_ell2 ρ)
    by (simp only: kf_partial_trace_left_def kf_apply_on_def kf_filter_map_inj
        aux1 kf_filter_comp kf_apply_map_inj inj_snd kf_filter_true
        kf_comp_apply o_def kf_of_op_apply aux2)
  also have  = sandwich_tc (tensor_ell2_left (ket x)*) ρ
    by (auto intro!: arg_cong[where f=λx. sandwich_tc x _]
        simp: kf_partial_trace_right_apply_singleton aux5
        simp flip: sandwich_tc_compose[unfolded o_def, THEN fun_cong])
  finally show ?thesis
    by -
qed

lemma kf_domain_partial_trace_right[simp]: kf_domain kf_partial_trace_right = UNIV
proof (intro Set.set_eqI iffI UNIV_I)
  fix x :: 'a and y :: 'b

  have kf_partial_trace_right *kr @{x} (tc_tensor (tc_butterfly (ket y) (ket y)) (tc_butterfly (ket x) (ket x)))
      = tc_butterfly (ket y) (ket y)
    by (simp add: kf_partial_trace_right_apply_singleton tensor_tc_butterfly sandwich_tc_butterfly)
  also have   0
  proof -
    have norm (tc_butterfly (ket y) (ket y)) = 1
      by (simp add: norm_tc_butterfly)
    then show ?thesis
      by auto
  qed
 finally have kf_apply_on (kf_partial_trace_right :: (('b×'a) ell2, 'b ell2, 'a) kraus_family) {x}  0
    by auto
  then show x  kf_domain (kf_partial_trace_right :: (('b×'a) ell2, 'b ell2, 'a) kraus_family)
    by (rule in_kf_domain_iff_apply_nonzero[THEN iffD2])
qed


lemma kf_domain_partial_trace_left[simp]: kf_domain kf_partial_trace_left = UNIV
proof (intro Set.set_eqI iffI UNIV_I)
  fix x :: 'a and y :: 'b

  have kf_partial_trace_left *kr @{x} (tc_tensor (tc_butterfly (ket x) (ket x)) (tc_butterfly (ket y) (ket y)))
      = tc_butterfly (ket y) (ket y)
    by (simp add: kf_partial_trace_left_apply_singleton tensor_tc_butterfly sandwich_tc_butterfly)
  also have   0
  proof -
    have norm (tc_butterfly (ket y) (ket y)) = 1
      by (simp add: norm_tc_butterfly)
    then show ?thesis
      by auto
  qed
 finally have kf_apply_on (kf_partial_trace_left :: (('a×'b) ell2, 'b ell2, 'a) kraus_family) {x}  0
    by auto
  then show x  kf_domain (kf_partial_trace_left :: (('a×'b) ell2, 'b ell2, 'a) kraus_family)
    by (rule in_kf_domain_iff_apply_nonzero[THEN iffD2])
qed


subsection ‹Complete measurement›

lemma complete_measurement_aux:
  fixes B and F :: ('a::chilbert_space CL 'a × 'a) set
  defines family  (λx. (selfbutter (sgn x), x)) ` B
  assumes finite F and FB: F  family and is_ortho_set B
  shows ((E, x)F. E* oCL E)  id_cblinfun
proof -
  obtain G where finite G and G  B and FG: F = (λx. (selfbutter (sgn x), x)) ` G
    apply atomize_elim
    using finite F and FB
    apply (simp add: family_def)
    by (meson finite_subset_image)
  from G  B have [simp]: is_ortho_set G
    by (simp add: is_ortho_set B is_ortho_set_antimono)
  then have [simp]: e  G  norm (sgn e) = 1 for e
    apply (simp add: is_ortho_set_def)
    by (metis norm_sgn)
  have [simp]: inj_on (λx. (selfbutter (sgn x), x)) G
    by (meson inj_onI prod.inject)
  have [simp]: inj_on sgn G
  proof (rule inj_onI, rule ccontr)
    fix x y assume x  G and y  G and sgn_eq: sgn x = sgn y and x  y
    with is_ortho_set G have is_orthogonal x y
      by (meson is_ortho_set_def)
    then have is_orthogonal (sgn x) (sgn y)
      by fastforce
    with sgn_eq have sgn x = 0
      by force
    with x  G is_ortho_set G show False
      by (metis x  y local.sgn_eq sgn_zero_iff)
  qed
  have ((E, x)F. E* oCL E) = (xG. selfbutter (sgn x))
    by (simp add: FG sum.reindex cdot_square_norm)
  also 
  have (xG. selfbutter (sgn x))  id_cblinfun
    apply (subst sum.reindex[where h=sgn, unfolded o_def, symmetric])
     apply simp
    apply (subgoal_tac x y. xG. yG. x  y  is_orthogonal x y 
           0  G  x  G  y  G  sgn x  sgn y  is_orthogonal x y)
    using is_ortho_set G
     apply (auto intro!: sum_butterfly_leq_id simp: is_ortho_set_def sgn_zero_iff)[1]
    by fast
  finally show ?thesis
    by -
qed

lemma complete_measurement_is_kraus_family:
  assumes is_ortho_set B
  shows kraus_family ((λx. (selfbutter (sgn x), x)) ` B)
proof (rule kraus_familyI)
  show bdd_above (sum (λ(E, x). E* oCL E) ` {F. finite F  F  (λx. (selfbutter (sgn x), x)) ` B})
    using complete_measurement_aux[OF _ _ assms]
    by (auto intro!: bdd_aboveI[where M=id_cblinfun] kraus_familyI)
  have selfbutter (sgn x) = 0  x = 0 for x
    by (smt (verit, best) mult_cancel_right1 norm_butterfly norm_sgn norm_zero)
  then show 0  fst ` (λx. (selfbutter (sgn x), x)) ` B
    using assms by (force simp: is_ortho_set_def)
qed

lift_definition kf_complete_measurement :: 'a set  ('a::chilbert_space, 'a, 'a) kraus_family is
  λB. if is_ortho_set B then (λx. (selfbutter (sgn x), x)) ` B else {}
  by (auto intro!: complete_measurement_is_kraus_family)

definition kf_complete_measurement_ket :: ('a ell2, 'a ell2, 'a) kraus_family where
  kf_complete_measurement_ket = kf_map_inj (inv ket) (kf_complete_measurement (range ket))

lemma kf_complete_measurement_domain[simp]:
  assumes is_ortho_set B
  shows kf_domain (kf_complete_measurement B) = B
  apply (transfer fixing: B)
  using assms by (auto simp: image_image)

lemma kf_complete_measurement_ket_domain[simp]:
  kf_domain kf_complete_measurement_ket = UNIV
  by (simp add: kf_complete_measurement_ket_def)

lemma kf_complete_measurement_ket_kf_map:
  kf_complete_measurement_ket kr kf_map (inv ket) (kf_complete_measurement (range ket))
  unfolding kf_complete_measurement_ket_def
  apply (rule kf_map_inj_eq_kf_map)
  using inj_on_inv_into by fastforce+



lemma kf_bound_complete_measurement:
  assumes is_ortho_set B
  shows kf_bound (kf_complete_measurement B)  id_cblinfun
  apply (rule kf_bound_leqI)
  by (simp add: assms complete_measurement_aux kf_complete_measurement.rep_eq)

lemma kf_norm_complete_measurement:
  assumes is_ortho_set B
  shows kf_norm (kf_complete_measurement B)  1
  by (smt (verit, ccfv_SIG) assms kf_norm_def kf_bound_complete_measurement kf_bound_pos norm_cblinfun_id_le norm_cblinfun_mono)

lemma kf_complete_measurement_invalid:
  assumes ¬ is_ortho_set B
  shows kf_complete_measurement B = 0
  apply (transfer' fixing: B)
  using assms by simp

lemma kf_complete_measurement_idem:
  kf_comp (kf_complete_measurement B) (kf_complete_measurement B)
      kr kf_map (λb. (b,b)) (kf_complete_measurement B)
proof -
  wlog [iff]: is_ortho_set B
    using negation 
    by (simp add: kf_complete_measurement_invalid)
  define f where f b = (selfbutter (sgn b), selfbutter (sgn b), b, b) for b :: 'a
  have b2: sgn b C sgn b = 1 if b  B for b
    by (metis b  B is_ortho_set B cnorm_eq_1 is_ortho_set_def norm_sgn)
  have 1: (E oCL F, F, E, b, c)  (λx. (selfbutter (sgn x), f x)) ` B
    if E oCL F  0 and (F, b)  Rep_kraus_family (kf_complete_measurement B) and (E, c)  Rep_kraus_family (kf_complete_measurement B)
    for E F b c
  proof -
    from that have b  B and c  B and E_def: E = selfbutter (sgn c) and F_def: F = selfbutter (sgn b)
      by (auto simp add: kf_complete_measurement.rep_eq)
    have E oCL F = (sgn c C sgn b) *C butterfly (sgn c) (sgn b)
      by (simp add: E_def F_def)
    with that have c C b  0
      by fastforce
    with b  B c  B is_ortho_set B
    have c = b
      using is_ortho_setD by blast
    then have (E oCL F, F, E, b, c) = (λx. (selfbutter (sgn x), f x)) c
      by (simp add: b2 b  B f_def E_def F_def)
    with c  B show ?thesis
      by blast
  qed
  have 2: x  B  selfbutter (sgn x)  0 for x
    by (smt (verit) is_ortho_set B inverse_1 is_ortho_set_def norm_butterfly norm_sgn norm_zero right_inverse)
  have 3: (selfbutter (sgn x), f x)  (λ((F,y),(E, x)). (E oCL F, F, E, y, x)) `
          (Rep_kraus_family (kf_complete_measurement B) × Rep_kraus_family (kf_complete_measurement B))
    if x  B and (E, F, c, b) = f x
    for E F c b x
    apply (rule image_eqI[where x=((selfbutter (sgn x),x),(selfbutter (sgn x),x))])
    by (auto intro!: simp: b2 that f_def kf_complete_measurement.rep_eq)
  have raw: (kf_comp_dependent_raw (λ_. kf_complete_measurement B) (kf_complete_measurement B))
      = kf_map_inj f (kf_complete_measurement B)
    apply (transfer' fixing: f B)
    using 1 2 3
    by (auto simp: image_image case_prod_unfold)
  have kf_comp (kf_complete_measurement B) (kf_complete_measurement B)
     kr kf_map (λ(F, E, y). y) ((kf_comp_dependent_raw (λ_. kf_complete_measurement B) (kf_complete_measurement B)))
    by (simp add: kf_comp_def kf_comp_dependent_def id_def)
  also have  kr kf_map (λ(F, E, y). y) (kf_map_inj f (kf_complete_measurement B))
    by (simp add: raw)
  also have  kr kf_map (λ(F, E, y). y) (kf_map f (kf_complete_measurement B))
    by (auto intro!: kf_map_cong kf_map_inj_eq_kf_map inj_onI simp: f_def)
  also have  kr kf_map (λb. (b, b)) (kf_complete_measurement B)
    apply (rule kf_map_twice[THEN kf_eq_trans])
    by (simp add: f_def o_def)
  finally show ?thesis
    by -
qed

lemma kf_complete_measurement_idem_weak:
  kf_comp (kf_complete_measurement B) (kf_complete_measurement B)
      =kr kf_complete_measurement B
  by (metis (no_types, lifting) kf_apply_map kf_complete_measurement_idem kf_eq_imp_eq_weak kf_eq_weak_def)

lemma kf_complete_measurement_ket_idem:
  kf_comp kf_complete_measurement_ket kf_complete_measurement_ket
      kr kf_map (λb. (b,b)) kf_complete_measurement_ket
proof -
  have kf_comp kf_complete_measurement_ket kf_complete_measurement_ket
    kr kf_comp (kf_map (inv ket) (kf_complete_measurement (range ket))) (kf_map (inv ket) (kf_complete_measurement (range ket)))
    by (intro kf_comp_cong kf_complete_measurement_ket_kf_map)
  also have  kr kf_map (λ(x, y). (x, inv ket y))
     (kf_map (λ(x, y). (inv ket x, y)) (kf_comp (kf_complete_measurement (range ket)) (kf_complete_measurement (range ket))))
    by (intro kf_comp_map_left[THEN kf_eq_trans] kf_map_cong kf_comp_map_right refl)
  also have  kr kf_map (λ(x, y). (x, inv ket y))
     (kf_map (λ(x, y). (inv ket x, y)) (kf_map (λb. (b, b)) (kf_complete_measurement (range ket))))
    by (intro kf_map_cong kf_complete_measurement_idem refl)
  also have  kr kf_map (λx. (inv ket x, inv ket x)) (kf_complete_measurement (range ket))
    apply (intro kf_map_twice[THEN kf_eq_trans])
    by (simp add: o_def)
  also have  kr kf_map (λb. (b, b)) (kf_map (inv ket) (kf_complete_measurement (range ket)))
    apply (rule kf_eq_sym)
    apply (rule kf_map_twice[THEN kf_eq_trans])
    by (simp add: o_def)
  also have  kr kf_map (λb. (b, b)) kf_complete_measurement_ket
    using kf_complete_measurement_ket_kf_map kf_eq_sym kf_map_cong by fastforce
  finally show ?thesis
    by -
qed

lemma kf_complete_measurement_ket_idem_weak:
  kf_comp kf_complete_measurement_ket kf_complete_measurement_ket
      =kr kf_complete_measurement_ket
  by (metis (no_types, lifting) kf_apply_map kf_complete_measurement_ket_idem kf_eq_imp_eq_weak kf_eq_weak_def)

(* lemma kf_complete_measurement_ket_idem[simp]:
  ‹kf_complete_measurement_ket *kr kf_complete_measurement_ket *kr ρ
      = kf_complete_measurement_ket *kr ρ›
try0
sledgehammer [dont_slice]
by -

  by (metis kf_apply_map kf_apply_on_UNIV kf_apply_on_eqI kf_complete_measurement_idem kf_complete_measurement_ket_kf_map) *)

lemma kf_complete_measurement_apply:
  assumes [simp]: is_ortho_set B
  shows kf_complete_measurement B *kr t = (xB. sandwich_tc (selfbutter (sgn x)) t)
proof -
  have kf_complete_measurement B *kr t = 
      (E(λx. (selfbutter (sgn x), x)) ` B. sandwich_tc (fst E) t)
    apply (transfer' fixing: B t)
    by simp
  also have  = (xB. sandwich_tc (selfbutter (sgn x)) t)
    apply (subst infsum_reindex)
    by (auto intro!: inj_onI simp: o_def)
  finally show ?thesis
    by -
qed

lemma kf_complete_measurement_has_sum:
  assumes is_ortho_set B
  shows ((λx. sandwich_tc (selfbutter (sgn x)) ρ) has_sum kf_complete_measurement B *kr ρ) B
  using kf_apply_has_sum[of _ kf_complete_measurement B] assms
  by (simp add: kf_complete_measurement_apply kf_complete_measurement.rep_eq
      has_sum_reindex inj_on_def o_def)

lemma kf_complete_measurement_has_sum_onb:
  assumes is_onb B
  shows ((λx. sandwich_tc (selfbutter x) ρ) has_sum kf_complete_measurement B *kr ρ) B
proof -
  have is_ortho_set B
    using assms by (simp add: is_onb_def)
  have sgnx: sgn x = x if x  B for x
    using assms that
    by (simp add: is_onb_def sgn_div_norm)
  from kf_complete_measurement_has_sum[OF is_ortho_set B]
  show ?thesis
    apply (rule has_sum_cong[THEN iffD1, rotated])
    by (simp add: sgnx)
qed

lemma kf_complete_measurement_ket_has_sum:
   ((λx. sandwich_tc (selfbutter (ket x)) ρ) has_sum kf_complete_measurement_ket *kr ρ) UNIV
proof -
  from kf_complete_measurement_has_sum_onb
  have ((λx. sandwich_tc (selfbutter x) ρ) has_sum kf_complete_measurement (range ket) *kr ρ) (range ket)
    by force
  then have ((λx. sandwich_tc (selfbutter (ket x)) ρ) has_sum kf_complete_measurement (range ket) *kr ρ) UNIV
    apply (subst (asm) has_sum_reindex)
    by (simp_all add: o_def)
  then have ((λx. sandwich_tc (selfbutter (ket x)) ρ) has_sum kf_map (inv ket) (kf_complete_measurement (range ket)) *kr ρ) UNIV
    by simp
  then show ?thesis
    by (metis (no_types, lifting) kf_apply_on_UNIV kf_apply_on_eqI kf_complete_measurement_ket_kf_map)
qed

lemma kf_complete_measurement_apply_onb:
  assumes is_onb B
  shows kf_complete_measurement B *kr t = (xB. sandwich_tc (selfbutter x) t)
  using kf_complete_measurement_has_sum_onb[OF assms]
  by (metis (lifting) infsumI)

lemma kf_complete_measurement_ket_apply: kf_complete_measurement_ket *kr t = (x. sandwich_tc (selfbutter (ket x)) t)
proof -
  have kf_complete_measurement_ket *kr t = kf_complete_measurement (range ket) *kr t
    by (metis kf_apply_map kf_apply_on_UNIV kf_apply_on_eqI kf_complete_measurement_ket_kf_map)
  also have  = (xrange ket. sandwich_tc (selfbutter x) t)
    by (simp add: kf_complete_measurement_apply_onb)
  also have   = (x. sandwich_tc (selfbutter (ket x)) t)
    by (simp add: infsum_reindex o_def)
  finally show ?thesis
    by -
qed


lemma kf_bound_complete_measurement_onb[simp]:
  assumes is_onb B
  shows kf_bound (kf_complete_measurement B) = id_cblinfun
proof (rule cblinfun_eq_gen_eqI[where G=B], rule cinner_extensionality_basis[where B=B])
  show ccspan B = 
    using assms is_onb_def by blast
  then show ccspan B = 
    by -
  fix x y assume x  B and y  B
  have aux1: j  x  j  B  sandwich_tc (selfbutter j) (tc_butterfly y x) = 0 for j
    apply (transfer' fixing: x B j y)
    by (smt (z3) x  B apply_id_cblinfun assms butterfly_0_right butterfly_adjoint butterfly_def cblinfun_apply_cblinfun_compose
        cblinfun_comp_butterfly is_onb_def is_ortho_setD of_complex_eq_id sandwich_apply vector_to_cblinfun_adj_apply)
  have aux2: trace_tc (sandwich_tc (selfbutter y) (tc_butterfly y y)) = 1
    apply (transfer' fixing: y)
    by (metis (no_types, lifting) ext y  B assms butterfly_adjoint butterfly_comp_butterfly cblinfun_comp_butterfly cinner_simps(6)
        is_onb_def norm_one one_cinner_a_scaleC_one one_cinner_one sandwich_apply selfbutter_pos trace_butterfly trace_norm_butterfly
        trace_norm_pos trace_scaleC)
  have aux3: x  y  trace_tc (sandwich_tc (selfbutter x) (tc_butterfly y x)) = 0
    apply (transfer' fixing: x y)
    by (metis (no_types, lifting) ext Trace_Class.trace_0 x  B y  B apply_id_cblinfun assms butterfly_0_left butterfly_def
        cblinfun.zero_right cblinfun_apply_cblinfun_compose cblinfun_comp_butterfly is_onb_def is_ortho_setD of_complex_eq_id
        sandwich_apply vector_to_cblinfun_adj_apply)

  have x C (kf_bound (kf_complete_measurement B) *V y) = trace_tc (kf_complete_measurement B *kr tc_butterfly y x)
    by (simp add: kf_bound_from_map)
  also have  = trace_tc (xaB. sandwich_tc (selfbutter xa) (tc_butterfly y x))
    by (simp add: kf_complete_measurement_apply_onb assms)
  also have  = trace_tc (if x  B then sandwich_tc (selfbutter x) (tc_butterfly y x) else 0)
    apply (subst infsum_single[where i=x])
    using aux1 by auto
  also have  = of_bool (x = y)
    using y  B aux2 aux3 by (auto intro!: simp: )
  also have  = x C (id_cblinfun *V y)
    using x  B y  B assms cnorm_eq_1 is_onb_def is_ortho_setD by fastforce
  finally show x C (kf_bound (kf_complete_measurement B) *V y) = x C (id_cblinfun *V y)
    by -
qed

lemma kf_bound_complete_measurement_ket[simp]:
  kf_bound kf_complete_measurement_ket = id_cblinfun
  by (metis is_onb_ket kf_bound_complete_measurement_onb kf_bound_cong kf_complete_measurement_ket_kf_map kf_eq_imp_eq_weak
      kf_map_bound)

lemma kf_norm_complete_measurement_onb[simp]:
  fixes B :: 'a::{not_singleton, chilbert_space} set
  assumes is_onb B
  shows kf_norm (kf_complete_measurement B) = 1
  by (simp add: kf_norm_def assms)

lemma kf_norm_complete_measurement_ket[simp]:
  kf_norm kf_complete_measurement_ket = 1
  by (simp add: kf_norm_def)


lemma kf_complete_measurement_ket_diagonal_operator[simp]:
  kf_complete_measurement_ket *kr diagonal_operator_tc f = diagonal_operator_tc f
proof (cases f abs_summable_on UNIV)
  case True
  have kf_complete_measurement_ket *kr diagonal_operator_tc f = (x. sandwich_tc (selfbutter (ket x)) (diagonal_operator_tc f))
    by (simp add: kf_complete_measurement_ket_apply)
  also have  = (x. sandwich_tc (selfbutter (ket x)) (y. f y *C tc_butterfly (ket y) (ket y)))
    by (simp add: flip: tc_butterfly_scaleC_infsum)
  also have  = (x. y. sandwich_tc (selfbutter (ket x)) (f y *C tc_butterfly (ket y) (ket y)))
    apply (rule infsum_cong)
    apply (rule infsum_bounded_linear[unfolded o_def, symmetric])
    by (auto intro!: bounded_clinear.bounded_linear bounded_clinear_sandwich_tc tc_butterfly_scaleC_summable True)
  also have  = (x. y. of_bool (y=x) *C f x *C tc_butterfly (ket x) (ket x))
    apply (rule infsum_cong)+
    apply (transfer' fixing: f)
    by (simp add: sandwich_apply)
  also have  = (x. f x *C tc_butterfly (ket x) (ket x))
    apply (subst infsum_of_bool_scaleC)
    by simp
  also have  = diagonal_operator_tc f
    by (simp add: flip: tc_butterfly_scaleC_infsum)
  finally show ?thesis
    by -
next
  case False
  then have diagonal_operator_tc f = 0
    by (rule diagonal_operator_tc_invalid)
  then show ?thesis
    by simp
qed

lemma kf_operators_complete_measurement:
  kf_operators (kf_complete_measurement B) = (selfbutter o sgn) ` B if is_ortho_set B
  apply (transfer' fixing: B)
  using that by force

lemma kf_operators_complete_measurement_invalid:
  kf_operators (kf_complete_measurement B) = {} if ¬ is_ortho_set B
  apply (transfer' fixing: B)
  using that by force

lemma kf_operators_complete_measurement_ket:
  kf_operators kf_complete_measurement_ket = range (λc. butterfly (ket c) (ket c))
  by (simp add: kf_complete_measurement_ket_def kf_operators_complete_measurement image_image)

lemma kf_complete_measurement_apply_butterfly:
  assumes is_ortho_set B and b  B
  shows kf_complete_measurement B *kr tc_butterfly b b = tc_butterfly b b
proof -
  have kf_complete_measurement B *kr tc_butterfly b b = (xB. sandwich_tc (selfbutter (sgn x)) (tc_butterfly b b))
    by (simp add: kf_complete_measurement_apply assms)
  also have  = (if b  B then sandwich_tc (selfbutter (sgn b)) (tc_butterfly b b) else 0)
  proof (rule infsum_single[where i=b])
    fix c assume c  B and c  b
    then have [iff]: c C b = 0
      using assms(1,2) is_ortho_setD by blast
    then have [iff]: sgn c C b = 0
      by auto
    then
    show sandwich_tc (selfbutter (sgn c)) (tc_butterfly b b) = 0
      by (auto intro!: simp: sandwich_tc_butterfly)
  qed
  also have  = tc_butterfly b b
    by (simp add: b  B sandwich_tc_butterfly)
  finally show ?thesis
    by -
qed

lemma kf_complete_measurement_ket_apply_butterfly:
  kf_complete_measurement_ket *kr tc_butterfly (ket x) (ket x) = tc_butterfly (ket x) (ket x)
  by (simp add: kf_complete_measurement_ket_def kf_apply_map_inj inj_on_def kf_complete_measurement_apply_butterfly)


(* lemma kf_unique_if_singleton:
  assumes ‹Rep_kraus_family 𝔈 ⊆ {Ee}›
  assumes ‹Rep_kraus_family 𝔉 ⊆ {Ff}›
  assumes ‹𝔈 ≡kr 𝔉›
  shows ‹𝔈 = 𝔉›
(* WRONG! Global phase *)
proof (cases ‹𝔈 = 0›)
  case True
  with ‹𝔈 ≡kr 𝔉› have ‹𝔉 = 0›
    using kf_eq_0_iff_eq_0 kf_eq_imp_eq_weak kf_eq_sym by blast
  with True show ?thesis
    by simp
next
  case False
  obtain E e F f where Ee: ‹Ee = (E,e)› and Ff: ‹Ff = (F,f)›
    apply atomize_elim
    by (auto split!: prod.split_asm)
  from False have ‹𝔉 ≠ 0›
    using ‹𝔈 ≡kr 𝔉› kf_eq_0_iff_eq_0 kf_eq_imp_eq_weak by blast
  from False have ‹Rep_kraus_family 𝔈 ≠ {}›
    using Rep_kraus_family_inject zero_kraus_family.rep_eq by auto
  with assms(1) have Rep𝔈: ‹Rep_kraus_family 𝔈 = {(E,e)}›
    using Ee by blast
  from ‹𝔉 ≠ 0› have ‹Rep_kraus_family 𝔉 ≠ {}›
    using Rep_kraus_family_inject zero_kraus_family.rep_eq by auto
  with assms(2) have Rep𝔉: ‹Rep_kraus_family 𝔉 = {(F,f)}›
    using Ff by blast
  from Rep𝔈 have ‹E ≠ 0›
    using Rep_kraus_family by (force simp: kraus_family_def)
  from Rep𝔉 have ‹F ≠ 0›
    using Rep_kraus_family by (force simp: kraus_family_def)
  have ‹E = F›
  proof (rule cblinfun_eqI)
    fix h
    have ‹tc_butterfly (E h) (E h) = 𝔈 *kr tc_butterfly h h›
      by (simp add: kf_apply.rep_eq Rep𝔈 sandwich_tc_butterfly)
    also have ‹… = 𝔉 *kr tc_butterfly h h›
      by (simp add: assms(3) kf_apply_eqI kf_eq_imp_eq_weak)
    also have ‹… = tc_butterfly (F h) (F h)›
      by (simp add: kf_apply.rep_eq Rep𝔉 sandwich_tc_butterfly)
    finally show ‹E h = F h›
      by (rule tc_selfbutter_inject)
  qed
  have ‹e = f›
  proof (rule ccontr)
    assume ‹e ≠ f›
    obtain h where neq0: ‹E h ≠ 0›
    proof (atomize_elim, rule ccontr)
      assume ‹∄h. E *V h ≠ 0›
      then have ‹E h = 0› for h
        by auto
      then have ‹E = 0›
        apply (rule_tac cblinfun_eqI)
        by simp
      with ‹E ≠ 0› show False
        by simp
    qed
    have ‹tc_butterfly (E h) (E h) = 𝔈 *kr @{e} tc_butterfly h h›
      by (simp add: kf_apply_on_def kf_filter.rep_eq kf_apply.rep_eq Rep𝔈 sandwich_tc_butterfly filter_insert_if)
  also have ‹… = 𝔉 *kr @{e} tc_butterfly h h›
  using assms(3) kf_apply_on_eqI by blast
  also from ‹e ≠ f› have ‹… = 0›
    by (simp add: kf_apply_on_def kf_filter.rep_eq kf_apply.rep_eq Rep𝔉 sandwich_tc_butterfly filter_insert_if)
  finally have ‹E h = 0›
    using tc_selfbutter_inject[where k=0]
    by auto
  with neq0 show False
    by simp
  qed
  from ‹E = F› ‹e = f› Rep𝔈 Rep𝔉 show ‹𝔈 = 𝔉›
    using Rep_kraus_family_inject by blast
qed *)



lemma kf_map_eq_kf_map_inj_singleton:
  assumes card_le_1 (Rep_kraus_family 𝔈)
  shows kf_map f 𝔈 = kf_map_inj f 𝔈
proof (cases 𝔈 = 0)
  case True
  then show ?thesis
    by simp
next
  case False
  then have Rep_kraus_family 𝔈  {}
  using Rep_kraus_family_inject zero_kraus_family.rep_eq by auto
  with assms obtain x where Rep𝔈: Rep_kraus_family 𝔈 = {x}
    by (meson card_le_1_def subset_singleton_iff)
  obtain E y where Ey: x = (E,y)
    by force
  have (E,y)  Rep_kraus_family 𝔈
    using Ey Rep𝔈 by force
  then have [iff]: E  0
    using Rep_kraus_family[of 𝔈]
    by (force simp: kraus_family_def)
  have *: {(F, xa). (F, xa) = x  f xa = f y  (r>0. E = r *R F)} = {(E,y)}
    apply (subgoal_tac r>0. E = r *R E)
     apply (auto intro!: simp: Ey)[1]
    by (metis scaleR_simps(12) verit_comp_simplify(28))
  have 1: (norm E)2 = kf_element_weight (kf_filter (λx. f x = f y) 𝔈) E
    by (auto simp add: kf_element_weight_def kf_similar_elements_def kf_filter.rep_eq * Rep𝔈)
  have 2: z = f y if (norm F)2 = kf_element_weight (kf_filter (λx. f x = z) 𝔈) F and F  0 for F z
  proof (rule ccontr)
    assume z  f y
    with Rep𝔈 have kf_filter (λx. f x = z) 𝔈 = 0
      apply (transfer' fixing: f z y x)
      by (auto simp: Ey)
    then have kf_element_weight (kf_filter (λx. f x = z) 𝔈) F = 0
      by simp
    with that show False
      by fastforce
  qed
  have 3: F = E if (norm F)2 = kf_element_weight (kf_filter (λx. f x = z) 𝔈) F and F  0 and z = f y for F z
  proof -
    from that Rep𝔈 have f𝔈: kf_filter (λx. f x = z) 𝔈 = 𝔈
      apply (transfer' fixing: F f z y x)
      using Ey Rep_kraus_family_inject kf_filter.rep_eq by fastforce
    have r>0. F = r *R E
    proof (rule ccontr)
      assume ¬ (r>0. F = r *R E)
      with Rep𝔈 have kf_similar_elements 𝔈 F = {}
        by (simp add: kf_similar_elements_def Ey)
      with that f𝔈 show False
        by (simp add: kf_element_weight_def)
    qed
    then obtain r where FrE: F = r *R E and rpos: r > 0
      by auto
    with Rep𝔈 have kf_similar_elements 𝔈 F = {(E,y)}
      by (force simp: kf_similar_elements_def Ey)
    then have kf_element_weight 𝔈 F = (norm E)2
      by (simp add: kf_element_weight_def)
    with that have norm E = norm F
      using f𝔈 by force
    with FrE rpos have r = 1
      by simp
    with FrE show F = E
      by simp
  qed
  show ?thesis
    apply (rule Rep_kraus_family_inject[THEN iffD1])
    by (auto intro!: 1 2 3 simp: kf_map.rep_eq kf_map_inj.rep_eq Rep𝔈 Ey)
qed

lemma kf_map_eq_kf_map_inj_singleton':
  assumes y. card_le_1 (Rep_kraus_family (kf_filter ((=)y) 𝔈))
  assumes inj_on f (kf_domain 𝔈)
  shows kf_map f 𝔈 = kf_map_inj f 𝔈
proof -
  have 1: card_le_1 (Rep_kraus_family (kf_filter (λz. x = f z) 𝔈)) for x
  proof (cases x  f ` kf_domain 𝔈)
    case True
    then have kf_filter (λz. x = f z) 𝔈 = kf_filter ((=) (inv_into (kf_domain 𝔈) f x)) 𝔈
      apply (rule_tac kf_filter_cong_eq)
      using assms(2)
      by auto
    with assms(1) show ?thesis
      by presburger
  next
    case False
    then have kf_filter (λz. x = f z) 𝔈 = 0
      by (simp add: image_iff)
    then show ?thesis
      by (simp add: zero_kraus_family.rep_eq)
  qed
  show ?thesis
    apply (rule kf_equal_if_filter_equal)
    unfolding kf_filter_map kf_filter_map_inj
    apply (rule kf_map_eq_kf_map_inj_singleton)
    by (rule 1)
qed

lemma kf_filter_singleton_kf_complete_measurement:
  assumes x  B and is_ortho_set B
  shows kf_filter ((=)x) (kf_complete_measurement B) = kf_map_inj (λ_. x) (kf_of_op (selfbutter (sgn x)))
proof -
  from assms have [iff]: selfbutter (sgn x)  0
    by (smt (verit, best) inverse_1 is_ortho_set_def norm_butterfly norm_sgn norm_zero right_inverse)
  show ?thesis
    apply (transfer' fixing: x B)
    by (auto intro!: simp: assms)
qed

lemma kf_filter_singleton_kf_complete_measurement':
  assumes x  B and is_ortho_set B
  shows kf_filter ((=)x) (kf_complete_measurement B) = kf_map (λ_. x) (kf_of_op (selfbutter (sgn x)))
  using kf_filter_singleton_kf_complete_measurement[OF assms]
  apply (subst kf_map_eq_kf_map_inj_singleton)
  by (auto simp: kf_of_op.rep_eq)

lemma kf_filter_disjoint:
  assumes x. x  kf_domain 𝔈  P x = False
  shows kf_filter P 𝔈 = 0
  using assms
  apply (transfer' fixing: P)
  by fastforce



lemma kf_complete_measurement_tensor:
  assumes is_ortho_set B and is_ortho_set C
  shows kf_map (λ(b,c). b s c) (kf_tensor (kf_complete_measurement B) (kf_complete_measurement C))
    = kf_complete_measurement ((λ(b,c). b s c) ` (B × C))
proof (rule kf_equal_if_filter_equal, rename_tac bc)
  fix bc :: ('a × 'b) ell2
  define BC where BC = ((λ(x, y). x s y) ` (B × C))
  consider (bc_tensor) b c where b  B c  C bc = b s c
      | (not_bc_tensor) ¬ (bB. cC. bc = b s c)
    by fastforce
  then show kf_filter ((=) bc) (kf_map (λ(b, c). b s c) (kf_tensor (kf_complete_measurement B) (kf_complete_measurement C)))
      = kf_filter ((=) bc) (kf_complete_measurement ((λ(b, c). b s c) ` (B × C)))
  proof cases
    case bc_tensor
    have uniq: b = b'  c = c' if b'  B and c'  C and b s c = b' s c' for b' c'
    proof -
      from bc_tensor have b  0
        using assms(1) is_ortho_set_def by blast
      with that(3) obtain γ where upto: c = γ *C c'
        apply atomize_elim
        by (rule tensor_ell2_almost_injective)
      from bc_tensor have c  0
        using assms(2) is_ortho_set_def by blast
      with upto have γ  0
        by auto
      with c  0 upto have c C c'  0
        by simp
      with c  C c'  C is_ortho_set C
      show c = c'
        using is_ortho_setD by blast
      with that have b s c = b' s c
        by simp
      with c  0 b  B b'  B is_ortho_set B show b = b'
        by (metis cinner_eq_zero_iff is_ortho_setD nonzero_mult_div_cancel_right tensor_ell2_inner_prod)
    qed

    have [iff]: selfbutter (sgn b)  0
      by (smt (verit, ccfv_SIG) bc_tensor assms inverse_1 is_ortho_set_def norm_butterfly norm_sgn norm_zero right_inverse)
    have [iff]: selfbutter (sgn c)  0
      by (smt (verit) bc_tensor assms inverse_1 is_ortho_set_def norm_butterfly norm_sgn norm_zero right_inverse)
    have [iff]: selfbutter (sgn bc)  0
      by (smt (verit, del_insts) selfbutter (sgn b)  0 selfbutter (sgn c)  0 bc_tensor(3) butterfly_0_right mult_cancel_right1 norm_butterfly norm_sgn sgn_zero
          tensor_ell2_nonzero)
    have kf_filter ((=) bc) (kf_map (λ(b, c). b s c) (kf_tensor (kf_complete_measurement B) (kf_complete_measurement C)))
      = kf_map (λ(x, y). x s y) (kf_filter (λx. bc = (case x of (b, c)  b s c))
                                 (kf_tensor (kf_complete_measurement B) (kf_complete_measurement C)))
      by (simp add: kf_filter_map)
    also have  = kf_map (λ(x, y). x s y) (kf_filter ((=)(b,c))
                                 (kf_tensor (kf_complete_measurement B) (kf_complete_measurement C)))
      apply (rule arg_cong[where f=kf_map _])
      apply (rule kf_filter_cong_eq[OF refl])
      by (auto intro!: uniq simp add: kf_domain_tensor kf_complete_measurement_domain assms case_prod_beta bc_tensor split!: prod.split)
    also have  = kf_map (λ(x, y). x s y) (kf_tensor (kf_filter ((=) b) (kf_complete_measurement B))
                 (kf_filter ((=) c) (kf_complete_measurement C)))
      by (simp add: kf_filter_tensor_singleton)
    also have  = kf_map (λ(x, y). x s y) (kf_map (λ(E, F, y). y) (kf_tensor_raw (kf_filter ((=) b) (kf_complete_measurement B))
                 (kf_filter ((=) c) (kf_complete_measurement C))))
      by (simp add: kf_tensor_def id_def)
    also have  = kf_map (λ(x, y). x s y) (kf_map (λ(E, F, y). y) (kf_tensor_raw (kf_map (λ_. b) (kf_of_op (selfbutter (sgn b)))) 
                  (kf_map (λ_. c) (kf_of_op (selfbutter (sgn c))))))
      by (simp add: kf_filter_singleton_kf_complete_measurement' bc_tensor assms)
    also have  = kf_map_inj (λ(x, y). x s y) (kf_map_inj (λ(E, F, y). y) (kf_tensor_raw (kf_map_inj (λ_. b) (kf_of_op (selfbutter (sgn b)))) 
                  (kf_map_inj (λ_. c) (kf_of_op (selfbutter (sgn c))))))
      apply (subst (4) kf_map_eq_kf_map_inj_singleton)
       apply (simp add: kf_of_op.rep_eq)
      apply (subst (3) kf_map_eq_kf_map_inj_singleton)
       apply (simp add: kf_of_op.rep_eq)
      apply (subst (2) kf_map_eq_kf_map_inj_singleton)
       apply (simp add: kf_of_op.rep_eq kf_map_inj.rep_eq kf_tensor_raw.rep_eq)
      apply (subst (1) kf_map_eq_kf_map_inj_singleton)
       apply (simp add: kf_of_op.rep_eq kf_map_inj.rep_eq kf_tensor_raw.rep_eq)
      by blast
    also have  = kf_map_inj (λ_. b s c) (kf_of_op (selfbutter (sgn bc)))
      apply (transfer' fixing: b c bc)
      by (auto intro!: bc_tensor simp: tensor_butterfly simp flip: sgn_tensor_ell2 bc_tensor)
    also have  = kf_map (λ_. bc) (kf_of_op (selfbutter (sgn bc)))
      apply (subst kf_map_eq_kf_map_inj_singleton)
      by (auto intro!: bc_tensor simp: kf_of_op.rep_eq kf_map_inj.rep_eq kf_tensor_raw.rep_eq simp flip: bc_tensor)
    also have  = kf_filter ((=) bc) (kf_complete_measurement ((λ(b, c). b s c) ` (B × C)))
      apply (subst kf_filter_singleton_kf_complete_measurement')
      by (auto simp: is_ortho_set_tensor bc_tensor assms)
    finally show ?thesis
      by -
  next
    case not_bc_tensor
    have ortho: is_ortho_set ((λx. case x of (x, xa)  x s xa) ` (B × C))
      by (metis is_ortho_set_tensor not_bc_tensor assms)
    show ?thesis
      apply (subst kf_filter_disjoint)
      using not_bc_tensor
       apply (simp add: kf_domain_tensor kf_complete_measurement_domain assms)
       apply fastforce
      apply (subst kf_filter_disjoint)
      using not_bc_tensor
       apply (simp add: kf_domain_tensor kf_complete_measurement_domain ortho)
       apply fastforce
      by simp
  qed
qed

lemma card_le_1_kf_filter: card_le_1 (Rep_kraus_family (kf_filter P 𝔈)) if card_le_1 (Rep_kraus_family 𝔈)
  by (metis (no_types, lifting) card_le_1_def filter_insert_if kf_filter.rep_eq kf_filter_0_right
      subset_singleton_iff that zero_kraus_family.rep_eq)

lemma card_le_1_kf_map_inj[iff]: card_le_1 (Rep_kraus_family (kf_map_inj f 𝔈)) if card_le_1 (Rep_kraus_family 𝔈)
  using that
  apply transfer'
  by (auto simp: card_le_1_def case_prod_unfold)


lemma card_le_1_kf_map[iff]: card_le_1 (Rep_kraus_family (kf_map f 𝔈)) if card_le_1 (Rep_kraus_family 𝔈)
  using kf_map_eq_kf_map_inj_singleton[OF that] card_le_1_kf_map_inj[OF that]
  by metis
  

lemma card_le_1_kf_tensor_raw[iff]: card_le_1 (Rep_kraus_family (kf_tensor_raw 𝔈 𝔉)) if card_le_1 (Rep_kraus_family 𝔈) and card_le_1 (Rep_kraus_family 𝔉)
  using that
  apply transfer'
  apply (simp add: card_le_1_def)
  by fast

lemma card_le_1_kf_tensor[iff]: card_le_1 (Rep_kraus_family (kf_tensor 𝔈 𝔉)) if card_le_1 (Rep_kraus_family 𝔈) and card_le_1 (Rep_kraus_family 𝔉)
  by (auto intro!: card_le_1_kf_map card_le_1_kf_tensor_raw that simp add: kf_tensor_def)

lemma card_le_1_kf_filter_complete_measurement: card_le_1 (Rep_kraus_family (kf_filter ((=)x) (kf_complete_measurement B)))
proof -
  consider (inB) is_ortho_set B  x  B | (not_ortho) ¬ is_ortho_set B | (notinB) x  B is_ortho_set B
    by auto
  then show ?thesis
  proof cases
    case inB
    then have Rep_kraus_family (kf_filter ((=)x) (kf_complete_measurement B)) = {(selfbutter (sgn x),x)}
      by (auto simp: kf_filter.rep_eq kf_complete_measurement.rep_eq)
    then show ?thesis
      by (simp add: card_le_1_def)
  next
    case not_ortho
    then show ?thesis
      by (simp add: kf_complete_measurement_invalid zero_kraus_family.rep_eq)
  next
    case notinB
    then have kf_filter ((=) x) (kf_complete_measurement B) = 0
      by (metis kf_complete_measurement_domain kf_filter_disjoint)
    then show ?thesis
      by (simp add: zero_kraus_family.rep_eq)
  qed
qed

lemma kf_complete_measurement_ket_tensor:
  shows kf_tensor (kf_complete_measurement_ket :: (_,_,'a) kraus_family) (kf_complete_measurement_ket :: (_,_,'b) kraus_family)
             = kf_complete_measurement_ket
proof -
  have 1: (λ(b, c). b s c) ` (range ket × range ket) = range ket
    by (auto intro!: image_eqI simp: tensor_ell2_ket case_prod_unfold)
  have 2: inj_on (inv ket  (λ(x, y). x s y)) (range ket × range ket)
    by (auto intro!: inj_onI simp: tensor_ell2_ket)
  have (kf_complete_measurement_ket :: (_,_,'a×'b) kraus_family)
      = kf_map_inj (inv ket) (kf_complete_measurement ((λ(b,c). b s c) ` (range ket × range ket)))
    by (simp add: 1 kf_complete_measurement_ket_def)
  also have  = kf_map_inj (inv ket)
     (kf_map (λ(x, y). x s y)
       (kf_tensor (kf_complete_measurement (range ket)) (kf_complete_measurement (range ket))))
    by (simp flip: kf_complete_measurement_tensor)
  also have  = kf_map (inv ket  (λ(x, y). x s y))
     (kf_tensor (kf_complete_measurement (range ket)) (kf_complete_measurement (range ket)))
    apply (subst kf_map_inj_kf_map_comp)
    by (auto intro!: simp: inj_on_def kf_domain_tensor tensor_ell2_ket)
  also have  = kf_map_inj (inv ket  (λ(x, y). x s y))
     (kf_tensor (kf_complete_measurement (range ket)) (kf_complete_measurement (range ket)))
    apply (rule kf_map_eq_kf_map_inj_singleton')
     apply (auto intro!: card_le_1_kf_filter_complete_measurement card_le_1_kf_tensor simp: kf_filter_tensor_singleton)[1]
    by (simp add: kf_domain_tensor 2)
  also have  = kf_map_inj (map_prod (inv ket) (inv ket))
      (kf_tensor (kf_complete_measurement (range ket)) (kf_complete_measurement (range ket)))
    apply (rule kf_map_inj_cong_eq)
    by (auto simp: kf_domain_tensor tensor_ell2_ket)
  also have  = kf_tensor (kf_map_inj (inv ket) (kf_complete_measurement (range ket)))
                            (kf_map_inj (inv ket) (kf_complete_measurement (range ket)))
    by (simp add: kf_tensor_map_inj_both inj_on_inv_into)
  also have  = kf_tensor kf_complete_measurement_ket kf_complete_measurement_ket
    by (simp add: kf_complete_measurement_ket_def)
  finally show ?thesis
    by simp
qed

subsection ‹Reconstruction›

lemma kf_reconstruction_is_bounded_clinear:
  assumes ρ. ((λa. sandwich_tc (f a) ρ) has_sum 𝔈 ρ) A
  shows bounded_clinear 𝔈
proof -
  have linear: clinear 𝔈
  proof (rule clinearI)
    fix ρ σ c
    have ((λa. sandwich_tc (f a) ρ + sandwich_tc (f a) σ) has_sum (𝔈 ρ + 𝔈 σ)) A
      by (intro has_sum_add assms)
    then have ((λa. sandwich_tc (f a) (ρ + σ)) has_sum (𝔈 ρ + 𝔈 σ)) A
      by (meson has_sum_cong sandwich_tc_plus)
    with assms[of ρ + σ]
    show 𝔈 (ρ + σ) = 𝔈 ρ + 𝔈 σ
      by (rule has_sum_unique)
    from assms[of ρ]
    have ((λa. sandwich_tc (f a) (c *C ρ)) has_sum c *C 𝔈 ρ) A
      using has_sum_scaleC_right[where A=A and s=𝔈 ρ]
      by (auto intro!: has_sum_scaleC_right simp: sandwich_tc_scaleC_right)
    with assms[of c *C ρ]
    show 𝔈 (c *C ρ) = c *C 𝔈 ρ
      by (rule has_sum_unique)
  qed
  have pos: 𝔈 ρ  0 if ρ  0 for ρ
    apply (rule has_sum_mono_traceclass[where f=λ_.0 and g=(λa. sandwich_tc (f a) ρ)])
    using assms
    by (auto intro!: sandwich_tc_pos simp: that)
  have mono: 𝔈 ρ  𝔈 σ if ρ  σ for ρ σ
  proof -
    have 𝔈 (σ - ρ)  0
      apply (rule pos)
      using that
      by auto
    then show ?thesis
      by (simp add: linear complex_vector.linear_diff)
  qed
  have bounded_pos: B0. ρ0. norm (𝔈 ρ)  B * norm ρ
  proof (rule ccontr)
    assume asm: ¬ (B0. ρ0. norm (𝔈 ρ)  B * norm ρ)
    obtain ρ0 where 𝔈_big0: norm (𝔈 (ρ0 i)) > 2^i * norm (ρ0 i) and ρ0_pos: ρ0 i  0 for i :: nat
    proof (atomize_elim, rule choice2, rule allI, rule ccontr)
      fix i
      define B :: real where B = 2^i
      have B  0
        by (simp add: B_def)
      assume ρ0. B * norm ρ0 < norm (𝔈 ρ0)  0  ρ0
      then have ρ0. norm (𝔈 ρ)  B * norm ρ
        by force
      with asm B  0 show False
        by blast
    qed
    have ρ0_neq0: ρ0 i  0 for i
      using 𝔈_big0[of i] linear complex_vector.linear_0 by force
    define ρ where ρ i = ρ0 i /R norm (ρ0 i) for i
    have ρ_pos: ρ i  0 for i
      by (simp add: ρ_def ρ0_pos scaleR_nonneg_nonneg)
    have norm_ρ: norm (ρ i) = 1 for i
      by (simp add: ρ0_neq0 ρ_def)
    from 𝔈_big0 have 𝔈_big: trace_tc (𝔈 (ρ i))  2^i for i :: nat
    proof -
      have trace_tc (𝔈 (ρ i)) = trace_tc (𝔈 (ρ0 i) /R norm (ρ0 i))
        by (simp add: ρ_def linear scaleR_scaleC clinear.scaleC 
          bounded_clinear_trace_tc[THEN bounded_clinear.clinear])
      also have  = norm (𝔈 (ρ0 i) /R norm (ρ0 i))
        using ρ0_pos pos
        by (metis linordered_field_class.inverse_nonnegative_iff_nonnegative norm_ge_zero norm_tc_pos scaleR_nonneg_nonneg)
      also have  = norm (𝔈 (ρ0 i)) / norm (ρ0 i)
        by (simp add: divide_inverse_commute)
      also have  > (2^i * norm (ρ0 i)) / norm (ρ0 i) (is _ > )
        using 𝔈_big0 ρ0_neq0
        by (smt (verit, best) complex_of_real_strict_mono_iff divide_le_eq norm_le_zero_iff)
      also have  = 2^i
        using ρ0_neq0 by force
      finally show ?thesis
        by simp
    qed
    define σ τ where σ n = (i<n. ρ i /R 2^i) and τ = (i. ρ i /R 2^i) for n :: nat
    have (λi. ρ i /R 2^i) abs_summable_on UNIV
    proof (rule infsum_tc_norm_bounded_abs_summable)
      from ρ_pos show ρ i /R 2^i  0 for i
        by (simp add: scaleR_nonneg_nonneg)
      show norm (iF. ρ i /R 2^i)  2 if finite F for F
      proof -
        from finite_nat_bounded[OF that]
        obtain n where i_leq_n: i  n if i  F for i
          apply atomize_elim
          by (auto intro!: order.strict_implies_order simp: lessThan_def Ball_def simp flip: Ball_Collect)
        have norm (iF. ρ i /R 2^i)  (iF. norm (ρ i /R 2^i))
          by (simp add: sum_norm_le)
        also have  = (iF. (1/2)^i)
          using norm_ρ 
          by (smt (verit, del_insts) Extra_Ordered_Fields.sign_simps(23) divide_inverse_commute linordered_field_class.inverse_nonnegative_iff_nonnegative
              norm_scaleR power_inverse power_one sum.cong zero_le_power)
        also have   (in. (1/2)^i)
          apply (rule sum_mono2)
          using i_leq_n
          by auto
        also have   (i. (1/2)^i)
          apply (rule sum_le_suminf)
          by auto
        also have ... = 2
          using suminf_geometric[of 1/2 :: real]
          by simp
        finally show ?thesis
          by -
      qed
    qed
    then have summable: (λi. ρ i /R 2^i) summable_on UNIV
      by (simp add: abs_summable_summable)
    have trace_tc (𝔈 τ)  n for n :: nat
    proof -
      have trace_tc (𝔈 τ)  trace_tc (𝔈 (σ n)) (is _  )
        by (auto intro!: trace_tc_mono mono infsum_mono_neutral_traceclass
            simp: τ_def σ_def summable ρ_pos scaleR_nonneg_nonneg simp flip: infsum_finite)
      moreover have  = (i<n. trace_tc (𝔈 (ρ i)) / 2^i)
        by (simp add: σ_def complex_vector.linear_sum linear scaleR_scaleC trace_scaleC
            bounded_clinear_trace_tc[THEN bounded_clinear.clinear] clinear.scaleC
            add.commute mult.commute divide_inverse)
      moreover have   (i<n. 2^i / 2^i) (is _  )
        apply (intro sum_mono divide_right_mono)
        using 𝔈_big
        by (simp_all add: less_eq_complex_def)
      moreover have  = (i<n. 1)
        by fastforce
      moreover have  = n
        by simp
      ultimately show ?thesis
        by order
    qed
    then have Re: Re (trace_tc (𝔈 τ))  n for n :: nat
      using Re_mono by fastforce
    obtain n :: nat where n > Re (trace_tc (𝔈 τ))
      apply atomize_elim
      by (rule reals_Archimedean2)
    with Re show False
      by (smt (verit, ccfv_threshold))
  qed
  then obtain B where bounded_B: norm (𝔈 ρ)  B * norm ρ and B_pos: B  0 if ρ  0 for ρ
    by auto
  have bounded: norm (𝔈 ρ)  (4*B) * norm ρ for ρ
  proof -
    obtain ρ1 ρ2 ρ3 ρ4 where ρ_decomp: ρ = ρ1 - ρ2 + 𝗂 *C ρ3 - 𝗂 *C ρ4
      and pos: ρ1  0 ρ2  0 ρ3  0 ρ4  0
      and norm: norm ρ1  norm ρ norm ρ2  norm ρ norm ρ3  norm ρ norm ρ4  norm ρ
      apply atomize_elim using trace_class_decomp_4pos'[of ρ] by blast
    have norm (𝔈 ρ)  norm (𝔈 ρ1) + norm (𝔈 ρ2) + norm (𝔈 ρ3) + norm (𝔈 ρ4)
      using linear
      by (auto intro!: norm_triangle_le norm_triangle_le_diff
          simp add: ρ_decomp kf_apply_plus_right kf_apply_minus_right
          kf_apply_scaleC complex_vector.linear_diff complex_vector.linear_add clinear.scaleC)
    also have   B * norm ρ1 + B * norm ρ2 + B * norm ρ3 + B * norm ρ4
      using pos by (auto intro!: add_mono simp add: pos bounded_B)
    also have  = B * (norm ρ1 + norm ρ2 + norm ρ3 + norm ρ4)
      by argo
    also have   B * (norm ρ + norm ρ + norm ρ + norm ρ)
      by (auto intro!: mult_left_mono add_mono pos B_pos
          simp only: norm)
    also have  = (4 * B) * norm ρ
      by argo
    finally show ?thesis
      by -
  qed
  show ?thesis
    apply (rule bounded_clinearI[where K=4*B])
      apply (simp add: complex_vector.linear_add linear) 
     apply (simp add: complex_vector.linear_scale linear) 
    using bounded by (metis Groups.mult_ac)
qed

lemma kf_reconstruction_is_kraus_family:
  assumes sum: ρ. ((λa. sandwich_tc (f a) ρ) has_sum 𝔈 ρ) A
  defines F  Set.filter (λ(E,_). E0) ((λa. (f a, a)) ` A)
  shows kraus_family F
proof -
  from sum have bounded_clinear 𝔈
    by (rule kf_reconstruction_is_bounded_clinear)
  then obtain B where B: norm (𝔈 ρ)  B * norm ρ for ρ
    apply atomize_elim
    by (simp add: bounded_clinear_axioms_def bounded_clinear_def mult.commute)
  show ?thesis
  proof (intro kraus_familyI bdd_aboveI2)
    fix S assume S  {S. finite S  S  F}
    then have S  F and finite S
      by auto
    then obtain A' where finite A' and A'  A and S_A': S = (λa. (f a,a)) ` A'
      by (metis (no_types, lifting) F_def finite_subset_filter_image)
    show ((E, x)S. E* oCL E)  B *C id_cblinfun
    proof (rule cblinfun_leI)
      fix h :: 'a assume norm h = 1
      have h C (((E, x)S. E* oCL E) h) = h C (aA'. (f a)* oCL f a) h
        by (simp add: S_A' sum.reindex inj_on_def)
      also have  = (aA'. h C ((f a)* oCL f a) h)
        apply (rule complex_vector.linear_sum)
        by (simp add: bounded_clinear.clinear bounded_clinear_cinner_right_comp) 
      also have  = (aA'. trace_tc (sandwich_tc (f a) (tc_butterfly h h)))
        by (auto intro!: sum.cong[OF refl]
            simp: trace_tc.rep_eq from_trace_class_sandwich_tc
            tc_butterfly.rep_eq cblinfun_comp_butterfly sandwich_apply trace_butterfly_comp)
      also have  = trace_tc (aA'. sandwich_tc (f a) (tc_butterfly h h))
        apply (rule complex_vector.linear_sum[symmetric])
        using clinearI trace_tc_plus trace_tc_scaleC by blast
      also have  = trace_tc (aA'. sandwich_tc (f a) (tc_butterfly h h))
        by (simp add: finite A')
      also have   trace_tc (aA. (sandwich_tc (f a) (tc_butterfly h h)))
        apply (intro trace_tc_mono infsum_mono_neutral_traceclass)
        using A'  A sum[of tc_butterfly h h]
        by (auto intro!: sandwich_tc_pos has_sum_imp_summable simp: finite A')
      also have  = trace_tc (𝔈 (tc_butterfly h h))
        by (metis sum infsumI)
      also have  = norm (𝔈 (tc_butterfly h h))
        by (metis (no_types, lifting) infsumI infsum_nonneg_traceclass norm_tc_pos sandwich_tc_pos sum tc_butterfly_pos)
      also from B have   B * norm (tc_butterfly h h)
        using complex_of_real_mono by blast
      also have  = B
        by (simp add: norm h = 1 norm_tc_butterfly)
      also have  = h C (complex_of_real B *C id_cblinfun *V h)
        using norm h = 1 cnorm_eq_1 by auto
      finally show h C (((E, x)S. E* oCL E) *V h)  h C (complex_of_real B *C id_cblinfun *V h)
        by -
    qed
  next
    show 0  fst ` F
      by (force simp: F_def)
  qed
qed

lemma kf_reconstruction:
  assumes sum: ρ. ((λa. sandwich_tc (f a) ρ) has_sum 𝔈 ρ) A
  defines F  Abs_kraus_family (Set.filter (λ(E,_). E0) ((λa. (f a, a)) ` A))
  shows kf_apply F = 𝔈
proof (rule ext)
  fix ρ :: ('a, 'a) trace_class
  have Rep_F: Rep_kraus_family F = (Set.filter (λ(E,_). E0) ((λa. (f a,a)) ` A))
    unfolding F_def
    apply (rule Abs_kraus_family_inverse)
    by (auto intro!: kf_reconstruction_is_kraus_family[of _ _ 𝔈] assms simp: F_def)
  have ((λ(E,x). sandwich_tc E ρ) has_sum kf_apply F ρ) (Rep_kraus_family F)
    by (auto intro!: kf_apply_has_sum)
  then have ((λ(E,x). sandwich_tc E ρ) has_sum kf_apply F ρ) ((λa. (f a,a)) ` A)
    apply (rule has_sum_cong_neutral[THEN iffD2, rotated -1])
    by (auto simp: Rep_F)
  then have ((λa. sandwich_tc (f a) ρ) has_sum kf_apply F ρ) A
    apply (subst (asm) has_sum_reindex)
    by (auto simp: inj_on_def o_def)
  with sum show kf_apply F ρ = 𝔈 ρ
    by (metis (no_types, lifting) infsumI)
qed

subsection ‹Cleanup›

unbundle no cblinfun_syntax
unbundle no kraus_map_syntax

end