# HG changeset patch # User traytel # Date 1389371092 -3600 # Node ID 7e0573a490ee83f48cf2b474488726403741fa49 # Parent d7593bfccf257a73acb3ea4c5b1b448656fcc7cc basic ordinal arithmetic and cardinals library extension (not relevant for BNFs) diff -r d7593bfccf25 -r 7e0573a490ee src/HOL/Cardinals/Cardinal_Arithmetic.thy --- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy Fri Jan 10 16:18:18 2014 +0100 +++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Fri Jan 10 17:24:52 2014 +0100 @@ -11,7 +11,6 @@ imports Cardinal_Arithmetic_FP Cardinal_Order_Relation begin - subsection {* Binary sum *} lemma csum_Cnotzero2: @@ -33,12 +32,40 @@ lemma cone_ordLeq_ctwo: "cone \o ctwo" unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto +lemma csum_czero1: "Card_order r \ r +c czero =o r" + unfolding czero_def csum_def Field_card_of + by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty1] card_of_Field_ordIso]) + +lemma csum_czero2: "Card_order r \ czero +c r =o r" + unfolding czero_def csum_def Field_card_of + by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty2] card_of_Field_ordIso]) + subsection {* Product *} lemma Times_cprod: "|A \ B| =o |A| *c |B|" by (simp only: cprod_def Field_card_of card_of_refl) +lemma card_of_Times_singleton: +fixes A :: "'a set" +shows "|A \ {x}| =o |A|" +proof - + def f \ "\(a::'a,b::'b). (a)" + have "A \ f ` (A <*> {x})" unfolding f_def by (auto simp: image_iff) + hence "bij_betw f (A <*> {x}) A" unfolding bij_betw_def inj_on_def f_def by fastforce + thus ?thesis using card_of_ordIso by blast +qed + +lemma cprod_assoc: "(r *c s) *c t =o r *c s *c t" + unfolding cprod_def Field_card_of by (rule card_of_Times_assoc) + +lemma cprod_czero: "r *c czero =o czero" + unfolding cprod_def czero_def Field_card_of by (simp add: card_of_empty_ordIso) + +lemma cprod_cone: "Card_order r \ r *c cone =o r" + unfolding cprod_def cone_def Field_card_of + by (drule card_of_Field_ordIso) (erule ordIso_transitive[OF card_of_Times_singleton]) + lemma cprod_cong2: "p2 =o r2 \ q *c p2 =o q *c r2" by (simp only: cprod_def ordIso_Times_cong2) @@ -48,6 +75,9 @@ lemma cprod_infinite: "Cinfinite r \ r *c r =o r" using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast +lemma cprod_mono: "\p1 \o r1; p2 \o r2\ \ p1 *c p2 \o r1 *c r2" + by (rule ordLeq_transitive[OF cprod_mono1 cprod_mono2]) + subsection {* Exponentiation *} @@ -112,7 +142,7 @@ qed lemma cone_ordLeq_cexp: "cone \o r1 \ cone \o r1 ^c r2" -by (simp add: cexp_def cone_def Func_non_emp card_of_singl_ordLeq cone_ordLeq_iff_Field) +by (simp add: cexp_def cone_def Func_non_emp cone_ordLeq_iff_Field) lemma Card_order_czero: "Card_order czero" by (simp only: card_of_Card_order czero_def) @@ -177,7 +207,7 @@ shows "card_order (r1 ^c r2)" proof - have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto - thus ?thesis unfolding cexp_def Func_def by (simp add: card_of_card_order_on) + thus ?thesis unfolding cexp_def Func_def by simp qed lemma Cinfinite_ordLess_cexp: @@ -195,6 +225,58 @@ shows "r \o r ^c r" by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]]) +lemma czero_cexp: "Cnotzero r \ czero ^c r =o czero" + by (drule Cnotzero_imp_not_empty) (simp add: cexp_def czero_def card_of_empty_ordIso) + +lemma Func_singleton: +fixes x :: 'b and A :: "'a set" +shows "|Func A {x}| =o |{x}|" +proof (rule ordIso_symmetric) + def f \ "\y a. if y = x \ a \ A then x else undefined" + have "Func A {x} \ f ` {x}" unfolding f_def Func_def by (force simp: fun_eq_iff) + hence "bij_betw f {x} (Func A {x})" unfolding bij_betw_def inj_on_def f_def Func_def + by (auto split: split_if_asm) + thus "|{x}| =o |Func A {x}|" using card_of_ordIso by blast +qed + +lemma cone_cexp: "cone ^c r =o cone" + unfolding cexp_def cone_def Field_card_of by (rule Func_singleton) + +lemma card_of_Func_squared: + fixes A :: "'a set" + shows "|Func (UNIV :: bool set) A| =o |A \ A|" +proof (rule ordIso_symmetric) + def f \ "\(x::'a,y) b. if A = {} then undefined else if b then x else y" + have "Func (UNIV :: bool set) A \ f ` (A \ A)" unfolding f_def Func_def + by (auto simp: image_iff fun_eq_iff split: option.splits split_if_asm) blast + hence "bij_betw f (A \ A) (Func (UNIV :: bool set) A)" + unfolding bij_betw_def inj_on_def f_def Func_def by (auto simp: fun_eq_iff) + thus "|A \ A| =o |Func (UNIV :: bool set) A|" using card_of_ordIso by blast +qed + +lemma cexp_ctwo: "r ^c ctwo =o r *c r" + unfolding cexp_def ctwo_def cprod_def Field_card_of by (rule card_of_Func_squared) + +lemma card_of_Func_Plus: + fixes A :: "'a set" and B :: "'b set" and C :: "'c set" + shows "|Func (A <+> B) C| =o |Func A C \ Func B C|" +proof (rule ordIso_symmetric) + def f \ "\(g :: 'a => 'c, h::'b \ 'c) ab. case ab of Inl a \ g a | Inr b \ h b" + def f' \ "\(f :: ('a + 'b) \ 'c). (\a. f (Inl a), \b. f (Inr b))" + have "f ` (Func A C \ Func B C) \ Func (A <+> B) C" + unfolding Func_def f_def by (force split: sum.splits) + moreover have "f' ` Func (A <+> B) C \ Func A C \ Func B C" unfolding Func_def f'_def by force + moreover have "\a \ Func A C \ Func B C. f' (f a) = a" unfolding f'_def f_def Func_def by auto + moreover have "\a' \ Func (A <+> B) C. f (f' a') = a'" unfolding f'_def f_def Func_def + by (auto split: sum.splits) + ultimately have "bij_betw f (Func A C \ Func B C) (Func (A <+> B) C)" + by (intro bij_betw_byWitness[of _ f' f]) + thus "|Func A C \ Func B C| =o |Func (A <+> B) C|" using card_of_ordIso by blast +qed + +lemma cexp_csum: "r ^c (s +c t) =o r ^c s *c r ^c t" + unfolding cexp_def cprod_def csum_def Field_card_of by (rule card_of_Func_Plus) + subsection {* Powerset *} @@ -218,4 +300,105 @@ lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r" unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func) +subsection {* Inverse image *} + +lemma vimage_ordLeq: +assumes "|A| \o k" and "\ a \ A. |vimage f {a}| \o k" and "Cinfinite k" +shows "|vimage f A| \o k" +proof- + have "vimage f A = (\ a \ A. vimage f {a})" by auto + also have "|\ a \ A. vimage f {a}| \o k" + using UNION_Cinfinite_bound[OF assms] . + finally show ?thesis . +qed + +subsection {* Maximum *} + +definition cmax where + "cmax r s = + (if cinfinite r \ cinfinite s then czero +c r +c s + else natLeq_on (max (card (Field r)) (card (Field s))) +c czero)" + +lemma cmax_com: "cmax r s =o cmax s r" + unfolding cmax_def + by (auto simp: max.commute intro: csum_cong2[OF csum_com] csum_cong2[OF czero_ordIso]) + +lemma cmax1: + assumes "Card_order r" "Card_order s" "s \o r" + shows "cmax r s =o r" +unfolding cmax_def proof (split if_splits, intro conjI impI) + assume "cinfinite r \ cinfinite s" + hence Cinf: "Cinfinite r" using assms(1,3) by (metis cinfinite_mono) + have "czero +c r +c s =o r +c s" by (rule csum_czero2[OF Card_order_csum]) + also have "r +c s =o r" by (rule csum_absorb1[OF Cinf assms(3)]) + finally show "czero +c r +c s =o r" . +next + assume "\ (cinfinite r \ cinfinite s)" + hence fin: "finite (Field r)" and "finite (Field s)" unfolding cinfinite_def by simp_all + moreover + { from assms(2) have "|Field s| =o s" by (rule card_of_Field_ordIso) + also from assms(3) have "s \o r" . + also from assms(1) have "r =o |Field r|" by (rule ordIso_symmetric[OF card_of_Field_ordIso]) + finally have "|Field s| \o |Field r|" . + } + ultimately have "card (Field s) \ card (Field r)" by (subst sym[OF finite_card_of_iff_card2]) + hence "max (card (Field r)) (card (Field s)) = card (Field r)" by (rule max_absorb1) + hence "natLeq_on (max (card (Field r)) (card (Field s))) +c czero = + natLeq_on (card (Field r)) +c czero" by simp + also have "\ =o natLeq_on (card (Field r))" by (rule csum_czero1[OF natLeq_on_Card_order]) + also have "natLeq_on (card (Field r)) =o |Field r|" + by (rule ordIso_symmetric[OF finite_imp_card_of_natLeq_on[OF fin]]) + also from assms(1) have "|Field r| =o r" by (rule card_of_Field_ordIso) + finally show "natLeq_on (max (card (Field r)) (card (Field s))) +c czero =o r" . +qed + +lemma cmax2: + assumes "Card_order r" "Card_order s" "r \o s" + shows "cmax r s =o s" + by (metis assms cmax1 cmax_com ordIso_transitive) + +lemma csum_absorb2: "Cinfinite r2 \ r1 \o r2 \ r1 +c r2 =o r2" + by (metis csum_absorb2') + +lemma cprod_infinite2': "\Cnotzero r1; Cinfinite r2; r1 \o r2\ \ r1 *c r2 =o r2" + unfolding ordIso_iff_ordLeq + by (intro conjI cprod_cinfinite_bound ordLeq_cprod2 ordLeq_refl) + (auto dest!: ordIso_imp_ordLeq not_ordLeq_ordLess simp: czero_def Card_order_empty) + +context + fixes r s + assumes r: "Cinfinite r" + and s: "Cinfinite s" +begin + +lemma cmax_csum: "cmax r s =o r +c s" +proof (cases "r \o s") + case True + hence "cmax r s =o s" by (metis cmax2 r s) + also have "s =o r +c s" by (metis True csum_absorb2 ordIso_symmetric s) + finally show ?thesis . +next + case False + hence "s \o r" by (metis ordLeq_total r s card_order_on_def) + hence "cmax r s =o r" by (metis cmax1 r s) + also have "r =o r +c s" by (metis `s \o r` csum_absorb1 ordIso_symmetric r) + finally show ?thesis . +qed + +lemma cmax_cprod: "cmax r s =o r *c s" +proof (cases "r \o s") + case True + hence "cmax r s =o s" by (metis cmax2 r s) + also have "s =o r *c s" by (metis Cinfinite_Cnotzero True cprod_infinite2' ordIso_symmetric r s) + finally show ?thesis . +next + case False + hence "s \o r" by (metis ordLeq_total r s card_order_on_def) + hence "cmax r s =o r" by (metis cmax1 r s) + also have "r =o r *c s" by (metis Cinfinite_Cnotzero `s \o r` cprod_infinite1' ordIso_symmetric r s) + finally show ?thesis . +qed + end + +end diff -r d7593bfccf25 -r 7e0573a490ee src/HOL/Cardinals/Cardinal_Order_Relation.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Fri Jan 10 16:18:18 2014 +0100 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Fri Jan 10 17:24:52 2014 +0100 @@ -67,8 +67,6 @@ cardSuc_finite[simp] card_of_Plus_ordLeq_infinite_Field[simp] curr_in[intro, simp] - Func_empty[simp] - Func_is_emp[simp] subsection {* Cardinal of a set *} @@ -418,8 +416,7 @@ proof- have "\i \ I. |A(f i)| =o |B(f i)|" using ISO BIJ unfolding bij_betw_def by blast - hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|" - using card_of_Sigma_cong1 by metis + hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|" by (rule card_of_Sigma_cong1) moreover have "|SIGMA i : I. B (f i)| =o |SIGMA j : J. B j|" using BIJ card_of_Sigma_cong2 by blast ultimately show ?thesis using ordIso_transitive by blast @@ -1173,9 +1170,8 @@ lemma natLeq_on_ordLeq_less: "((natLeq_on m) o natLeq_on n" @@ -1298,7 +1294,7 @@ hence 4: "m \ n" using 2 by force (* *) have "bij_betw f (Field r') (f ` (Field r'))" - using 1 WELL embed_inj_on unfolding bij_betw_def by force + using WELL embed_inj_on[OF _ 1] unfolding bij_betw_def by blast moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force ultimately have 5: "finite (Field r') \ card(Field r') = card (f ` (Field r'))" using bij_betw_same_card bij_betw_finite by metis @@ -1547,8 +1543,6 @@ ultimately show ?thesis by (metis ordIso_ordLeq_trans) qed -lemma Func_emp2[simp]: "A \ {} \ Func A {} = {}" by auto - lemma empty_in_Func[simp]: "B \ {} \ (\x. undefined) \ Func {} B" unfolding Func_def by auto @@ -1583,4 +1577,333 @@ shows "\finite (Func A B)" using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite) +section {* Infinite cardinals are limit ordinals *} + +lemma card_order_infinite_isLimOrd: +assumes c: "Card_order r" and i: "\finite (Field r)" +shows "isLimOrd r" +proof- + have 0: "wo_rel r" and 00: "Well_order r" + using c unfolding card_order_on_def wo_rel_def by auto + hence rr: "Refl r" by (metis wo_rel.REFL) + show ?thesis unfolding wo_rel.isLimOrd_def[OF 0] wo_rel.isSuccOrd_def[OF 0] proof safe + fix j assume j: "j \ Field r" and jm: "\i\Field r. (i, j) \ r" + def p \ "Restr r (Field r - {j})" + have fp: "Field p = Field r - {j}" + unfolding p_def apply(rule Refl_Field_Restr2[OF rr]) by auto + have of: "ofilter r (Field p)" unfolding wo_rel.ofilter_def[OF 0] proof safe + fix a x assume a: "a \ Field p" and "x \ under r a" + hence x: "(x,a) \ r" "x \ Field r" unfolding rel.under_def Field_def by auto + moreover have a: "a \ j" "a \ Field r" "(a,j) \ r" using a jm unfolding fp by auto + ultimately have "x \ j" using j jm by (metis 0 wo_rel.max2_def wo_rel.max2_equals1) + thus "x \ Field p" using x unfolding fp by auto + qed(unfold p_def Field_def, auto) + have "p Chains r" and "i \ Field r" and "\j. j \ C \ (j,i) \ r \ (i,j) \ r" +shows "insert i C \ Chains r" +using assms unfolding Chains_def by (auto dest: refl_onD) + +lemma Collect_insert: "{R j |j. j \ insert j1 J} = insert (R j1) {R j |j. j \ J}" +by auto + +lemma Field_init_seg_of[simp]: +"Field init_seg_of = UNIV" +unfolding Field_def init_seg_of_def by auto + +lemma refl_init_seg_of[intro, simp]: "refl init_seg_of" +unfolding refl_on_def Field_def by auto + +lemma regular_all_ex: +assumes r: "Card_order r" "regular r" +and As: "\ i j b. b \ B \ (i,j) \ r \ P i b \ P j b" +and Bsub: "\ b \ B. \ i \ Field r. P i b" +and cardB: "|B| i \ Field r. \ b \ B. P i b" +proof- + let ?As = "\i. {b \ B. P i b}" + have "EX i : Field r. B \ ?As i" + apply(rule regular_UNION) using assms unfolding relChain_def by auto + thus ?thesis by auto +qed + +lemma relChain_stabilize: +assumes rc: "relChain r As" and AsB: "(\ i \ Field r. As i) \ B" and Br: "|B| finite (Field r)" and cr: "Card_order r" +shows "\ i \ Field r. As (succ r i) = As i" +proof(rule ccontr, auto) + have 0: "wo_rel r" and 00: "Well_order r" + unfolding wo_rel_def by (metis card_order_on_well_order_on cr)+ + have L: "isLimOrd r" using ir cr by (metis card_order_infinite_isLimOrd) + have AsBs: "(\ i \ Field r. As (succ r i)) \ B" + using AsB L apply safe by (metis "0" UN_I set_mp wo_rel.isLimOrd_succ) + assume As_s: "\i\Field r. As (succ r i) \ As i" + have 1: "\i j. (i,j) \ r \ i \ j \ As i \ As j" + proof safe + fix i j assume 1: "(i, j) \ r" "i \ j" and Asij: "As i = As j" + hence rij: "(succ r i, j) \ r" by (metis "0" wo_rel.succ_smallest) + hence "As (succ r i) \ As j" using rc unfolding relChain_def by auto + moreover + {have "(i,succ r i) \ r" apply(rule wo_rel.succ_in[OF 0]) + using 1 unfolding rel.aboveS_def by auto + hence "As i \ As (succ r i)" using As_s rc rij unfolding relChain_def Field_def by auto + } + ultimately show False unfolding Asij by auto + qed (insert rc, unfold relChain_def, auto) + hence "\ i \ Field r. \ a. a \ As (succ r i) - As i" + using wo_rel.succ_in[OF 0] AsB + by(metis 0 card_order_infinite_isLimOrd cr ir psubset_imp_ex_mem + wo_rel.isLimOrd_aboveS wo_rel.succ_diff) + then obtain f where f: "\ i \ Field r. f i \ As (succ r i) - As i" by metis + have "inj_on f (Field r)" unfolding inj_on_def proof safe + fix i j assume ij: "i \ Field r" "j \ Field r" and fij: "f i = f j" + show "i = j" + proof(cases rule: wo_rel.cases_Total3[OF 0], safe) + assume "(i, j) \ r" and ijd: "i \ j" + hence rij: "(succ r i, j) \ r" by (metis "0" wo_rel.succ_smallest) + hence "As (succ r i) \ As j" using rc unfolding relChain_def by auto + thus "i = j" using ij ijd fij f by auto + next + assume "(j, i) \ r" and ijd: "i \ j" + hence rij: "(succ r j, i) \ r" by (metis "0" wo_rel.succ_smallest) + hence "As (succ r j) \ As i" using rc unfolding relChain_def by auto + thus "j = i" using ij ijd fij f by fastforce + qed(insert ij, auto) + qed + moreover have "f ` (Field r) \ B" using f AsBs by auto + moreover have "|B| bool" +where +"stable r \ \(A::'a set) (F :: 'a \ 'a set). + |A| (\a \ A. |F a| |SIGMA a : A. F a| finite (Field r)" and reg: "regular r" +shows "stable r" +unfolding stable_def proof safe + fix A :: "'a set" and F :: "'a \ 'a set" assume A: "|A| a\A. |F a| o |Sigma A F|" + hence "|Field r| \o |Sigma A F|" using card_of_Field_ordIso[OF cr] + by (metis Field_card_of card_of_cong ordLeq_iff_ordLess_or_ordIso ordLeq_ordLess_trans) + moreover have Fi: "Field r \ {}" using ir by auto + ultimately obtain f where f: "f ` Sigma A F = Field r" using card_of_ordLeq2 by metis + have r: "wo_rel r" using cr unfolding card_order_on_def wo_rel_def by auto + {fix a assume a: "a \ A" + def L \ "{(a,u) | u. u \ F a}" + have fL: "f ` L \ Field r" using f a unfolding L_def by auto + have "|L| =o |F a|" unfolding L_def card_of_ordIso[symmetric] + apply(rule exI[of _ snd]) unfolding bij_betw_def inj_on_def by (auto simp: image_def) + hence "|L| cofinal (f ` L) r" using reg fL unfolding regular_def by (metis not_ordLess_ordIso) + then obtain k where k: "k \ Field r" and "\ l \ L. \ (f l \ k \ (k, f l) \ r)" + unfolding cofinal_def image_def by auto + hence "\ k \ Field r. \ l \ L. (f l, k) \ r" using r by (metis fL image_subset_iff wo_rel.in_notinI) + hence "\ k \ Field r. \ u \ F a. (f (a,u), k) \ r" unfolding L_def by auto + } + then obtain gg where gg: "\ a \ A. \ u \ F a. (f (a,u), gg a) \ r" by metis + obtain j0 where j0: "j0 \ Field r" using Fi by auto + def g \ "\ a. if F a \ {} then gg a else j0" + have g: "\ a \ A. \ u \ F a. (f (a,u),g a) \ r" using gg unfolding g_def by auto + hence 1: "Field r \ (\ a \ A. under r (g a))" + using f[symmetric] unfolding rel.under_def image_def by auto + have gA: "g ` A \ Field r" using gg j0 unfolding Field_def g_def by auto + moreover have "cofinal (g ` A) r" unfolding cofinal_def proof safe + fix i assume "i \ Field r" + then obtain j where ij: "(i,j) \ r" "i \ j" using cr ir by (metis infinite_Card_order_limit) + hence "j \ Field r" by (metis card_order_on_def cr rel.well_order_on_domain) + then obtain a where a: "a \ A" and j: "(j, g a) \ r" using 1 unfolding rel.under_def by auto + hence "(i, g a) \ r" using ij wo_rel.TRANS[OF r] unfolding trans_def by blast + moreover have "i \ g a" + using ij j r unfolding wo_rel_def unfolding well_order_on_def linear_order_on_def + partial_order_on_def antisym_def by auto + ultimately show "\j\g ` A. i \ j \ (i, j) \ r" using a by auto + qed + ultimately have "|g ` A| =o r" using reg unfolding regular_def by auto + moreover have "|g ` A| \o |A|" by (metis card_of_image) + ultimately have False using A by (metis not_ordLess_ordIso ordLeq_ordLess_trans) + } + thus "|Sigma A F| finite (Field r)" and st: "stable r" +shows "regular r" +unfolding regular_def proof safe + fix K assume K: "K \ Field r" and cof: "cofinal K r" + have "|K| \o r" using K by (metis card_of_Field_ordIso card_of_mono1 cr ordLeq_ordIso_trans) + moreover + {assume Kr: "|K| a \ Field r. f a \ K \ a \ f a \ (a, f a) \ r" + using cof unfolding cofinal_def by metis + hence "Field r \ (\ a \ K. underS r a)" unfolding rel.underS_def by auto + hence "r \o |\ a \ K. underS r a|" using cr + by (metis Card_order_iff_ordLeq_card_of card_of_mono1 ordLeq_transitive) + also have "|\ a \ K. underS r a| \o |SIGMA a: K. underS r a|" by (rule card_of_UNION_Sigma) + also + {have "\ a \ K. |underS r a| a. a \ A \ |F a| |A'| A" + hence "\B'. B' \ Field r \ |F a| =o |B'| \ |B'| a \ A. F' a \ Field r \ |F a| =o |F' a| \ |F' a| a B'. B' \ Field r \ |F a| =o |B'| \ |B'| a \ A. F' a \ Field r \ |F' a| a \ A. |F' a| =o |F a|" using temp ordIso_symmetric by auto + (* *) + have "\a' \ A'. F'(G a') \ Field r \ |F'(G a')| 'a set" + assume "|A| a\A. |F a| (\a \ A. finite(F a))" + by (auto simp add: finite_iff_ordLess_natLeq) + hence "finite(Sigma A F)" by (simp only: finite_SigmaI[of A F]) + thus "|Sigma A F | finite (Field r)" +shows "stable(cardSuc r)" +using infinite_cardSuc_regular regular_stable +by (metis CARD INF cardSuc_Card_order cardSuc_finite) + +lemma stable_UNION: +assumes ST: "stable r" and A_LESS: "|A| a. a \ A \ |F a| a \ A. F a| a \ A. F a| \o |SIGMA a : A. F a|" + using card_of_UNION_Sigma by blast + thus ?thesis using assms stable_elim ordLeq_ordLess_trans by blast +qed + +lemma stable_ordIso1: +assumes ST: "stable r" and ISO: "r' =o r" +shows "stable r'" +proof(unfold stable_def, auto) + fix A::"'b set" and F::"'b \ 'b set" + assume "|A| a \ A. |F a| (\a \ A. |F a| finite sets of stable cardinality exist. *} + +lemma infinite_stable_exists: +assumes CARD: "\r \ R. Card_order (r::'a rel)" +shows "\(A :: (nat + 'a set)set). + \finite A \ stable |A| \ (\r \ R. r (A :: (nat + 'a set)set). + \finite A \ stable |A| \ |UNIV::'a set| finite(?B <+> {})" using finite_Plus_iff by blast + moreover + have "stable |?B|" using stable_natLeq card_of_nat stable_ordIso1 by blast + hence "stable |?B <+> {}|" using stable_ordIso card_of_Plus_empty1 by blast + moreover + have "\finite(Field |?B| ) \ finite(Field |UNIV::'a set| )" using Case1 by simp + hence "|UNIV::'a set| {}|" using card_of_Plus_empty1 ordLess_ordIso_trans by blast + ultimately show ?thesis by blast + next + assume Case1: "\finite (UNIV::'a set)" + hence *: "\finite(Field |UNIV::'a set| )" by simp + let ?B = "Field(cardSuc |UNIV::'a set| )" + have 0: "|?B| =o |{} <+> ?B|" using card_of_Plus_empty2 by blast + have 1: "\finite ?B" using Case1 card_of_cardSuc_finite by blast + hence 2: "\finite({} <+> ?B)" using 0 card_of_ordIso_finite by blast + have "|?B| =o cardSuc |UNIV::'a set|" + using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast + moreover have "stable(cardSuc |UNIV::'a set| )" + using stable_cardSuc * card_of_Card_order by blast + ultimately have "stable |?B|" using stable_ordIso by blast + hence 3: "stable |{} <+> ?B|" using stable_ordIso 0 by blast + have "|UNIV::'a set| ?B|" using 0 ordLess_ordIso_trans by blast + thus ?thesis using 2 3 by blast + qed + thus ?thesis using CARD card_of_UNIV2 ordLeq_ordLess_trans by blast +qed + +corollary infinite_regular_exists: +assumes CARD: "\r \ R. Card_order (r::'a rel)" +shows "\(A :: (nat + 'a set)set). + \finite A \ regular |A| \ (\r \ R. r a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" - -lemma Func_empty: -"Func {} B = {\x. undefined}" -unfolding Func_def by auto - -lemma Func_elim: -assumes "g \ Func A B" and "a \ A" -shows "\ b. b \ B \ g a = b" -using assms unfolding Func_def by (cases "g a = undefined") auto - -definition curr where -"curr A f \ \ a. if a \ A then \b. f (a,b) else undefined" - -lemma curr_in: -assumes f: "f \ Func (A <*> B) C" -shows "curr A f \ Func A (Func B C)" -using assms unfolding curr_def Func_def by auto - -lemma curr_inj: -assumes "f1 \ Func (A <*> B) C" and "f2 \ Func (A <*> B) C" -shows "curr A f1 = curr A f2 \ f1 = f2" -proof safe - assume c: "curr A f1 = curr A f2" - show "f1 = f2" - proof (rule ext, clarify) - fix a b show "f1 (a, b) = f2 (a, b)" - proof (cases "(a,b) \ A <*> B") - case False - thus ?thesis using assms unfolding Func_def by auto - next - case True hence a: "a \ A" and b: "b \ B" by auto - thus ?thesis - using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp - qed - qed -qed - -lemma curr_surj: -assumes "g \ Func A (Func B C)" -shows "\ f \ Func (A <*> B) C. curr A f = g" -proof - let ?f = "\ ab. if fst ab \ A \ snd ab \ B then g (fst ab) (snd ab) else undefined" - show "curr A ?f = g" - proof (rule ext) - fix a show "curr A ?f a = g a" - proof (cases "a \ A") - case False - hence "g a = undefined" using assms unfolding Func_def by auto - thus ?thesis unfolding curr_def using False by simp - next - case True - obtain g1 where "g1 \ Func B C" and "g a = g1" - using assms using Func_elim[OF assms True] by blast - thus ?thesis using True unfolding Func_def curr_def by auto - qed - qed - show "?f \ Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto -qed - -lemma bij_betw_curr: -"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))" -unfolding bij_betw_def inj_on_def image_def -apply (intro impI conjI ballI) -apply (erule curr_inj[THEN iffD1], assumption+) -apply auto -apply (erule curr_in) -using curr_surj by blast - lemma card_of_Func_Times: "|Func (A <*> B) C| =o |Func A (Func B C)|" unfolding card_of_ordIso[symmetric] using bij_betw_curr by blast -definition Func_map where -"Func_map B2 f1 f2 g b2 \ if b2 \ B2 then f1 (g (f2 b2)) else undefined" - -lemma Func_map: -assumes g: "g \ Func A2 A1" and f1: "f1 ` A1 \ B1" and f2: "f2 ` B2 \ A2" -shows "Func_map B2 f1 f2 g \ Func B2 B1" -using assms unfolding Func_def Func_map_def mem_Collect_eq by auto - -lemma Func_non_emp: -assumes "B \ {}" -shows "Func A B \ {}" -proof- - obtain b where b: "b \ B" using assms by auto - hence "(\ a. if a \ A then b else undefined) \ Func A B" unfolding Func_def by auto - thus ?thesis by blast -qed - -lemma Func_is_emp: -"Func A B = {} \ A \ {} \ B = {}" (is "?L \ ?R") -proof - assume L: ?L - moreover {assume "A = {}" hence False using L Func_empty by auto} - moreover {assume "B \ {}" hence False using L Func_non_emp by metis} - ultimately show ?R by blast -next - assume R: ?R - moreover - {fix f assume "f \ Func A B" - moreover obtain a where "a \ A" using R by blast - ultimately obtain b where "b \ B" unfolding Func_def by blast - with R have False by blast - } - thus ?L by blast -qed - -lemma Func_map_surj: -assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \ A2" -and B2A2: "B2 = {} \ A2 = {}" -shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1" -proof(cases "B2 = {}") - case True - thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def) -next - case False note B2 = False - show ?thesis - proof safe - fix h assume h: "h \ Func B2 B1" - def j1 \ "inv_into A1 f1" - have "\ a2 \ f2 ` B2. \ b2. b2 \ B2 \ f2 b2 = a2" by blast - then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by metis - {fix b2 assume b2: "b2 \ B2" - hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto - moreover have "k (f2 b2) \ B2" using b2 A2(2) k by auto - ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast - } note kk = this - obtain b22 where b22: "b22 \ B2" using B2 by auto - def j2 \ "\ a2. if a2 \ f2 ` B2 then k a2 else b22" - have j2A2: "j2 ` A2 \ B2" unfolding j2_def using k b22 by auto - have j2: "\ b2. b2 \ B2 \ j2 (f2 b2) = b2" - using kk unfolding j2_def by auto - def g \ "Func_map A2 j1 j2 h" - have "Func_map B2 f1 f2 g = h" - proof (rule ext) - fix b2 show "Func_map B2 f1 f2 g b2 = h b2" - proof(cases "b2 \ B2") - case True - show ?thesis - proof (cases "h b2 = undefined") - case True - hence b1: "h b2 \ f1 ` A1" using h `b2 \ B2` unfolding B1 Func_def by auto - show ?thesis using A2 f_inv_into_f[OF b1] - unfolding True g_def Func_map_def j1_def j2[OF `b2 \ B2`] by auto - qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def, - auto intro: f_inv_into_f) - qed(insert h, unfold Func_def Func_map_def, auto) - qed - moreover have "g \ Func A2 A1" unfolding g_def apply(rule Func_map[OF h]) - using inv_into_into j2A2 B1 A2 inv_into_into - unfolding j1_def image_def by fast+ - ultimately show "h \ Func_map B2 f1 f2 ` Func A2 A1" - unfolding Func_map_def[abs_def] unfolding image_def by auto - qed(insert B1 Func_map[OF _ _ A2(2)], auto) -qed - lemma card_of_Pow_Func: "|Pow A| =o |Func A (UNIV::bool set)|" proof- diff -r d7593bfccf25 -r 7e0573a490ee src/HOL/Cardinals/Cardinals.thy --- a/src/HOL/Cardinals/Cardinals.thy Fri Jan 10 16:18:18 2014 +0100 +++ b/src/HOL/Cardinals/Cardinals.thy Fri Jan 10 17:24:52 2014 +0100 @@ -9,7 +9,7 @@ header {* Theory of Ordinals and Cardinals *} theory Cardinals -imports Cardinal_Order_Relation Cardinal_Arithmetic Wellorder_Extension +imports Ordinal_Arithmetic Cardinal_Arithmetic Wellorder_Extension begin end diff -r d7593bfccf25 -r 7e0573a490ee src/HOL/Cardinals/Constructions_on_Wellorders.thy --- a/src/HOL/Cardinals/Constructions_on_Wellorders.thy Fri Jan 10 16:18:18 2014 +0100 +++ b/src/HOL/Cardinals/Constructions_on_Wellorders.thy Fri Jan 10 17:24:52 2014 +0100 @@ -15,7 +15,10 @@ ordLeq_Well_order_simp[simp] not_ordLeq_iff_ordLess[simp] not_ordLess_iff_ordLeq[simp] + Func_empty[simp] + Func_is_emp[simp] +lemma Func_emp2[simp]: "A \ {} \ Func A {} = {}" by auto subsection {* Restriction to a set *} @@ -430,4 +433,648 @@ using NE_P P omax_ordLess by blast qed + + +section {* Limit and Succesor Ordinals *} + +lemma embed_underS2: +assumes r: "Well_order r" and s: "Well_order s" and g: "embed r s g" and a: "a \ Field r" +shows "g ` underS r a = underS s (g a)" +using embed_underS[OF assms] unfolding bij_betw_def by auto + +lemma bij_betw_insert: +assumes "b \ A" and "f b \ A'" and "bij_betw f A A'" +shows "bij_betw f (insert b A) (insert (f b) A')" +using notIn_Un_bij_betw[OF assms] by auto + +context wo_rel +begin + +lemma underS_induct: + assumes "\a. (\ a1. a1 \ underS a \ P a1) \ P a" + shows "P a" + by (induct rule: well_order_induct) (rule assms, simp add: underS_def) + +lemma suc_underS: +assumes B: "B \ Field r" and A: "AboveS B \ {}" and b: "b \ B" +shows "b \ underS (suc B)" +using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto + +lemma underS_supr: +assumes bA: "b \ underS (supr A)" and A: "A \ Field r" +shows "\ a \ A. b \ underS a" +proof(rule ccontr, auto) + have bb: "b \ Field r" using bA unfolding underS_def Field_def by auto + assume "\a\A. b \ underS a" + hence 0: "\a \ A. (a,b) \ r" using A bA unfolding underS_def + by simp (metis REFL in_mono max2_def max2_greater refl_on_domain) + have "(supr A, b) \ r" apply(rule supr_least[OF A bb]) using 0 by auto + thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD) +qed + +lemma underS_suc: +assumes bA: "b \ underS (suc A)" and A: "A \ Field r" +shows "\ a \ A. b \ under a" +proof(rule ccontr, auto) + have bb: "b \ Field r" using bA unfolding underS_def Field_def by auto + assume "\a\A. b \ under a" + hence 0: "\a \ A. a \ underS b" using A bA unfolding underS_def + by simp (metis (lifting) bb max2_def max2_greater mem_Collect_eq rel.under_def set_rev_mp) + have "(suc A, b) \ r" apply(rule suc_least[OF A bb]) using 0 unfolding rel.underS_def by auto + thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD) +qed + +lemma (in rel) AboveS_underS: +assumes "i \ Field r" +shows "i \ AboveS (underS i)" +using assms unfolding AboveS_def underS_def by auto + +lemma (in wo_rel) in_underS_supr: +assumes j: "j \ underS i" and i: "i \ A" and A: "A \ Field r" and AA: "Above A \ {}" +shows "j \ underS (supr A)" +proof- + have "(i,supr A) \ r" using supr_greater[OF A AA i] . + thus ?thesis using j unfolding underS_def + by simp (metis REFL TRANS max2_def max2_equals1 rel.refl_on_domain transD) +qed + +lemma inj_on_Field: +assumes A: "A \ Field r" and f: "\ a b. \a \ A; b \ A; a \ underS b\ \ f a \ f b" +shows "inj_on f A" +unfolding inj_on_def proof safe + fix a b assume a: "a \ A" and b: "b \ A" and fab: "f a = f b" + {assume "a \ underS b" + hence False using f[OF a b] fab by auto + } + moreover + {assume "b \ underS a" + hence False using f[OF b a] fab by auto + } + ultimately show "a = b" using TOTALS A a b unfolding underS_def by auto +qed + +lemma in_notinI: +assumes "(j,i) \ r \ j = i" and "i \ Field r" and "j \ Field r" +shows "(i,j) \ r" by (metis assms max2_def max2_greater_among) + +lemma ofilter_init_seg_of: +assumes "ofilter F" +shows "Restr r F initial_segment_of r" +using assms unfolding ofilter_def init_seg_of_def under_def by auto + +lemma underS_init_seg_of_Collect: +assumes "\j1 j2. \j2 \ underS i; (j1, j2) \ r\ \ R j1 initial_segment_of R j2" +shows "{R j |j. j \ underS i} \ Chains init_seg_of" +unfolding Chains_def proof safe + fix j ja assume jS: "j \ underS i" and jaS: "ja \ underS i" + and init: "(R ja, R j) \ init_seg_of" + hence jja: "{j,ja} \ Field r" and j: "j \ Field r" and ja: "ja \ Field r" + and jjai: "(j,i) \ r" "(ja,i) \ r" + and i: "i \ {j,ja}" unfolding Field_def rel.underS_def by auto + have jj: "(j,j) \ r" and jaja: "(ja,ja) \ r" using j ja by (metis in_notinI)+ + show "R j initial_segment_of R ja" + using jja init jjai i + by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: underS_def) +qed + +lemma (in wo_rel) Field_init_seg_of_Collect: +assumes "\j1 j2. \j2 \ Field r; (j1, j2) \ r\ \ R j1 initial_segment_of R j2" +shows "{R j |j. j \ Field r} \ Chains init_seg_of" +unfolding Chains_def proof safe + fix j ja assume jS: "j \ Field r" and jaS: "ja \ Field r" + and init: "(R ja, R j) \ init_seg_of" + hence jja: "{j,ja} \ Field r" and j: "j \ Field r" and ja: "ja \ Field r" + unfolding Field_def rel.underS_def by auto + have jj: "(j,j) \ r" and jaja: "(ja,ja) \ r" using j ja by (metis in_notinI)+ + show "R j initial_segment_of R ja" + using jja init + by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: Field_def) +qed + +subsection {* Successor and limit elements of an ordinal *} + +definition "succ i \ suc {i}" + +definition "isSucc i \ \ j. aboveS j \ {} \ i = succ j" + +definition "zero = minim (Field r)" + +definition "isLim i \ \ isSucc i" + +lemma zero_smallest[simp]: +assumes "j \ Field r" shows "(zero, j) \ r" +unfolding zero_def +by (metis AboveS_Field assms subset_AboveS_UnderS subset_antisym subset_refl suc_def suc_least_AboveS) + +lemma zero_in_Field: assumes "Field r \ {}" shows "zero \ Field r" +using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def) + +lemma leq_zero_imp[simp]: +"(x, zero) \ r \ x = zero" +by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest) + +lemma leq_zero[simp]: +assumes "Field r \ {}" shows "(x, zero) \ r \ x = zero" +using zero_in_Field[OF assms] in_notinI[of x zero] by auto + +lemma under_zero[simp]: +assumes "Field r \ {}" shows "under zero = {zero}" +using assms unfolding under_def by auto + +lemma underS_zero[simp,intro]: "underS zero = {}" +unfolding underS_def by auto + +lemma isSucc_succ: "aboveS i \ {} \ isSucc (succ i)" +unfolding isSucc_def succ_def by auto + +lemma succ_in_diff: +assumes "aboveS i \ {}" shows "(i,succ i) \ r \ succ i \ i" +using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto + +lemmas succ_in[simp] = succ_in_diff[THEN conjunct1] +lemmas succ_diff[simp] = succ_in_diff[THEN conjunct2] + +lemma succ_in_Field[simp]: +assumes "aboveS i \ {}" shows "succ i \ Field r" +using succ_in[OF assms] unfolding Field_def by auto + +lemma succ_not_in: +assumes "aboveS i \ {}" shows "(succ i, i) \ r" +proof + assume 1: "(succ i, i) \ r" + hence "succ i \ Field r \ i \ Field r" unfolding Field_def by auto + hence "(i, succ i) \ r \ succ i \ i" using assms by auto + thus False using 1 by (metis ANTISYM antisymD) +qed + +lemma not_isSucc_zero: "\ isSucc zero" +proof + assume "isSucc zero" + moreover + then obtain i where "aboveS i \ {}" and 1: "minim (Field r) = succ i" + unfolding isSucc_def zero_def by auto + hence "succ i \ Field r" by auto + ultimately show False by (metis REFL isSucc_def minim_least refl_on_domain + subset_refl succ_in succ_not_in zero_def) +qed + +lemma isLim_zero[simp]: "isLim zero" + by (metis isLim_def not_isSucc_zero) + +lemma succ_smallest: +assumes "(i,j) \ r" and "i \ j" +shows "(succ i, j) \ r" +unfolding succ_def apply(rule suc_least) +using assms unfolding Field_def by auto + +lemma isLim_supr: +assumes f: "i \ Field r" and l: "isLim i" +shows "i = supr (underS i)" +proof(rule equals_supr) + fix j assume j: "j \ Field r" and 1: "\ j'. j' \ underS i \ (j',j) \ r" + show "(i,j) \ r" proof(intro in_notinI[OF _ f j], safe) + assume ji: "(j,i) \ r" "j \ i" + hence a: "aboveS j \ {}" unfolding aboveS_def by auto + hence "i \ succ j" using l unfolding isLim_def isSucc_def by auto + moreover have "(succ j, i) \ r" using succ_smallest[OF ji] by auto + ultimately have "succ j \ underS i" unfolding underS_def by auto + hence "(succ j, j) \ r" using 1 by auto + thus False using succ_not_in[OF a] by simp + qed +qed(insert f, unfold underS_def Field_def, auto) + +definition "pred i \ SOME j. j \ Field r \ aboveS j \ {} \ succ j = i" + +lemma pred_Field_succ: +assumes "isSucc i" shows "pred i \ Field r \ aboveS (pred i) \ {} \ succ (pred i) = i" +proof- + obtain j where a: "aboveS j \ {}" and i: "i = succ j" using assms unfolding isSucc_def by auto + have 1: "j \ Field r" "j \ i" unfolding Field_def i + using succ_diff[OF a] a unfolding aboveS_def by auto + show ?thesis unfolding pred_def apply(rule someI_ex) using 1 i a by auto +qed + +lemmas pred_Field[simp] = pred_Field_succ[THEN conjunct1] +lemmas aboveS_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct1] +lemmas succ_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct2] + +lemma isSucc_pred_in: +assumes "isSucc i" shows "(pred i, i) \ r" +proof- + def j \ "pred i" + have i: "i = succ j" using assms unfolding j_def by simp + have a: "aboveS j \ {}" unfolding j_def using assms by auto + show ?thesis unfolding j_def[symmetric] unfolding i using succ_in[OF a] . +qed + +lemma isSucc_pred_diff: +assumes "isSucc i" shows "pred i \ i" +by (metis aboveS_pred assms succ_diff succ_pred) + +(* todo: pred maximal, pred injective? *) + +lemma succ_inj[simp]: +assumes "aboveS i \ {}" and "aboveS j \ {}" +shows "succ i = succ j \ i = j" +proof safe + assume s: "succ i = succ j" + {assume "i \ j" and "(i,j) \ r" + hence "(succ i, j) \ r" using assms by (metis succ_smallest) + hence False using s assms by (metis succ_not_in) + } + moreover + {assume "i \ j" and "(j,i) \ r" + hence "(succ j, i) \ r" using assms by (metis succ_smallest) + hence False using s assms by (metis succ_not_in) + } + ultimately show "i = j" by (metis TOTALS WELL assms(1) assms(2) succ_in_diff well_order_on_domain) +qed + +lemma pred_succ[simp]: +assumes "aboveS j \ {}" shows "pred (succ j) = j" +unfolding pred_def apply(rule some_equality) +using assms apply(force simp: Field_def aboveS_def) +by (metis assms succ_inj) + +lemma less_succ[simp]: +assumes "aboveS i \ {}" +shows "(j, succ i) \ r \ (j,i) \ r \ j = succ i" +apply safe + apply (metis WELL assms in_notinI rel.well_order_on_domain suc_singl_pred succ_def succ_in_diff) + apply (metis (hide_lams, full_types) REFL TRANS assms max2_def max2_equals1 rel.refl_on_domain succ_in_Field succ_not_in transD) + apply (metis assms in_notinI succ_in_Field) +done + +lemma underS_succ[simp]: +assumes "aboveS i \ {}" +shows "underS (succ i) = under i" +unfolding underS_def under_def by (auto simp: assms succ_not_in) + +lemma succ_mono: +assumes "aboveS j \ {}" and "(i,j) \ r" +shows "(succ i, succ j) \ r" +by (metis (full_types) assms less_succ succ_smallest) + +lemma under_succ[simp]: +assumes "aboveS i \ {}" +shows "under (succ i) = insert (succ i) (under i)" +using less_succ[OF assms] unfolding under_def by auto + +definition mergeSL :: "('a \ 'b \ 'b) \ (('a \ 'b) \ 'a \ 'b) \ ('a \ 'b) \ 'a \ 'b" +where +"mergeSL S L f i \ + if isSucc i then S (pred i) (f (pred i)) + else L f i" + + +subsection {* Well-order recursion with (zero), succesor, and limit *} + +definition worecSL :: "('a \ 'b \ 'b) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" +where "worecSL S L \ worec (mergeSL S L)" + +definition "adm_woL L \ \f g i. isLim i \ (\j\underS i. f j = g j) \ L f i = L g i" + +lemma mergeSL: +assumes "adm_woL L" shows "adm_wo (mergeSL S L)" +unfolding adm_wo_def proof safe + fix f g :: "'a => 'b" and i :: 'a + assume 1: "\j\underS i. f j = g j" + show "mergeSL S L f i = mergeSL S L g i" + proof(cases "isSucc i") + case True + hence "pred i \ underS i" unfolding underS_def using isSucc_pred_in isSucc_pred_diff by auto + thus ?thesis using True 1 unfolding mergeSL_def by auto + next + case False hence "isLim i" unfolding isLim_def by auto + thus ?thesis using assms False 1 unfolding mergeSL_def adm_woL_def by auto + qed +qed + +lemma worec_fixpoint1: "adm_wo H \ worec H i = H (worec H) i" +by (metis worec_fixpoint) + +lemma worecSL_isSucc: +assumes a: "adm_woL L" and i: "isSucc i" +shows "worecSL S L i = S (pred i) (worecSL S L (pred i))" +proof- + let ?H = "mergeSL S L" + have "worecSL S L i = ?H (worec ?H) i" + unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] . + also have "... = S (pred i) (worecSL S L (pred i))" + unfolding worecSL_def mergeSL_def using i by simp + finally show ?thesis . +qed + +lemma worecSL_succ: +assumes a: "adm_woL L" and i: "aboveS j \ {}" +shows "worecSL S L (succ j) = S j (worecSL S L j)" +proof- + def i \ "succ j" + have i: "isSucc i" by (metis i i_def isSucc_succ) + have ij: "j = pred i" unfolding i_def using assms by simp + have 0: "succ (pred i) = i" using i by simp + show ?thesis unfolding ij using worecSL_isSucc[OF a i] unfolding 0 . +qed + +lemma worecSL_isLim: +assumes a: "adm_woL L" and i: "isLim i" +shows "worecSL S L i = L (worecSL S L) i" +proof- + let ?H = "mergeSL S L" + have "worecSL S L i = ?H (worec ?H) i" + unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] . + also have "... = L (worecSL S L) i" + using i unfolding worecSL_def mergeSL_def isLim_def by simp + finally show ?thesis . +qed + +definition worecZSL :: "'b \ ('a \ 'b \ 'b) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" +where "worecZSL Z S L \ worecSL S (\ f a. if a = zero then Z else L f a)" + +lemma worecZSL_zero: +assumes a: "adm_woL L" +shows "worecZSL Z S L zero = Z" +proof- + let ?L = "\ f a. if a = zero then Z else L f a" + have "worecZSL Z S L zero = ?L (worecZSL Z S L) zero" + unfolding worecZSL_def apply(rule worecSL_isLim) + using assms unfolding adm_woL_def by auto + also have "... = Z" by simp + finally show ?thesis . +qed + +lemma worecZSL_succ: +assumes a: "adm_woL L" and i: "aboveS j \ {}" +shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)" +unfolding worecZSL_def apply(rule worecSL_succ) +using assms unfolding adm_woL_def by auto + +lemma worecZSL_isLim: +assumes a: "adm_woL L" and "isLim i" and "i \ zero" +shows "worecZSL Z S L i = L (worecZSL Z S L) i" +proof- + let ?L = "\ f a. if a = zero then Z else L f a" + have "worecZSL Z S L i = ?L (worecZSL Z S L) i" + unfolding worecZSL_def apply(rule worecSL_isLim) + using assms unfolding adm_woL_def by auto + also have "... = L (worecZSL Z S L) i" using assms by simp + finally show ?thesis . +qed + + +subsection {* Well-order succ-lim induction: *} + +lemma ord_cases: +obtains j where "i = succ j" and "aboveS j \ {}" | "isLim i" +by (metis isLim_def isSucc_def) + +lemma well_order_inductSL[case_names Suc Lim]: +assumes SUCC: "\i. \aboveS i \ {}; P i\ \ P (succ i)" and +LIM: "\i. \isLim i; \j. j \ underS i \ P j\ \ P i" +shows "P i" +proof(induction rule: well_order_induct) + fix i assume 0: "\j. j \ i \ (j, i) \ r \ P j" + show "P i" proof(cases i rule: ord_cases) + fix j assume i: "i = succ j" and j: "aboveS j \ {}" + hence "j \ i \ (j, i) \ r" by (metis succ_diff succ_in) + hence 1: "P j" using 0 by simp + show "P i" unfolding i apply(rule SUCC) using 1 j by auto + qed(insert 0 LIM, unfold underS_def, auto) +qed + +lemma well_order_inductZSL[case_names Zero Suc Lim]: +assumes ZERO: "P zero" +and SUCC: "\i. \aboveS i \ {}; P i\ \ P (succ i)" and +LIM: "\i. \isLim i; i \ zero; \j. j \ underS i \ P j\ \ P i" +shows "P i" +apply(induction rule: well_order_inductSL) using assms by auto + +(* Succesor and limit ordinals *) +definition "isSuccOrd \ \ j \ Field r. \ i \ Field r. (i,j) \ r" +definition "isLimOrd \ \ isSuccOrd" + +lemma isLimOrd_succ: +assumes isLimOrd and "i \ Field r" +shows "succ i \ Field r" +using assms unfolding isLimOrd_def isSuccOrd_def +by (metis REFL in_notinI refl_on_domain succ_smallest) + +lemma isLimOrd_aboveS: +assumes l: isLimOrd and i: "i \ Field r" +shows "aboveS i \ {}" +proof- + obtain j where "j \ Field r" and "(j,i) \ r" + using assms unfolding isLimOrd_def isSuccOrd_def by auto + hence "(i,j) \ r \ j \ i" by (metis i max2_def max2_greater) + thus ?thesis unfolding aboveS_def by auto +qed + +lemma succ_aboveS_isLimOrd: +assumes "\ i \ Field r. aboveS i \ {} \ succ i \ Field r" +shows isLimOrd +unfolding isLimOrd_def isSuccOrd_def proof safe + fix j assume j: "j \ Field r" "\i\Field r. (i, j) \ r" + hence "(succ j, j) \ r" using assms by auto + moreover have "aboveS j \ {}" using assms j unfolding aboveS_def by auto + ultimately show False by (metis succ_not_in) +qed + +lemma isLim_iff: +assumes l: "isLim i" and j: "j \ underS i" +shows "\ k. k \ underS i \ j \ underS k" +proof- + have a: "aboveS j \ {}" using j unfolding underS_def aboveS_def by auto + show ?thesis apply(rule exI[of _ "succ j"]) apply safe + using assms a unfolding underS_def isLim_def + apply (metis (lifting, full_types) isSucc_def mem_Collect_eq succ_smallest) + by (metis (lifting, full_types) a mem_Collect_eq succ_diff succ_in) +qed + + +end (* context wo_rel *) + +abbreviation "zero \ wo_rel.zero" +abbreviation "succ \ wo_rel.succ" +abbreviation "pred \ wo_rel.pred" +abbreviation "isSucc \ wo_rel.isSucc" +abbreviation "isLim \ wo_rel.isLim" +abbreviation "isLimOrd \ wo_rel.isLimOrd" +abbreviation "isSuccOrd \ wo_rel.isSuccOrd" +abbreviation "adm_woL \ wo_rel.adm_woL" +abbreviation "worecSL \ wo_rel.worecSL" +abbreviation "worecZSL \ wo_rel.worecZSL" + +section {* Projections of Wellorders *} + +definition "oproj r s f \ Field s \ f ` (Field r) \ compat r s f" + +lemma oproj_in: +assumes "oproj r s f" and "(a,a') \ r" +shows "(f a, f a') \ s" +using assms unfolding oproj_def compat_def by auto + +lemma oproj_Field: +assumes f: "oproj r s f" and a: "a \ Field r" +shows "f a \ Field s" +using oproj_in[OF f] a unfolding Field_def by auto + +lemma oproj_Field2: +assumes f: "oproj r s f" and a: "b \ Field s" +shows "\ a \ Field r. f a = b" +using assms unfolding oproj_def by auto + +lemma oproj_under: +assumes f: "oproj r s f" and a: "a \ under r a'" +shows "f a \ under s (f a')" +using oproj_in[OF f] a unfolding rel.under_def by auto + +(* An ordinal is embedded in another whenever it is embedded as an order +(not necessarily as initial segment):*) +theorem embedI: +assumes r: "Well_order r" and s: "Well_order s" +and f: "\ a. a \ Field r \ f a \ Field s \ f ` underS r a \ underS s (f a)" +shows "\ g. embed r s g" +proof- + interpret r!: wo_rel r by unfold_locales (rule r) + interpret s!: wo_rel s by unfold_locales (rule s) + let ?G = "\ g a. suc s (g ` underS r a)" + def g \ "worec r ?G" + have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto + (* *) + {fix a assume "a \ Field r" + hence "bij_betw g (under r a) (under s (g a)) \ + g a \ under s (f a)" + proof(induction a rule: r.underS_induct) + case (1 a) + hence a: "a \ Field r" + and IH1a: "\ a1. a1 \ underS r a \ inj_on g (under r a1)" + and IH1b: "\ a1. a1 \ underS r a \ g ` under r a1 = under s (g a1)" + and IH2: "\ a1. a1 \ underS r a \ g a1 \ under s (f a1)" + unfolding rel.underS_def Field_def bij_betw_def by auto + have fa: "f a \ Field s" using f[OF a] by auto + have g: "g a = suc s (g ` underS r a)" + using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by simp + have A0: "g ` underS r a \ Field s" + using IH1b by (metis IH2 image_subsetI in_mono rel.under_Field) + {fix a1 assume a1: "a1 \ underS r a" + from IH2[OF this] have "g a1 \ under s (f a1)" . + moreover have "f a1 \ underS s (f a)" using f[OF a] a1 by auto + ultimately have "g a1 \ underS s (f a)" by (metis s.ANTISYM s.TRANS rel.under_underS_trans) + } + hence "f a \ AboveS s (g ` underS r a)" unfolding rel.AboveS_def + using fa by simp (metis (lifting, full_types) mem_Collect_eq rel.underS_def) + hence A: "AboveS s (g ` underS r a) \ {}" by auto + have B: "\ a1. a1 \ underS r a \ g a1 \ underS s (g a)" + unfolding g apply(rule s.suc_underS[OF A0 A]) by auto + {fix a1 a2 assume a2: "a2 \ underS r a" and 1: "a1 \ underS r a2" + hence a12: "{a1,a2} \ under r a2" and "a1 \ a2" using r.REFL a + unfolding rel.underS_def rel.under_def refl_on_def Field_def by auto + hence "g a1 \ g a2" using IH1a[OF a2] unfolding inj_on_def by auto + hence "g a1 \ underS s (g a2)" using IH1b[OF a2] a12 + unfolding rel.underS_def rel.under_def by auto + } note C = this + have ga: "g a \ Field s" unfolding g using s.suc_inField[OF A0 A] . + have aa: "a \ under r a" + using a r.REFL unfolding rel.under_def rel.underS_def refl_on_def by auto + show ?case proof safe + show "bij_betw g (under r a) (under s (g a))" unfolding bij_betw_def proof safe + show "inj_on g (under r a)" proof(rule r.inj_on_Field) + fix a1 a2 assume "a1 \ under r a" and a2: "a2 \ under r a" and a1: "a1 \ underS r a2" + hence a22: "a2 \ under r a2" and a12: "a1 \ under r a2" "a1 \ a2" + using a r.REFL unfolding rel.under_def rel.underS_def refl_on_def by auto + show "g a1 \ g a2" + proof(cases "a2 = a") + case False hence "a2 \ underS r a" + using a2 unfolding rel.underS_def rel.under_def by auto + from IH1a[OF this] show ?thesis using a12 a22 unfolding inj_on_def by auto + qed(insert B a1, unfold rel.underS_def, auto) + qed(unfold rel.under_def Field_def, auto) + next + fix a1 assume a1: "a1 \ under r a" + show "g a1 \ under s (g a)" + proof(cases "a1 = a") + case True thus ?thesis + using ga s.REFL unfolding refl_on_def rel.under_def by auto + next + case False + hence a1: "a1 \ underS r a" using a1 unfolding rel.underS_def rel.under_def by auto + thus ?thesis using B unfolding rel.underS_def rel.under_def by auto + qed + next + fix b1 assume b1: "b1 \ under s (g a)" + show "b1 \ g ` under r a" + proof(cases "b1 = g a") + case True thus ?thesis using aa by auto + next + case False + hence "b1 \ underS s (g a)" using b1 unfolding rel.underS_def rel.under_def by auto + from s.underS_suc[OF this[unfolded g] A0] + obtain a1 where a1: "a1 \ underS r a" and b1: "b1 \ under s (g a1)" by auto + obtain a2 where "a2 \ under r a1" and b1: "b1 = g a2" using IH1b[OF a1] b1 by auto + hence "a2 \ under r a" using a1 + by (metis r.ANTISYM r.TRANS in_mono rel.underS_subset_under rel.under_underS_trans) + thus ?thesis using b1 by auto + qed + qed + next + have "(g a, f a) \ s" unfolding g proof(rule s.suc_least[OF A0]) + fix b1 assume "b1 \ g ` underS r a" + then obtain a1 where a1: "b1 = g a1" and a1: "a1 \ underS r a" by auto + hence "b1 \ underS s (f a)" + using a by (metis `\a1. a1 \ underS r a \ g a1 \ underS s (f a)`) + thus "f a \ b1 \ (b1, f a) \ s" unfolding rel.underS_def by auto + qed(insert fa, auto) + thus "g a \ under s (f a)" unfolding rel.under_def by auto + qed + qed + } + thus ?thesis unfolding embed_def by auto +qed + +corollary ordLeq_def2: + "r \o s \ Well_order r \ Well_order s \ + (\ f. \ a \ Field r. f a \ Field s \ f ` underS r a \ underS s (f a))" +using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s] +unfolding ordLeq_def by fast + +lemma iso_oproj: + assumes r: "Well_order r" and s: "Well_order s" and f: "iso r s f" + shows "oproj r s f" +using assms unfolding oproj_def +by (subst (asm) iso_iff3) (auto simp: bij_betw_def) + +theorem oproj_embed: +assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f" +shows "\ g. embed s r g" +proof (rule embedI[OF s r, of "inv_into (Field r) f"], unfold rel.underS_def, safe) + fix b assume "b \ Field s" + thus "inv_into (Field r) f b \ Field r" using oproj_Field2[OF f] by (metis imageI inv_into_into) +next + fix a b assume "b \ Field s" "a \ b" "(a, b) \ s" + "inv_into (Field r) f a = inv_into (Field r) f b" + with f show False by (auto dest!: inv_into_injective simp: Field_def oproj_def) +next + fix a b assume *: "b \ Field s" "a \ b" "(a, b) \ s" + { assume "(inv_into (Field r) f a, inv_into (Field r) f b) \ r" + moreover + from *(3) have "a \ Field s" unfolding Field_def by auto + with *(1,2) have + "inv_into (Field r) f a \ Field r" "inv_into (Field r) f b \ Field r" + "inv_into (Field r) f a \ inv_into (Field r) f b" + by (auto dest!: oproj_Field2[OF f] inv_into_injective intro!: inv_into_into) + ultimately have "(inv_into (Field r) f b, inv_into (Field r) f a) \ r" + using r by (auto simp: well_order_on_def linear_order_on_def total_on_def) + with f[unfolded oproj_def compat_def] *(1) `a \ Field s` + f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"] + have "(b, a) \ s" by (metis in_mono) + with *(2,3) s have False + by (auto simp: well_order_on_def linear_order_on_def partial_order_on_def antisym_def) + } thus "(inv_into (Field r) f a, inv_into (Field r) f b) \ r" by blast +qed + +corollary oproj_ordLeq: +assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f" +shows "s \o r" +using oproj_embed[OF assms] r s unfolding ordLeq_def by auto + end diff -r d7593bfccf25 -r 7e0573a490ee src/HOL/Cardinals/Constructions_on_Wellorders_FP.thy --- a/src/HOL/Cardinals/Constructions_on_Wellorders_FP.thy Fri Jan 10 16:18:18 2014 +0100 +++ b/src/HOL/Cardinals/Constructions_on_Wellorders_FP.thy Fri Jan 10 17:24:52 2014 +0100 @@ -1617,5 +1617,158 @@ ultimately show ?thesis by blast qed +definition Func where +"Func A B = {f . (\ a \ A. f a \ B) \ (\ a. a \ A \ f a = undefined)}" + +lemma Func_empty: +"Func {} B = {\x. undefined}" +unfolding Func_def by auto + +lemma Func_elim: +assumes "g \ Func A B" and "a \ A" +shows "\ b. b \ B \ g a = b" +using assms unfolding Func_def by (cases "g a = undefined") auto + +definition curr where +"curr A f \ \ a. if a \ A then \b. f (a,b) else undefined" + +lemma curr_in: +assumes f: "f \ Func (A <*> B) C" +shows "curr A f \ Func A (Func B C)" +using assms unfolding curr_def Func_def by auto + +lemma curr_inj: +assumes "f1 \ Func (A <*> B) C" and "f2 \ Func (A <*> B) C" +shows "curr A f1 = curr A f2 \ f1 = f2" +proof safe + assume c: "curr A f1 = curr A f2" + show "f1 = f2" + proof (rule ext, clarify) + fix a b show "f1 (a, b) = f2 (a, b)" + proof (cases "(a,b) \ A <*> B") + case False + thus ?thesis using assms unfolding Func_def by auto + next + case True hence a: "a \ A" and b: "b \ B" by auto + thus ?thesis + using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp + qed + qed +qed + +lemma curr_surj: +assumes "g \ Func A (Func B C)" +shows "\ f \ Func (A <*> B) C. curr A f = g" +proof + let ?f = "\ ab. if fst ab \ A \ snd ab \ B then g (fst ab) (snd ab) else undefined" + show "curr A ?f = g" + proof (rule ext) + fix a show "curr A ?f a = g a" + proof (cases "a \ A") + case False + hence "g a = undefined" using assms unfolding Func_def by auto + thus ?thesis unfolding curr_def using False by simp + next + case True + obtain g1 where "g1 \ Func B C" and "g a = g1" + using assms using Func_elim[OF assms True] by blast + thus ?thesis using True unfolding Func_def curr_def by auto + qed + qed + show "?f \ Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto +qed + +lemma bij_betw_curr: +"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))" +unfolding bij_betw_def inj_on_def image_def +apply (intro impI conjI ballI) +apply (erule curr_inj[THEN iffD1], assumption+) +apply auto +apply (erule curr_in) +using curr_surj by blast + +definition Func_map where +"Func_map B2 f1 f2 g b2 \ if b2 \ B2 then f1 (g (f2 b2)) else undefined" + +lemma Func_map: +assumes g: "g \ Func A2 A1" and f1: "f1 ` A1 \ B1" and f2: "f2 ` B2 \ A2" +shows "Func_map B2 f1 f2 g \ Func B2 B1" +using assms unfolding Func_def Func_map_def mem_Collect_eq by auto + +lemma Func_non_emp: +assumes "B \ {}" +shows "Func A B \ {}" +proof- + obtain b where b: "b \ B" using assms by auto + hence "(\ a. if a \ A then b else undefined) \ Func A B" unfolding Func_def by auto + thus ?thesis by blast +qed + +lemma Func_is_emp: +"Func A B = {} \ A \ {} \ B = {}" (is "?L \ ?R") +proof + assume L: ?L + moreover {assume "A = {}" hence False using L Func_empty by auto} + moreover {assume "B \ {}" hence False using L Func_non_emp by metis} + ultimately show ?R by blast +next + assume R: ?R + moreover + {fix f assume "f \ Func A B" + moreover obtain a where "a \ A" using R by blast + ultimately obtain b where "b \ B" unfolding Func_def by blast + with R have False by blast + } + thus ?L by blast +qed + +lemma Func_map_surj: +assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \ A2" +and B2A2: "B2 = {} \ A2 = {}" +shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1" +proof(cases "B2 = {}") + case True + thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def) +next + case False note B2 = False + show ?thesis + proof safe + fix h assume h: "h \ Func B2 B1" + def j1 \ "inv_into A1 f1" + have "\ a2 \ f2 ` B2. \ b2. b2 \ B2 \ f2 b2 = a2" by blast + then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by metis + {fix b2 assume b2: "b2 \ B2" + hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto + moreover have "k (f2 b2) \ B2" using b2 A2(2) k by auto + ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast + } note kk = this + obtain b22 where b22: "b22 \ B2" using B2 by auto + def j2 \ "\ a2. if a2 \ f2 ` B2 then k a2 else b22" + have j2A2: "j2 ` A2 \ B2" unfolding j2_def using k b22 by auto + have j2: "\ b2. b2 \ B2 \ j2 (f2 b2) = b2" + using kk unfolding j2_def by auto + def g \ "Func_map A2 j1 j2 h" + have "Func_map B2 f1 f2 g = h" + proof (rule ext) + fix b2 show "Func_map B2 f1 f2 g b2 = h b2" + proof(cases "b2 \ B2") + case True + show ?thesis + proof (cases "h b2 = undefined") + case True + hence b1: "h b2 \ f1 ` A1" using h `b2 \ B2` unfolding B1 Func_def by auto + show ?thesis using A2 f_inv_into_f[OF b1] + unfolding True g_def Func_map_def j1_def j2[OF `b2 \ B2`] by auto + qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def, + auto intro: f_inv_into_f) + qed(insert h, unfold Func_def Func_map_def, auto) + qed + moreover have "g \ Func A2 A1" unfolding g_def apply(rule Func_map[OF h]) + using inv_into_into j2A2 B1 A2 inv_into_into + unfolding j1_def image_def by fast+ + ultimately show "h \ Func_map B2 f1 f2 ` Func A2 A1" + unfolding Func_map_def[abs_def] unfolding image_def by auto + qed(insert B1 Func_map[OF _ _ A2(2)], auto) +qed end diff -r d7593bfccf25 -r 7e0573a490ee src/HOL/Cardinals/Ordinal_Arithmetic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy Fri Jan 10 17:24:52 2014 +0100 @@ -0,0 +1,1728 @@ +(* Title: HOL/Cardinals/Ordinal_Arithmetic.thy + Author: Dmitriy Traytel, TU Muenchen + Copyright 2014 + +Ordinal arithmetic. +*) + +header {* Ordinal Arithmetic *} + +theory Ordinal_Arithmetic +imports Constructions_on_Wellorders +begin + +definition osum :: "'a rel \ 'b rel \ ('a + 'b) rel" (infixr "+o" 70) +where + "r +o r' = map_pair Inl Inl ` r \ map_pair Inr Inr ` r' \ + {(Inl a, Inr a') | a a' . a \ Field r \ a' \ Field r'}" + +lemma Field_osum: "Field(r +o r') = Inl ` Field r \ Inr ` Field r'" + unfolding osum_def Field_def by auto + +lemma osum_Refl:"\Refl r; Refl r'\ \ Refl (r +o r')" + (*Need first unfold Field_osum, only then osum_def*) + unfolding refl_on_def Field_osum unfolding osum_def by blast + +lemma osum_trans: +assumes TRANS: "trans r" and TRANS': "trans r'" +shows "trans (r +o r')" +proof(unfold trans_def, safe) + fix x y z assume *: "(x, y) \ r +o r'" "(y, z) \ r +o r'" + thus "(x, z) \ r +o r'" + proof (cases x y z rule: sum.exhaust[case_product sum.exhaust sum.exhaust]) + case (Inl_Inl_Inl a b c) + with * have "(a,b) \ r" "(b,c) \ r" unfolding osum_def by auto + with TRANS have "(a,c) \ r" unfolding trans_def by blast + with Inl_Inl_Inl show ?thesis unfolding osum_def by auto + next + case (Inl_Inl_Inr a b c) + with * have "a \ Field r" "c \ Field r'" unfolding osum_def Field_def by auto + with Inl_Inl_Inr show ?thesis unfolding osum_def by auto + next + case (Inl_Inr_Inr a b c) + with * have "a \ Field r" "c \ Field r'" unfolding osum_def Field_def by auto + with Inl_Inr_Inr show ?thesis unfolding osum_def by auto + next + case (Inr_Inr_Inr a b c) + with * have "(a,b) \ r'" "(b,c) \ r'" unfolding osum_def by auto + with TRANS' have "(a,c) \ r'" unfolding trans_def by blast + with Inr_Inr_Inr show ?thesis unfolding osum_def by auto + qed (auto simp: osum_def) +qed + +lemma osum_Preorder: "\Preorder r; Preorder r'\ \ Preorder (r +o r')" + unfolding preorder_on_def using osum_Refl osum_trans by blast + +lemma osum_antisym: "\antisym r; antisym r'\ \ antisym (r +o r')" + unfolding antisym_def osum_def by auto + +lemma osum_Partial_order: "\Partial_order r; Partial_order r'\ \ Partial_order (r +o r')" + unfolding partial_order_on_def using osum_Preorder osum_antisym by blast + +lemma osum_Total: "\Total r; Total r'\ \ Total (r +o r')" + unfolding total_on_def Field_osum unfolding osum_def by blast + +lemma osum_Linear_order: "\Linear_order r; Linear_order r'\ \ Linear_order (r +o r')" + unfolding linear_order_on_def using osum_Partial_order osum_Total by blast + +lemma osum_wf: +assumes WF: "wf r" and WF': "wf r'" +shows "wf (r +o r')" +unfolding wf_eq_minimal2 unfolding Field_osum +proof(intro allI impI, elim conjE) + fix A assume *: "A \ Inl ` Field r \ Inr ` Field r'" and **: "A \ {}" + obtain B where B_def: "B = A Int Inl ` Field r" by blast + show "\a\A. \a'\A. (a', a) \ r +o r'" + proof(cases "B = {}") + case False + hence "B \ {}" "B \ Inl ` Field r" using B_def by auto + hence "Inl -` B \ {}" "Inl -` B \ Field r" unfolding vimage_def by auto + then obtain a where 1: "a \ Inl -` B" and "\a1 \ Inl -` B. (a1, a) \ r" + using WF unfolding wf_eq_minimal2 by metis + hence "\a1 \ A. (a1, Inl a) \ r +o r'" + unfolding osum_def using B_def ** by (auto simp: vimage_def Field_def) + thus ?thesis using 1 unfolding B_def by auto + next + case True + hence 1: "A \ Inr ` Field r'" using * B_def by auto + with ** have "Inr -`A \ {}" "Inr -` A \ Field r'" unfolding vimage_def by auto + with ** obtain a' where 2: "a' \ Inr -` A" and "\a1' \ Inr -` A. (a1',a') \ r'" + using WF' unfolding wf_eq_minimal2 by metis + hence "\a1' \ A. (a1', Inr a') \ r +o r'" + unfolding osum_def using ** 1 by (auto simp: vimage_def Field_def) + thus ?thesis using 2 by blast + qed +qed + +lemma osum_minus_Id: + assumes r: "Total r" "\ (r \ Id)" and r': "Total r'" "\ (r' \ Id)" + shows "(r +o r') - Id \ (r - Id) +o (r' - Id)" + unfolding osum_def Total_Id_Field[OF r] Total_Id_Field[OF r'] by auto + +lemma osum_minus_Id1: + "r \ Id \ (r +o r') - Id \ (Inl ` Field r \ Inr ` Field r') \ (map_pair Inr Inr ` (r' - Id))" + unfolding osum_def by auto + +lemma osum_minus_Id2: + "r' \ Id \ (r +o r') - Id \ (map_pair Inl Inl ` (r - Id)) \ (Inl ` Field r \ Inr ` Field r')" + unfolding osum_def by auto + +lemma osum_wf_Id: + assumes TOT: "Total r" and TOT': "Total r'" and WF: "wf(r - Id)" and WF': "wf(r' - Id)" + shows "wf ((r +o r') - Id)" +proof(cases "r \ Id \ r' \ Id") + case False + thus ?thesis + using osum_minus_Id[of r r'] assms osum_wf[of "r - Id" "r' - Id"] + wf_subset[of "(r - Id) +o (r' - Id)" "(r +o r') - Id"] by auto +next + have 1: "wf (Inl ` Field r \ Inr ` Field r')" by (rule wf_Int_Times) auto + case True + thus ?thesis + proof (elim disjE) + assume "r \ Id" + thus "wf ((r +o r') - Id)" + by (rule wf_subset[rotated, OF osum_minus_Id1 wf_Un[OF 1 wf_map_pair_image[OF WF']]]) auto + next + assume "r' \ Id" + thus "wf ((r +o r') - Id)" + by (rule wf_subset[rotated, OF osum_minus_Id2 wf_Un[OF wf_map_pair_image[OF WF] 1]]) auto + qed +qed + +lemma osum_Well_order: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "Well_order (r +o r')" +proof- + have "Total r \ Total r'" using WELL WELL' by (auto simp add: order_on_defs) + thus ?thesis using assms unfolding well_order_on_def + using osum_Linear_order osum_wf_Id by blast +qed + +lemma osum_embedL: + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "embed r (r +o r') Inl" +proof - + have 1: "Well_order (r +o r')" using assms by (auto simp add: osum_Well_order) + moreover + have "compat r (r +o r') Inl" unfolding compat_def osum_def by auto + moreover + have "ofilter (r +o r') (Inl ` Field r)" + unfolding wo_rel.ofilter_def[unfolded wo_rel_def, OF 1] Field_osum rel.under_def + unfolding osum_def Field_def by auto + ultimately show ?thesis using assms by (auto simp add: embed_iff_compat_inj_on_ofilter) +qed + +corollary osum_ordLeqL: + assumes WELL: "Well_order r" and WELL': "Well_order r'" + shows "r \o r +o r'" + using assms osum_embedL osum_Well_order unfolding ordLeq_def by blast + +lemma dir_image_alt: "dir_image r f = map_pair f f ` r" + unfolding dir_image_def map_pair_def by auto + +lemma map_pair_ordIso: "\Well_order r; inj_on f (Field r)\ \ map_pair f f ` r =o r" + unfolding dir_image_alt[symmetric] by (rule ordIso_symmetric[OF dir_image_ordIso]) + +definition oprod :: "'a rel \ 'b rel \ ('a \ 'b) rel" (infixr "*o" 80) +where "r *o r' = {((x1, y1), (x2, y2)). + (((y1, y2) \ r' - Id \ x1 \ Field r \ x2 \ Field r) \ + ((y1, y2) \ Restr Id (Field r') \ (x1, x2) \ r))}" + +lemma Field_oprod: "Field (r *o r') = Field r \ Field r'" + unfolding oprod_def Field_def by auto blast+ + +lemma oprod_Refl:"\Refl r; Refl r'\ \ Refl (r *o r')" + unfolding refl_on_def Field_oprod unfolding oprod_def by auto + +lemma oprod_trans: + assumes "trans r" "trans r'" "antisym r" "antisym r'" + shows "trans (r *o r')" +proof(unfold trans_def, safe) + fix x y z assume *: "(x, y) \ r *o r'" "(y, z) \ r *o r'" + thus "(x, z) \ r *o r'" + unfolding oprod_def + apply safe + apply (metis assms(2) transE) + apply (metis assms(2) transE) + apply (metis assms(2) transE) + apply (metis assms(4) antisymD) + apply (metis assms(4) antisymD) + apply (metis assms(2) transE) + apply (metis assms(4) antisymD) + apply (metis Field_def Range_iff Un_iff) + apply (metis Field_def Range_iff Un_iff) + apply (metis Field_def Range_iff Un_iff) + apply (metis Field_def Domain_iff Un_iff) + apply (metis Field_def Domain_iff Un_iff) + apply (metis Field_def Domain_iff Un_iff) + apply (metis assms(1) transE) + apply (metis assms(1) transE) + apply (metis assms(1) transE) + apply (metis assms(1) transE) + done +qed + +lemma oprod_Preorder: "\Preorder r; Preorder r'; antisym r; antisym r'\ \ Preorder (r *o r')" + unfolding preorder_on_def using oprod_Refl oprod_trans by blast + +lemma oprod_antisym: "\antisym r; antisym r'\ \ antisym (r *o r')" + unfolding antisym_def oprod_def by auto + +lemma oprod_Partial_order: "\Partial_order r; Partial_order r'\ \ Partial_order (r *o r')" + unfolding partial_order_on_def using oprod_Preorder oprod_antisym by blast + +lemma oprod_Total: "\Total r; Total r'\ \ Total (r *o r')" + unfolding total_on_def Field_oprod unfolding oprod_def by auto + +lemma oprod_Linear_order: "\Linear_order r; Linear_order r'\ \ Linear_order (r *o r')" + unfolding linear_order_on_def using oprod_Partial_order oprod_Total by blast + +lemma oprod_wf: +assumes WF: "wf r" and WF': "wf r'" +shows "wf (r *o r')" +unfolding wf_eq_minimal2 unfolding Field_oprod +proof(intro allI impI, elim conjE) + fix A assume *: "A \ Field r \ Field r'" and **: "A \ {}" + then obtain y where y: "y \ snd ` A" "\y'\snd ` A. (y', y) \ r'" + using spec[OF WF'[unfolded wf_eq_minimal2], of "snd ` A"] by auto + let ?A = "fst ` A \ {x. (x, y) \ A}" + from * y have "?A \ {}" "?A \ Field r" by auto + then obtain x where x: "x \ ?A" and "\x'\ ?A. (x', x) \ r" + using spec[OF WF[unfolded wf_eq_minimal2], of "?A"] by auto + with y have "\a'\A. (a', (x, y)) \ r *o r'" + unfolding oprod_def mem_Collect_eq split_beta fst_conv snd_conv Id_def by auto + moreover from x have "(x, y) \ A" by auto + ultimately show "\a\A. \a'\A. (a', a) \ r *o r'" by blast +qed + +lemma oprod_minus_Id: + assumes r: "Total r" "\ (r \ Id)" and r': "Total r'" "\ (r' \ Id)" + shows "(r *o r') - Id \ (r - Id) *o (r' - Id)" + unfolding oprod_def Total_Id_Field[OF r] Total_Id_Field[OF r'] by auto + +lemma oprod_minus_Id1: + "r \ Id \ r *o r' - Id \ {((x,y1), (x,y2)). x \ Field r \ (y1, y2) \ (r' - Id)}" + unfolding oprod_def by auto + +lemma wf_extend_oprod1: + assumes "wf r" + shows "wf {((x,y1), (x,y2)) . x \ A \ (y1, y2) \ r}" +proof (unfold wf_eq_minimal2, intro allI impI, elim conjE) + fix B + assume *: "B \ Field {((x,y1), (x,y2)) . x \ A \ (y1, y2) \ r}" and "B \ {}" + from image_mono[OF *, of snd] have "snd ` B \ Field r" unfolding Field_def by force + with `B \ {}` obtain x where x: "x \ snd ` B" "\x'\snd ` B. (x', x) \ r" + using spec[OF assms[unfolded wf_eq_minimal2], of "snd ` B"] by auto + then obtain a where "(a, x) \ B" by auto + moreover + from * x have "\a'\B. (a', (a, x)) \ {((x,y1), (x,y2)) . x \ A \ (y1, y2) \ r}" by auto + ultimately show "\ax\B. \a'\B. (a', ax) \ {((x,y1), (x,y2)) . x \ A \ (y1, y2) \ r}" by blast +qed + +lemma oprod_minus_Id2: + "r' \ Id \ r *o r' - Id \ {((x1,y), (x2,y)). (x1, x2) \ (r - Id) \ y \ Field r'}" + unfolding oprod_def by auto + +lemma wf_extend_oprod2: + assumes "wf r" + shows "wf {((x1,y), (x2,y)) . (x1, x2) \ r \ y \ A}" +proof (unfold wf_eq_minimal2, intro allI impI, elim conjE) + fix B + assume *: "B \ Field {((x1, y), (x2, y)). (x1, x2) \ r \ y \ A}" and "B \ {}" + from image_mono[OF *, of fst] have "fst ` B \ Field r" unfolding Field_def by force + with `B \ {}` obtain x where x: "x \ fst ` B" "\x'\fst ` B. (x', x) \ r" + using spec[OF assms[unfolded wf_eq_minimal2], of "fst ` B"] by auto + then obtain a where "(x, a) \ B" by auto + moreover + from * x have "\a'\B. (a', (x, a)) \ {((x1, y), x2, y). (x1, x2) \ r \ y \ A}" by auto + ultimately show "\xa\B. \a'\B. (a', xa) \ {((x1, y), x2, y). (x1, x2) \ r \ y \ A}" by blast +qed + +lemma oprod_wf_Id: + assumes TOT: "Total r" and TOT': "Total r'" and WF: "wf(r - Id)" and WF': "wf(r' - Id)" + shows "wf ((r *o r') - Id)" +proof(cases "r \ Id \ r' \ Id") + case False + thus ?thesis + using oprod_minus_Id[of r r'] assms oprod_wf[of "r - Id" "r' - Id"] + wf_subset[of "(r - Id) *o (r' - Id)" "(r *o r') - Id"] by auto +next + case True + thus ?thesis using wf_subset[OF wf_extend_oprod1[OF WF'] oprod_minus_Id1] + wf_subset[OF wf_extend_oprod2[OF WF] oprod_minus_Id2] by auto +qed + +lemma oprod_Well_order: +assumes WELL: "Well_order r" and WELL': "Well_order r'" +shows "Well_order (r *o r')" +proof- + have "Total r \ Total r'" using WELL WELL' by (auto simp add: order_on_defs) + thus ?thesis using assms unfolding well_order_on_def + using oprod_Linear_order oprod_wf_Id by blast +qed + +lemma oprod_embed: + assumes WELL: "Well_order r" and WELL': "Well_order r'" and "r' \ {}" + shows "embed r (r *o r') (\x. (x, minim r' (Field r')))" (is "embed _ _ ?f") +proof - + from assms(3) have r': "Field r' \ {}" unfolding Field_def by auto + have minim[simp]: "minim r' (Field r') \ Field r'" + using wo_rel.minim_inField[unfolded wo_rel_def, OF WELL' _ r'] by auto + { fix b + assume "(b, minim r' (Field r')) \ r'" + moreover hence "b \ Field r'" unfolding Field_def by auto + hence "(minim r' (Field r'), b) \ r'" + using wo_rel.minim_least[unfolded wo_rel_def, OF WELL' subset_refl] r' by auto + ultimately have "b = minim r' (Field r')" + by (metis WELL' antisym_def linear_order_on_def partial_order_on_def well_order_on_def) + } note * = this + have 1: "Well_order (r *o r')" using assms by (auto simp add: oprod_Well_order) + moreover + from r' have "compat r (r *o r') ?f" unfolding compat_def oprod_def by auto + moreover + from * have "ofilter (r *o r') (?f ` Field r)" + unfolding wo_rel.ofilter_def[unfolded wo_rel_def, OF 1] Field_oprod rel.under_def + unfolding oprod_def by auto (auto simp: image_iff Field_def) + moreover have "inj_on ?f (Field r)" unfolding inj_on_def by auto + ultimately show ?thesis using assms by (auto simp add: embed_iff_compat_inj_on_ofilter) +qed + +corollary oprod_ordLeq: "\Well_order r; Well_order r'; r' \ {}\ \ r \o r *o r'" + using oprod_embed oprod_Well_order unfolding ordLeq_def by blast + +definition "support z A f = {x \ A. f x \ z}" + +lemma support_Un[simp]: "support z (A \ B) f = support z A f \ support z B f" + unfolding support_def by auto + +lemma support_upd[simp]: "support z A (f(x := z)) = support z A f - {x}" + unfolding support_def by auto + +lemma support_upd_subset[simp]: "support z A (f(x := y)) \ support z A f \ {x}" + unfolding support_def by auto + +lemma fun_unequal_in_support: + assumes "f \ g" "f \ Func A B" "g \ Func A C" + shows "(support z A f \ support z A g) \ {a. f a \ g a} \ {}" (is "?L \ ?R \ {}") +proof - + from assms(1) obtain x where x: "f x \ g x" by blast + hence "x \ ?R" by simp + moreover from assms(2-3) x have "x \ A" unfolding Func_def by fastforce + with x have "x \ ?L" unfolding support_def by auto + ultimately show ?thesis by auto +qed + +definition fin_support where + "fin_support z A = {f. finite (support z A f)}" + +lemma finite_support: "f \ fin_support z A \ finite (support z A f)" + unfolding support_def fin_support_def by auto + +lemma fin_support_Field_osum: + "f \ fin_support z (Inl ` A \ Inr ` B) \ + (f o Inl) \ fin_support z A \ (f o Inr) \ fin_support z B" (is "?L \ ?R1 \ ?R2") +proof safe + assume ?L + from `?L` show ?R1 unfolding fin_support_def support_def + by (fastforce simp: image_iff elim: finite_surj[of _ _ "sum_case id undefined"]) + from `?L` show ?R2 unfolding fin_support_def support_def + by (fastforce simp: image_iff elim: finite_surj[of _ _ "sum_case undefined id"]) +next + assume ?R1 ?R2 + thus ?L unfolding fin_support_def support_Un + by (auto simp: support_def elim: finite_surj[of _ _ Inl] finite_surj[of _ _ Inr]) +qed + +lemma Func_upd: "\f \ Func A B; x \ A; y \ B\ \ f(x := y) \ Func A B" + unfolding Func_def by auto + +context wo_rel +begin + +definition isMaxim :: "'a set \ 'a \ bool" +where "isMaxim A b \ b \ A \ (\a \ A. (a,b) \ r)" + +definition maxim :: "'a set \ 'a" +where "maxim A \ THE b. isMaxim A b" + +lemma isMaxim_unique[intro]: "\isMaxim A x; isMaxim A y\ \ x = y" + unfolding isMaxim_def using antisymD[OF ANTISYM, of x y] by auto + +lemma maxim_isMaxim: "\finite A; A \ {}; A \ Field r\ \ isMaxim A (maxim A)" +unfolding maxim_def +proof (rule theI', rule ex_ex1I[OF _ isMaxim_unique, rotated], assumption+, + induct A rule: finite_induct) + case (insert x A) + thus ?case + proof (cases "A = {}") + case True + moreover have "isMaxim {x} x" unfolding isMaxim_def using refl_onD[OF REFL] insert(5) by auto + ultimately show ?thesis by blast + next + case False + with insert(3,5) obtain y where "isMaxim A y" by blast + with insert(2,5) have "if (y, x) \ r then isMaxim (insert x A) x else isMaxim (insert x A) y" + unfolding isMaxim_def subset_eq by (metis insert_iff max2_def max2_equals1 max2_iff) + thus ?thesis by metis + qed +qed simp + +lemma maxim_in: "\finite A; A \ {}; A \ Field r\ \ maxim A \ A" + using maxim_isMaxim unfolding isMaxim_def by auto + +lemma maxim_greatest: "\finite A; x \ A; A \ Field r\ \ (x, maxim A) \ r" + using maxim_isMaxim unfolding isMaxim_def by auto + +lemma isMaxim_zero: "isMaxim A zero \ A = {zero}" + unfolding isMaxim_def by auto + +lemma maxim_insert: + assumes "finite A" "A \ {}" "A \ Field r" "x \ Field r" + shows "maxim (insert x A) = max2 x (maxim A)" +proof - + from assms have *: "isMaxim (insert x A) (maxim (insert x A))" "isMaxim A (maxim A)" + using maxim_isMaxim by auto + show ?thesis + proof (cases "(x, maxim A) \ r") + case True + with *(2) have "isMaxim (insert x A) (maxim A)" unfolding isMaxim_def + using transD[OF TRANS, of _ x "maxim A"] by blast + with *(1) True show ?thesis unfolding max2_def by (metis isMaxim_unique) + next + case False + hence "(maxim A, x) \ r" by (metis *(2) assms(3,4) in_mono in_notinI isMaxim_def) + with *(2) assms(4) have "isMaxim (insert x A) x" unfolding isMaxim_def + using transD[OF TRANS, of _ "maxim A" x] refl_onD[OF REFL, of x] by blast + with *(1) False show ?thesis unfolding max2_def by (metis isMaxim_unique) + qed +qed + +lemma maxim_Un: + assumes "finite A" "A \ {}" "A \ Field r" "finite B" "B \ {}" "B \ Field r" + shows "maxim (A \ B) = max2 (maxim A) (maxim B)" +proof - + from assms have *: "isMaxim (A \ B) (maxim (A \ B))" "isMaxim A (maxim A)" "isMaxim B (maxim B)" + using maxim_isMaxim by auto + show ?thesis + proof (cases "(maxim A, maxim B) \ r") + case True + with *(2,3) have "isMaxim (A \ B) (maxim B)" unfolding isMaxim_def + using transD[OF TRANS, of _ "maxim A" "maxim B"] by blast + with *(1) True show ?thesis unfolding max2_def by (metis isMaxim_unique) + next + case False + hence "(maxim B, maxim A) \ r" by (metis *(2,3) assms(3,6) in_mono in_notinI isMaxim_def) + with *(2,3) have "isMaxim (A \ B) (maxim A)" unfolding isMaxim_def + using transD[OF TRANS, of _ "maxim B" "maxim A"] by blast + with *(1) False show ?thesis unfolding max2_def by (metis isMaxim_unique) + qed +qed + +lemma maxim_insert_zero: + assumes "finite A" "A \ {}" "A \ Field r" + shows "maxim (insert zero A) = maxim A" +using assms zero_in_Field maxim_in[OF assms] by (subst maxim_insert[unfolded max2_def]) auto + +lemma maxim_equality: "isMaxim A x \ maxim A = x" + unfolding maxim_def by (rule the_equality) auto + +lemma maxim_singleton: + "x \ Field r \ maxim {x} = x" + using refl_onD[OF REFL] by (intro maxim_equality) (simp add: isMaxim_def) + +lemma maxim_Int: "\finite A; A \ {}; A \ Field r; maxim A \ B\ \ maxim (A \ B) = maxim A" + by (rule maxim_equality) (auto simp: isMaxim_def intro: maxim_in maxim_greatest) + +lemma maxim_mono: "\X \ Y; finite Y; X \ {}; Y \ Field r\ \ (maxim X, maxim Y) \ r" + using maxim_in[OF finite_subset, of X Y] by (auto intro: maxim_greatest) + +definition "max_fun_diff f g \ maxim ({a \ Field r. f a \ g a})" + +lemma max_fun_diff_commute: "max_fun_diff f g = max_fun_diff g f" + unfolding max_fun_diff_def by metis + +lemma zero_under: "x \ Field r \ zero \ under x" + unfolding under_def by (auto intro: zero_smallest) + +end + +definition "FinFunc r s = Func (Field s) (Field r) \ fin_support (zero r) (Field s)" + +lemma FinFuncD: "\f \ FinFunc r s; x \ Field s\ \ f x \ Field r" + unfolding FinFunc_def Func_def by (fastforce split: option.splits) + +locale wo_rel2 = + fixes r s + assumes rWELL: "Well_order r" + and sWELL: "Well_order s" +begin + +interpretation r!: wo_rel r by unfold_locales (rule rWELL) +interpretation s!: wo_rel s by unfold_locales (rule sWELL) + +abbreviation "SUPP \ support r.zero (Field s)" +abbreviation "FINFUNC \ FinFunc r s" +lemmas FINFUNCD = FinFuncD[of _ r s] + +lemma fun_diff_alt: "{a \ Field s. f a \ g a} = (SUPP f \ SUPP g) \ {a. f a \ g a}" + by (auto simp: support_def) + +lemma max_fun_diff_alt: + "s.max_fun_diff f g = s.maxim ((SUPP f \ SUPP g) \ {a. f a \ g a})" + unfolding s.max_fun_diff_def fun_diff_alt .. + +lemma isMaxim_max_fun_diff: "\f \ g; f \ FINFUNC; g \ FINFUNC\ \ + s.isMaxim {a \ Field s. f a \ g a} (s.max_fun_diff f g)" + using fun_unequal_in_support[of f g] unfolding max_fun_diff_alt fun_diff_alt fun_eq_iff + by (intro s.maxim_isMaxim) (auto simp: FinFunc_def fin_support_def support_def) + +lemma max_fun_diff_in: "\f \ g; f \ FINFUNC; g \ FINFUNC\ \ + s.max_fun_diff f g \ {a \ Field s. f a \ g a}" + using isMaxim_max_fun_diff unfolding s.isMaxim_def by blast + +lemma max_fun_diff_max: "\f \ g; f \ FINFUNC; g \ FINFUNC; x \ {a \ Field s. f a \ g a}\ \ + (x, s.max_fun_diff f g) \ s" + using isMaxim_max_fun_diff unfolding s.isMaxim_def by blast + +lemma max_fun_diff: + "\f \ g; f \ FINFUNC; g \ FINFUNC\ \ + (\a b. a \ b \ a \ Field r \ b \ Field r \ + f (s.max_fun_diff f g) = a \ g (s.max_fun_diff f g) = b)" + using isMaxim_max_fun_diff[of f g] unfolding s.isMaxim_def FinFunc_def Func_def by auto + +lemma max_fun_diff_le_eq: + "\(s.max_fun_diff f g, x) \ s; f \ g; f \ FINFUNC; g \ FINFUNC; x \ s.max_fun_diff f g\ \ + f x = g x" + using max_fun_diff_max[of f g x] antisymD[OF s.ANTISYM, of "s.max_fun_diff f g" x] + by (auto simp: Field_def) + +lemma max_fun_diff_max2: + assumes ineq: "s.max_fun_diff f g = s.max_fun_diff g h \ + f (s.max_fun_diff f g) \ h (s.max_fun_diff g h)" and + fg: "f \ g" and gh: "g \ h" and fh: "f \ h" and + f: "f \ FINFUNC" and g: "g \ FINFUNC" and h: "h \ FINFUNC" + shows "s.max_fun_diff f h = s.max2 (s.max_fun_diff f g) (s.max_fun_diff g h)" + (is "?fh = s.max2 ?fg ?gh") +proof (cases "?fg = ?gh") + case True + with ineq have "f ?fg \ h ?fg" by simp + moreover + { fix x assume x: "x \ {a \ Field s. f a \ h a}" + hence "(x, ?fg) \ s" + proof (cases "x = ?fg") + case False show ?thesis + proof (rule ccontr) + assume "(x, ?fg) \ s" + with max_fun_diff_in[OF fg f g] x False have *: "(?fg, x) \ s" by (blast intro: s.in_notinI) + hence "f x = g x" by (rule max_fun_diff_le_eq[OF _ fg f g False]) + moreover have "g x = h x" using max_fun_diff_le_eq[OF _ gh g h] False True * by simp + ultimately show False using x by simp + qed + qed (simp add: refl_onD[OF s.REFL]) + } + ultimately have "s.isMaxim {a \ Field s. f a \ h a} ?fg" + unfolding s.isMaxim_def using max_fun_diff_in[OF fg f g] by simp + hence "?fh = ?fg" using isMaxim_max_fun_diff[OF fh f h] by blast + thus ?thesis unfolding True s.max2_def by simp +next + case False note * = this + show ?thesis + proof (cases "(?fg, ?gh) \ s") + case True + hence *: "f ?gh = g ?gh" by (rule max_fun_diff_le_eq[OF _ fg f g *[symmetric]]) + hence "s.isMaxim {a \ Field s. f a \ h a} ?gh" using isMaxim_max_fun_diff[OF gh g h] + isMaxim_max_fun_diff[OF fg f g] transD[OF s.TRANS _ True] + unfolding s.isMaxim_def by auto + hence "?fh = ?gh" using isMaxim_max_fun_diff[OF fh f h] by blast + thus ?thesis using True unfolding s.max2_def by simp + next + case False + with max_fun_diff_in[OF fg f g] max_fun_diff_in[OF gh g h] have True: "(?gh, ?fg) \ s" + by (blast intro: s.in_notinI) + hence *: "g ?fg = h ?fg" by (rule max_fun_diff_le_eq[OF _ gh g h *]) + hence "s.isMaxim {a \ Field s. f a \ h a} ?fg" using isMaxim_max_fun_diff[OF gh g h] + isMaxim_max_fun_diff[OF fg f g] True transD[OF s.TRANS, of _ _ ?fg] + unfolding s.isMaxim_def by auto + hence "?fh = ?fg" using isMaxim_max_fun_diff[OF fh f h] by blast + thus ?thesis using False unfolding s.max2_def by simp + qed +qed + + +definition oexp where + "oexp = {(f, g) . f \ FINFUNC \ g \ FINFUNC \ + ((let m = s.max_fun_diff f g in (f m, g m) \ r) \ f = g)}" + +lemma Field_oexp: "Field oexp = FINFUNC" + unfolding oexp_def FinFunc_def by (auto simp: Let_def Field_def) + +lemma oexp_Refl: "Refl oexp" + unfolding refl_on_def Field_oexp unfolding oexp_def by (auto simp: Let_def) + +lemma oexp_trans: "trans oexp" +proof (unfold trans_def, safe) + fix f g h :: "'b \ 'a" + let ?fg = "s.max_fun_diff f g" + and ?gh = "s.max_fun_diff g h" + and ?fh = "s.max_fun_diff f h" + assume oexp: "(f, g) \ oexp" "(g, h) \ oexp" + thus "(f, h) \ oexp" + proof (cases "f = g \ g = h") + case False + with oexp have "f \ FINFUNC" "g \ FINFUNC" "h \ FINFUNC" + "(f ?fg, g ?fg) \ r" "(g ?gh, h ?gh) \ r" unfolding oexp_def Let_def by auto + note * = this False + show ?thesis + proof (cases "f \ h") + case True + show ?thesis + proof (cases "?fg = ?gh \ f ?fg \ h ?gh") + case True + show ?thesis using max_fun_diff_max2[of f g h, OF True] * `f \ h` max_fun_diff_in + r.max2_iff[OF FINFUNCD FINFUNCD] r.max2_equals1[OF FINFUNCD FINFUNCD] max_fun_diff_le_eq + s.in_notinI[OF disjI1] unfolding oexp_def Let_def s.max2_def mem_Collect_eq by safe metis + next + case False with * show ?thesis unfolding oexp_def Let_def + using antisymD[OF r.ANTISYM, of "g ?gh" "h ?gh"] max_fun_diff_in[of g h] by auto + qed + qed (auto simp: oexp_def *(3)) + qed auto +qed + +lemma oexp_Preorder: "Preorder oexp" + unfolding preorder_on_def using oexp_Refl oexp_trans by blast + +lemma oexp_antisym: "antisym oexp" +proof (unfold antisym_def, safe, rule ccontr) + fix f g assume "(f, g) \ oexp" "(g, f) \ oexp" "g \ f" + thus False using refl_onD[OF r.REFL FINFUNCD] max_fun_diff_in unfolding oexp_def Let_def + by (auto dest!: antisymD[OF r.ANTISYM] simp: s.max_fun_diff_commute) +qed + +lemma oexp_Partial_order: "Partial_order oexp" + unfolding partial_order_on_def using oexp_Preorder oexp_antisym by blast + +lemma oexp_Total: "Total oexp" + unfolding total_on_def Field_oexp unfolding oexp_def using FINFUNCD max_fun_diff_in + by (auto simp: Let_def s.max_fun_diff_commute intro!: r.in_notinI) + +lemma oexp_Linear_order: "Linear_order oexp" + unfolding linear_order_on_def using oexp_Partial_order oexp_Total by blast + +definition "const = (\x. if x \ Field s then r.zero else undefined)" + +lemma const_in[simp]: "x \ Field s \ const x = r.zero" + unfolding const_def by auto + +lemma const_notin[simp]: "x \ Field s \ const x = undefined" + unfolding const_def by auto + +lemma const_Int_Field[simp]: "Field s \ - {x. const x = r.zero} = {}" + by auto + +lemma const_FINFUNC[simp]: "Field r \ {} \ const \ FINFUNC" + unfolding FinFunc_def Func_def fin_support_def support_def const_def Int_iff mem_Collect_eq + using r.zero_in_Field by (metis (lifting) Collect_empty_eq finite.emptyI) + +lemma const_least: + assumes "Field r \ {}" "f \ FINFUNC" + shows "(const, f) \ oexp" +proof (cases "f = const") + case True thus ?thesis using refl_onD[OF oexp_Refl] assms(2) unfolding Field_oexp by auto +next + case False + with assms show ?thesis using max_fun_diff_in[of f const] + unfolding oexp_def Let_def by (auto intro: r.zero_smallest FinFuncD simp: s.max_fun_diff_commute) +qed + +lemma support_not_const: + assumes "F \ FINFUNC" and "const \ F" + shows "\f \ F. finite (SUPP f) \ SUPP f \ {} \ SUPP f \ Field s" +proof (intro ballI conjI) + fix f assume "f \ F" + thus "finite (SUPP f)" "SUPP f \ Field s" + using assms(1) unfolding FinFunc_def fin_support_def support_def by auto + show "SUPP f \ {}" + proof (rule ccontr, unfold not_not) + assume "SUPP f = {}" + moreover from `f \ F` assms(1) have "f \ FINFUNC" by blast + ultimately have "f = const" + by (auto simp: fun_eq_iff support_def FinFunc_def Func_def const_def) + with assms(2) `f \ F` show False by blast + qed +qed + +lemma maxim_isMaxim_support: + assumes f: "F \ FINFUNC" and "const \ F" + shows "\f \ F. s.isMaxim (SUPP f) (s.maxim (SUPP f))" + using support_not_const[OF assms] by (auto intro!: s.maxim_isMaxim) + +lemma oexp_empty2: "Field s = {} \ oexp = {(\x. undefined, \x. undefined)}" + unfolding oexp_def FinFunc_def fin_support_def support_def by auto + +lemma oexp_empty: "\Field r = {}; Field s \ {}\ \ oexp = {}" + unfolding oexp_def FinFunc_def Let_def by auto + +lemma fun_upd_FINFUNC: "\f \ FINFUNC; x \ Field s; y \ Field r\ \ f(x := y) \ FINFUNC" + unfolding FinFunc_def Func_def fin_support_def + by (auto intro: finite_subset[OF support_upd_subset]) + +lemma fun_upd_same_oexp: + assumes "(f, g) \ oexp" "f x = g x" "x \ Field s" "y \ Field r" + shows "(f(x := y), g(x := y)) \ oexp" +proof - + from assms(1) fun_upd_FINFUNC[OF _ assms(3,4)] have fg: "f(x := y) \ FINFUNC" "g(x := y) \ FINFUNC" + unfolding oexp_def by auto + moreover from assms(2) have "s.max_fun_diff (f(x := y)) (g(x := y)) = s.max_fun_diff f g" + unfolding s.max_fun_diff_def by auto metis + ultimately show ?thesis using assms refl_onD[OF r.REFL] unfolding oexp_def Let_def by auto +qed + +lemma fun_upd_smaller_oexp: + assumes "f \ FINFUNC" "x \ Field s" "y \ Field r" "(y, f x) \ r" + shows "(f(x := y), f) \ oexp" + using assms fun_upd_FINFUNC[OF assms(1-3)] s.maxim_singleton[of "x"] + unfolding oexp_def FinFunc_def Let_def fin_support_def s.max_fun_diff_def by (auto simp: fun_eq_iff) + +lemma oexp_wf_Id: "wf (oexp - Id)" +proof (cases "Field r = {} \ Field s = {}") + case True thus ?thesis using oexp_empty oexp_empty2 by fastforce +next + case False + hence Fields: "Field s \ {}" "Field r \ {}" by simp_all + hence [simp]: "r.zero \ Field r" by (intro r.zero_in_Field) + have const[simp]: "\F. \const \ F; F \ FINFUNC\ \ \f0\F. \f\F. (f0, f) \ oexp" + using const_least[OF Fields(2)] by auto + show ?thesis + unfolding Linear_order_wf_diff_Id[OF oexp_Linear_order] Field_oexp + proof (intro allI impI) + fix A assume A: "A \ FINFUNC" "A \ {}" + { fix y F + have "F \ FINFUNC \ (\f \ F. y = s.maxim (SUPP f)) \ + (\f0 \ F. \f \ F. (f0, f) \ oexp)" (is "?P F y") + proof (induct y arbitrary: F rule: s.well_order_induct) + case (1 y) + show ?case + proof (intro impI, elim conjE bexE) + fix f assume F: "F \ FINFUNC" "f \ F" "y = s.maxim (SUPP f)" + thus "\f0\F. \f\F. (f0, f) \ oexp" + proof (cases "const \ F") + case False + with F have maxF: "\f \ F. s.isMaxim (SUPP f) (s.maxim (SUPP f))" + and SUPPF: "\f \ F. finite (SUPP f) \ SUPP f \ {} \ SUPP f \ Field s" + using maxim_isMaxim_support support_not_const by auto + def z \ "s.minim {s.maxim (SUPP f) | f. f \ F}" + from F SUPPF maxF have zmin: "s.isMinim {s.maxim (SUPP f) | f. f \ F} z" + unfolding z_def by (intro s.minim_isMinim) (auto simp: s.isMaxim_def) + with F have zy: "(z, y) \ s" unfolding s.isMinim_def by auto + hence zField: "z \ Field s" unfolding Field_def by auto + def x0 \ "r.minim {f z | f. f \ F \ z = s.maxim (SUPP f)}" + from F(1,2) maxF(1) SUPPF zmin + have x0min: "r.isMinim {f z | f. f \ F \ z = s.maxim (SUPP f)} x0" + unfolding x0_def s.isMaxim_def s.isMinim_def + by (blast intro!: r.minim_isMinim FinFuncD[of _ r s]) + with maxF(1) SUPPF F(1) have x0Field: "x0 \ Field r" + unfolding r.isMinim_def s.isMaxim_def by (auto intro!: FINFUNCD) + from x0min maxF(1) SUPPF F(1) have x0notzero: "x0 \ r.zero" + unfolding r.isMinim_def s.isMaxim_def FinFunc_def Func_def support_def + by fastforce + def G \ "{f(z := r.zero) | f. f \ F \ z = s.maxim (SUPP f) \ f z = x0}" + from zmin x0min have "G \ {}" unfolding G_def z_def s.isMinim_def r.isMinim_def by blast + have GF: "G \ (\f. f(z := r.zero)) ` F" unfolding G_def by auto + have "G \ fin_support r.zero (Field s)" + unfolding FinFunc_def fin_support_def proof safe + fix g assume "g \ G" + with GF obtain f where "f \ F" "g = f(z := r.zero)" by auto + moreover with SUPPF have "finite (SUPP f)" by blast + ultimately show "finite (SUPP g)" + by (elim finite_subset[rotated]) (auto simp: support_def) + qed + moreover from F GF zField have "G \ Func (Field s) (Field r)" + using Func_upd[of _ "Field s" "Field r" z r.zero] unfolding FinFunc_def by auto + ultimately have G: "G \ FINFUNC" unfolding FinFunc_def by blast + hence "\g0\G. \g\G. (g0, g) \ oexp" + proof (cases "const \ G") + case False + with G have maxG: "\g \ G. s.isMaxim (SUPP g) (s.maxim (SUPP g))" + and SUPPG: "\g \ G. finite (SUPP g) \ SUPP g \ {} \ SUPP g \ Field s" + using maxim_isMaxim_support support_not_const by auto + def y' \ "s.minim {s.maxim (SUPP f) | f. f \ G}" + from G SUPPG maxG `G \ {}` have y'min: "s.isMinim {s.maxim (SUPP f) | f. f \ G} y'" + unfolding y'_def by (intro s.minim_isMinim) (auto simp: s.isMaxim_def) + moreover + have "\g \ G. z \ SUPP g" unfolding support_def G_def by auto + moreover + { fix g assume g: "g \ G" + then obtain f where "f \ F" "g = f(z := r.zero)" and z: "z = s.maxim (SUPP f)" + unfolding G_def by auto + with SUPPF bspec[OF SUPPG g] have "(s.maxim (SUPP g), z) \ s" + unfolding z by (intro s.maxim_mono) auto + } + moreover from y'min have "\g. g \ G \ (y', s.maxim (SUPP g)) \ s" + unfolding s.isMinim_def by auto + ultimately have "y' \ z" "(y', z) \ s" using maxG + unfolding s.isMinim_def s.isMaxim_def by auto + with zy have "y' \ y" "(y', y) \ s" using antisymD[OF s.ANTISYM] transD[OF s.TRANS] + by blast+ + moreover from `G \ {}` have "\g \ G. y' = wo_rel.maxim s (SUPP g)" using y'min + by (auto simp: G_def s.isMinim_def) + ultimately show ?thesis using mp[OF spec[OF mp[OF spec[OF 1]]], of y' G] G by auto + qed simp + then obtain g0 where g0: "g0 \ G" "\g \ G. (g0, g) \ oexp" by blast + hence g0z: "g0 z = r.zero" unfolding G_def by auto + def f0 \ "g0(z := x0)" + with x0notzero zField have SUPP: "SUPP f0 = SUPP g0 \ {z}" unfolding support_def by auto + from g0z have f0z: "f0(z := r.zero) = g0" unfolding f0_def fun_upd_upd by auto + have f0: "f0 \ F" using x0min g0(1) + Func_elim[OF set_mp[OF subset_trans[OF F(1)[unfolded FinFunc_def] Int_lower1]] zField] + unfolding f0_def r.isMinim_def G_def by (force simp: fun_upd_idem) + from g0(1) maxF(1) have maxf0: "s.maxim (SUPP f0) = z" unfolding SUPP G_def + by (intro s.maxim_equality) (auto simp: s.isMaxim_def) + show ?thesis + proof (intro bexI[OF _ f0] ballI) + fix f assume f: "f \ F" + show "(f0, f) \ oexp" + proof (cases "f0 = f") + case True thus ?thesis by (metis F(1) Field_oexp f0 in_mono oexp_Refl refl_onD) + next + case False + thus ?thesis + proof (cases "s.maxim (SUPP f) = z \ f z = x0") + case True + with f have "f(z := r.zero) \ G" unfolding G_def by blast + with g0(2) f0z have "(f0(z := r.zero), f(z := r.zero)) \ oexp" by auto + hence "(f0(z := r.zero, z := x0), f(z := r.zero, z := x0)) \ oexp" + by (elim fun_upd_same_oexp[OF _ _ zField x0Field]) simp + moreover + with f F(1) x0min True + have "(f(z := x0), f) \ oexp" unfolding G_def r.isMinim_def + by (intro fun_upd_smaller_oexp[OF _ zField x0Field]) auto + ultimately show ?thesis using transD[OF oexp_trans, of f0 "f(z := x0)" f] + unfolding f0_def by auto + next + case False note notG = this + thus ?thesis + proof (cases "s.maxim (SUPP f) = z") + case True + with notG have "f0 z \ f z" unfolding f0_def by auto + hence "f0 z \ f z" by metis + with True maxf0 f0 f SUPPF have "s.max_fun_diff f0 f = z" + using s.maxim_Un[of "SUPP f0" "SUPP f", unfolded s.max2_def] + unfolding max_fun_diff_alt by (intro trans[OF s.maxim_Int]) auto + moreover + from x0min True f have "(x0, f z) \ r" unfolding r.isMinim_def by auto + ultimately show ?thesis using f f0 F(1) unfolding oexp_def f0_def by auto + next + case False + with notG have *: "(z, s.maxim (SUPP f)) \ s" "z \ s.maxim (SUPP f)" + using zmin f unfolding s.isMinim_def G_def by auto + have f0f: "f0 (s.maxim (SUPP f)) = r.zero" + proof (rule ccontr) + assume "f0 (s.maxim (SUPP f)) \ r.zero" + with f SUPPF maxF(1) have "s.maxim (SUPP f) \ SUPP f0" + unfolding support_def[of _ _ f0] s.isMaxim_def by auto + with SUPPF f0 have "(s.maxim (SUPP f), z) \ s" unfolding maxf0[symmetric] + by (auto intro: s.maxim_greatest) + with * antisymD[OF s.ANTISYM] show False by simp + qed + moreover + have "f (s.maxim (SUPP f)) \ r.zero" + using bspec[OF maxF(1) f, unfolded s.isMaxim_def] by (auto simp: support_def) + with f0f * f f0 maxf0 SUPPF + have "s.max_fun_diff f0 f = s.maxim (SUPP f0 \ SUPP f)" + unfolding max_fun_diff_alt using s.maxim_Un[of "SUPP f0" "SUPP f"] + by (intro s.maxim_Int) (auto simp: s.max2_def) + moreover have "s.maxim (SUPP f0 \ SUPP f) = s.maxim (SUPP f)" + using s.maxim_Un[of "SUPP f0" "SUPP f"] * maxf0 SUPPF f0 f + by (auto simp: s.max2_def) + ultimately show ?thesis using f f0 F(1) maxF(1) SUPPF unfolding oexp_def Let_def + by (fastforce simp: s.isMaxim_def intro!: r.zero_smallest FINFUNCD) + qed + qed + qed + qed + qed simp + qed + qed + } note * = mp[OF this] + from A(2) obtain f where f: "f \ A" by blast + with A(1) show "\a\A. \a'\A. (a, a') \ oexp" + proof (cases "f = const") + case False with f A(1) show ?thesis using maxim_isMaxim_support[of "{f}"] + by (intro *[of _ "s.maxim (SUPP f)"]) (auto simp: s.isMaxim_def support_def) + qed simp + qed +qed + +lemma oexp_Well_order: "Well_order oexp" + unfolding well_order_on_def using oexp_Linear_order oexp_wf_Id by blast + +interpretation o: wo_rel oexp by unfold_locales (rule oexp_Well_order) + +lemma zero_oexp: "Field r \ {} \ o.zero = const" + by (rule sym[OF o.leq_zero_imp[OF const_least]]) + (auto intro!: o.zero_in_Field[unfolded Field_oexp] dest!: const_FINFUNC) + +end + +notation wo_rel2.oexp (infixl "^o" 90) +lemmas oexp_def = wo_rel2.oexp_def[unfolded wo_rel2_def, OF conjI] +lemmas oexp_Well_order = wo_rel2.oexp_Well_order[unfolded wo_rel2_def, OF conjI] +lemmas Field_oexp = wo_rel2.Field_oexp[unfolded wo_rel2_def, OF conjI] + +definition "ozero = {}" + +lemma ozero_Well_order[simp]: "Well_order ozero" + unfolding ozero_def by simp + +lemma ozero_ordIso[simp]: "ozero =o ozero" + unfolding ozero_def ordIso_def iso_def[abs_def] embed_def bij_betw_def by auto + +lemma Field_ozero[simp]: "Field ozero = {}" + unfolding ozero_def by simp + +lemma iso_ozero_empty[simp]: "r =o ozero = (r = {})" + unfolding ozero_def ordIso_def iso_def[abs_def] embed_def bij_betw_def + by (auto dest: rel.well_order_on_domain) + +lemma ozero_ordLeq: +assumes "Well_order r" shows "ozero \o r" +using assms unfolding ozero_def ordLeq_def embed_def[abs_def] rel.under_def by auto + +definition "oone = {((),())}" + +lemma oone_Well_order[simp]: "Well_order oone" + unfolding oone_def unfolding well_order_on_def linear_order_on_def partial_order_on_def + preorder_on_def total_on_def refl_on_def trans_def antisym_def by auto + +lemma Field_oone[simp]: "Field oone = {()}" + unfolding oone_def by simp + +lemma oone_ordIso: "oone =o {(x,x)}" + unfolding ordIso_def oone_def well_order_on_def linear_order_on_def partial_order_on_def + preorder_on_def total_on_def refl_on_def trans_def antisym_def + by (auto simp: iso_def embed_def bij_betw_def rel.under_def inj_on_def intro!: exI[of _ "\_. x"]) + +lemma osum_ordLeqR: "Well_order r \ Well_order s \ s \o r +o s" + unfolding ordLeq_def2 rel.underS_def + by (auto intro!: exI[of _ Inr] osum_Well_order) (auto simp add: osum_def Field_def) + +lemma osum_congL: + assumes "r =o s" and t: "Well_order t" + shows "r +o t =o s +o t" (is "?L =o ?R") +proof - + from assms(1) obtain f where r: "Well_order r" and s: "Well_order s" and f: "iso r s f" + unfolding ordIso_def by blast + let ?f = "sum_map f id" + from f have "inj_on ?f (Field ?L)" + unfolding Field_osum iso_def bij_betw_def inj_on_def by fastforce + with f have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding Field_osum iso_def bij_betw_def image_image image_Un by auto + moreover from f have "compat ?L ?R ?f" + unfolding osum_def iso_iff3[OF r s] compat_def bij_betw_def + by (auto simp: map_pair_imageI) + ultimately have "iso ?L ?R ?f" by (subst iso_iff3) (auto intro: osum_Well_order r s t) + thus ?thesis unfolding ordIso_def by (auto intro: osum_Well_order r s t) +qed + +lemma osum_congR: + assumes "r =o s" and t: "Well_order t" + shows "t +o r =o t +o s" (is "?L =o ?R") +proof - + from assms(1) obtain f where r: "Well_order r" and s: "Well_order s" and f: "iso r s f" + unfolding ordIso_def by blast + let ?f = "sum_map id f" + from f have "inj_on ?f (Field ?L)" + unfolding Field_osum iso_def bij_betw_def inj_on_def by fastforce + with f have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding Field_osum iso_def bij_betw_def image_image image_Un by auto + moreover from f have "compat ?L ?R ?f" + unfolding osum_def iso_iff3[OF r s] compat_def bij_betw_def + by (auto simp: map_pair_imageI) + ultimately have "iso ?L ?R ?f" by (subst iso_iff3) (auto intro: osum_Well_order r s t) + thus ?thesis unfolding ordIso_def by (auto intro: osum_Well_order r s t) +qed + +lemma osum_cong: + assumes "t =o u" and "r =o s" + shows "t +o r =o u +o s" +using ordIso_transitive[OF osum_congL[OF assms(1)] osum_congR[OF assms(2)]] + assms[unfolded ordIso_def] by auto + +lemma Well_order_empty[simp]: "Well_order {}" + unfolding Field_empty by (rule well_order_on_empty) + +lemma well_order_on_singleton[simp]: "well_order_on {x} {(x, x)}" + unfolding well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def total_on_def + Field_def refl_on_def trans_def antisym_def by auto + +lemma oexp_empty[simp]: + assumes "Well_order r" + shows "r ^o {} = {(\x. undefined, \x. undefined)}" + unfolding oexp_def[OF assms Well_order_empty] FinFunc_def fin_support_def support_def by auto + +lemma oexp_empty2[simp]: + assumes "Well_order r" "r \ {}" + shows "{} ^o r = {}" +proof - + from assms(2) have "Field r \ {}" unfolding Field_def by auto + thus ?thesis unfolding oexp_def[OF Well_order_empty assms(1)] FinFunc_def fin_support_def support_def + by auto +qed + +lemma oprod_zero[simp]: "{} *o r = {}" "r *o {} = {}" + unfolding oprod_def by simp_all + +lemma oprod_congL: + assumes "r =o s" and t: "Well_order t" + shows "r *o t =o s *o t" (is "?L =o ?R") +proof - + from assms(1) obtain f where r: "Well_order r" and s: "Well_order s" and f: "iso r s f" + unfolding ordIso_def by blast + let ?f = "map_pair f id" + from f have "inj_on ?f (Field ?L)" + unfolding Field_oprod iso_def bij_betw_def inj_on_def by fastforce + with f have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding Field_oprod iso_def bij_betw_def by (auto intro!: map_pair_surj_on) + moreover from f have "compat ?L ?R ?f" + unfolding iso_iff3[OF r s] compat_def oprod_def bij_betw_def + by (auto simp: map_pair_imageI) + ultimately have "iso ?L ?R ?f" by (subst iso_iff3) (auto intro: oprod_Well_order r s t) + thus ?thesis unfolding ordIso_def by (auto intro: oprod_Well_order r s t) +qed + +lemma oprod_congR: + assumes "r =o s" and t: "Well_order t" + shows "t *o r =o t *o s" (is "?L =o ?R") +proof - + from assms(1) obtain f where r: "Well_order r" and s: "Well_order s" and f: "iso r s f" + unfolding ordIso_def by blast + let ?f = "map_pair id f" + from f have "inj_on ?f (Field ?L)" + unfolding Field_oprod iso_def bij_betw_def inj_on_def by fastforce + with f have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding Field_oprod iso_def bij_betw_def by (auto intro!: map_pair_surj_on) + moreover from f rel.well_order_on_domain[OF r] have "compat ?L ?R ?f" + unfolding iso_iff3[OF r s] compat_def oprod_def bij_betw_def + by (auto simp: map_pair_imageI dest: inj_onD) + ultimately have "iso ?L ?R ?f" by (subst iso_iff3) (auto intro: oprod_Well_order r s t) + thus ?thesis unfolding ordIso_def by (auto intro: oprod_Well_order r s t) +qed + +lemma oprod_cong: + assumes "t =o u" and "r =o s" + shows "t *o r =o u *o s" +using ordIso_transitive[OF oprod_congL[OF assms(1)] oprod_congR[OF assms(2)]] + assms[unfolded ordIso_def] by auto + +lemma Field_singleton[simp]: "Field {(z,z)} = {z}" + by (metis rel.well_order_on_Field well_order_on_singleton) + +lemma zero_singleton[simp]: "zero {(z,z)} = z" + using wo_rel.zero_in_Field[unfolded wo_rel_def, of "{(z, z)}"] well_order_on_singleton[of z] + by auto + +lemma FinFunc_singleton: "FinFunc {(z,z)} s = {\x. if x \ Field s then z else undefined}" + unfolding FinFunc_def Func_def fin_support_def support_def + by (auto simp: fun_eq_iff split: split_if_asm intro!: finite_subset[of _ "{}"]) + +lemma oone_ordIso_oexp: + assumes "r =o oone" and s: "Well_order s" + shows "r ^o s =o oone" (is "?L =o ?R") +proof - + from assms obtain f where *: "\x\Field r. \y\Field r. x = y" and "f ` Field r = {()}" + and r: "Well_order r" + unfolding ordIso_def oone_def by (auto simp: iso_def bij_betw_def inj_on_def) + then obtain x where "x \ Field r" by auto + with * have Fr: "Field r = {x}" by auto + interpret r: wo_rel r by unfold_locales (rule r) + from Fr r.well_order_on_domain[OF r] refl_onD[OF r.REFL, of x] have r_def: "r = {(x, x)}" by fast + interpret wo_rel2 r s by unfold_locales (rule r, rule s) + have "bij_betw (\x. ()) (Field ?L) (Field ?R)" + unfolding bij_betw_def Field_oexp by (auto simp: r_def FinFunc_singleton) + moreover have "compat ?L ?R (\x. ())" unfolding compat_def oone_def by auto + ultimately have "iso ?L ?R (\x. ())" using s oone_Well_order + by (subst iso_iff3) (auto intro: oexp_Well_order) + thus ?thesis using s oone_Well_order unfolding ordIso_def by (auto intro: oexp_Well_order) +qed + +(*Lemma 1.4.3 from Holz et al.*) +context + fixes r s t + assumes r: "Well_order r" + assumes s: "Well_order s" + assumes t: "Well_order t" +begin + +lemma osum_ozeroL: "ozero +o r =o r" + using r unfolding osum_def ozero_def by (auto intro: map_pair_ordIso) + +lemma osum_ozeroR: "r +o ozero =o r" + using r unfolding osum_def ozero_def by (auto intro: map_pair_ordIso) + +lemma osum_assoc: "(r +o s) +o t =o r +o s +o t" (is "?L =o ?R") +proof - + let ?f = + "\rst. case rst of Inl (Inl r) \ Inl r | Inl (Inr s) \ Inr (Inl s) | Inr t \ Inr (Inr t)" + have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding Field_osum bij_betw_def inj_on_def by (auto simp: image_Un image_iff) + moreover + have "compat ?L ?R ?f" + proof (unfold compat_def, safe) + fix a b + assume "(a, b) \ ?L" + thus "(?f a, ?f b) \ ?R" + unfolding osum_def[of "r +o s" t] osum_def[of r "s +o t"] Field_osum + unfolding osum_def Field_osum image_iff image_Un map_pair_def + by fastforce + qed + ultimately have "iso ?L ?R ?f" using r s t by (subst iso_iff3) (auto intro: osum_Well_order) + thus ?thesis using r s t unfolding ordIso_def by (auto intro: osum_Well_order) +qed + +lemma osum_monoR: + assumes "s Field t" + using embed_iff_compat_inj_on_ofilter[OF s t, of f] embedS_iff[OF s, of t f] + unfolding embedS_def by auto + let ?f = "sum_map id f" + from *(1) have "inj_on ?f (Field ?L)" unfolding Field_osum inj_on_def by fastforce + moreover + from *(2,4) have "compat ?L ?R ?f" unfolding compat_def osum_def map_pair_def by fastforce + moreover + interpret t!: wo_rel t by unfold_locales (rule t) + interpret rt!: wo_rel ?R by unfold_locales (rule osum_Well_order[OF r t]) + from *(3) have "ofilter ?R (?f ` Field ?L)" + unfolding t.ofilter_def rt.ofilter_def Field_osum image_Un image_image rel.under_def + by (auto simp: osum_def intro!: imageI) (auto simp: Field_def) + ultimately have "embed ?L ?R ?f" using embed_iff_compat_inj_on_ofilter[of ?L ?R ?f] + by (auto intro: osum_Well_order r s t) + moreover + from *(4) have "?f ` Field ?L \ Field ?R" unfolding Field_osum image_Un image_image by auto + ultimately have "embedS ?L ?R ?f" using embedS_iff[OF osum_Well_order[OF r s], of ?R ?f] by auto + thus ?thesis unfolding ordLess_def by (auto intro: osum_Well_order r s t) +qed + +lemma osum_monoL: + assumes "r \o s" + shows "r +o t \o s +o t" +proof - + from assms obtain f where f: "\a\Field r. f a \ Field s \ f ` underS r a \ underS s (f a)" + unfolding ordLeq_def2 by blast + let ?f = "sum_map f id" + from f have "\a\Field (r +o t). + ?f a \ Field (s +o t) \ ?f ` underS (r +o t) a \ underS (s +o t) (?f a)" + unfolding Field_osum rel.underS_def by (fastforce simp: osum_def) + thus ?thesis unfolding ordLeq_def2 by (auto intro: osum_Well_order r s t) +qed + +lemma oprod_ozeroL: "ozero *o r =o ozero" + using ozero_ordIso unfolding ozero_def by simp + +lemma oprod_ozeroR: "r *o ozero =o ozero" + using ozero_ordIso unfolding ozero_def by simp + +lemma oprod_ooneR: "r *o oone =o r" (is "?L =o ?R") +proof - + have "bij_betw fst (Field ?L) (Field ?R)" unfolding Field_oprod bij_betw_def inj_on_def by simp + moreover have "compat ?L ?R fst" unfolding compat_def oprod_def by auto + ultimately have "iso ?L ?R fst" using r oone_Well_order + by (subst iso_iff3) (auto intro: oprod_Well_order) + thus ?thesis using r oone_Well_order unfolding ordIso_def by (auto intro: oprod_Well_order) +qed + +lemma oprod_ooneL: "oone *o r =o r" (is "?L =o ?R") +proof - + have "bij_betw snd (Field ?L) (Field ?R)" unfolding Field_oprod bij_betw_def inj_on_def by simp + moreover have "Refl r" by (rule wo_rel.REFL[unfolded wo_rel_def, OF r]) + hence "compat ?L ?R snd" unfolding compat_def oprod_def refl_on_def by auto + ultimately have "iso ?L ?R snd" using r oone_Well_order + by (subst iso_iff3) (auto intro: oprod_Well_order) + thus ?thesis using r oone_Well_order unfolding ordIso_def by (auto intro: oprod_Well_order) +qed + +lemma oprod_monoR: + assumes "ozero Field t" + using embed_iff_compat_inj_on_ofilter[OF s t, of f] embedS_iff[OF s, of t f] + unfolding embedS_def by auto + let ?f = "map_pair id f" + from *(1) have "inj_on ?f (Field ?L)" unfolding Field_oprod inj_on_def by fastforce + moreover + from *(2,4) the_inv_into_f_f[OF *(1)] have "compat ?L ?R ?f" unfolding compat_def oprod_def + by auto (metis rel.well_order_on_domain t, metis rel.well_order_on_domain s) + moreover + interpret t!: wo_rel t by unfold_locales (rule t) + interpret rt!: wo_rel ?R by unfold_locales (rule oprod_Well_order[OF r t]) + from *(3) have "ofilter ?R (?f ` Field ?L)" + unfolding t.ofilter_def rt.ofilter_def Field_oprod rel.under_def + by (auto simp: oprod_def image_iff) (fast | metis r rel.well_order_on_domain)+ + ultimately have "embed ?L ?R ?f" using embed_iff_compat_inj_on_ofilter[of ?L ?R ?f] + by (auto intro: oprod_Well_order r s t) + moreover + from not_ordLess_ordIso[OF assms(1)] have "r \ {}" by (metis ozero_def ozero_ordIso) + hence "Field r \ {}" unfolding Field_def by auto + with *(4) have "?f ` Field ?L \ Field ?R" unfolding Field_oprod + by auto (metis SigmaD2 SigmaI map_pair_surj_on) + ultimately have "embedS ?L ?R ?f" using embedS_iff[OF oprod_Well_order[OF r s], of ?R ?f] by auto + thus ?thesis unfolding ordLess_def by (auto intro: oprod_Well_order r s t) +qed + +lemma oprod_monoL: + assumes "r \o s" + shows "r *o t \o s *o t" +proof - + from assms obtain f where f: "\a\Field r. f a \ Field s \ f ` underS r a \ underS s (f a)" + unfolding ordLeq_def2 by blast + let ?f = "map_pair f id" + from f have "\a\Field (r *o t). + ?f a \ Field (s *o t) \ ?f ` underS (r *o t) a \ underS (s *o t) (?f a)" + unfolding Field_oprod rel.underS_def unfolding map_pair_def oprod_def by auto + thus ?thesis unfolding ordLeq_def2 by (auto intro: oprod_Well_order r s t) +qed + +lemma oprod_assoc: "(r *o s) *o t =o r *o s *o t" (is "?L =o ?R") +proof - + let ?f = "\((a,b),c). (a,b,c)" + have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding Field_oprod bij_betw_def inj_on_def by (auto simp: image_Un image_iff) + moreover + have "compat ?L ?R ?f" + proof (unfold compat_def, safe) + fix a1 a2 a3 b1 b2 b3 + assume "(((a1, a2), a3), ((b1, b2), b3)) \ ?L" + thus "((a1, a2, a3), (b1, b2, b3)) \ ?R" + unfolding oprod_def[of "r *o s" t] oprod_def[of r "s *o t"] Field_oprod + unfolding oprod_def Field_oprod image_iff image_Un by fast + qed + ultimately have "iso ?L ?R ?f" using r s t by (subst iso_iff3) (auto intro: oprod_Well_order) + thus ?thesis using r s t unfolding ordIso_def by (auto intro: oprod_Well_order) +qed + +lemma oprod_osum: "r *o (s +o t) =o r *o s +o r *o t" (is "?L =o ?R") +proof - + let ?f = "\(a,bc). case bc of Inl b \ Inl (a, b) | Inr c \ Inr (a, c)" + have "bij_betw ?f (Field ?L) (Field ?R)" unfolding Field_oprod Field_osum bij_betw_def inj_on_def + by (fastforce simp: image_Un image_iff split: sum.splits) + moreover + have "compat ?L ?R ?f" + proof (unfold compat_def, intro allI impI) + fix a b + assume "(a, b) \ ?L" + thus "(?f a, ?f b) \ ?R" + unfolding oprod_def[of r "s +o t"] osum_def[of "r *o s" "r *o t"] Field_oprod Field_osum + unfolding oprod_def osum_def Field_oprod Field_osum image_iff image_Un by auto + qed + ultimately have "iso ?L ?R ?f" using r s t + by (subst iso_iff3) (auto intro: oprod_Well_order osum_Well_order) + thus ?thesis using r s t unfolding ordIso_def by (auto intro: oprod_Well_order osum_Well_order) +qed + +lemma ozero_oexp: "\ (s =o ozero) \ ozero ^o s =o ozero" + unfolding oexp_def[OF ozero_Well_order s] FinFunc_def + by simp (metis Func_emp2 bot.extremum_uniqueI emptyE rel.well_order_on_domain s subrelI) + +lemma oone_oexp: "oone ^o s =o oone" (is "?L =o ?R") + by (rule oone_ordIso_oexp[OF ordIso_reflexive[OF oone_Well_order] s]) + +lemma oexp_monoR: + assumes "oone {}" by (metis assms(1) internalize_ordLess not_psubset_empty) + moreover + { assume "Field r = {r.zero}" + hence "r = {(r.zero, r.zero)}" using refl_onD[OF r.REFL, of r.zero] unfolding Field_def by auto + hence "r =o oone" by (metis oone_ordIso ordIso_symmetric) + with not_ordLess_ordIso[OF assms(1)] have False by (metis ordIso_symmetric) + } + ultimately obtain x where x: "x \ Field r" "r.zero \ Field r" "x \ r.zero" + by (metis insert_iff r.zero_in_Field subsetI subset_singletonD) + moreover from assms(2) obtain f where "embedS s t f" unfolding ordLess_def by blast + hence *: "inj_on f (Field s)" "compat s t f" "ofilter t (f ` Field s)" "f ` Field s \ Field t" + using embed_iff_compat_inj_on_ofilter[OF s t, of f] embedS_iff[OF s, of t f] + unfolding embedS_def by auto + note invff = the_inv_into_f_f[OF *(1)] and injfD = inj_onD[OF *(1)] + def F \ "\g z. if z \ f ` Field s then g (the_inv_into (Field s) f z) else + if z \ Field t then r.zero else undefined" + from *(4) x(2) the_inv_into_f_eq[OF *(1)] have FLR: "F ` Field ?L \ Field ?R" + unfolding rt.Field_oexp rs.Field_oexp FinFunc_def Func_def fin_support_def support_def F_def + by (fastforce split: option.splits split_if_asm elim!: finite_surj[of _ _ f]) + have "inj_on F (Field ?L)" unfolding rs.Field_oexp inj_on_def fun_eq_iff + proof safe + fix g h x assume "g \ FinFunc r s" "h \ FinFunc r s" "\y. F g y = F h y" + with invff show "g x = h x" unfolding F_def fun_eq_iff FinFunc_def Func_def + by auto (metis image_eqI) + qed + moreover + have "compat ?L ?R F" unfolding compat_def rs.oexp_def rt.oexp_def + proof (safe elim!: bspec[OF iffD1[OF image_subset_iff FLR[unfolded rs.Field_oexp rt.Field_oexp]]]) + fix g h assume gh: "g \ FinFunc r s" "h \ FinFunc r s" "F g \ F h" + "let m = s.max_fun_diff g h in (g m, h m) \ r" + hence "g \ h" by auto + note max_fun_diff_in = rs.max_fun_diff_in[OF `g \ h` gh(1,2)] + and max_fun_diff_max = rs.max_fun_diff_max[OF `g \ h` gh(1,2)] + with *(4) invff *(2) have "t.max_fun_diff (F g) (F h) = f (s.max_fun_diff g h)" + unfolding t.max_fun_diff_def compat_def + by (intro t.maxim_equality) (auto simp: t.isMaxim_def F_def dest: injfD) + with gh invff max_fun_diff_in + show "let m = t.max_fun_diff (F g) (F h) in (F g m, F h m) \ r" + unfolding F_def Let_def by (auto simp: dest: injfD) + qed + moreover + from FLR have "ofilter ?R (F ` Field ?L)" + unfolding rexpt.ofilter_def rel.under_def rs.Field_oexp rt.Field_oexp unfolding rt.oexp_def + proof (safe elim!: imageI) + fix g h assume gh: "g \ FinFunc r s" "h \ FinFunc r t" "F g \ FinFunc r t" + "let m = t.max_fun_diff h (F g) in (h m, F g m) \ r" + thus "h \ F ` FinFunc r s" + proof (cases "h = F g") + case False + hence max_Field: "t.max_fun_diff h (F g) \ {a \ Field t. h a \ F g a}" + by (rule rt.max_fun_diff_in[OF _ gh(2,3)]) + { assume *: "t.max_fun_diff h (F g) \ f ` Field s" + with max_Field have "F g (t.max_fun_diff h (F g)) = r.zero" unfolding F_def by auto + moreover + with * gh(4) have "h (t.max_fun_diff h (F g)) = r.zero" unfolding Let_def by auto + ultimately have False using max_Field gh(2,3) unfolding FinFunc_def Func_def by auto + } + hence max_f_Field: "t.max_fun_diff h (F g) \ f ` Field s" by blast + { fix z assume z: "z \ Field t - f ` Field s" + have "(t.max_fun_diff h (F g), z) \ t" + proof (rule ccontr) + assume "(t.max_fun_diff h (F g), z) \ t" + hence "(z, t.max_fun_diff h (F g)) \ t" using t.in_notinI[of "t.max_fun_diff h (F g)" z] + z max_Field by auto + hence "z \ f ` Field s" using *(3) max_f_Field unfolding t.ofilter_def rel.under_def + by fastforce + with z show False by blast + qed + hence "h z = r.zero" using rt.max_fun_diff_le_eq[OF _ False gh(2,3), of z] + z max_f_Field unfolding F_def by auto + } note ** = this + with *(3) gh(2) have "h = F (\x. if x \ Field s then h (f x) else undefined)" using invff + unfolding F_def fun_eq_iff FinFunc_def Func_def Let_def t.ofilter_def rel.under_def by auto + moreover from gh(2) *(1,3) have "(\x. if x \ Field s then h (f x) else undefined) \ FinFunc r s" + unfolding FinFunc_def Func_def fin_support_def support_def t.ofilter_def rel.under_def + by (auto intro: subset_inj_on elim!: finite_imageD[OF finite_subset[rotated]]) + ultimately show "?thesis" by (rule image_eqI) + qed simp + qed + ultimately have "embed ?L ?R F" using embed_iff_compat_inj_on_ofilter[of ?L ?R F] + by (auto intro: oexp_Well_order r s t) + moreover + from FLR have "F ` Field ?L \ Field ?R" + proof (intro psubsetI) + from *(4) obtain z where z: "z \ Field t" "z \ f ` Field s" by auto + def h \ "\z'. if z' \ Field t then if z' = z then x else r.zero else undefined" + from z x(3) have "rt.SUPP h = {z}" unfolding support_def h_def by simp + with x have "h \ Field ?R" unfolding h_def rt.Field_oexp FinFunc_def Func_def fin_support_def + by auto + moreover + { fix g + from z have "F g z = r.zero" "h z = x" unfolding support_def h_def F_def by auto + with x(3) have "F g \ h" unfolding fun_eq_iff by fastforce + } + hence "h \ F ` Field ?L" by blast + ultimately show "F ` Field ?L \ Field ?R" by blast + qed + ultimately have "embedS ?L ?R F" using embedS_iff[OF rs.oexp_Well_order, of ?R F] by auto + thus ?thesis unfolding ordLess_def using r s t by (auto intro: oexp_Well_order) +qed + +lemma oexp_monoL: + assumes "r \o s" + shows "r ^o t \o s ^o t" +proof - + interpret rt!: wo_rel2 r t by unfold_locales (rule r, rule t) + interpret st!: wo_rel2 s t by unfold_locales (rule s, rule t) + interpret r!: wo_rel r by unfold_locales (rule r) + interpret s!: wo_rel s by unfold_locales (rule s) + interpret t!: wo_rel t by unfold_locales (rule t) + show ?thesis + proof (cases "t = {}") + case True thus ?thesis using r s unfolding ordLeq_def2 rel.underS_def by auto + next + case False thus ?thesis + proof (cases "r = {}") + case True thus ?thesis using t `t \ {}` st.oexp_Well_order ozero_ordLeq[unfolded ozero_def] + by auto + next + case False + from assms obtain f where f: "embed r s f" unfolding ordLeq_def by blast + hence f_underS: "\a\Field r. f a \ Field s \ f ` underS r a \ underS s (f a)" + using embed_in_Field[OF rt.rWELL f] embed_underS2[OF rt.rWELL st.rWELL f] by auto + from f `t \ {}` False have *: "Field r \ {}" "Field s \ {}" "Field t \ {}" + unfolding Field_def embed_def rel.under_def bij_betw_def by auto + with f obtain x where "s.zero = f x" "x \ Field r" unfolding embed_def bij_betw_def + using embed_in_Field[OF r.WELL f] s.zero_under set_mp[OF r.under_Field] by blast + with f have fz: "f r.zero = s.zero" and inj: "inj_on f (Field r)" and compat: "compat r s f" + unfolding embed_iff_compat_inj_on_ofilter[OF r s] compat_def + by (fastforce intro: s.leq_zero_imp)+ + let ?f = "\g x. if x \ Field t then f (g x) else undefined" + { fix g assume g: "g \ Field (r ^o t)" + with fz f_underS have Field_fg: "?f g \ Field (s ^o t)" + unfolding st.Field_oexp rt.Field_oexp FinFunc_def Func_def fin_support_def support_def + by (auto elim!: finite_subset[rotated]) + moreover + have "?f ` underS (r ^o t) g \ underS (s ^o t) (?f g)" + proof safe + fix h + assume h_underS: "h \ underS (r ^o t) g" + hence "h \ Field (r ^o t)" unfolding rel.underS_def Field_def by auto + with fz f_underS have Field_fh: "?f h \ Field (s ^o t)" + unfolding st.Field_oexp rt.Field_oexp FinFunc_def Func_def fin_support_def support_def + by (auto elim!: finite_subset[rotated]) + from h_underS have "h \ g" and hg: "(h, g) \ rt.oexp" unfolding rel.underS_def by auto + with f inj have neq: "?f h \ ?f g" + unfolding fun_eq_iff inj_on_def rt.oexp_def Option.map_def FinFunc_def Func_def Let_def + by simp metis + moreover + with hg have "t.max_fun_diff (?f h) (?f g) = t.max_fun_diff h g" unfolding rt.oexp_def + using rt.max_fun_diff[OF `h \ g`] rt.max_fun_diff_in[OF `h \ g`] + by (subst t.max_fun_diff_def, intro t.maxim_equality) + (auto simp: t.isMaxim_def intro: inj_onD[OF inj] intro!: rt.max_fun_diff_max) + with Field_fg Field_fh hg fz f_underS compat neq have "(?f h, ?f g) \ st.oexp" + using rt.max_fun_diff[OF `h \ g`] rt.max_fun_diff_in[OF `h \ g`] unfolding st.Field_oexp + unfolding rt.oexp_def st.oexp_def Let_def compat_def by auto + ultimately show "?f h \ underS (s ^o t) (?f g)" unfolding rel.underS_def by auto + qed + ultimately have "?f g \ Field (s ^o t) \ ?f ` underS (r ^o t) g \ underS (s ^o t) (?f g)" + by blast + } + thus ?thesis unfolding ordLeq_def2 by (fastforce intro: oexp_Well_order r s t) + qed + qed +qed + +lemma ordLeq_oexp2: + assumes "oone o r ^o s" +proof - + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret r!: wo_rel r by unfold_locales (rule r) + interpret s!: wo_rel s by unfold_locales (rule s) + from assms rel.well_order_on_domain[OF r] obtain x where + x: "x \ Field r" "r.zero \ Field r" "x \ r.zero" + unfolding ordLess_def oone_def embedS_def[abs_def] bij_betw_def embed_def rel.under_def + by (auto simp: image_def) + (metis equals0D insert_not_empty r.under_def r.zero_in_Field rel.under_empty) + let ?f = "\a b. if b \ Field s then if b = a then x else r.zero else undefined" + from x(3) have SUPP: "\y. y \ Field s \ rs.SUPP (?f y) = {y}" unfolding support_def by auto + { fix y assume y: "y \ Field s" + with x(1,2) SUPP have "?f y \ Field (r ^o s)" unfolding rs.Field_oexp + by (auto simp: FinFunc_def Func_def fin_support_def) + moreover + have "?f ` underS s y \ underS (r ^o s) (?f y)" + proof safe + fix z + assume "z \ underS s y" + hence z: "z \ y" "(z, y) \ s" "z \ Field s" unfolding rel.underS_def Field_def by auto + from x(3) y z(1,3) have "?f z \ ?f y" unfolding fun_eq_iff by auto + moreover + { from x(1,2) have "?f z \ FinFunc r s" "?f y \ FinFunc r s" + unfolding FinFunc_def Func_def fin_support_def by (auto simp: SUPP[OF z(3)] SUPP[OF y]) + moreover + from x(3) y z(1,2) refl_onD[OF s.REFL] have "s.max_fun_diff (?f z) (?f y) = y" + unfolding rs.max_fun_diff_alt SUPP[OF z(3)] SUPP[OF y] + by (intro s.maxim_equality) (auto simp: s.isMaxim_def) + ultimately have "(?f z, ?f y) \ rs.oexp" using y x(1) + unfolding rs.oexp_def Let_def by auto + } + ultimately show "?f z \ underS (r ^o s) (?f y)" unfolding rel.underS_def by blast + qed + ultimately have "?f y \ Field (r ^o s) \ ?f ` underS s y \ underS (r ^o s) (?f y)" by blast + } + thus ?thesis unfolding ordLeq_def2 by (fast intro: oexp_Well_order r s) +qed + +lemma FinFunc_osum: + "fg \ FinFunc r (s +o t) = (fg o Inl \ FinFunc r s \ fg o Inr \ FinFunc r t)" + (is "?L = (?R1 \ ?R2)") +proof safe + assume ?L + from `?L` show ?R1 unfolding FinFunc_def Field_osum Func_def Int_iff fin_support_Field_osum o_def + by (auto split: sum.splits) + from `?L` show ?R2 unfolding FinFunc_def Field_osum Func_def Int_iff fin_support_Field_osum o_def + by (auto split: sum.splits) +next + assume ?R1 ?R2 + thus "?L" unfolding FinFunc_def Field_osum Func_def + by (auto simp: fin_support_Field_osum o_def image_iff split: sum.splits) (metis sumE) +qed + +lemma max_fun_diff_eq_Inl: + assumes "wo_rel.max_fun_diff (s +o t) (sum_case f1 g1) (sum_case f2 g2) = Inl x" + "sum_case f1 g1 \ sum_case f2 g2" + "sum_case f1 g1 \ FinFunc r (s +o t)" "sum_case f2 g2 \ FinFunc r (s +o t)" + shows "wo_rel.max_fun_diff s f1 f2 = x" (is ?P) "g1 = g2" (is ?Q) +proof - + interpret st!: wo_rel "s +o t" by unfold_locales (rule osum_Well_order[OF s t]) + interpret s!: wo_rel s by unfold_locales (rule s) + interpret rst!: wo_rel2 r "s +o t" by unfold_locales (rule r, rule osum_Well_order[OF s t]) + from assms(1) have *: "st.isMaxim {a \ Field (s +o t). sum_case f1 g1 a \ sum_case f2 g2 a} (Inl x)" + using rst.isMaxim_max_fun_diff[OF assms(2-4)] by simp + hence "s.isMaxim {a \ Field s. f1 a \ f2 a} x" + unfolding st.isMaxim_def s.isMaxim_def Field_osum by (auto simp: osum_def) + thus ?P unfolding s.max_fun_diff_def by (rule s.maxim_equality) + from assms(3,4) have **: "g1 \ FinFunc r t" "g2 \ FinFunc r t" unfolding FinFunc_osum + by (auto simp: o_def) + show ?Q + proof + fix x + from * ** show "g1 x = g2 x" unfolding st.isMaxim_def Field_osum FinFunc_def Func_def fun_eq_iff + unfolding osum_def by (case_tac "x \ Field t") auto + qed +qed + +lemma max_fun_diff_eq_Inr: + assumes "wo_rel.max_fun_diff (s +o t) (sum_case f1 g1) (sum_case f2 g2) = Inr x" + "sum_case f1 g1 \ sum_case f2 g2" + "sum_case f1 g1 \ FinFunc r (s +o t)" "sum_case f2 g2 \ FinFunc r (s +o t)" + shows "wo_rel.max_fun_diff t g1 g2 = x" (is ?P) "g1 \ g2" (is ?Q) +proof - + interpret st!: wo_rel "s +o t" by unfold_locales (rule osum_Well_order[OF s t]) + interpret t!: wo_rel t by unfold_locales (rule t) + interpret rst!: wo_rel2 r "s +o t" by unfold_locales (rule r, rule osum_Well_order[OF s t]) + from assms(1) have *: "st.isMaxim {a \ Field (s +o t). sum_case f1 g1 a \ sum_case f2 g2 a} (Inr x)" + using rst.isMaxim_max_fun_diff[OF assms(2-4)] by simp + hence "t.isMaxim {a \ Field t. g1 a \ g2 a} x" + unfolding st.isMaxim_def t.isMaxim_def Field_osum by (auto simp: osum_def) + thus ?P ?Q unfolding t.max_fun_diff_def fun_eq_iff + by (auto intro: t.maxim_equality simp: t.isMaxim_def) +qed + +lemma oexp_osum: "r ^o (s +o t) =o (r ^o s) *o (r ^o t)" (is "?R =o ?L") +proof (rule ordIso_symmetric) + interpret rst!: wo_rel2 r "s +o t" by unfold_locales (rule r, rule osum_Well_order[OF s t]) + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret rt!: wo_rel2 r t by unfold_locales (rule r, rule t) + let ?f = "\(f, g). sum_case f g" + have "bij_betw ?f (Field ?L) (Field ?R)" + unfolding bij_betw_def rst.Field_oexp rs.Field_oexp rt.Field_oexp Field_oprod proof (intro conjI) + show "inj_on ?f (FinFunc r s \ FinFunc r t)" unfolding inj_on_def + by (auto simp: fun_eq_iff split: sum.splits) + show "?f ` (FinFunc r s \ FinFunc r t) = FinFunc r (s +o t)" + proof safe + fix fg assume "fg \ FinFunc r (s +o t)" + thus "fg \ ?f ` (FinFunc r s \ FinFunc r t)" + by (intro image_eqI[of _ _ "(fg o Inl, fg o Inr)"]) + (auto simp: FinFunc_osum fun_eq_iff split: sum.splits) + qed (auto simp: FinFunc_osum o_def) + qed + moreover have "compat ?L ?R ?f" + unfolding compat_def rst.Field_oexp rs.Field_oexp rt.Field_oexp oprod_def + unfolding rst.oexp_def Let_def rs.oexp_def rt.oexp_def + by (fastforce simp: Field_osum FinFunc_osum o_def split: sum.splits + dest: max_fun_diff_eq_Inl max_fun_diff_eq_Inr) + ultimately have "iso ?L ?R ?f" using r s t + by (subst iso_iff3) (auto intro: oexp_Well_order oprod_Well_order osum_Well_order) + thus "?L =o ?R" using r s t unfolding ordIso_def + by (auto intro: oexp_Well_order oprod_Well_order osum_Well_order) +qed + +definition "rev_curr f b = (if b \ Field t then \a. f (a, b) else undefined)" + +lemma rev_curr_FinFunc: + assumes Field: "Field r \ {}" + shows "rev_curr ` (FinFunc r (s *o t)) = FinFunc (r ^o s) t" +proof safe + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret rst!: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) + fix g assume g: "g \ FinFunc r (s *o t)" + hence "finite (rst.SUPP (rev_curr g))" "\x \ Field t. finite (rs.SUPP (rev_curr g x))" + unfolding FinFunc_def Field_oprod rs.Field_oexp Func_def fin_support_def support_def + rs.zero_oexp[OF Field] rev_curr_def by (auto simp: fun_eq_iff rs.const_def elim!: finite_surj) + with g show "rev_curr g \ FinFunc (r ^o s) t" + unfolding FinFunc_def Field_oprod rs.Field_oexp Func_def + by (auto simp: rev_curr_def fin_support_def) +next + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret rst!: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) + fix fg assume *: "fg \ FinFunc (r ^o s) t" + let ?g = "\(a, b). if (a, b) \ Field (s *o t) then fg b a else undefined" + show "fg \ rev_curr ` FinFunc r (s *o t)" + proof (rule image_eqI[of _ _ ?g]) + show "fg = rev_curr ?g" + proof + fix x + from * show "fg x = rev_curr ?g x" + unfolding FinFunc_def rs.Field_oexp Func_def rev_curr_def Field_oprod by auto + qed + next + have **: "(\g \ fg ` Field t. rs.SUPP g) = + (\g \ fg ` Field t - {rs.const}. rs.SUPP g)" + unfolding support_def by auto + from * have "\g \ fg ` Field t. finite (rs.SUPP g)" "finite (rst.SUPP fg)" + unfolding rs.Field_oexp FinFunc_def Func_def fin_support_def these_def by force+ + moreover hence "finite (fg ` Field t - {rs.const})" using * + unfolding support_def rs.zero_oexp[OF Field] FinFunc_def Func_def + by (elim finite_surj[of _ _ fg]) (fastforce simp: image_iff these_def) + ultimately have "finite ((\g \ fg ` Field t. rs.SUPP g) \ rst.SUPP fg)" + by (subst **) (auto intro!: finite_cartesian_product) + with * show "?g \ FinFunc r (s *o t)" + unfolding Field_oprod rs.Field_oexp FinFunc_def Func_def fin_support_def these_def + support_def rs.zero_oexp[OF Field] by (auto elim!: finite_subset[rotated]) + qed +qed + +lemma rev_curr_app_FinFunc[elim!]: + "\f \ FinFunc r (s *o t); z \ Field t\ \ rev_curr f z \ FinFunc r s" + unfolding rev_curr_def FinFunc_def Func_def Field_oprod fin_support_def support_def + by (auto elim: finite_surj) + +lemma max_fun_diff_oprod: + assumes Field: "Field r \ {}" and "f \ g" "f \ FinFunc r (s *o t)" "g \ FinFunc r (s *o t)" + defines "m \ wo_rel.max_fun_diff t (rev_curr f) (rev_curr g)" + shows "wo_rel.max_fun_diff (s *o t) f g = + (wo_rel.max_fun_diff s (rev_curr f m) (rev_curr g m), m)" +proof - + interpret st!: wo_rel "s *o t" by unfold_locales (rule oprod_Well_order[OF s t]) + interpret s!: wo_rel s by unfold_locales (rule s) + interpret t!: wo_rel t by unfold_locales (rule t) + interpret r_st!: wo_rel2 r "s *o t" by unfold_locales (rule r, rule oprod_Well_order[OF s t]) + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret rst!: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) + from fun_unequal_in_support[OF assms(2), of "Field (s *o t)" "Field r" "Field r"] assms(3,4) + have diff1: "rev_curr f \ rev_curr g" + "rev_curr f \ FinFunc (r ^o s) t" "rev_curr g \ FinFunc (r ^o s) t" using rev_curr_FinFunc[OF Field] + unfolding fun_eq_iff rev_curr_def[abs_def] FinFunc_def support_def Field_oprod + by auto fast + hence diff2: "rev_curr f m \ rev_curr g m" "rev_curr f m \ FinFunc r s" "rev_curr g m \ FinFunc r s" + using rst.max_fun_diff[OF diff1] assms(3,4) rst.max_fun_diff_in unfolding m_def by auto + show ?thesis unfolding st.max_fun_diff_def + proof (intro st.maxim_equality, unfold st.isMaxim_def Field_oprod, safe) + show "s.max_fun_diff (rev_curr f m) (rev_curr g m) \ Field s" + using rs.max_fun_diff_in[OF diff2] by auto + next + show "m \ Field t" using rst.max_fun_diff_in[OF diff1] unfolding m_def by auto + next + assume "f (s.max_fun_diff (rev_curr f m) (rev_curr g m), m) = + g (s.max_fun_diff (rev_curr f m) (rev_curr g m), m)" + (is "f (?x, m) = g (?x, m)") + hence "rev_curr f m ?x = rev_curr g m ?x" unfolding rev_curr_def by auto + with rs.max_fun_diff[OF diff2] show False by auto + next + fix x y assume "f (x, y) \ g (x, y)" "x \ Field s" "y \ Field t" + thus "((x, y), (s.max_fun_diff (rev_curr f m) (rev_curr g m), m)) \ s *o t" + using rst.max_fun_diff_in[OF diff1] rs.max_fun_diff_in[OF diff2] diff1 diff2 + rst.max_fun_diff_max[OF diff1, of y] rs.max_fun_diff_le_eq[OF _ diff2, of x] + unfolding oprod_def m_def rev_curr_def fun_eq_iff by auto (metis s.in_notinI) + qed +qed + +lemma oexp_oexp: "(r ^o s) ^o t =o r ^o (s *o t)" (is "?R =o ?L") +proof (cases "r = {}") + case True + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret rst!: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) + show ?thesis + proof (cases "s = {} \ t = {}") + case True with `r = {}` show ?thesis + by (auto simp: oexp_empty[OF oexp_Well_order[OF Well_order_empty s]] + intro!: ordIso_transitive[OF ordIso_symmetric[OF oone_ordIso] oone_ordIso] + ordIso_transitive[OF oone_ordIso_oexp[OF ordIso_symmetric[OF oone_ordIso] t] oone_ordIso]) + next + case False + moreover hence "s *o t \ {}" unfolding oprod_def Field_def by fastforce + ultimately show ?thesis using `r = {}` ozero_ordIso + by (auto simp add: s t oprod_Well_order ozero_def) + qed +next + case False + hence Field: "Field r \ {}" by (metis Field_def Range_empty_iff Un_empty) + show ?thesis + proof (rule ordIso_symmetric) + interpret r_st!: wo_rel2 r "s *o t" by unfold_locales (rule r, rule oprod_Well_order[OF s t]) + interpret rs!: wo_rel2 r s by unfold_locales (rule r, rule s) + interpret rst!: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t) + have bij: "bij_betw rev_curr (Field ?L) (Field ?R)" + unfolding bij_betw_def r_st.Field_oexp rst.Field_oexp Field_oprod proof (intro conjI) + show "inj_on rev_curr (FinFunc r (s *o t))" + unfolding inj_on_def FinFunc_def Func_def Field_oprod rs.Field_oexp rev_curr_def[abs_def] + by (auto simp: fun_eq_iff) metis + show "rev_curr ` (FinFunc r (s *o t)) = FinFunc (r ^o s) t" by (rule rev_curr_FinFunc[OF Field]) + qed + moreover + have "compat ?L ?R rev_curr" + unfolding compat_def proof safe + fix fg1 fg2 assume fg: "(fg1, fg2) \ r ^o (s *o t)" + show "(rev_curr fg1, rev_curr fg2) \ r ^o s ^o t" + proof (cases "fg1 = fg2") + assume "fg1 \ fg2" + with fg show ?thesis + using rst.max_fun_diff_in[of "rev_curr fg1" "rev_curr fg2"] + max_fun_diff_oprod[OF Field, of fg1 fg2] rev_curr_FinFunc[OF Field, symmetric] + unfolding r_st.Field_oexp rs.Field_oexp rst.Field_oexp unfolding r_st.oexp_def rst.oexp_def + by (auto simp: rs.oexp_def Let_def) (auto simp: rev_curr_def[abs_def]) + next + assume "fg1 = fg2" + with fg bij show ?thesis unfolding r_st.Field_oexp rs.Field_oexp rst.Field_oexp bij_betw_def + by (auto simp: r_st.oexp_def rst.oexp_def) + qed + qed + ultimately have "iso ?L ?R rev_curr" using r s t + by (subst iso_iff3) (auto intro: oexp_Well_order oprod_Well_order) + thus "?L =o ?R" using r s t unfolding ordIso_def + by (auto intro: oexp_Well_order oprod_Well_order) + qed +qed + +end (* context with 3 wellorders *) + +end