Theory HOL.BNF_Def
section ‹Definition of Bounded Natural Functors›
theory BNF_Def
imports BNF_Cardinal_Arithmetic Fun_Def_Base
keywords
"print_bnfs" :: diag and
"bnf" :: thy_goal_defn
begin
lemma Collect_case_prodD: "x ∈ Collect (case_prod A) ⟹ A (fst x) (snd x)"
by auto
inductive
rel_sum :: "('a ⇒ 'c ⇒ bool) ⇒ ('b ⇒ 'd ⇒ bool) ⇒ 'a + 'b ⇒ 'c + 'd ⇒ bool" for R1 R2
where
"R1 a c ⟹ rel_sum R1 R2 (Inl a) (Inl c)"
| "R2 b d ⟹ rel_sum R1 R2 (Inr b) (Inr d)"
definition
rel_fun :: "('a ⇒ 'c ⇒ bool) ⇒ ('b ⇒ 'd ⇒ bool) ⇒ ('a ⇒ 'b) ⇒ ('c ⇒ 'd) ⇒ bool"
where
"rel_fun A B = (λf g. ∀x y. A x y ⟶ B (f x) (g y))"
lemma rel_funI [intro]:
assumes "⋀x y. A x y ⟹ B (f x) (g y)"
shows "rel_fun A B f g"
using assms by (simp add: rel_fun_def)
lemma rel_funD:
assumes "rel_fun A B f g" and "A x y"
shows "B (f x) (g y)"
using assms by (simp add: rel_fun_def)
lemma rel_fun_mono:
"⟦ rel_fun X A f g; ⋀x y. Y x y ⟶ X x y; ⋀x y. A x y ⟹ B x y ⟧ ⟹ rel_fun Y B f g"
by(simp add: rel_fun_def)
lemma rel_fun_mono' [mono]:
"⟦ ⋀x y. Y x y ⟶ X x y; ⋀x y. A x y ⟶ B x y ⟧ ⟹ rel_fun X A f g ⟶ rel_fun Y B f g"
by(simp add: rel_fun_def)
definition rel_set :: "('a ⇒ 'b ⇒ bool) ⇒ 'a set ⇒ 'b set ⇒ bool"
where "rel_set R = (λA B. (∀x∈A. ∃y∈B. R x y) ∧ (∀y∈B. ∃x∈A. R x y))"
lemma rel_setI:
assumes "⋀x. x ∈ A ⟹ ∃y∈B. R x y"
assumes "⋀y. y ∈ B ⟹ ∃x∈A. R x y"
shows "rel_set R A B"
using assms unfolding rel_set_def by simp
lemma predicate2_transferD:
"⟦rel_fun R1 (rel_fun R2 (=)) P Q; a ∈ A; b ∈ B; A ⊆ {(x, y). R1 x y}; B ⊆ {(x, y). R2 x y}⟧ ⟹
P (fst a) (fst b) ⟷ Q (snd a) (snd b)"
unfolding rel_fun_def by (blast dest!: Collect_case_prodD)
definition collect where
"collect F x = (⋃f ∈ F. f x)"
lemma fstI: "x = (y, z) ⟹ fst x = y"
by simp
lemma sndI: "x = (y, z) ⟹ snd x = z"
by simp
lemma bijI': "⟦⋀x y. (f x = f y) = (x = y); ⋀y. ∃x. y = f x⟧ ⟹ bij f"
unfolding bij_def inj_on_def by auto blast
definition "Gr A f = {(a, f a) | a. a ∈ A}"
definition "Grp A f = (λa b. b = f a ∧ a ∈ A)"
definition vimage2p where
"vimage2p f g R = (λx y. R (f x) (g y))"
lemma collect_comp: "collect F ∘ g = collect ((λf. f ∘ g) ` F)"
by (rule ext) (simp add: collect_def)
definition convol ("⟨(_,/ _)⟩") where
"⟨f, g⟩ ≡ λa. (f a, g a)"
lemma fst_convol: "fst ∘ ⟨f, g⟩ = f"
apply(rule ext)
unfolding convol_def by simp
lemma snd_convol: "snd ∘ ⟨f, g⟩ = g"
apply(rule ext)
unfolding convol_def by simp
lemma convol_mem_GrpI:
"x ∈ A ⟹ ⟨id, g⟩ x ∈ (Collect (case_prod (Grp A g)))"
unfolding convol_def Grp_def by auto
definition csquare where
"csquare A f1 f2 p1 p2 ⟷ (∀ a ∈ A. f1 (p1 a) = f2 (p2 a))"
lemma eq_alt: "(=) = Grp UNIV id"
unfolding Grp_def by auto
lemma leq_conversepI: "R = (=) ⟹ R ≤ R¯¯"
by auto
lemma leq_OOI: "R = (=) ⟹ R ≤ R OO R"
by auto
lemma OO_Grp_alt: "(Grp A f)¯¯ OO Grp A g = (λx y. ∃z. z ∈ A ∧ f z = x ∧ g z = y)"
unfolding Grp_def by auto
lemma Grp_UNIV_id: "f = id ⟹ (Grp UNIV f)¯¯ OO Grp UNIV f = Grp UNIV f"
unfolding Grp_def by auto
lemma Grp_UNIV_idI: "x = y ⟹ Grp UNIV id x y"
unfolding Grp_def by auto
lemma Grp_mono: "A ≤ B ⟹ Grp A f ≤ Grp B f"
unfolding Grp_def by auto
lemma GrpI: "⟦f x = y; x ∈ A⟧ ⟹ Grp A f x y"
unfolding Grp_def by auto
lemma GrpE: "Grp A f x y ⟹ (⟦f x = y; x ∈ A⟧ ⟹ R) ⟹ R"
unfolding Grp_def by auto
lemma Collect_case_prod_Grp_eqD: "z ∈ Collect (case_prod (Grp A f)) ⟹ (f ∘ fst) z = snd z"
unfolding Grp_def comp_def by auto
lemma Collect_case_prod_Grp_in: "z ∈ Collect (case_prod (Grp A f)) ⟹ fst z ∈ A"
unfolding Grp_def comp_def by auto
definition "pick_middlep P Q a c = (SOME b. P a b ∧ Q b c)"
lemma pick_middlep:
"(P OO Q) a c ⟹ P a (pick_middlep P Q a c) ∧ Q (pick_middlep P Q a c) c"
unfolding pick_middlep_def by (rule someI_ex) auto
definition fstOp where
"fstOp P Q ac = (fst ac, pick_middlep P Q (fst ac) (snd ac))"
definition sndOp where
"sndOp P Q ac = (pick_middlep P Q (fst ac) (snd ac), (snd ac))"
lemma fstOp_in: "ac ∈ Collect (case_prod (P OO Q)) ⟹ fstOp P Q ac ∈ Collect (case_prod P)"
unfolding fstOp_def mem_Collect_eq
by (subst (asm) surjective_pairing, unfold prod.case) (erule pick_middlep[THEN conjunct1])
lemma fst_fstOp: "fst bc = (fst ∘ fstOp P Q) bc"
unfolding comp_def fstOp_def by simp
lemma snd_sndOp: "snd bc = (snd ∘ sndOp P Q) bc"
unfolding comp_def sndOp_def by simp
lemma sndOp_in: "ac ∈ Collect (case_prod (P OO Q)) ⟹ sndOp P Q ac ∈ Collect (case_prod Q)"
unfolding sndOp_def mem_Collect_eq
by (subst (asm) surjective_pairing, unfold prod.case) (erule pick_middlep[THEN conjunct2])
lemma csquare_fstOp_sndOp:
"csquare (Collect (f (P OO Q))) snd fst (fstOp P Q) (sndOp P Q)"
unfolding csquare_def fstOp_def sndOp_def using pick_middlep by simp
lemma snd_fst_flip: "snd xy = (fst ∘ (%(x, y). (y, x))) xy"
by (simp split: prod.split)
lemma fst_snd_flip: "fst xy = (snd ∘ (%(x, y). (y, x))) xy"
by (simp split: prod.split)
lemma flip_pred: "A ⊆ Collect (case_prod (R ¯¯)) ⟹ (%(x, y). (y, x)) ` A ⊆ Collect (case_prod R)"
by auto
lemma predicate2_eqD: "A = B ⟹ A a b ⟷ B a b"
by simp
lemma case_sum_o_inj: "case_sum f g ∘ Inl = f" "case_sum f g ∘ Inr = g"
by auto
lemma map_sum_o_inj: "map_sum f g ∘ Inl = Inl ∘ f" "map_sum f g ∘ Inr = Inr ∘ g"
by auto
lemma card_order_csum_cone_cexp_def:
"card_order r ⟹ ( |A1| +c cone) ^c r = |Func UNIV (Inl ` A1 ∪ {Inr ()})|"
unfolding cexp_def cone_def Field_csum Field_card_of by (auto dest: Field_card_order)
lemma If_the_inv_into_in_Func:
"⟦inj_on g C; C ⊆ B ∪ {x}⟧ ⟹
(λi. if i ∈ g ` C then the_inv_into C g i else x) ∈ Func UNIV (B ∪ {x})"
unfolding Func_def by (auto dest: the_inv_into_into)
lemma If_the_inv_into_f_f:
"⟦i ∈ C; inj_on g C⟧ ⟹ ((λi. if i ∈ g ` C then the_inv_into C g i else x) ∘ g) i = id i"
unfolding Func_def by (auto elim: the_inv_into_f_f)
lemma the_inv_f_o_f_id: "inj f ⟹ (the_inv f ∘ f) z = id z"
by (simp add: the_inv_f_f)
lemma vimage2pI: "R (f x) (g y) ⟹ vimage2p f g R x y"
unfolding vimage2p_def .
lemma rel_fun_iff_leq_vimage2p: "(rel_fun R S) f g = (R ≤ vimage2p f g S)"
unfolding rel_fun_def vimage2p_def by auto
lemma convol_image_vimage2p: "⟨f ∘ fst, g ∘ snd⟩ ` Collect (case_prod (vimage2p f g R)) ⊆ Collect (case_prod R)"
unfolding vimage2p_def convol_def by auto
lemma vimage2p_Grp: "vimage2p f g P = Grp UNIV f OO P OO (Grp UNIV g)¯¯"
unfolding vimage2p_def Grp_def by auto
lemma subst_Pair: "P x y ⟹ a = (x, y) ⟹ P (fst a) (snd a)"
by simp
lemma comp_apply_eq: "f (g x) = h (k x) ⟹ (f ∘ g) x = (h ∘ k) x"
unfolding comp_apply by assumption
lemma refl_ge_eq: "(⋀x. R x x) ⟹ (=) ≤ R"
by auto
lemma ge_eq_refl: "(=) ≤ R ⟹ R x x"
by auto
lemma reflp_eq: "reflp R = ((=) ≤ R)"
by (auto simp: reflp_def fun_eq_iff)
lemma transp_relcompp: "transp r ⟷ r OO r ≤ r"
by (auto simp: transp_def)
lemma symp_conversep: "symp R = (R¯¯ ≤ R)"
by (auto simp: symp_def fun_eq_iff)
lemma diag_imp_eq_le: "(⋀x. x ∈ A ⟹ R x x) ⟹ ∀x y. x ∈ A ⟶ y ∈ A ⟶ x = y ⟶ R x y"
by blast
definition eq_onp :: "('a ⇒ bool) ⇒ 'a ⇒ 'a ⇒ bool"
where "eq_onp R = (λx y. R x ∧ x = y)"
lemma eq_onp_Grp: "eq_onp P = BNF_Def.Grp (Collect P) id"
unfolding eq_onp_def Grp_def by auto
lemma eq_onp_to_eq: "eq_onp P x y ⟹ x = y"
by (simp add: eq_onp_def)
lemma eq_onp_top_eq_eq: "eq_onp top = (=)"
by (simp add: eq_onp_def)
lemma eq_onp_same_args: "eq_onp P x x = P x"
by (auto simp add: eq_onp_def)
lemma eq_onp_eqD: "eq_onp P = Q ⟹ P x = Q x x"
unfolding eq_onp_def by blast
lemma Ball_Collect: "Ball A P = (A ⊆ (Collect P))"
by auto
lemma eq_onp_mono0: "∀x∈A. P x ⟶ Q x ⟹ ∀x∈A. ∀y∈A. eq_onp P x y ⟶ eq_onp Q x y"
unfolding eq_onp_def by auto
lemma eq_onp_True: "eq_onp (λ_. True) = (=)"
unfolding eq_onp_def by simp
lemma Ball_image_comp: "Ball (f ` A) g = Ball A (g ∘ f)"
by auto
lemma rel_fun_Collect_case_prodD:
"rel_fun A B f g ⟹ X ⊆ Collect (case_prod A) ⟹ x ∈ X ⟹ B ((f ∘ fst) x) ((g ∘ snd) x)"
unfolding rel_fun_def by auto
lemma eq_onp_mono_iff: "eq_onp P ≤ eq_onp Q ⟷ P ≤ Q"
unfolding eq_onp_def by auto
ML_file ‹Tools/BNF/bnf_util.ML›
ML_file ‹Tools/BNF/bnf_tactics.ML›
ML_file ‹Tools/BNF/bnf_def_tactics.ML›
ML_file ‹Tools/BNF/bnf_def.ML›
end