Theory Misc_CryptHOL

(* Title: Misc_CryptHOL.thy
  Author: Andreas Lochbihler, ETH Zurich *)

section ‹Miscellaneous library additions›

theory Misc_CryptHOL imports 
  Probabilistic_While.While_SPMF
  "HOL-Library.Rewrite"
  "HOL-Library.Simps_Case_Conv"
  "HOL-Library.Type_Length"
  "HOL-Eisbach.Eisbach"
  Coinductive.TLList
  Monad_Normalisation.Monad_Normalisation
  Monomorphic_Monad.Monomorphic_Monad
  Applicative_Lifting.Applicative
begin

hide_const (open) Henstock_Kurzweil_Integration.negligible

declare eq_on_def [simp del]

subsection ‹HOL›

lemma asm_rl_conv: "(PROP P  PROP P)  Trueprop True"
by(rule equal_intr_rule) iprover+

named_theorems if_distribs "Distributivity theorems for If"

lemma if_mono_cong: "b  x  x'; ¬ b  y  y'   If b x y  If b x' y'"
by simp

lemma if_cong_then: " b = b'; b'  t = t'; e = e'   If b t e = If b' t' e'"
by simp

lemma if_False_eq: " b  False; e = e'   If b t e = e'"
by auto

lemma imp_OO_imp [simp]: "(⟶) OO (⟶) = (⟶)"
by auto

lemma inj_on_fun_updD: " inj_on (f(x := y)) A; x  A   inj_on f A"
by(auto simp add: inj_on_def split: if_split_asm)

lemma disjoint_notin1: " A  B = {}; x  B   x  A" by auto

lemma Least_le_Least:
  fixes x :: "'a :: wellorder"
  assumes "Q x"
  and Q: "x. Q x  yx. P y"
  shows "Least P  Least Q"
  by (metis assms order_trans wellorder_Least_lemma)

lemma is_empty_image [simp]: "Set.is_empty (f ` A) = Set.is_empty A"
  by(auto simp add: Set.is_empty_def)

subsection ‹Relations›

inductive Imagep :: "('a  'b  bool)  ('a  bool)  'b  bool"
  for R P
where ImagepI: " P x; R x y   Imagep R P y"

lemma r_r_into_tranclp: " r x y; r y z   r^++ x z"
by(rule tranclp.trancl_into_trancl)(rule tranclp.r_into_trancl)

lemma transp_tranclp_id:
  assumes "transp R"
  shows "tranclp R = R"
proof(intro ext iffI)
  fix x y
  assume "R^++ x y"
  thus "R x y" by induction(blast dest: transpD[OF assms])+
qed simp

lemma transp_inv_image: "transp r  transp (λx y. r (f x) (f y))"
using trans_inv_image[where r="{(x, y). r x y}" and f = f]
by(simp add: transp_trans inv_image_def)

lemma Domainp_conversep: "Domainp R¯¯ = Rangep R"
by(auto)

lemma bi_unique_rel_set_bij_betw:
  assumes unique: "bi_unique R"
  and rel: "rel_set R A B"
  shows "f. bij_betw f A B  (xA. R x (f x))"
proof -
  from assms obtain f where f: "x. x  A  R x (f x)" and B: "x. x  A  f x  B"
    apply(atomize_elim)
    apply(fold all_conj_distrib)
    apply(subst choice_iff[symmetric])
    apply(auto dest: rel_setD1)
    done
  have "inj_on f A" by(rule inj_onI)(auto dest!: f dest: bi_uniqueDl[OF unique])
  moreover have "f ` A = B" using rel
    by(auto 4 3 intro: B dest: rel_setD2 f bi_uniqueDr[OF unique])
  ultimately have "bij_betw f A B" unfolding bij_betw_def ..
  thus ?thesis using f by blast
qed

definition restrict_relp :: "('a  'b  bool)  ('a  bool)  ('b  bool)  'a  'b  bool"
  (‹_  (_  _) [53, 54, 54] 53)
where "restrict_relp R P Q = (λx y. R x y  P x  Q y)"

lemma restrict_relp_apply [simp]: "(R  P  Q) x y  R x y  P x  Q y"
by(simp add: restrict_relp_def)

lemma restrict_relpI [intro?]: " R x y; P x; Q y   (R  P  Q) x y"
by(simp add: restrict_relp_def)

lemma restrict_relpE [elim?, cases pred]:
  assumes "(R  P  Q) x y"
  obtains (restrict_relp) "R x y" "P x" "Q y"
using assms by(simp add: restrict_relp_def)

lemma conversep_restrict_relp [simp]: "(R  P  Q)¯¯ = R¯¯  Q  P"
by(auto simp add: fun_eq_iff)

lemma restrict_relp_restrict_relp [simp]: "R  P  Q  P'  Q' = R  inf P P'  inf Q Q'"
by(auto simp add: fun_eq_iff)

lemma restrict_relp_cong:
  " P = P'; Q = Q'; x y.  P x; Q y   R x y = R' x y   R  P  Q = R'  P'  Q'"
by(auto simp add: fun_eq_iff)

lemma restrict_relp_cong_simp:
  " P = P'; Q = Q'; x y. P x =simp=> Q y =simp=> R x y = R' x y   R  P  Q = R'  P'  Q'"
by(rule restrict_relp_cong; simp add: simp_implies_def)

lemma restrict_relp_parametric [transfer_rule]:
  includes lifting_syntax shows
  "((A ===> B ===> (=)) ===> (A ===> (=)) ===> (B ===> (=)) ===> A ===> B ===> (=)) restrict_relp restrict_relp"
unfolding restrict_relp_def[abs_def] by transfer_prover

lemma restrict_relp_mono: " R  R'; P  P'; Q  Q'   R  P  Q  R'  P'  Q'"
by(simp add: le_fun_def)

lemma restrict_relp_mono': 
  " (R  P  Q) x y;  R x y; P x; Q y   R' x y &&& P' x &&& Q' y 
   (R'  P'  Q') x y"
by(auto dest: conjunctionD1 conjunctionD2)

lemma restrict_relp_DomainpD: "Domainp (R  P  Q) x  Domainp R x  P x"
by(auto simp add: Domainp.simps)

lemma restrict_relp_True: "R  (λ_. True)  (λ_. True) = R"
by(simp add: fun_eq_iff)

lemma restrict_relp_False1: "R  (λ_. False)  Q = bot"
by(simp add: fun_eq_iff)

lemma restrict_relp_False2: "R  P  (λ_. False) = bot"
by(simp add: fun_eq_iff)

definition rel_prod2 :: "('a  'b  bool)  'a  ('c × 'b)  bool"
where "rel_prod2 R a = (λ(c, b). R a b)"

lemma rel_prod2_simps [simp]: "rel_prod2 R a (c, b)  R a b"
by(simp add: rel_prod2_def)

lemma restrict_rel_prod:
  "rel_prod (R  I1  I2) (S  I1'  I2') = rel_prod R S  pred_prod I1 I1'  pred_prod I2 I2'"
by(auto simp add: fun_eq_iff)

lemma restrict_rel_prod1:
  "rel_prod (R  I1  I2) S = rel_prod R S  pred_prod I1 (λ_. True)  pred_prod I2 (λ_. True)"
by(simp add: restrict_rel_prod[symmetric] restrict_relp_True)

lemma restrict_rel_prod2:
  "rel_prod R (S  I1  I2) = rel_prod R S  pred_prod (λ_. True) I1  pred_prod (λ_. True) I2"
by(simp add: restrict_rel_prod[symmetric] restrict_relp_True)

consts relcompp_witness :: "('a  'b  bool)  ('b  'c  bool)  'a × 'c  'b"
specification (relcompp_witness)
  relcompp_witness1: "(A OO B) (fst xy) (snd xy)  A (fst xy) (relcompp_witness A B xy)"
  relcompp_witness2: "(A OO B) (fst xy) (snd xy)  B (relcompp_witness A B xy) (snd xy)"
  apply(fold all_conj_distrib)
  apply(rule choice allI)+
  by(auto intro: choice allI)

lemmas relcompp_witness[of _ _ "(x, y)" for x y, simplified] = relcompp_witness1 relcompp_witness2

hide_fact (open) relcompp_witness1 relcompp_witness2

lemma relcompp_witness_eq [simp]: "relcompp_witness (=) (=) (x, x) = x"
  using relcompp_witness(1)[of "(=)" "(=)" x x] by(simp add: eq_OO)

subsection ‹Pairs›

lemma split_apfst [simp]: "case_prod h (apfst f xy) = case_prod (h  f) xy"
by(cases xy) simp

definition corec_prod :: "('s  'a)  ('s  'b)  's  'a × 'b"
where "corec_prod f g = (λs. (f s, g s))"

lemma corec_prod_apply: "corec_prod f g s = (f s, g s)"
by(simp add: corec_prod_def)

lemma corec_prod_sel [simp]:
  shows fst_corec_prod: "fst (corec_prod f g s) = f s"
  and snd_corec_prod: "snd (corec_prod f g s) = g s"
by(simp_all add: corec_prod_apply)

lemma apfst_corec_prod [simp]: "apfst h (corec_prod f g s) = corec_prod (h  f) g s"
by(simp add: corec_prod_apply)

lemma apsnd_corec_prod [simp]: "apsnd h (corec_prod f g s) = corec_prod f (h  g) s"
by(simp add: corec_prod_apply)

lemma map_corec_prod [simp]: "map_prod f g (corec_prod h k s) = corec_prod (f  h) (g  k) s"
by(simp add: corec_prod_apply)

lemma split_corec_prod [simp]: "case_prod h (corec_prod f g s) = h (f s) (g s)"
by(simp add: corec_prod_apply)

lemma Pair_fst_Unity: "(fst x, ()) = x"
  by(cases x) simp

definition rprodl :: "('a × 'b) × 'c  'a × ('b × 'c)" where "rprodl = (λ((a, b), c). (a, (b, c)))"

lemma rprodl_simps [simp]: "rprodl ((a, b), c) = (a, (b, c))"
  by(simp add: rprodl_def)

lemma rprodl_parametric [transfer_rule]: includes lifting_syntax shows
  "(rel_prod (rel_prod A B) C ===> rel_prod A (rel_prod B C)) rprodl rprodl"
  unfolding rprodl_def by transfer_prover

definition lprodr :: "'a × ('b × 'c)  ('a × 'b) × 'c" where "lprodr = (λ(a, b, c). ((a, b), c))"

lemma lprodr_simps [simp]: "lprodr (a, b, c) = ((a, b), c)"
  by(simp add: lprodr_def)

lemma lprodr_parametric [transfer_rule]: includes lifting_syntax shows
  "(rel_prod A (rel_prod B C) ===> rel_prod (rel_prod A B) C) lprodr lprodr"
  unfolding lprodr_def by transfer_prover

lemma lprodr_inverse [simp]: "rprodl (lprodr x) = x"
  by(cases x) auto

lemma rprodl_inverse [simp]: "lprodr (rprodl x) = x"
  by(cases x) auto

lemma pred_prod_mono' [mono]:
  "pred_prod A B xy  pred_prod A' B' xy"
  if "x. A x  A' x" "y. B y  B' y"
  using that by(cases xy) auto

fun rel_witness_prod :: "('a × 'b) × ('c × 'd)  (('a × 'c) × ('b × 'd))" where
  "rel_witness_prod ((a, b), (c, d)) = ((a, c), (b, d))"

subsection ‹Sums›

lemma islE:
  assumes "isl x"
  obtains l where "x = Inl l"
using assms by(cases x) auto

lemma Inl_in_Plus [simp]: "Inl x  A <+> B  x  A"
by auto

lemma Inr_in_Plus [simp]: "Inr x  A <+> B  x  B"
by auto

lemma Inl_eq_map_sum_iff: "Inl x = map_sum f g y  (z. y = Inl z  x = f z)"
by(cases y) auto

lemma Inr_eq_map_sum_iff: "Inr x = map_sum f g y  (z. y = Inr z  x = g z)"
by(cases y) auto

lemma inj_on_map_sum [simp]:
  " inj_on f A; inj_on g B   inj_on (map_sum f g) (A <+> B)"
proof(rule inj_onI, goal_cases)
  case (1 x y)
  then show ?case by(cases x; cases y; auto simp add: inj_on_def)
qed

lemma inv_into_map_sum:
  "inv_into (A <+> B) (map_sum f g) x = map_sum (inv_into A f) (inv_into B g) x"
  if "x  f ` A <+> g ` B" "inj_on f A" "inj_on g B"
  using that by(cases rule: PlusE[consumes 1])(auto simp add: inv_into_f_eq f_inv_into_f)

fun rsuml :: "('a + 'b) + 'c  'a + ('b + 'c)" where
  "rsuml (Inl (Inl a)) = Inl a"
| "rsuml (Inl (Inr b)) = Inr (Inl b)"
| "rsuml (Inr c) = Inr (Inr c)"

fun lsumr :: "'a + ('b + 'c)  ('a + 'b) + 'c" where
  "lsumr (Inl a) = Inl (Inl a)"
| "lsumr (Inr (Inl b)) = Inl (Inr b)"
| "lsumr (Inr (Inr c)) = Inr c"

lemma rsuml_lsumr [simp]: "rsuml (lsumr x) = x"
  by(cases x rule: lsumr.cases) simp_all

lemma lsumr_rsuml [simp]: "lsumr (rsuml x) = x"
  by(cases x rule: rsuml.cases) simp_all

subsection ‹Option›

declare is_none_bind [simp]

lemma case_option_collapse: "case_option x (λ_. x) y = x"
by(simp split: option.split)

lemma indicator_single_Some: "indicator {Some x} (Some y) = indicator {x} y"
by(simp split: split_indicator)

subsubsection ‹Predicator and relator›

lemma option_pred_mono_strong:
  " pred_option P x; a.  a  set_option x; P a   P' a   pred_option P' x"
by(fact option.pred_mono_strong)

lemma option_pred_map [simp]: "pred_option P (map_option f x) = pred_option (P  f) x"
by(fact option.pred_map)

lemma option_pred_o_map [simp]: "pred_option P  map_option f = pred_option (P  f)"
by(simp add: fun_eq_iff)

lemma option_pred_bind [simp]: "pred_option P (Option.bind x f) = pred_option (pred_option P  f) x"
by(simp add: pred_option_def)

lemma pred_option_conj [simp]:
  "pred_option (λx. P x  Q x) = (λx. pred_option P x  pred_option Q x)"
by(auto simp add: pred_option_def)

lemma pred_option_top [simp]:
  "pred_option (λ_. True) = (λ_. True)"
by(fact option.pred_True)

lemma rel_option_restrict_relpI [intro?]:
  " rel_option R x y; pred_option P x; pred_option Q y   rel_option (R  P  Q) x y"
by(erule option.rel_mono_strong) simp

lemma rel_option_restrict_relpE [elim?]:
  assumes "rel_option (R  P  Q) x y"
  obtains "rel_option R x y" "pred_option P x" "pred_option Q y"
proof
  show "rel_option R x y" using assms by(auto elim!: option.rel_mono_strong)
  have "pred_option (Domainp (R  P  Q)) x" using assms by(fold option.Domainp_rel) blast
  then show "pred_option P x" by(rule option_pred_mono_strong)(blast dest!: restrict_relp_DomainpD)
  have "pred_option (Domainp (R  P  Q)¯¯) y" using assms
    by(fold option.Domainp_rel)(auto simp only: option.rel_conversep Domainp_conversep)
  then show "pred_option Q y" by(rule option_pred_mono_strong)(auto dest!: restrict_relp_DomainpD)
qed

lemma rel_option_restrict_relp_iff:
  "rel_option (R  P  Q) x y  rel_option R x y  pred_option P x  pred_option Q y"
by(blast intro: rel_option_restrict_relpI elim: rel_option_restrict_relpE)

lemma option_rel_map_restrict_relp:
  shows option_rel_map_restrict_relp1:
  "rel_option (R  P  Q) (map_option f x) = rel_option (R  f  P  f  Q) x"
  and option_rel_map_restrict_relp2:
  "rel_option (R  P  Q) x (map_option g y) = rel_option ((λx. R x  g)  P  Q  g) x y"
by(simp_all add: option.rel_map restrict_relp_def fun_eq_iff)

fun rel_witness_option :: "'a option × 'b option  ('a × 'b) option" where
  "rel_witness_option (Some x, Some y) = Some (x, y)"
| "rel_witness_option (None, None) = None"
| "rel_witness_option _ = None" ― ‹Just to make the definition complete›

lemma rel_witness_option:
  shows set_rel_witness_option: " rel_option A x y; (a, b)  set_option (rel_witness_option (x, y))   A a b"
    and map1_rel_witness_option: "rel_option A x y  map_option fst (rel_witness_option (x, y)) = x"
    and map2_rel_witness_option: "rel_option A x y  map_option snd (rel_witness_option (x, y)) = y"
  by(cases "(x, y)" rule: rel_witness_option.cases; simp; fail)+

lemma rel_witness_option1:
  assumes "rel_option A x y"
  shows "rel_option (λa (a', b). a = a'  A a' b) x (rel_witness_option (x, y))"
  using map1_rel_witness_option[OF assms, symmetric]
  unfolding option.rel_eq[symmetric] option.rel_map
  by(rule option.rel_mono_strong)(auto intro: set_rel_witness_option[OF assms])

lemma rel_witness_option2:
  assumes "rel_option A x y"
  shows "rel_option (λ(a, b') b. b = b'  A a b') (rel_witness_option (x, y)) y"
  using map2_rel_witness_option[OF assms]
  unfolding option.rel_eq[symmetric] option.rel_map
  by(rule option.rel_mono_strong)(auto intro: set_rel_witness_option[OF assms])

subsubsection ‹Orders on option›

abbreviation le_option :: "'a option  'a option  bool"
where "le_option  ord_option (=)"

lemma le_option_bind_mono:
  " le_option x y; a. a  set_option x  le_option (f a) (g a) 
   le_option (Option.bind x f) (Option.bind y g)"
by(cases x) simp_all

lemma le_option_refl [simp]: "le_option x x"
by(cases x) simp_all


lemma le_option_conv_option_ord: "le_option = option_ord"
by(auto simp add: fun_eq_iff flat_ord_def elim: ord_option.cases)

definition pcr_Some :: "('a  'b  bool)  'a  'b option  bool"
where "pcr_Some R x y  (z. y = Some z  R x z)"

lemma pcr_Some_simps [simp]: "pcr_Some R x (Some y)  R x y"
by(simp add: pcr_Some_def)

lemma pcr_SomeE [cases pred]:
  assumes "pcr_Some R x y"
  obtains (pcr_Some) z where "y = Some z" "R x z"
using assms by(auto simp add: pcr_Some_def)

subsubsection ‹Filter for option›

fun filter_option :: "('a  bool)  'a option  'a option"
where
  "filter_option P None = None"
| "filter_option P (Some x) = (if P x then Some x else None)"

lemma set_filter_option [simp]: "set_option (filter_option P x) = {y  set_option x. P y}"
by(cases x) auto

lemma filter_map_option: "filter_option P (map_option f x) = map_option f (filter_option (P  f) x)"
by(cases x) simp_all

lemma is_none_filter_option [simp]: "Option.is_none (filter_option P x)  Option.is_none x  ¬ P (the x)"
by(cases x) simp_all

lemma filter_option_eq_Some_iff [simp]: "filter_option P x = Some y  x = Some y  P y"
by(cases x) auto

lemma Some_eq_filter_option_iff [simp]: "Some y = filter_option P x  x = Some y  P y"
by(cases x) auto

lemma filter_conv_bind_option: "filter_option P x = Option.bind x (λy. if P y then Some y else None)"
by(cases x) simp_all

subsubsection ‹Assert for option›

primrec assert_option :: "bool  unit option" where
  "assert_option True = Some ()"
| "assert_option False = None"

lemma set_assert_option_conv: "set_option (assert_option b) = (if b then {()} else {})"
by(simp)

lemma in_set_assert_option [simp]: "x  set_option (assert_option b)  b"
by(cases b) simp_all


subsubsection ‹Join on options›

definition join_option :: "'a option option  'a option"
where "join_option x = (case x of Some y  y | None  None)"

simps_of_case join_simps [simp, code]: join_option_def

lemma set_join_option [simp]: "set_option (join_option x) = (set_option ` set_option x)"
by(cases x)(simp_all)

lemma in_set_join_option: "x  set_option (join_option (Some (Some x)))"
by simp

lemma map_join_option: "map_option f (join_option x) = join_option (map_option (map_option f) x)"
by(cases x) simp_all

lemma bind_conv_join_option: "Option.bind x f = join_option (map_option f x)"
by(cases x) simp_all

lemma join_conv_bind_option: "join_option x = Option.bind x id"
by(cases x) simp_all

lemma join_option_parametric [transfer_rule]:
  includes lifting_syntax shows
  "(rel_option (rel_option R) ===> rel_option R) join_option join_option"
unfolding join_conv_bind_option[abs_def] by transfer_prover

lemma join_option_eq_Some [simp]: "join_option x = Some y  x = Some (Some y)"
by(cases x) simp_all

lemma Some_eq_join_option [simp]: "Some y = join_option x  x = Some (Some y)"
by(cases x) auto

lemma join_option_eq_None: "join_option x = None  x = None  x = Some None"
by(cases x) simp_all

lemma None_eq_join_option: "None = join_option x  x = None  x = Some None"
by(cases x) auto

subsubsection ‹Zip on options›

function zip_option :: "'a option  'b option  ('a × 'b) option"
where
  "zip_option (Some x) (Some y) = Some (x, y)"
| "zip_option _ None = None"
| "zip_option None _ = None"
by pat_completeness auto
termination by lexicographic_order

lemma zip_option_eq_Some_iff [iff]:
  "zip_option x y = Some (a, b)  x = Some a  y = Some b"
by(cases "(x, y)" rule: zip_option.cases) simp_all

lemma set_zip_option [simp]:
  "set_option (zip_option x y) = set_option x × set_option y"
by auto

lemma zip_map_option1: "zip_option (map_option f x) y = map_option (apfst f) (zip_option x y)"
by(cases "(x, y)" rule: zip_option.cases) simp_all

lemma zip_map_option2: "zip_option x (map_option g y) = map_option (apsnd g) (zip_option x y)"
by(cases "(x, y)" rule: zip_option.cases) simp_all

lemma map_zip_option:
  "map_option (map_prod f g) (zip_option x y) = zip_option (map_option f x) (map_option g y)"
by(simp add: zip_map_option1 zip_map_option2 option.map_comp apfst_def apsnd_def o_def prod.map_comp)

lemma zip_conv_bind_option:
  "zip_option x y = Option.bind x (λx. Option.bind y (λy. Some (x, y)))"
by(cases "(x, y)" rule: zip_option.cases) simp_all

lemma zip_option_parametric [transfer_rule]:
  includes lifting_syntax shows
  "(rel_option R ===> rel_option Q ===> rel_option (rel_prod R Q)) zip_option zip_option"
unfolding zip_conv_bind_option[abs_def] by transfer_prover

lemma rel_option_eqI [simp]: "rel_option (=) x x"
by(simp add: option.rel_eq)

subsubsection ‹Binary supremum on @{typ "'a option"}

primrec sup_option :: "'a option  'a option  'a option"
where
  "sup_option x None = x"
| "sup_option x (Some y) = (Some y)"

lemma sup_option_idem [simp]: "sup_option x x = x"
by(cases x) simp_all

lemma sup_option_assoc: "sup_option (sup_option x y) z = sup_option x (sup_option y z)"
by(cases z) simp_all

lemma sup_option_left_idem: "sup_option x (sup_option x y) = sup_option x y"
by(rewrite sup_option_assoc[symmetric])(simp)

lemmas sup_option_ai = sup_option_assoc sup_option_left_idem

lemma sup_option_None [simp]: "sup_option None y = y"
by(cases y) simp_all

subsubsection ‹Restriction on @{typ "'a option"}

primrec (transfer) enforce_option :: "('a  bool)  'a option  'a option" where
  "enforce_option P (Some x) = (if P x then Some x else None)"
| "enforce_option P None = None"

lemma set_enforce_option [simp]: "set_option (enforce_option P x) = {a  set_option x. P a}"
  by(cases x) auto

lemma enforce_map_option: "enforce_option P (map_option f x) = map_option f (enforce_option (P  f) x)"
  by(cases x) auto

lemma enforce_bind_option [simp]:
  "enforce_option P (Option.bind x f) = Option.bind x (enforce_option P  f)"
  by(cases x) auto

lemma enforce_option_alt_def:
  "enforce_option P x = Option.bind x (λa. Option.bind (assert_option (P a)) (λ_ :: unit. Some a))"
  by(cases x) simp_all

lemma enforce_option_eq_None_iff [simp]:
  "enforce_option P x = None  (a. x = Some a  ¬ P a)"
  by(cases x) auto

lemma enforce_option_eq_Some_iff [simp]:
  "enforce_option P x = Some y  x = Some y  P y"
  by(cases x) auto

lemma Some_eq_enforce_option_iff [simp]:
  "Some y = enforce_option P x  x = Some y  P y"
  by(cases x) auto

lemma enforce_option_top [simp]: "enforce_option  = id"
  by(rule ext; rename_tac x; case_tac x; simp)

lemma enforce_option_K_True [simp]: "enforce_option (λ_. True) x = x"
  by(cases x) simp_all

lemma enforce_option_bot [simp]: "enforce_option  = (λ_. None)"
  by(simp add: fun_eq_iff)

lemma enforce_option_K_False [simp]: "enforce_option (λ_. False) x = None"
  by simp

lemma enforce_pred_id_option: "pred_option P x  enforce_option P x = x"
  by(cases x) auto

subsubsection ‹Maps›

lemma map_add_apply: "(m1 ++ m2) x = sup_option (m1 x) (m2 x)"
by(simp add: map_add_def split: option.split)

lemma map_le_map_upd2: " f m g; y'. f x = Some y'  y' = y   f m g(x  y)"
by(cases "x  dom f")(auto simp add: map_le_def Ball_def)

lemma eq_None_iff_not_dom: "f x = None  x  dom f"
by auto

lemma card_ran_le_dom: "finite (dom m)  card (ran m)  card (dom m)"
by(simp add: ran_alt_def card_image_le)

lemma dom_subset_ran_iff:
  assumes "finite (ran m)"
  shows "dom m  ran m  dom m = ran m"
proof
  assume le: "dom m  ran m"
  then have "card (dom m)  card (ran m)" by(simp add: card_mono assms)
  moreover have "card (ran m)  card (dom m)" by(simp add: finite_subset[OF le assms] card_ran_le_dom)
  ultimately show "dom m = ran m" using card_subset_eq[OF assms le] by simp
qed simp

text ‹
  We need a polymorphic constant for the empty map such that transfer_prover›
  can use a custom transfer rule for @{const Map.empty}
definition Map_empty where [simp]: "Map_empty  Map.empty"

lemma map_le_Some1D: " m m m'; m x = Some y   m' x = Some y"
by(auto simp add: map_le_def Ball_def)

lemma map_le_fun_upd2: " f m g; x  dom f   f m g(x := y)"
by(auto simp add: map_le_def)

lemma map_eqI: "xdom m  dom m'. m x = m' x  m = m'"
by(auto simp add: fun_eq_iff domIff intro: option.expand)


subsection ‹Countable›

lemma countable_lfp:
  assumes step: "Y. countable Y  countable (F Y)"
  and cont: "Order_Continuity.sup_continuous F"
  shows "countable (lfp F)"
by(subst sup_continuous_lfp[OF cont])(simp add: countable_funpow[OF step])

lemma countable_lfp_apply:
  assumes step: "Y x. (x. countable (Y x))  countable (F Y x)"
  and cont: "Order_Continuity.sup_continuous F"
  shows "countable (lfp F x)"
proof -
  { fix n
    have "x. countable ((F ^^ n) bot x)"
      by(induct n)(auto intro: step) }
  thus ?thesis using cont by(simp add: sup_continuous_lfp)
qed


subsection ‹ Extended naturals ›

lemma idiff_enat_eq_enat_iff: "x - enat n = enat m  (k. x = enat k  k - n = m)"
  by (cases x) simp_all

lemma eSuc_SUP: "A  {}  eSuc ( (f ` A)) = (xA. eSuc (f x))"
  by (subst eSuc_Sup) (simp_all add: image_comp)

lemma ereal_of_enat_1: "ereal_of_enat 1 = ereal 1"
  by (simp add: one_enat_def)

lemma ennreal_real_conv_ennreal_of_enat: "ennreal (real n) = ennreal_of_enat n"
  by (simp add: ennreal_of_nat_eq_real_of_nat)

lemma enat_add_sub_same2: "b    a + b - b = (a :: enat)"
  by (cases a; cases b) simp_all

lemma enat_sub_add: "y  x  x - y + z = x + z - (y :: enat)"
  by (cases x; cases y; cases z) simp_all

lemma SUP_enat_eq_0_iff [simp]: " (f ` A) = (0 :: enat)  (xA. f x = 0)"
  by (simp add: bot_enat_def [symmetric])

lemma SUP_enat_add_left:
  assumes "I  {}"
  shows "(SUP iI. f i + c :: enat) = (SUP iI. f i) + c" (is "?lhs = ?rhs")
proof(cases "c", rule antisym)
  case (enat n)
  show "?lhs  ?rhs" by(auto 4 3 intro: SUP_upper intro: SUP_least)
  have "(SUP iI. f i)  ?lhs - c" using enat 
    by(auto simp add: enat_add_sub_same2 intro!: SUP_least order_trans[OF _ SUP_upper[THEN enat_minus_mono1]])
  note add_right_mono[OF this, of c]
  also have " + c  ?lhs" using assms
    by(subst enat_sub_add)(auto intro: SUP_upper2 simp add: enat_add_sub_same2 enat)
  finally show "?rhs  ?lhs" .
qed(simp add: assms SUP_constant)

lemma SUP_enat_add_right:
  assumes "I  {}"
  shows "(SUP iI. c + f i :: enat) = c + (SUP iI. f i)"
using SUP_enat_add_left[OF assms, of f c]
by(simp add: add.commute)

lemma iadd_SUP_le_iff: "n + (SUP xA. f x :: enat)  y  (if A = {} then n  y else xA. n + f x  y)"
by(simp add: bot_enat_def SUP_enat_add_right[symmetric] SUP_le_iff)

lemma SUP_iadd_le_iff: "(SUP xA. f x :: enat) + n  y  (if A = {} then n  y else xA. f x + n  y)"
using iadd_SUP_le_iff[of n f A y] by(simp add: add.commute)


subsection ‹Extended non-negative reals›

lemma (in finite_measure) nn_integral_indicator_neq_infty: 
  "f -` A  sets M  (+ x. indicator A (f x) M)  "
unfolding ennreal_indicator[symmetric]
apply(rule integrableD)
apply(rule integrable_const_bound[where B=1])
apply(simp_all add: indicator_vimage[symmetric])
done

lemma (in finite_measure) nn_integral_indicator_neq_top: 
  "f -` A  sets M  (+ x. indicator A (f x) M)  "
by(drule nn_integral_indicator_neq_infty) simp

lemma nn_integral_indicator_map:
  assumes [measurable]: "f  measurable M N" "{xspace N. P x}  sets N"
  shows "(+x. indicator {xspace N. P x} (f x) M) = emeasure M {xspace M. P (f x)}"
  using assms(1)[THEN measurable_space] 
  by (subst nn_integral_indicator[symmetric])
     (auto intro!: nn_integral_cong split: split_indicator simp del: nn_integral_indicator)


subsection ‹BNF material›

lemma transp_rel_fun: " is_equality Q; transp R   transp (rel_fun Q R)"
by(rule transpI)(auto dest: transpD rel_funD simp add: is_equality_def)

lemma rel_fun_inf: "inf (rel_fun Q R) (rel_fun Q R') = rel_fun Q (inf R R')"
by(rule antisym)(auto elim: rel_fun_mono dest: rel_funD)

lemma reflp_fun1: includes lifting_syntax shows " is_equality A; reflp B   reflp (A ===> B)"
by(simp add: reflp_def rel_fun_def is_equality_def)

lemma type_copy_id': "type_definition (λx. x) (λx. x) UNIV"
by unfold_locales simp_all

lemma type_copy_id: "type_definition id id UNIV"
by(simp add: id_def type_copy_id')

lemma GrpE [cases pred]:
  assumes "BNF_Def.Grp A f x y"
  obtains (Grp) "y = f x" "x  A"
using assms
by(simp add: Grp_def)

lemma rel_fun_Grp_copy_Abs:
  includes lifting_syntax
  assumes "type_definition Rep Abs A"
  shows "rel_fun (BNF_Def.Grp A Abs) (BNF_Def.Grp B g) = BNF_Def.Grp {f. f ` A  B} (Rep ---> g)"
proof -
  interpret type_definition Rep Abs A by fact
  show ?thesis
    by(auto simp add: rel_fun_def Grp_def fun_eq_iff Abs_inverse Rep_inverse intro!: Rep)
qed

lemma rel_set_Grp:
  "rel_set (BNF_Def.Grp A f) = BNF_Def.Grp {B. B  A} (image f)"
by(auto simp add: rel_set_def BNF_Def.Grp_def fun_eq_iff)

lemma rel_set_comp_Grp:
  "rel_set R = (BNF_Def.Grp {x. x  {(x, y). R x y}} ((`) fst))¯¯ OO BNF_Def.Grp {x. x  {(x, y). R x y}} ((`) snd)"
apply(auto 4 4 del: ext intro!: ext simp add: BNF_Def.Grp_def intro!: rel_setI intro: rev_bexI)
apply(simp add: relcompp_apply)
subgoal for A B
  apply(rule exI[where x="A × B  {(x, y). R x y}"])
  apply(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI)
  done
done

lemma Domainp_Grp: "Domainp (BNF_Def.Grp A f) = (λx. x  A)"
by(auto simp add: fun_eq_iff Grp_def)

lemma pred_prod_conj [simp]:
  shows pred_prod_conj1: "P Q R. pred_prod (λx. P x  Q x) R = (λx. pred_prod P R x  pred_prod Q R x)"
  and pred_prod_conj2: "P Q R. pred_prod P (λx. Q x  R x) = (λx. pred_prod P Q x  pred_prod P R x)"
by(auto simp add: pred_prod.simps)

lemma pred_sum_conj [simp]:
  shows pred_sum_conj1: "P Q R. pred_sum (λx. P x  Q x) R = (λx. pred_sum P R x  pred_sum Q R x)"
  and pred_sum_conj2: "P Q R. pred_sum P (λx. Q x  R x) = (λx. pred_sum P Q x  pred_sum P R x)"
by(auto simp add: pred_sum.simps fun_eq_iff)

lemma pred_list_conj [simp]: "list_all (λx. P x  Q x) = (λx. list_all P x  list_all Q x)"
by(auto simp add: list_all_def)

lemma pred_prod_top [simp]:
  "pred_prod (λ_. True) (λ_. True) = (λ_. True)"
by(simp add: pred_prod.simps fun_eq_iff)

lemma rel_fun_conversep: includes lifting_syntax shows
  "(A^--1 ===> B^--1) = (A ===> B)^--1"
by(auto simp add: rel_fun_def fun_eq_iff)

lemma left_unique_Grp [iff]:
  "left_unique (BNF_Def.Grp A f)  inj_on f A"
unfolding Grp_def left_unique_def by(auto simp add: inj_on_def)

lemma right_unique_Grp [simp, intro!]: "right_unique (BNF_Def.Grp A f)"
by(simp add: Grp_def right_unique_def)

lemma bi_unique_Grp [iff]:
  "bi_unique (BNF_Def.Grp A f)  inj_on f A"
by(simp add: bi_unique_alt_def)

lemma left_total_Grp [iff]:
  "left_total (BNF_Def.Grp A f)  A = UNIV"
by(auto simp add: left_total_def Grp_def)

lemma right_total_Grp [iff]:
  "right_total (BNF_Def.Grp A f)  f ` A = UNIV"
by(auto simp add: right_total_def BNF_Def.Grp_def image_def)

lemma bi_total_Grp [iff]:
  "bi_total (BNF_Def.Grp A f)  A = UNIV  surj f"
by(auto simp add: bi_total_alt_def)

lemma left_unique_vimage2p [simp]:
  " left_unique P; inj f   left_unique (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro left_unique_OO) simp_all

lemma right_unique_vimage2p [simp]:
  " right_unique P; inj g   right_unique (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro right_unique_OO) simp_all

lemma bi_unique_vimage2p [simp]:
  " bi_unique P; inj f; inj g   bi_unique (BNF_Def.vimage2p f g P)"
unfolding bi_unique_alt_def by simp

lemma left_total_vimage2p [simp]:
  " left_total P; surj g   left_total (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro left_total_OO) simp_all

lemma right_total_vimage2p [simp]:
  " right_total P; surj f   right_total (BNF_Def.vimage2p f g P)"
unfolding vimage2p_Grp by(intro right_total_OO) simp_all

lemma bi_total_vimage2p [simp]:
  " bi_total P; surj f; surj g   bi_total (BNF_Def.vimage2p f g P)"
unfolding bi_total_alt_def by simp

lemma vimage2p_eq [simp]:
  "inj f  BNF_Def.vimage2p f f (=) = (=)"
by(auto simp add: vimage2p_def fun_eq_iff inj_on_def)

lemma vimage2p_conversep: "BNF_Def.vimage2p f g R^--1 = (BNF_Def.vimage2p g f R)^--1"
by(simp add: vimage2p_def fun_eq_iff)

lemma rel_fun_refl: " A  (=); (=)  B   (=)  rel_fun A B"
  by(subst fun.rel_eq[symmetric])(rule fun_mono)

lemma rel_fun_mono_strong:
  " rel_fun A B f g; A'  A; x y.  x  f ` {x. Domainp A' x}; y  g ` {x. Rangep A' x}; B x y   B' x y   rel_fun A' B' f g"
  by(auto simp add: rel_fun_def) fastforce

lemma rel_fun_refl_strong: 
  assumes "A  (=)" "x. x  f ` {x. Domainp A x}  B x x"
  shows "rel_fun A B f f"
proof -
  have "rel_fun (=) (=) f f" by(simp add: rel_fun_eq)
  then show ?thesis using assms(1)
    by(rule rel_fun_mono_strong) (auto intro: assms(2))
qed

lemma Grp_iff: "BNF_Def.Grp B g x y  y = g x  x  B" by(simp add: Grp_def)

lemma Rangep_Grp: "Rangep (BNF_Def.Grp A f) = (λx. x  f ` A)"
  by(auto simp add: fun_eq_iff Grp_iff)

lemma rel_fun_Grp:
  "rel_fun (BNF_Def.Grp UNIV h)¯¯ (BNF_Def.Grp A g) = BNF_Def.Grp {f. f ` range h  A} (map_fun h g)"
  by(auto simp add: rel_fun_def fun_eq_iff Grp_iff)

subsection ‹Transfer and lifting material›

context includes lifting_syntax begin

lemma monotone_parametric [transfer_rule]:
  assumes [transfer_rule]: "bi_total A"
  shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (A ===> B) ===> (=)) monotone monotone"
unfolding monotone_def[abs_def] by transfer_prover

lemma fun_ord_parametric [transfer_rule]:
  assumes [transfer_rule]: "bi_total C"
  shows "((A ===> B ===> (=)) ===> (C ===> A) ===> (C ===> B) ===> (=)) fun_ord fun_ord"
unfolding fun_ord_def[abs_def] by transfer_prover

lemma Plus_parametric [transfer_rule]:
  "(rel_set A ===> rel_set B ===> rel_set (rel_sum A B)) (<+>) (<+>)"
unfolding Plus_def[abs_def] by transfer_prover

lemma pred_fun_parametric [transfer_rule]:
  assumes [transfer_rule]: "bi_total A"
  shows "((A ===> (=)) ===> (B ===> (=)) ===> (A ===> B) ===> (=)) pred_fun pred_fun"
unfolding pred_fun_def by(transfer_prover)

lemma rel_fun_eq_OO: "((=) ===> A) OO ((=) ===> B) = ((=) ===> A OO B)"
by(clarsimp simp add: rel_fun_def fun_eq_iff relcompp.simps) metis

end

lemma Quotient_set_rel_eq:
  includes lifting_syntax
  assumes "Quotient R Abs Rep T"
  shows "(rel_set T ===> rel_set T ===> (=)) (rel_set R) (=)"
proof(rule rel_funI iffI)+
  fix A B C D
  assume AB: "rel_set T A B" and CD: "rel_set T C D"
  have *: "x y. R x y = (T x (Abs x)  T y (Abs y)  Abs x = Abs y)"
    "a b. T a b  Abs a = b"
    using assms unfolding Quotient_alt_def by simp_all

  { assume [simp]: "B = D"
    thus "rel_set R A C"
      by(auto 4 4 intro!: rel_setI dest: rel_setD1[OF AB, simplified] rel_setD2[OF AB, simplified] rel_setD2[OF CD] rel_setD1[OF CD] simp add: * elim!: rev_bexI)
  next
    assume AC: "rel_set R A C"
    show "B = D"
      apply safe
       apply(drule rel_setD2[OF AB], erule bexE)
       apply(drule rel_setD1[OF AC], erule bexE)
       apply(drule rel_setD1[OF CD], erule bexE)
       apply(simp add: *)
      apply(drule rel_setD2[OF CD], erule bexE)
      apply(drule rel_setD2[OF AC], erule bexE)
      apply(drule rel_setD1[OF AB], erule bexE)
      apply(simp add: *)
      done
  }
qed

lemma Domainp_eq: "Domainp (=) = (λ_. True)"
by(simp add: Domainp.simps fun_eq_iff)

lemma rel_fun_eq_onpI: "eq_onp (pred_fun P Q) f g  rel_fun (eq_onp P) (eq_onp Q) f g"
by(auto simp add: eq_onp_def rel_fun_def)

lemma bi_unique_eq_onp: "bi_unique (eq_onp P)"
by(simp add: bi_unique_def eq_onp_def)

lemma rel_fun_eq_conversep: includes lifting_syntax shows "(A¯¯ ===> (=)) = (A ===> (=))¯¯"
by(auto simp add: fun_eq_iff rel_fun_def)

lemma rel_fun_comp:
  "f g h. rel_fun A B (f  g) h = rel_fun A (λx. B (f x)) g h"
  "f g h. rel_fun A B f (g  h) = rel_fun A (λx y. B x (g y)) f h"
  by(auto simp add: rel_fun_def)

lemma rel_fun_map_fun1: "rel_fun (BNF_Def.Grp UNIV h)¯¯ A f g  rel_fun (=) A (map_fun h id f) g"
  by(auto simp add: rel_fun_def Grp_def)

lemma map_fun2_id: "map_fun f g x = g  map_fun f id x"
  by(simp add: map_fun_def o_assoc)

lemma map_fun_id2_in: "map_fun g h f = map_fun g id (h  f)"
  by(simp add: map_fun_def)

lemma Domainp_rel_fun_le: "Domainp (rel_fun A B)  pred_fun (Domainp A) (Domainp B)"
  by(auto dest: rel_funD)

definition rel_witness_fun :: "('a  'b  bool)  ('b  'c  bool)  ('a  'd) × ('c  'e)  ('b  'd × 'e)" where
  "rel_witness_fun A A' = (λ(f, g) b. (f (THE a. A a b), g (THE c. A' b c)))"

lemma
  assumes fg: "rel_fun (A OO A') B f g"
    and A: "left_unique A" "right_total A"
    and A': "right_unique A'" "left_total A'"
  shows rel_witness_fun1: "rel_fun A (λx (x', y). x = x'  B x' y) f (rel_witness_fun A A' (f, g))"
    and rel_witness_fun2: "rel_fun A' (λ(x, y') y. y = y'  B x y') (rel_witness_fun A A' (f, g)) g"
proof (goal_cases)
  case 1
  have "A x y  f x = f (THE a. A a y)  B (f (THE a. A a y)) (g (The (A' y)))" for x y 
    by(rule left_totalE[OF A'(2)]; erule meta_allE[of _ y]; erule exE; frule (1) fg[THEN rel_funD, OF relcomppI])
      (auto intro!: arg_cong[where f=f] arg_cong[where f=g] rel_funI the_equality the_equality[symmetric] dest: left_uniqueD[OF A(1)] right_uniqueD[OF A'(1)] elim!: arg_cong2[where f=B, THEN iffD2, rotated -1])

  with 1 show ?case by(clarsimp simp add: rel_fun_def rel_witness_fun_def)
next
  case 2
  have "A' x y  g y = g (The (A' x))  B (f (THE a. A a x)) (g (The (A' x)))" for x y
    by(rule right_totalE[OF A(2), of x]; frule (1) fg[THEN rel_funD, OF relcomppI])
      (auto intro!: arg_cong[where f=f] arg_cong[where f=g] rel_funI the_equality the_equality[symmetric] dest: left_uniqueD[OF A(1)] right_uniqueD[OF A'(1)] elim!: arg_cong2[where f=B, THEN iffD2, rotated -1])

  with 2 show ?case by(clarsimp simp add: rel_fun_def rel_witness_fun_def)    
qed

lemma rel_witness_fun_eq [simp]: "rel_witness_fun (=) (=) (f, g) = (λx. (f x, g x))"
  by(simp add: rel_witness_fun_def)

subsection ‹Arithmetic›

lemma abs_diff_triangle_ineq2: "¦a - b :: _ :: ordered_ab_group_add_abs¦  ¦a - c¦ + ¦c - b¦"
by(rule order_trans[OF _ abs_diff_triangle_ineq]) simp

lemma (in ordered_ab_semigroup_add) add_left_mono_trans:
  " x  a + b; b  c   x  a + c"
by(erule order_trans)(rule add_left_mono)

lemma of_nat_le_one_cancel_iff [simp]:
  fixes n :: nat shows "real n  1  n  1"
by linarith

lemma (in linordered_semidom) mult_right_le: "c  1  0  a  c * a  a"
by(subst mult.commute)(rule mult_left_le)

subsection ‹Chain-complete partial orders and partial_function›

lemma fun_ordD: "fun_ord ord f g  ord (f x) (g x)"
by(simp add: fun_ord_def)

lemma parallel_fixp_induct_strong:
  assumes ccpo1: "class.ccpo luba orda (mk_less orda)"
  and ccpo2: "class.ccpo lubb ordb (mk_less ordb)"
  and adm: "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (λx. P (fst x) (snd x))"
  and f: "monotone orda orda f"
  and g: "monotone ordb ordb g"
  and bot: "P (luba {}) (lubb {})"
  and step: "x y.  orda x (ccpo.fixp luba orda f); ordb y (ccpo.fixp lubb ordb g); P x y   P (f x) (g y)"
  shows "P (ccpo.fixp luba orda f) (ccpo.fixp lubb ordb g)"
proof -
  let ?P="λx y. orda x (ccpo.fixp luba orda f)  ordb y (ccpo.fixp lubb ordb g)  P x y"
  show ?thesis using ccpo1 ccpo2 _ f g
  proof(rule parallel_fixp_induct[where P="?P", THEN conjunct2, THEN conjunct2])
    note [cont_intro] = 
      admissible_leI[OF ccpo1] ccpo.mcont_const[OF ccpo1]
      admissible_leI[OF ccpo2] ccpo.mcont_const[OF ccpo2]
    show "ccpo.admissible (prod_lub luba lubb) (rel_prod orda ordb) (λxy. ?P (fst xy) (snd xy))"
      using adm by simp
    show "?P (luba {}) (lubb {})" using bot by(auto intro: ccpo.ccpo_Sup_least ccpo1 ccpo2 chain_empty)
    show "?P (f x) (g y)" if "?P x y" for x y using that
      apply(subst ccpo.fixp_unfold[OF ccpo1 f])
      apply(subst ccpo.fixp_unfold[OF ccpo2 g])
      apply(auto intro: step monotoneD[OF f] monotoneD[OF g])
      done
  qed
qed

lemma parallel_fixp_induct_strong_uc:
  assumes a: "partial_function_definitions orda luba"
  and b: "partial_function_definitions ordb lubb"
  and F: "x. monotone (fun_ord orda) orda (λf. U1 (F (C1 f)) x)"
  and G: "y. monotone (fun_ord ordb) ordb (λg. U2 (G (C2 g)) y)"
  and eq1: "f  C1 (ccpo.fixp (fun_lub luba) (fun_ord orda) (λf. U1 (F (C1 f))))"
  and eq2: "g  C2 (ccpo.fixp (fun_lub lubb) (fun_ord ordb) (λg. U2 (G (C2 g))))"
  and inverse: "f. U1 (C1 f) = f"
  and inverse2: "g. U2 (C2 g) = g"
  and adm: "ccpo.admissible (prod_lub (fun_lub luba) (fun_lub lubb)) (rel_prod (fun_ord orda) (fun_ord ordb)) (λx. P (fst x) (snd x))"
  and bot: "P (λ_. luba {}) (λ_. lubb {})"
  and step: "f' g'.  x. orda (U1 f' x) (U1 f x); y. ordb (U2 g' y) (U2 g y); P (U1 f') (U2 g')   P (U1 (F f')) (U2 (G g'))"
  shows "P (U1 f) (U2 g)"
apply(unfold eq1 eq2 inverse inverse2)
apply(rule parallel_fixp_induct_strong[OF partial_function_definitions.ccpo[OF a] partial_function_definitions.ccpo[OF b] adm])
using F apply(simp add: monotone_def fun_ord_def)
using G apply(simp add: monotone_def fun_ord_def)
apply(simp add: fun_lub_def bot)
apply(rule step; simp add: inverse inverse2 eq1 eq2 fun_ordD)
done

lemmas parallel_fixp_induct_strong_1_1 = parallel_fixp_induct_strong_uc[
  of _ _ _ _ "λx. x" _ "λx. x" "λx. x" _ "λx. x",
  OF _ _ _ _ _ _ refl refl]

lemmas parallel_fixp_induct_strong_2_2 = parallel_fixp_induct_strong_uc[
  of _ _ _ _ "case_prod" _ "curry" "case_prod" _ "curry",
  where P="λf g. P (curry f) (curry g)",
  unfolded case_prod_curry curry_case_prod curry_K,
  OF _ _ _ _ _ _ refl refl,
  split_format (complete), unfolded prod.case]
  for P

lemma fixp_induct_option': ― ‹Stronger induction rule›
  fixes F :: "'c  'c" and
    U :: "'c  'b  'a option" and
    C :: "('b  'a option)  'c" and
    P :: "'b  'a  bool"
  assumes mono: "x. mono_option (λf. U (F (C f)) x)"
  assumes eq: "f  C (ccpo.fixp (fun_lub (flat_lub None)) (fun_ord option_ord) (λf. U (F (C f))))"
  assumes inverse2: "f. U (C f) = f"
  assumes step: "g x y.  x y. U g x = Some y  P x y; U (F g) x = Some y; x. option_ord (U g x) (U f x)   P x y"
  assumes defined: "U f x = Some y"
  shows "P x y"
using step defined option.fixp_strong_induct_uc[of U F C, OF mono eq inverse2 option_admissible, of P]
unfolding fun_lub_def flat_lub_def fun_ord_def
by(simp (no_asm_use)) blast

declaration Partial_Function.init "option'" @{term option.fixp_fun}
  @{term option.mono_body} @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc}
  (SOME @{thm fixp_induct_option'})

lemma bot_fun_least [simp]: "(λ_. bot :: 'a :: order_bot)  x"
by(fold bot_fun_def) simp

lemma fun_ord_conv_rel_fun: "fun_ord = rel_fun (=)"
by(simp add: fun_ord_def fun_eq_iff rel_fun_def)

inductive finite_chains :: "('a  'a  bool)  bool"
  for ord
where finite_chainsI: "(Y. Complete_Partial_Order.chain ord Y  finite Y)  finite_chains ord"

lemma finite_chainsD: " finite_chains ord; Complete_Partial_Order.chain ord Y   finite Y"
by(rule finite_chains.cases)

lemma finite_chains_flat_ord [simp, intro!]: "finite_chains (flat_ord x)"
proof
  fix Y
  assume chain: "Complete_Partial_Order.chain (flat_ord x) Y"
  show "finite Y"
  proof(cases "y  Y. y  x")
    case True
    then obtain y where y: "y  Y" and yx: "y  x" by blast
    hence "Y  {x, y}" by(auto dest: chainD[OF chain] simp add: flat_ord_def)
    thus ?thesis by(rule finite_subset) simp
  next
    case False
    hence "Y  {x}" by auto
    thus ?thesis by(rule finite_subset) simp
  qed
qed    

lemma mcont_finite_chains:
  assumes finite: "finite_chains ord"
  and mono: "monotone ord ord' f"
  and ccpo: "class.ccpo lub ord (mk_less ord)"
  and ccpo': "class.ccpo lub' ord' (mk_less ord')"
  shows "mcont lub ord lub' ord' f"
proof(intro mcontI contI)
  fix Y
  assume chain: "Complete_Partial_Order.chain ord Y" and Y: "Y  {}"
  from finite chain have fin: "finite Y" by(rule finite_chainsD)
  from ccpo chain fin Y have lub: "lub Y  Y" by(rule ccpo.in_chain_finite)

  interpret ccpo': ccpo lub' ord' "mk_less ord'" by(rule ccpo')

  have chain': "Complete_Partial_Order.chain ord' (f ` Y)" using chain
    by(rule chain_imageI)(rule monotoneD[OF mono])

  have "ord' (f (lub Y)) (lub' (f ` Y))" using chain'
    by(rule ccpo'.ccpo_Sup_upper)(simp add: lub)
  moreover
  have "ord' (lub' (f ` Y)) (f (lub Y))" using chain'
    by(rule ccpo'.ccpo_Sup_least)(blast intro: monotoneD[OF mono] ccpo.ccpo_Sup_upper[OF ccpo chain])
  ultimately show "f (lub Y) = lub' (f ` Y)" by(rule ccpo'.order.antisym)
qed(fact mono)  

lemma rel_fun_curry: includes lifting_syntax shows
  "(A ===> B ===> C) f g  (rel_prod A B ===> C) (case_prod f) (case_prod g)"
by(auto simp add: rel_fun_def)

lemma (in ccpo) Sup_image_mono:
  assumes ccpo: "class.ccpo luba orda lessa"
  and mono: "monotone orda (≤) f"
  and chain: "Complete_Partial_Order.chain orda A"
  and "A  {}"
  shows "Sup (f ` A)  (f (luba A))"
proof(rule ccpo_Sup_least)
  from chain show "Complete_Partial_Order.chain (≤) (f ` A)"
    by(rule chain_imageI)(rule monotoneD[OF mono])
  fix x
  assume "x  f ` A"
  then obtain y where "x = f y" "y  A" by blast
  from y  A have "orda y (luba A)" by(rule ccpo.ccpo_Sup_upper[OF ccpo chain])
  hence "f y  f (luba A)" by(rule monotoneD[OF mono])
  thus "x  f (luba A)" using x = f y by simp
qed

lemma (in ccpo) admissible_le_mono:
  assumes "monotone (≤) (≤) f"
  shows "ccpo.admissible Sup (≤) (λx. x  f x)"
proof(rule ccpo.admissibleI)
  fix Y
  assume chain: "Complete_Partial_Order.chain (≤) Y"
    and Y: "Y  {}"
    and le [rule_format]: "xY. x  f x"
  have "Y  (f ` Y)" using chain
    by(rule ccpo_Sup_least)(rule order_trans[OF le]; blast intro!: ccpo_Sup_upper chain_imageI[OF chain] intro: monotoneD[OF assms])
  also have "  f (Y)"
    by(rule Sup_image_mono[OF _ assms chain Y, where lessa="(<)"]) unfold_locales
  finally show "Y  " .
qed

lemma (in ccpo) fixp_induct_strong2:
  assumes adm: "ccpo.admissible Sup (≤) P"
  and mono: "monotone (≤) (≤) f"
  and bot: "P ({})"
  and step: "x.  x  ccpo_class.fixp f; x  f x; P x   P (f x)"
  shows "P (ccpo_class.fixp f)"
proof(rule fixp_strong_induct[where P="λx. x  f x  P x", THEN conjunct2])
  show "ccpo.admissible Sup (≤) (λx. x  f x  P x)"
    using admissible_le_mono adm by(rule admissible_conj)(rule mono)
next
  show "{}  f ({})  P ({})"
    by(auto simp add: bot chain_empty intro: ccpo_Sup_least)
next
  fix x
  assume "x  ccpo_class.fixp f" "x  f x  P x"
  thus "f x  f (f x)  P (f x)"
    by(auto dest: monotoneD[OF mono] intro: step)
qed(rule mono)

context partial_function_definitions begin

lemma fixp_induct_strong2_uc:
  fixes F :: "'c  'c"
    and U :: "'c  'b  'a"
    and C :: "('b  'a)  'c"
    and P :: "('b  'a)  bool"
  assumes mono: "x. mono_body (λf. U (F (C f)) x)"
    and eq: "f  C (fixp_fun (λf. U (F (C f))))"
    and inverse: "f. U (C f) = f"
    and adm: "ccpo.admissible lub_fun le_fun P"
    and bot: "P (λ_. lub {})"
    and step: "f'.  le_fun (U f') (U f); le_fun (U f') (U (F f')); P (U f')   P (U (F f'))"
  shows "P (U f)"
unfolding eq inverse
apply (rule ccpo.fixp_induct_strong2[OF ccpo adm])
apply (insert mono, auto simp: monotone_def fun_ord_def bot fun_lub_def)[2]
apply (rule_tac f'5="C x" in step)
apply (simp_all add: inverse eq)
done

end

lemmas parallel_fixp_induct_2_4 = parallel_fixp_induct_uc[
  of _ _ _ _ "case_prod" _ "curry" "λf. case_prod (case_prod (case_prod f))" _ "λf. curry (curry (curry f))",
  where P="λf g. P (curry f) (curry (curry (curry g)))",
  unfolded case_prod_curry curry_case_prod curry_K,
  OF _ _ _ _ _ _ refl refl]
  for P
  
lemma (in ccpo) fixp_greatest:
  assumes f: "monotone (≤) (≤) f"
    and ge: "y. f y  y  x  y"
  shows "x  ccpo.fixp Sup (≤) f"
  by(rule ge)(simp add: fixp_unfold[OF f, symmetric])

lemma fixp_rolling:
  assumes "class.ccpo lub1 leq1 (mk_less leq1)"
    and "class.ccpo lub2 leq2 (mk_less leq2)"
    and f: "monotone leq1 leq2 f"
    and g: "monotone leq2 leq1 g"
  shows "ccpo.fixp lub1 leq1 (λx. g (f x)) = g (ccpo.fixp lub2 leq2 (λx. f (g x)))"
proof -
  interpret c1: ccpo lub1 leq1 "mk_less leq1" by fact
  interpret c2: ccpo lub2 leq2 "mk_less leq2" by fact
  show ?thesis
  proof(rule c1.order.antisym)
    have fg: "monotone leq2 leq2 (λx. f (g x))" using f g by(rule monotone2monotone) simp_all
    have gf: "monotone leq1 leq1 (λx. g (f x))" using g f by(rule monotone2monotone) simp_all
    show "leq1 (c1.fixp (λx. g (f x))) (g (c2.fixp (λx. f (g x))))" using gf
      by(rule c1.fixp_lowerbound)(subst (2) c2.fixp_unfold[OF fg], simp)
    show "leq1 (g (c2.fixp (λx. f (g x)))) (c1.fixp (λx. g (f x)))" using gf
    proof(rule c1.fixp_greatest)
      fix u
      assume u: "leq1 (g (f u)) u"
      have "leq1 (g (c2.fixp (λx. f (g x)))) (g (f u))"
        by(intro monotoneD[OF g] c2.fixp_lowerbound[OF fg] monotoneD[OF f u])
      then show "leq1 (g (c2.fixp (λx. f (g x)))) u" using u by(rule c1.order_trans)
    qed
  qed
qed

lemma fixp_lfp_parametric_eq:
  includes lifting_syntax
  assumes f: "x. lfp.mono_body (λf. F f x)"
  and g: "x. lfp.mono_body (λf. G f x)"
  and param: "((A ===> (=)) ===> A ===> (=)) F G"
  shows "(A ===> (=)) (lfp.fixp_fun F) (lfp.fixp_fun G)"
using f g
proof(rule parallel_fixp_induct_1_1[OF complete_lattice_partial_function_definitions complete_lattice_partial_function_definitions _ _ reflexive reflexive, where P="(A ===> (=))"])
  show "ccpo.admissible (prod_lub lfp.lub_fun lfp.lub_fun) (rel_prod lfp.le_fun lfp.le_fun) (λx. (A ===> (=)) (fst x) (snd x))"
    unfolding rel_fun_def by simp
  show "(A ===> (=)) (λ_. {}) (λ_. {})" by auto
  show "(A ===> (=)) (F f) (G g)" if "(A ===> (=)) f g" for f g
    using that by(rule rel_funD[OF param])
qed

lemma mono2mono_map_option[THEN option.mono2mono, simp, cont_intro]:
  shows monotone_map_option: "monotone option_ord option_ord (map_option f)"
by(rule monotoneI)(auto simp add: flat_ord_def)

lemma mcont2mcont_map_option[THEN option.mcont2mcont, simp, cont_intro]:
  shows mcont_map_option: "mcont (flat_lub None) option_ord (flat_lub None) option_ord (map_option f)"
by(rule mcont_finite_chains[OF _ _ flat_interpretation[THEN ccpo] flat_interpretation[THEN ccpo]]) simp_all

lemma mono2mono_set_option [THEN lfp.mono2mono]:
  shows monotone_set_option: "monotone option_ord (⊆) set_option"
by(auto intro!: monotoneI simp add: option_ord_Some1_iff)

lemma mcont2mcont_set_option [THEN lfp.mcont2mcont, cont_intro, simp]:
  shows mcont_set_option: "mcont (flat_lub None) option_ord Union (⊆) set_option"
by(rule mcont_finite_chains)(simp_all add: monotone_set_option ccpo option.partial_function_definitions_axioms)

lemma eadd_gfp_partial_function_mono [partial_function_mono]:
  " monotone (fun_ord (≥)) (≥) f; monotone (fun_ord (≥)) (≥) g 
   monotone (fun_ord (≥)) (≥) (λx. f x + g x :: enat)"
by(rule mono2mono_gfp_eadd)

lemma map_option_mono [partial_function_mono]:
  "mono_option B  mono_option (λf. map_option g (B f))"
unfolding map_conv_bind_option by(rule bind_mono) simp_all


subsection ‹Folding over finite sets›

lemma (in comp_fun_commute) fold_invariant_remove [consumes 1, case_names start step]:
  assumes fin: "finite A"
  and start: "I A s"
  and step: "x s A'.  x  A'; I A' s; A'  A   I (A' - {x}) (f x s)"
  shows "I {} (Finite_Set.fold f s A)"
proof -
  define A' where "A' == A"
  with fin start have "finite A'" "A'  A" "I A' s" by simp_all
  thus "I {} (Finite_Set.fold f s A')"
  proof(induction arbitrary: s)
    case empty thus ?case by simp
  next
    case (insert x A')
    let ?A' = "insert x A'"
    have "x  ?A'" "I ?A' s" "?A'  A" using insert by auto
    hence "I (?A' - {x}) (f x s)" by(rule step)
    with insert have "A'  A" "I A' (f x s)" by auto
    hence "I {} (Finite_Set.fold f (f x s) A')" by(rule insert.IH)
    thus ?case using insert by(simp add: fold_insert2 del: fold_insert)
  qed
qed

lemma (in comp_fun_commute) fold_invariant_insert [consumes 1, case_names start step]:
  assumes fin: "finite A"
  and start: "I {} s"
  and step: "x s A'.  I A' s; x  A'; x  A; A'  A   I (insert x A') (f x s)"
  shows "I A (Finite_Set.fold f s A)"
using fin start
proof(rule fold_invariant_remove[where I="λA'. I (A - A')" and A=A and s=s, simplified])
  fix x s A'
  assume *: "x  A'" "I (A - A') s" "A'  A"
  hence "x  A - A'" "x  A" "A - A'  A" by auto
  with I (A