diff --git a/src/ZF/AC/AC15_WO6.thy b/src/ZF/AC/AC15_WO6.thy --- a/src/ZF/AC/AC15_WO6.thy +++ b/src/ZF/AC/AC15_WO6.thy @@ -1,289 +1,289 @@ (* Title: ZF/AC/AC15_WO6.thy Author: Krzysztof Grabczewski The proofs needed to state that AC10, ..., AC15 are equivalent to the rest. We need the following: WO1 \ AC10(n) \ AC11 \ AC12 \ AC15 \ WO6 In order to add the formulations AC13 and AC14 we need: AC10(succ(n)) \ AC13(n) \ AC14 \ AC15 or AC1 \ AC13(1); AC13(m) \ AC13(n) \ AC14 \ AC15 (m\n) So we don't have to prove all implications of both cases. Moreover we don't need to prove AC13(1) \ AC1 and AC11 \ AC14 as Rubin \ Rubin do. *) theory AC15_WO6 imports HH Cardinal_aux begin (* ********************************************************************** *) (* Lemmas used in the proofs in which the conclusion is AC13, AC14 *) (* or AC15 *) (* - cons_times_nat_not_Finite *) (* - ex_fun_AC13_AC15 *) (* ********************************************************************** *) lemma lepoll_Sigma: "A\0 \ B \ A*B" unfolding lepoll_def apply (erule not_emptyE) apply (rule_tac x = "\z \ B. \x,z\" in exI) apply (fast intro!: snd_conv lam_injective) done lemma cons_times_nat_not_Finite: "0\A \ \B \ {cons(0,x*nat). x \ A}. \Finite(B)" apply clarify apply (rule nat_not_Finite [THEN notE] ) apply (subgoal_tac "x \ 0") apply (blast intro: lepoll_Sigma [THEN lepoll_Finite])+ done lemma lemma1: "\\(C)=A; a \ A\ \ \B \ C. a \ B \ B \ A" by fast lemma lemma2: "\pairwise_disjoint(A); B \ A; C \ A; a \ B; a \ C\ \ B=C" by (unfold pairwise_disjoint_def, blast) lemma lemma3: "\B \ {cons(0, x*nat). x \ A}. pairwise_disjoint(f`B) \ sets_of_size_between(f`B, 2, n) \ \(f`B)=B \ \B \ A. \! u. u \ f`cons(0, B*nat) \ u \ cons(0, B*nat) \ 0 \ u \ 2 \ u \ u \ n" unfolding sets_of_size_between_def apply (rule ballI) apply (erule_tac x="cons(0, B*nat)" in ballE) apply (blast dest: lemma1 intro!: lemma2, blast) done lemma lemma4: "\A \ i; Ord(i)\ \ {P(a). a \ A} \ i" unfolding lepoll_def apply (erule exE) apply (rule_tac x = "\x \ RepFun(A,P). \ j. \a\A. x=P(a) \ f`a=j" in exI) apply (rule_tac d = "\y. P (converse (f) `y) " in lam_injective) apply (erule RepFunE) apply (frule inj_is_fun [THEN apply_type], assumption) apply (fast intro: LeastI2 elim!: Ord_in_Ord inj_is_fun [THEN apply_type]) apply (erule RepFunE) apply (rule LeastI2) apply fast apply (fast elim!: Ord_in_Ord inj_is_fun [THEN apply_type]) apply (fast elim: sym left_inverse [THEN ssubst]) done lemma lemma5_1: "\B \ A; 2 \ u(B)\ \ (\x \ A. {fst(x). x \ u(x)-{0}})`B \ 0" apply simp apply (fast dest: lepoll_Diff_sing elim: lepoll_trans [THEN succ_lepoll_natE] ssubst intro!: lepoll_refl) done lemma lemma5_2: "\B \ A; u(B) \ cons(0, B*nat)\ \ (\x \ A. {fst(x). x \ u(x)-{0}})`B \ B" apply auto done lemma lemma5_3: "\n \ nat; B \ A; 0 \ u(B); u(B) \ succ(n)\ \ (\x \ A. {fst(x). x \ u(x)-{0}})`B \ n" apply simp apply (fast elim!: Diff_lepoll [THEN lemma4 [OF _ nat_into_Ord]]) done lemma ex_fun_AC13_AC15: "\\B \ {cons(0, x*nat). x \ A}. pairwise_disjoint(f`B) \ sets_of_size_between(f`B, 2, succ(n)) \ \(f`B)=B; n \ nat\ \ \f. \B \ A. f`B \ 0 \ f`B \ B \ f`B \ n" by (fast del: subsetI notI dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3) (* ********************************************************************** *) (* The target proofs *) (* ********************************************************************** *) (* ********************************************************************** *) (* AC10(n) \ AC11 *) (* ********************************************************************** *) theorem AC10_AC11: "\n \ nat; 1\n; AC10(n)\ \ AC11" by (unfold AC10_def AC11_def, blast) (* ********************************************************************** *) (* AC11 \ AC12 *) (* ********************************************************************** *) theorem AC11_AC12: "AC11 \ AC12" by (unfold AC10_def AC11_def AC11_def AC12_def, blast) (* ********************************************************************** *) (* AC12 \ AC15 *) (* ********************************************************************** *) theorem AC12_AC15: "AC12 \ AC15" -apply (unfold AC12_def AC15_def) + unfolding AC12_def AC15_def apply (blast del: ballI intro!: cons_times_nat_not_Finite ex_fun_AC13_AC15) done (* ********************************************************************** *) (* AC15 \ WO6 *) (* ********************************************************************** *) lemma OUN_eq_UN: "Ord(x) \ (\aa \ x. F(a))" by (fast intro!: ltI dest!: ltD) lemma AC15_WO6_aux1: "\x \ Pow(A)-{0}. f`x\0 \ f`x \ x \ f`x \ m \ (\i<\ x. HH(f,A,x)={A}. HH(f,A,i)) = A" apply (simp add: Ord_Least [THEN OUN_eq_UN]) apply (rule equalityI) apply (fast dest!: less_Least_subset_x) apply (blast del: subsetI intro!: f_subsets_imp_UN_HH_eq_x [THEN Diff_eq_0_iff [THEN iffD1]]) done lemma AC15_WO6_aux2: "\x \ Pow(A)-{0}. f`x\0 \ f`x \ x \ f`x \ m \ \x < (\ x. HH(f,A,x)={A}). HH(f,A,x) \ m" apply (rule oallI) apply (drule ltD [THEN less_Least_subset_x]) apply (frule HH_subset_imp_eq) apply (erule ssubst) apply (blast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2]) (*but can't use del: DiffE despite the obvious conflict*) done theorem AC15_WO6: "AC15 \ WO6" -apply (unfold AC15_def WO6_def) + unfolding AC15_def WO6_def apply (rule allI) apply (erule_tac x = "Pow (A) -{0}" in allE) apply (erule impE, fast) apply (elim bexE conjE exE) apply (rule bexI) apply (rule conjI, assumption) apply (rule_tac x = "\ i. HH (f,A,i) ={A}" in exI) apply (rule_tac x = "\j \ (\ i. HH (f,A,i) ={A}) . HH (f,A,j) " in exI) apply (simp_all add: ltD) apply (fast intro!: Ord_Least lam_type [THEN domain_of_fun] elim!: less_Least_subset_x AC15_WO6_aux1 AC15_WO6_aux2) done (* ********************************************************************** *) (* The proof needed in the first case, not in the second *) (* ********************************************************************** *) (* ********************************************************************** *) (* AC10(n) \ AC13(n-1) if 2\n *) (* *) (* Because of the change to the formal definition of AC10(n) we prove *) (* the following obviously equivalent theorem \ *) (* AC10(n) implies AC13(n) for (1\n) *) (* ********************************************************************** *) theorem AC10_AC13: "\n \ nat; 1\n; AC10(n)\ \ AC13(n)" apply (unfold AC10_def AC13_def, safe) apply (erule allE) apply (erule impE [OF _ cons_times_nat_not_Finite], assumption) apply (fast elim!: impE [OF _ cons_times_nat_not_Finite] dest!: ex_fun_AC13_AC15) done (* ********************************************************************** *) (* The proofs needed in the second case, not in the first *) (* ********************************************************************** *) (* ********************************************************************** *) (* AC1 \ AC13(1) *) (* ********************************************************************** *) lemma AC1_AC13: "AC1 \ AC13(1)" -apply (unfold AC1_def AC13_def) + unfolding AC1_def AC13_def apply (rule allI) apply (erule allE) apply (rule impI) apply (drule mp, assumption) apply (elim exE) apply (rule_tac x = "\x \ A. {f`x}" in exI) apply (simp add: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll]) done (* ********************************************************************** *) (* AC13(m) \ AC13(n) for m \ n *) (* ********************************************************************** *) lemma AC13_mono: "\m\n; AC13(m)\ \ AC13(n)" unfolding AC13_def apply (drule le_imp_lepoll) apply (fast elim!: lepoll_trans) done (* ********************************************************************** *) (* The proofs necessary for both cases *) (* ********************************************************************** *) (* ********************************************************************** *) (* AC13(n) \ AC14 if 1 \ n *) (* ********************************************************************** *) theorem AC13_AC14: "\n \ nat; 1\n; AC13(n)\ \ AC14" by (unfold AC13_def AC14_def, auto) (* ********************************************************************** *) (* AC14 \ AC15 *) (* ********************************************************************** *) theorem AC14_AC15: "AC14 \ AC15" by (unfold AC13_def AC14_def AC15_def, fast) (* ********************************************************************** *) (* The redundant proofs; however cited by Rubin \ Rubin *) (* ********************************************************************** *) (* ********************************************************************** *) (* AC13(1) \ AC1 *) (* ********************************************************************** *) lemma lemma_aux: "\A\0; A \ 1\ \ \a. A={a}" by (fast elim!: not_emptyE lepoll_1_is_sing) lemma AC13_AC1_lemma: "\B \ A. f(B)\0 \ f(B)<=B \ f(B) \ 1 \ (\x \ A. THE y. f(x)={y}) \ (\X \ A. X)" apply (rule lam_type) apply (drule bspec, assumption) apply (elim conjE) apply (erule lemma_aux [THEN exE], assumption) apply (simp add: the_equality) done theorem AC13_AC1: "AC13(1) \ AC1" -apply (unfold AC13_def AC1_def) + unfolding AC13_def AC1_def apply (fast elim!: AC13_AC1_lemma) done (* ********************************************************************** *) (* AC11 \ AC14 *) (* ********************************************************************** *) theorem AC11_AC14: "AC11 \ AC14" -apply (unfold AC11_def AC14_def) + unfolding AC11_def AC14_def apply (fast intro!: AC10_AC13) done end diff --git a/src/ZF/AC/AC16_WO4.thy b/src/ZF/AC/AC16_WO4.thy --- a/src/ZF/AC/AC16_WO4.thy +++ b/src/ZF/AC/AC16_WO4.thy @@ -1,578 +1,578 @@ (* Title: ZF/AC/AC16_WO4.thy Author: Krzysztof Grabczewski The proof of AC16(n, k) \ WO4(n-k) Tidied (using locales) by LCP *) theory AC16_WO4 imports AC16_lemmas begin (* ********************************************************************** *) (* The case of finite set *) (* ********************************************************************** *) lemma lemma1: "\Finite(A); 0 nat\ \ \a f. Ord(a) \ domain(f) = a \ (\b (\b m)" unfolding Finite_def apply (erule bexE) apply (drule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]]) apply (erule exE) apply (rule_tac x = n in exI) apply (rule_tac x = "\i \ n. {f`i}" in exI) apply (simp add: ltD bij_def surj_def) apply (fast intro!: ltI nat_into_Ord lam_funtype [THEN domain_of_fun] singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans] nat_1_lepoll_iff [THEN iffD2] elim!: apply_type ltE) done (* ********************************************************************** *) (* The case of infinite set *) (* ********************************************************************** *) (* well_ord(x,r) \ well_ord({{y,z}. y \ x}, Something(x,z)) **) lemmas well_ord_paired = paired_bij [THEN bij_is_inj, THEN well_ord_rvimage] lemma lepoll_trans1: "\A \ B; \ A \ C\ \ \ B \ C" by (blast intro: lepoll_trans) (* ********************************************************************** *) (* There exists a well ordered set y such that ... *) (* ********************************************************************** *) lemmas lepoll_paired = paired_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll] lemma lemma2: "\y R. well_ord(y,R) \ x \ y = 0 \ \y \ z \ \Finite(y)" apply (rule_tac x = "{{a,x}. a \ nat \ Hartog (z) }" in exI) apply (rule well_ord_Un [OF Ord_nat [THEN well_ord_Memrel] Ord_Hartog [THEN well_ord_Memrel], THEN exE]) apply (blast intro!: Ord_Hartog well_ord_Memrel well_ord_paired lepoll_trans1 [OF _ not_Hartog_lepoll_self] lepoll_trans [OF subset_imp_lepoll lepoll_paired] elim!: nat_not_Finite [THEN notE] elim: mem_asym dest!: Un_upper1 [THEN subset_imp_lepoll, THEN lepoll_Finite] lepoll_paired [THEN lepoll_Finite]) done lemma infinite_Un: "\Finite(B) \ \Finite(A \ B)" by (blast intro: subset_Finite) (* ********************************************************************** *) (* There is a v \ s(u) such that k \ x \ y (in our case succ(k)) *) (* The idea of the proof is the following \ *) (* Suppose not, i.e. every element of s(u) has exactly k-1 elements of y *) (* Thence y is less than or equipollent to {v \ Pow(x). v \ n#-k} *) (* We have obtained this result in two steps \ *) (* 1. y is less than or equipollent to {v \ s(u). a \ v} *) (* where a is certain k-2 element subset of y *) (* 2. {v \ s(u). a \ v} is less than or equipollent *) (* to {v \ Pow(x). v \ n-k} *) (* ********************************************************************** *) (*Proof simplified by LCP*) lemma succ_not_lepoll_lemma: "\\(\x \ A. f`x=y); f \ inj(A, B); y \ B\ \ (\a \ succ(A). if(a=A, y, f`a)) \ inj(succ(A), B)" apply (rule_tac d = "\z. if (z=y, A, converse (f) `z) " in lam_injective) apply (force simp add: inj_is_fun [THEN apply_type]) (*this preliminary simplification prevents looping somehow*) apply (simp (no_asm_simp)) apply force done lemma succ_not_lepoll_imp_eqpoll: "\\A \ B; A \ B\ \ succ(A) \ B" -apply (unfold lepoll_def eqpoll_def bij_def surj_def) + unfolding lepoll_def eqpoll_def bij_def surj_def apply (fast elim!: succ_not_lepoll_lemma inj_is_fun) done (* ********************************************************************** *) (* There is a k-2 element subset of y *) (* ********************************************************************** *) lemmas ordertype_eqpoll = ordermap_bij [THEN exI [THEN eqpoll_def [THEN def_imp_iff, THEN iffD2]]] lemma cons_cons_subset: "\a \ y; b \ y-a; u \ x\ \ cons(b, cons(u, a)) \ Pow(x \ y)" by fast lemma cons_cons_eqpoll: "\a \ k; a \ y; b \ y-a; u \ x; x \ y = 0\ \ cons(b, cons(u, a)) \ succ(succ(k))" by (fast intro!: cons_eqpoll_succ) lemma set_eq_cons: "\succ(k) \ A; k \ B; B \ A; a \ A-B; k \ nat\ \ A = cons(a, B)" apply (rule equalityI) prefer 2 apply fast apply (rule Diff_eq_0_iff [THEN iffD1]) apply (rule equals0I) apply (drule eqpoll_sym [THEN eqpoll_imp_lepoll]) apply (drule eqpoll_sym [THEN cons_eqpoll_succ], fast) apply (drule cons_eqpoll_succ, fast) apply (fast elim!: lepoll_trans [THEN lepoll_trans, THEN succ_lepoll_natE, OF eqpoll_sym [THEN eqpoll_imp_lepoll] subset_imp_lepoll]) done lemma cons_eqE: "\cons(x,a) = cons(y,a); x \ a\ \ x = y " by (fast elim!: equalityE) lemma eq_imp_Int_eq: "A = B \ A \ C = B \ C" by blast (* ********************************************************************** *) (* some arithmetic *) (* ********************************************************************** *) lemma eqpoll_sum_imp_Diff_lepoll_lemma [rule_format]: "\k \ nat; m \ nat\ \ \A B. A \ k #+ m \ k \ B \ B \ A \ A-B \ m" apply (induct_tac "k") apply (simp add: add_0) apply (blast intro: eqpoll_imp_lepoll lepoll_trans Diff_subset [THEN subset_imp_lepoll]) apply (intro allI impI) apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], fast) apply (erule_tac x = "A - {xa}" in allE) apply (erule_tac x = "B - {xa}" in allE) apply (erule impE) apply (simp add: add_succ) apply (fast intro!: Diff_sing_eqpoll lepoll_Diff_sing) apply (subgoal_tac "A - {xa} - (B - {xa}) = A - B", simp) apply blast done lemma eqpoll_sum_imp_Diff_lepoll: "\A \ succ(k #+ m); B \ A; succ(k) \ B; k \ nat; m \ nat\ \ A-B \ m" apply (simp only: add_succ [symmetric]) apply (blast intro: eqpoll_sum_imp_Diff_lepoll_lemma) done (* ********************************************************************** *) (* similar properties for \ *) (* ********************************************************************** *) lemma eqpoll_sum_imp_Diff_eqpoll_lemma [rule_format]: "\k \ nat; m \ nat\ \ \A B. A \ k #+ m \ k \ B \ B \ A \ A-B \ m" apply (induct_tac "k") apply (force dest!: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_0_is_0]) apply (intro allI impI) apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE]) apply (fast elim!: eqpoll_imp_lepoll) apply (erule_tac x = "A - {xa}" in allE) apply (erule_tac x = "B - {xa}" in allE) apply (erule impE) apply (force intro: eqpoll_sym intro!: Diff_sing_eqpoll) apply (subgoal_tac "A - {xa} - (B - {xa}) = A - B", simp) apply blast done lemma eqpoll_sum_imp_Diff_eqpoll: "\A \ succ(k #+ m); B \ A; succ(k) \ B; k \ nat; m \ nat\ \ A-B \ m" apply (simp only: add_succ [symmetric]) apply (blast intro: eqpoll_sum_imp_Diff_eqpoll_lemma) done (* ********************************************************************** *) (* LL can be well ordered *) (* ********************************************************************** *) lemma subsets_lepoll_0_eq_unit: "{x \ Pow(X). x \ 0} = {0}" by (fast dest!: lepoll_0_is_0 intro!: lepoll_refl) lemma subsets_lepoll_succ: "n \ nat \ {z \ Pow(y). z \ succ(n)} = {z \ Pow(y). z \ n} \ {z \ Pow(y). z \ succ(n)}" by (blast intro: leI le_imp_lepoll nat_into_Ord lepoll_trans eqpoll_imp_lepoll dest!: lepoll_succ_disj) lemma Int_empty: "n \ nat \ {z \ Pow(y). z \ n} \ {z \ Pow(y). z \ succ(n)} = 0" by (blast intro: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans] succ_lepoll_natE) locale AC16 = fixes x and y and k and l and m and t_n and R and MM and LL and GG and s defines k_def: "k \ succ(l)" and MM_def: "MM \ {v \ t_n. succ(k) \ v \ y}" and LL_def: "LL \ {v \ y. v \ MM}" and GG_def: "GG \ \v \ LL. (THE w. w \ MM \ v \ w) - v" and s_def: "s(u) \ {v \ t_n. u \ v \ k \ v \ y}" assumes all_ex: "\z \ {z \ Pow(x \ y) . z \ succ(k)}. \! w. w \ t_n \ z \ w " and disjoint[iff]: "x \ y = 0" and "includes": "t_n \ {v \ Pow(x \ y). v \ succ(k #+ m)}" and WO_R[iff]: "well_ord(y,R)" and lnat[iff]: "l \ nat" and mnat[iff]: "m \ nat" and mpos[iff]: "0 Finite(y)" and noLepoll: "\ y \ {v \ Pow(x). v \ m}" begin lemma knat [iff]: "k \ nat" by (simp add: k_def) (* ********************************************************************** *) (* 1. y is less than or equipollent to {v \ s(u). a \ v} *) (* where a is certain k-2 element subset of y *) (* ********************************************************************** *) lemma Diff_Finite_eqpoll: "\l \ a; a \ y\ \ y - a \ y" apply (insert WO_R Infinite lnat) apply (rule eqpoll_trans) apply (rule Diff_lesspoll_eqpoll_Card) apply (erule well_ord_cardinal_eqpoll [THEN eqpoll_sym]) apply (blast intro: lesspoll_trans1 intro!: Card_cardinal Card_cardinal [THEN Card_is_Ord, THEN nat_le_infinite_Ord, THEN le_imp_lepoll] dest: well_ord_cardinal_eqpoll eqpoll_sym eqpoll_imp_lepoll n_lesspoll_nat [THEN lesspoll_trans2] well_ord_cardinal_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll, THEN lepoll_infinite])+ done lemma s_subset: "s(u) \ t_n" by (unfold s_def, blast) lemma sI: "\w \ t_n; cons(b,cons(u,a)) \ w; a \ y; b \ y-a; l \ a\ \ w \ s(u)" -apply (unfold s_def succ_def k_def) + unfolding s_def succ_def k_def apply (blast intro!: eqpoll_imp_lepoll [THEN cons_lepoll_cong] intro: subset_imp_lepoll lepoll_trans) done lemma in_s_imp_u_in: "v \ s(u) \ u \ v" by (unfold s_def, blast) lemma ex1_superset_a: "\l \ a; a \ y; b \ y - a; u \ x\ \ \! c. c \ s(u) \ a \ c \ b \ c" apply (rule all_ex [simplified k_def, THEN ballE]) apply (erule ex1E) apply (rule_tac a = w in ex1I, blast intro: sI) apply (blast dest: s_subset [THEN subsetD] in_s_imp_u_in) apply (blast del: PowI intro!: cons_cons_subset eqpoll_sym [THEN cons_cons_eqpoll]) done lemma the_eq_cons: "\\v \ s(u). succ(l) \ v \ y; l \ a; a \ y; b \ y - a; u \ x\ \ (THE c. c \ s(u) \ a \ c \ b \ c) \ y = cons(b, a)" apply (frule ex1_superset_a [THEN theI], assumption+) apply (rule set_eq_cons) apply (fast+) done lemma y_lepoll_subset_s: "\\v \ s(u). succ(l) \ v \ y; l \ a; a \ y; u \ x\ \ y \ {v \ s(u). a \ v}" apply (rule Diff_Finite_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll, THEN lepoll_trans], fast+) apply (rule_tac f3 = "\b \ y-a. THE c. c \ s (u) \ a \ c \ b \ c" in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]]) apply (simp add: inj_def) apply (rule conjI) apply (rule lam_type) apply (frule ex1_superset_a [THEN theI], fast+, clarify) apply (rule cons_eqE [of _ a]) apply (drule_tac A = "THE c. P (c)" and C = y for P in eq_imp_Int_eq) apply (simp_all add: the_eq_cons) done (* ********************************************************************** *) (* back to the second part *) (* ********************************************************************** *) (*relies on the disjointness of x, y*) lemma x_imp_not_y [dest]: "a \ x \ a \ y" by (blast dest: disjoint [THEN equalityD1, THEN subsetD, OF IntI]) lemma w_Int_eq_w_Diff: "w \ x \ y \ w \ (x - {u}) = w - cons(u, w \ y)" by blast lemma w_Int_eqpoll_m: "\w \ {v \ s(u). a \ v}; l \ a; u \ x; \v \ s(u). succ(l) \ v \ y\ \ w \ (x - {u}) \ m" apply (erule CollectE) apply (subst w_Int_eq_w_Diff) apply (fast dest!: s_subset [THEN subsetD] "includes" [simplified k_def, THEN subsetD]) apply (blast dest: s_subset [THEN subsetD] "includes" [simplified k_def, THEN subsetD] dest: eqpoll_sym [THEN cons_eqpoll_succ, THEN eqpoll_sym] in_s_imp_u_in intro!: eqpoll_sum_imp_Diff_eqpoll) done (* ********************************************************************** *) (* 2. {v \ s(u). a \ v} is less than or equipollent *) (* to {v \ Pow(x). v \ n-k} *) (* ********************************************************************** *) lemma eqpoll_m_not_empty: "a \ m \ a \ 0" apply (insert mpos) apply (fast elim!: zero_lt_natE dest!: eqpoll_succ_imp_not_empty) done lemma cons_cons_in: "\z \ xa \ (x - {u}); l \ a; a \ y; u \ x\ \ \! w. w \ t_n \ cons(z, cons(u, a)) \ w" apply (rule all_ex [THEN bspec]) unfolding k_def apply (fast intro!: cons_eqpoll_succ elim: eqpoll_sym) done lemma subset_s_lepoll_w: "\\v \ s(u). succ(l) \ v \ y; a \ y; l \ a; u \ x\ \ {v \ s(u). a \ v} \ {v \ Pow(x). v \ m}" apply (rule_tac f3 = "\w \ {v \ s (u) . a \ v}. w \ (x - {u})" in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]]) apply (simp add: inj_def) apply (intro conjI lam_type CollectI) apply fast apply (blast intro: w_Int_eqpoll_m) apply (intro ballI impI) (** LEVEL 8 **) apply (rule w_Int_eqpoll_m [THEN eqpoll_m_not_empty, THEN not_emptyE]) apply (blast, assumption+) apply (drule equalityD1 [THEN subsetD], assumption) apply (frule cons_cons_in, assumption+) apply (blast dest: ex1_two_eq intro: s_subset [THEN subsetD] in_s_imp_u_in)+ done (* ********************************************************************** *) (* well_ord_subsets_lepoll_n *) (* ********************************************************************** *) lemma well_ord_subsets_eqpoll_n: "n \ nat \ \S. well_ord({z \ Pow(y) . z \ succ(n)}, S)" apply (rule WO_R [THEN well_ord_infinite_subsets_eqpoll_X, THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE]) apply (fast intro: bij_is_inj [THEN well_ord_rvimage])+ done lemma well_ord_subsets_lepoll_n: "n \ nat \ \R. well_ord({z \ Pow(y). z \ n}, R)" apply (induct_tac "n") apply (force intro!: well_ord_unit simp add: subsets_lepoll_0_eq_unit) apply (erule exE) apply (rule well_ord_subsets_eqpoll_n [THEN exE], assumption) apply (simp add: subsets_lepoll_succ) apply (drule well_ord_radd, assumption) apply (erule Int_empty [THEN disj_Un_eqpoll_sum, THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE]) apply (fast elim!: bij_is_inj [THEN well_ord_rvimage]) done lemma LL_subset: "LL \ {z \ Pow(y). z \ succ(k #+ m)}" -apply (unfold LL_def MM_def) + unfolding LL_def MM_def apply (insert "includes") apply (blast intro: subset_imp_lepoll eqpoll_imp_lepoll lepoll_trans) done lemma well_ord_LL: "\S. well_ord(LL,S)" apply (rule well_ord_subsets_lepoll_n [THEN exE, of "succ(k#+m)"]) apply simp apply (blast intro: well_ord_subset [OF _ LL_subset]) done (* ********************************************************************** *) (* every element of LL is a contained in exactly one element of MM *) (* ********************************************************************** *) lemma unique_superset_in_MM: "v \ LL \ \! w. w \ MM \ v \ w" apply (unfold MM_def LL_def, safe, fast) apply (rule lepoll_imp_eqpoll_subset [THEN exE], assumption) apply (rule_tac x = x in all_ex [THEN ballE]) apply (blast intro: eqpoll_sym)+ done (* ********************************************************************** *) (* The function GG satisfies the conditions of WO4 *) (* ********************************************************************** *) (* ********************************************************************** *) (* The union of appropriate values is the whole x *) (* ********************************************************************** *) lemma Int_in_LL: "w \ MM \ w \ y \ LL" by (unfold LL_def, fast) lemma in_LL_eq_Int: "v \ LL \ v = (THE x. x \ MM \ v \ x) \ y" apply (unfold LL_def, clarify) apply (subst unique_superset_in_MM [THEN the_equality2]) apply (auto simp add: Int_in_LL) done lemma unique_superset1: "a \ LL \ (THE x. x \ MM \ a \ x) \ MM" by (erule unique_superset_in_MM [THEN theI, THEN conjunct1]) lemma the_in_MM_subset: "v \ LL \ (THE x. x \ MM \ v \ x) \ x \ y" apply (drule unique_superset1) unfolding MM_def apply (fast dest!: unique_superset1 "includes" [THEN subsetD]) done lemma GG_subset: "v \ LL \ GG ` v \ x" unfolding GG_def apply (frule the_in_MM_subset) apply (frule in_LL_eq_Int) apply (force elim: equalityE) done lemma nat_lepoll_ordertype: "nat \ ordertype(y, R)" apply (rule nat_le_infinite_Ord [THEN le_imp_lepoll]) apply (rule Ord_ordertype [OF WO_R]) apply (rule ordertype_eqpoll [THEN eqpoll_imp_lepoll, THEN lepoll_infinite]) apply (rule WO_R) apply (rule Infinite) done lemma ex_subset_eqpoll_n: "n \ nat \ \z. z \ y \ n \ z" apply (erule nat_lepoll_imp_ex_eqpoll_n) apply (rule lepoll_trans [OF nat_lepoll_ordertype]) apply (rule ordertype_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]) apply (rule WO_R) done lemma exists_proper_in_s: "u \ x \ \v \ s(u). succ(k) \ v \ y" apply (rule ccontr) apply (subgoal_tac "\v \ s (u) . k \ v \ y") prefer 2 apply (simp add: s_def, blast intro: succ_not_lepoll_imp_eqpoll) unfolding k_def apply (insert all_ex "includes" lnat) apply (rule ex_subset_eqpoll_n [THEN exE], assumption) apply (rule noLepoll [THEN notE]) apply (blast intro: lepoll_trans [OF y_lepoll_subset_s subset_s_lepoll_w]) done lemma exists_in_MM: "u \ x \ \w \ MM. u \ w" apply (erule exists_proper_in_s [THEN bexE]) apply (unfold MM_def s_def, fast) done lemma exists_in_LL: "u \ x \ \w \ LL. u \ GG`w" apply (rule exists_in_MM [THEN bexE], assumption) apply (rule bexI) apply (erule_tac [2] Int_in_LL) unfolding GG_def apply (simp add: Int_in_LL) apply (subst unique_superset_in_MM [THEN the_equality2]) apply (fast elim!: Int_in_LL)+ done lemma OUN_eq_x: "well_ord(LL,S) \ (\b MM \ w \ succ(k #+ m)" unfolding MM_def apply (fast dest: "includes" [THEN subsetD]) done lemma in_LL_eqpoll_n: "w \ LL \ succ(k) \ w" by (unfold LL_def MM_def, fast) lemma in_LL: "w \ LL \ w \ (THE x. x \ MM \ w \ x)" by (erule subset_trans [OF in_LL_eq_Int [THEN equalityD1] Int_lower1]) lemma all_in_lepoll_m: "well_ord(LL,S) \ \b m" unfolding GG_def apply (rule oallI) apply (simp add: ltD ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN apply_type]) apply (insert "includes") apply (rule eqpoll_sum_imp_Diff_lepoll) apply (blast del: subsetI dest!: ltD intro!: eqpoll_sum_imp_Diff_lepoll in_LL_eqpoll_n intro: in_LL unique_superset1 [THEN in_MM_eqpoll_n] ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN apply_type])+ done lemma "conclusion": "\a f. Ord(a) \ domain(f) = a \ (\b (\b m)" apply (rule well_ord_LL [THEN exE]) apply (rename_tac S) apply (rule_tac x = "ordertype (LL,S)" in exI) apply (rule_tac x = "\b \ ordertype(LL,S). GG ` (converse (ordermap (LL,S)) ` b)" in exI) apply (simp add: ltD) apply (blast intro: lam_funtype [THEN domain_of_fun] Ord_ordertype OUN_eq_x all_in_lepoll_m [THEN ospec]) done end (* ********************************************************************** *) (* The main theorem AC16(n, k) \ WO4(n-k) *) (* ********************************************************************** *) theorem AC16_WO4: "\AC_Equiv.AC16(k #+ m, k); 0 < k; 0 < m; k \ nat; m \ nat\ \ WO4(m)" -apply (unfold AC_Equiv.AC16_def WO4_def) + unfolding AC_Equiv.AC16_def WO4_def apply (rule allI) apply (case_tac "Finite (A)") apply (rule lemma1, assumption+) apply (cut_tac lemma2) apply (elim exE conjE) apply (erule_tac x = "A \ y" in allE) apply (frule infinite_Un, drule mp, assumption) apply (erule zero_lt_natE, assumption, clarify) apply (blast intro: AC16.conclusion [OF AC16.intro]) done end diff --git a/src/ZF/AC/AC16_lemmas.thy b/src/ZF/AC/AC16_lemmas.thy --- a/src/ZF/AC/AC16_lemmas.thy +++ b/src/ZF/AC/AC16_lemmas.thy @@ -1,236 +1,236 @@ (* Title: ZF/AC/AC16_lemmas.thy Author: Krzysztof Grabczewski Lemmas used in the proofs concerning AC16 *) theory AC16_lemmas imports AC_Equiv Hartog Cardinal_aux begin lemma cons_Diff_eq: "a\A \ cons(a,A)-{a}=A" by fast lemma nat_1_lepoll_iff: "1\X \ (\x. x \ X)" unfolding lepoll_def apply (rule iffI) apply (fast intro: inj_is_fun [THEN apply_type]) apply (erule exE) apply (rule_tac x = "\a \ 1. x" in exI) apply (fast intro!: lam_injective) done lemma eqpoll_1_iff_singleton: "X\1 \ (\x. X={x})" apply (rule iffI) apply (erule eqpollE) apply (drule nat_1_lepoll_iff [THEN iffD1]) apply (fast intro!: lepoll_1_is_sing) apply (fast intro!: singleton_eqpoll_1) done lemma cons_eqpoll_succ: "\x\n; y\x\ \ cons(y,x)\succ(n)" unfolding succ_def apply (fast elim!: cons_eqpoll_cong mem_irrefl) done lemma subsets_eqpoll_1_eq: "{Y \ Pow(X). Y\1} = {{x}. x \ X}" apply (rule equalityI) apply (rule subsetI) apply (erule CollectE) apply (drule eqpoll_1_iff_singleton [THEN iffD1]) apply (fast intro!: RepFunI) apply (rule subsetI) apply (erule RepFunE) apply (rule CollectI, fast) apply (fast intro!: singleton_eqpoll_1) done lemma eqpoll_RepFun_sing: "X\{{x}. x \ X}" -apply (unfold eqpoll_def bij_def) + unfolding eqpoll_def bij_def apply (rule_tac x = "\x \ X. {x}" in exI) apply (rule IntI) apply (unfold inj_def surj_def, simp) apply (fast intro!: lam_type RepFunI intro: singleton_eq_iff [THEN iffD1], simp) apply (fast intro!: lam_type) done lemma subsets_eqpoll_1_eqpoll: "{Y \ Pow(X). Y\1}\X" apply (rule subsets_eqpoll_1_eq [THEN ssubst]) apply (rule eqpoll_RepFun_sing [THEN eqpoll_sym]) done lemma InfCard_Least_in: "\InfCard(x); y \ x; y \ succ(z)\ \ (\ i. i \ y) \ y" apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, THEN succ_lepoll_imp_not_empty, THEN not_emptyE]) apply (fast intro: LeastI dest!: InfCard_is_Card [THEN Card_is_Ord] elim: Ord_in_Ord) done lemma subsets_lepoll_lemma1: "\InfCard(x); n \ nat\ \ {y \ Pow(x). y\succ(succ(n))} \ x*{y \ Pow(x). y\succ(n)}" unfolding lepoll_def apply (rule_tac x = "\y \ {y \ Pow(x) . y\succ (succ (n))}. <\ i. i \ y, y-{\ i. i \ y}>" in exI) apply (rule_tac d = "\z. cons (fst(z), snd(z))" in lam_injective) apply (blast intro!: Diff_sing_eqpoll intro: InfCard_Least_in) apply (simp, blast intro: InfCard_Least_in) done lemma set_of_Ord_succ_Union: "(\y \ z. Ord(y)) \ z \ succ(\(z))" apply (rule subsetI) apply (case_tac "\y \ z. y \ x", blast ) apply (simp, erule bexE) apply (rule_tac i=y and j=x in Ord_linear_le) apply (blast dest: le_imp_subset elim: leE ltE)+ done lemma subset_not_mem: "j \ i \ i \ j" by (fast elim!: mem_irrefl) lemma succ_Union_not_mem: "(\y. y \ z \ Ord(y)) \ succ(\(z)) \ z" apply (rule set_of_Ord_succ_Union [THEN subset_not_mem], blast) done lemma Union_cons_eq_succ_Union: "\(cons(succ(\(z)),z)) = succ(\(z))" by fast lemma Un_Ord_disj: "\Ord(i); Ord(j)\ \ i \ j = i | i \ j = j" by (fast dest!: le_imp_subset elim: Ord_linear_le) lemma Union_eq_Un: "x \ X \ \(X) = x \ \(X-{x})" by fast lemma Union_in_lemma [rule_format]: "n \ nat \ \z. (\y \ z. Ord(y)) \ z\n \ z\0 \ \(z) \ z" apply (induct_tac "n") apply (fast dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0]) apply (intro allI impI) apply (erule natE) apply (fast dest!: eqpoll_1_iff_singleton [THEN iffD1] intro!: Union_singleton, clarify) apply (elim not_emptyE) apply (erule_tac x = "z-{xb}" in allE) apply (erule impE) apply (fast elim!: Diff_sing_eqpoll Diff_sing_eqpoll [THEN eqpoll_succ_imp_not_empty]) apply (subgoal_tac "xb \ \(z - {xb}) \ z") apply (simp add: Union_eq_Un [symmetric]) apply (frule bspec, assumption) apply (drule bspec) apply (erule Diff_subset [THEN subsetD]) apply (drule Un_Ord_disj, assumption, auto) done lemma Union_in: "\\x \ z. Ord(x); z\n; z\0; n \ nat\ \ \(z) \ z" by (blast intro: Union_in_lemma) lemma succ_Union_in_x: "\InfCard(x); z \ Pow(x); z\n; n \ nat\ \ succ(\(z)) \ x" apply (rule Limit_has_succ [THEN ltE]) prefer 3 apply assumption apply (erule InfCard_is_Limit) apply (case_tac "z=0") apply (simp, fast intro!: InfCard_is_Limit [THEN Limit_has_0]) apply (rule ltI [OF PowD [THEN subsetD] InfCard_is_Card [THEN Card_is_Ord]], assumption) apply (blast intro: Union_in InfCard_is_Card [THEN Card_is_Ord, THEN Ord_in_Ord])+ done lemma succ_lepoll_succ_succ: "\InfCard(x); n \ nat\ \ {y \ Pow(x). y\succ(n)} \ {y \ Pow(x). y\succ(succ(n))}" unfolding lepoll_def apply (rule_tac x = "\z \ {y\Pow(x). y\succ(n)}. cons(succ(\(z)), z)" in exI) apply (rule_tac d = "\z. z-{\(z) }" in lam_injective) apply (blast intro!: succ_Union_in_x succ_Union_not_mem intro: cons_eqpoll_succ Ord_in_Ord dest!: InfCard_is_Card [THEN Card_is_Ord]) apply (simp only: Union_cons_eq_succ_Union) apply (rule cons_Diff_eq) apply (fast dest!: InfCard_is_Card [THEN Card_is_Ord] elim: Ord_in_Ord intro!: succ_Union_not_mem) done lemma subsets_eqpoll_X: "\InfCard(X); n \ nat\ \ {Y \ Pow(X). Y\succ(n)} \ X" apply (induct_tac "n") apply (rule subsets_eqpoll_1_eqpoll) apply (rule eqpollI) apply (rule subsets_lepoll_lemma1 [THEN lepoll_trans], assumption+) apply (rule eqpoll_trans [THEN eqpoll_imp_lepoll]) apply (erule eqpoll_refl [THEN prod_eqpoll_cong]) apply (erule InfCard_square_eqpoll) apply (fast elim: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans] intro!: succ_lepoll_succ_succ) done lemma image_vimage_eq: "\f \ surj(A,B); y \ B\ \ f``(converse(f)``y) = y" unfolding surj_def apply (fast dest: apply_equality2 elim: apply_iff [THEN iffD2]) done lemma vimage_image_eq: "\f \ inj(A,B); y \ A\ \ converse(f)``(f``y) = y" by (fast elim!: inj_is_fun [THEN apply_Pair] dest: inj_equality) lemma subsets_eqpoll: "A\B \ {Y \ Pow(A). Y\n}\{Y \ Pow(B). Y\n}" unfolding eqpoll_def apply (erule exE) apply (rule_tac x = "\X \ {Y \ Pow (A) . \f. f \ bij (Y, n) }. f``X" in exI) apply (rule_tac d = "\Z. converse (f) ``Z" in lam_bijective) apply (fast intro!: bij_is_inj [THEN restrict_bij, THEN bij_converse_bij, THEN comp_bij] elim!: bij_is_fun [THEN fun_is_rel, THEN image_subset]) apply (blast intro!: bij_is_inj [THEN restrict_bij] comp_bij bij_converse_bij bij_is_fun [THEN fun_is_rel, THEN image_subset]) apply (fast elim!: bij_is_inj [THEN vimage_image_eq]) apply (fast elim!: bij_is_surj [THEN image_vimage_eq]) done lemma WO2_imp_ex_Card: "WO2 \ \a. Card(a) \ X\a" unfolding WO2_def apply (drule spec [of _ X]) apply (blast intro: Card_cardinal eqpoll_trans well_ord_Memrel [THEN well_ord_cardinal_eqpoll, THEN eqpoll_sym]) done lemma lepoll_infinite: "\X\Y; \Finite(X)\ \ \Finite(Y)" by (blast intro: lepoll_Finite) lemma infinite_Card_is_InfCard: "\\Finite(X); Card(X)\ \ InfCard(X)" unfolding InfCard_def apply (fast elim!: Card_is_Ord [THEN nat_le_infinite_Ord]) done lemma WO2_infinite_subsets_eqpoll_X: "\WO2; n \ nat; \Finite(X)\ \ {Y \ Pow(X). Y\succ(n)}\X" apply (drule WO2_imp_ex_Card) apply (elim allE exE conjE) apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) apply (drule infinite_Card_is_InfCard, assumption) apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans) done lemma well_ord_imp_ex_Card: "well_ord(X,R) \ \a. Card(a) \ X\a" by (fast elim!: well_ord_cardinal_eqpoll [THEN eqpoll_sym] intro!: Card_cardinal) lemma well_ord_infinite_subsets_eqpoll_X: "\well_ord(X,R); n \ nat; \Finite(X)\ \ {Y \ Pow(X). Y\succ(n)}\X" apply (drule well_ord_imp_ex_Card) apply (elim allE exE conjE) apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) apply (drule infinite_Card_is_InfCard, assumption) apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans) done end diff --git a/src/ZF/AC/AC17_AC1.thy b/src/ZF/AC/AC17_AC1.thy --- a/src/ZF/AC/AC17_AC1.thy +++ b/src/ZF/AC/AC17_AC1.thy @@ -1,301 +1,301 @@ (* Title: ZF/AC/AC17_AC1.thy Author: Krzysztof Grabczewski The equivalence of AC0, AC1 and AC17 Also, the proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent to AC0 and AC1. *) theory AC17_AC1 imports HH begin (** AC0 is equivalent to AC1. AC0 comes from Suppes, AC1 from Rubin \ Rubin **) lemma AC0_AC1_lemma: "\f:(\X \ A. X); D \ A\ \ \g. g:(\X \ D. X)" by (fast intro!: lam_type apply_type) lemma AC0_AC1: "AC0 \ AC1" -apply (unfold AC0_def AC1_def) + unfolding AC0_def AC1_def apply (blast intro: AC0_AC1_lemma) done lemma AC1_AC0: "AC1 \ AC0" by (unfold AC0_def AC1_def, blast) (**** The proof of AC1 \ AC17 ****) lemma AC1_AC17_lemma: "f \ (\X \ Pow(A) - {0}. X) \ f \ (Pow(A) - {0} -> A)" apply (rule Pi_type, assumption) apply (drule apply_type, assumption, fast) done lemma AC1_AC17: "AC1 \ AC17" -apply (unfold AC1_def AC17_def) + unfolding AC1_def AC17_def apply (rule allI) apply (rule ballI) apply (erule_tac x = "Pow (A) -{0}" in allE) apply (erule impE, fast) apply (erule exE) apply (rule bexI) apply (erule_tac [2] AC1_AC17_lemma) apply (rule apply_type, assumption) apply (fast dest!: AC1_AC17_lemma elim!: apply_type) done (**** The proof of AC17 \ AC1 ****) (* *********************************************************************** *) (* more properties of HH *) (* *********************************************************************** *) lemma UN_eq_imp_well_ord: "\x - (\j \ \ i. HH(\X \ Pow(x)-{0}. {f`X}, x, i) = {x}. HH(\X \ Pow(x)-{0}. {f`X}, x, j)) = 0; f \ Pow(x)-{0} -> x\ \ \r. well_ord(x,r)" apply (rule exI) apply (erule well_ord_rvimage [OF bij_Least_HH_x [THEN bij_converse_bij, THEN bij_is_inj] Ord_Least [THEN well_ord_Memrel]], assumption) done (* *********************************************************************** *) (* theorems closer to the proof *) (* *********************************************************************** *) lemma not_AC1_imp_ex: "\AC1 \ \A. \f \ Pow(A)-{0} -> A. \u \ Pow(A)-{0}. f`u \ u" unfolding AC1_def apply (erule swap) apply (rule allI) apply (erule swap) apply (rule_tac x = "\(A)" in exI) apply (blast intro: lam_type) done lemma AC17_AC1_aux1: "\\f \ Pow(x) - {0} -> x. \u \ Pow(x) - {0}. f`u\u; \f \ Pow(x)-{0}->x. x - (\a \ (\ i. HH(\X \ Pow(x)-{0}. {f`X},x,i)={x}). HH(\X \ Pow(x)-{0}. {f`X},x,a)) = 0\ \ P" apply (erule bexE) apply (erule UN_eq_imp_well_ord [THEN exE], assumption) apply (erule ex_choice_fun_Pow [THEN exE]) apply (erule ballE) apply (fast intro: apply_type del: DiffE) apply (erule notE) apply (rule Pi_type, assumption) apply (blast dest: apply_type) done lemma AC17_AC1_aux2: "\ (\f \ Pow(x)-{0}->x. x - F(f) = 0) \ (\f \ Pow(x)-{0}->x . x - F(f)) \ (Pow(x) -{0} -> x) -> Pow(x) - {0}" by (fast intro!: lam_type dest!: Diff_eq_0_iff [THEN iffD1]) lemma AC17_AC1_aux3: "\f`Z \ Z; Z \ Pow(x)-{0}\ \ (\X \ Pow(x)-{0}. {f`X})`Z \ Pow(Z)-{0}" by auto lemma AC17_AC1_aux4: "\f \ F. f`((\f \ F. Q(f))`f) \ (\f \ F. Q(f))`f \ \f \ F. f`Q(f) \ Q(f)" by simp lemma AC17_AC1: "AC17 \ AC1" unfolding AC17_def apply (rule classical) apply (erule not_AC1_imp_ex [THEN exE]) apply (case_tac "\f \ Pow(x)-{0} -> x. x - (\a \ (\ i. HH (\X \ Pow (x) -{0}. {f`X},x,i) ={x}) . HH (\X \ Pow (x) -{0}. {f`X},x,a)) = 0") apply (erule AC17_AC1_aux1, assumption) apply (drule AC17_AC1_aux2) apply (erule allE) apply (drule bspec, assumption) apply (drule AC17_AC1_aux4) apply (erule bexE) apply (drule apply_type, assumption) apply (simp add: HH_Least_eq_x del: Diff_iff ) apply (drule AC17_AC1_aux3, assumption) apply (fast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]] f_subset_imp_HH_subset elim!: mem_irrefl) done (* ********************************************************************** AC1 \ AC2 \ AC1 AC1 \ AC4 \ AC3 \ AC1 AC4 \ AC5 \ AC4 AC1 \ AC6 ************************************************************************* *) (* ********************************************************************** *) (* AC1 \ AC2 *) (* ********************************************************************** *) lemma AC1_AC2_aux1: "\f:(\X \ A. X); B \ A; 0\A\ \ {f`B} \ B \ {f`C. C \ A}" by (fast elim!: apply_type) lemma AC1_AC2_aux2: "\pairwise_disjoint(A); B \ A; C \ A; D \ B; D \ C\ \ f`B = f`C" by (unfold pairwise_disjoint_def, fast) lemma AC1_AC2: "AC1 \ AC2" -apply (unfold AC1_def AC2_def) + unfolding AC1_def AC2_def apply (rule allI) apply (rule impI) apply (elim asm_rl conjE allE exE impE, assumption) apply (intro exI ballI equalityI) prefer 2 apply (rule AC1_AC2_aux1, assumption+) apply (fast elim!: AC1_AC2_aux2 elim: apply_type) done (* ********************************************************************** *) (* AC2 \ AC1 *) (* ********************************************************************** *) lemma AC2_AC1_aux1: "0\A \ 0 \ {B*{B}. B \ A}" by (fast dest!: sym [THEN Sigma_empty_iff [THEN iffD1]]) lemma AC2_AC1_aux2: "\X*{X} \ C = {y}; X \ A\ \ (THE y. X*{X} \ C = {y}): X*A" apply (rule subst_elem [of y]) apply (blast elim!: equalityE) apply (auto simp add: singleton_eq_iff) done lemma AC2_AC1_aux3: "\D \ {E*{E}. E \ A}. \y. D \ C = {y} \ (\x \ A. fst(THE z. (x*{x} \ C = {z}))) \ (\X \ A. X)" apply (rule lam_type) apply (drule bspec, blast) apply (blast intro: AC2_AC1_aux2 fst_type) done lemma AC2_AC1: "AC2 \ AC1" -apply (unfold AC1_def AC2_def pairwise_disjoint_def) + unfolding AC1_def AC2_def pairwise_disjoint_def apply (intro allI impI) apply (elim allE impE) prefer 2 apply (fast elim!: AC2_AC1_aux3) apply (blast intro!: AC2_AC1_aux1) done (* ********************************************************************** *) (* AC1 \ AC4 *) (* ********************************************************************** *) lemma empty_notin_images: "0 \ {R``{x}. x \ domain(R)}" by blast lemma AC1_AC4: "AC1 \ AC4" -apply (unfold AC1_def AC4_def) + unfolding AC1_def AC4_def apply (intro allI impI) apply (drule spec, drule mp [OF _ empty_notin_images]) apply (best intro!: lam_type elim!: apply_type) done (* ********************************************************************** *) (* AC4 \ AC3 *) (* ********************************************************************** *) lemma AC4_AC3_aux1: "f \ A->B \ (\z \ A. {z}*f`z) \ A*\(B)" by (fast dest!: apply_type) lemma AC4_AC3_aux2: "domain(\z \ A. {z}*f(z)) = {a \ A. f(a)\0}" by blast lemma AC4_AC3_aux3: "x \ A \ (\z \ A. {z}*f(z))``{x} = f(x)" by fast lemma AC4_AC3: "AC4 \ AC3" -apply (unfold AC3_def AC4_def) + unfolding AC3_def AC4_def apply (intro allI ballI) apply (elim allE impE) apply (erule AC4_AC3_aux1) apply (simp add: AC4_AC3_aux2 AC4_AC3_aux3 cong add: Pi_cong) done (* ********************************************************************** *) (* AC3 \ AC1 *) (* ********************************************************************** *) lemma AC3_AC1_lemma: "b\A \ (\x \ {a \ A. id(A)`a\b}. id(A)`x) = (\x \ A. x)" apply (simp add: id_def cong add: Pi_cong) apply (rule_tac b = A in subst_context, fast) done lemma AC3_AC1: "AC3 \ AC1" -apply (unfold AC1_def AC3_def) + unfolding AC1_def AC3_def apply (fast intro!: id_type elim: AC3_AC1_lemma [THEN subst]) done (* ********************************************************************** *) (* AC4 \ AC5 *) (* ********************************************************************** *) lemma AC4_AC5: "AC4 \ AC5" -apply (unfold range_def AC4_def AC5_def) + unfolding range_def AC4_def AC5_def apply (intro allI ballI) apply (elim allE impE) apply (erule fun_is_rel [THEN converse_type]) apply (erule exE) apply (rename_tac g) apply (rule_tac x=g in bexI) apply (blast dest: apply_equality range_type) apply (blast intro: Pi_type dest: apply_type fun_is_rel) done (* ********************************************************************** *) (* AC5 \ AC4, Rubin \ Rubin, p. 11 *) (* ********************************************************************** *) lemma AC5_AC4_aux1: "R \ A*B \ (\x \ R. fst(x)) \ R -> A" by (fast intro!: lam_type fst_type) lemma AC5_AC4_aux2: "R \ A*B \ range(\x \ R. fst(x)) = domain(R)" by (unfold lam_def, force) lemma AC5_AC4_aux3: "\\f \ A->C. P(f,domain(f)); A=B\ \ \f \ B->C. P(f,B)" apply (erule bexE) apply (frule domain_of_fun, fast) done lemma AC5_AC4_aux4: "\R \ A*B; g \ C->R; \x \ C. (\z \ R. fst(z))` (g`x) = x\ \ (\x \ C. snd(g`x)): (\x \ C. R``{x})" apply (rule lam_type) apply (force dest: apply_type) done lemma AC5_AC4: "AC5 \ AC4" apply (unfold AC4_def AC5_def, clarify) apply (elim allE ballE) apply (drule AC5_AC4_aux3 [OF _ AC5_AC4_aux2], assumption) apply (fast elim!: AC5_AC4_aux4) apply (blast intro: AC5_AC4_aux1) done (* ********************************************************************** *) (* AC1 \ AC6 *) (* ********************************************************************** *) lemma AC1_iff_AC6: "AC1 \ AC6" by (unfold AC1_def AC6_def, blast) end diff --git a/src/ZF/AC/AC7_AC9.thy b/src/ZF/AC/AC7_AC9.thy --- a/src/ZF/AC/AC7_AC9.thy +++ b/src/ZF/AC/AC7_AC9.thy @@ -1,169 +1,169 @@ (* Title: ZF/AC/AC7_AC9.thy Author: Krzysztof Grabczewski The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous instances of AC. *) theory AC7_AC9 imports AC_Equiv begin (* ********************************************************************** *) (* Lemmas used in the proofs AC7 \ AC6 and AC9 \ AC1 *) (* - Sigma_fun_space_not0 *) (* - Sigma_fun_space_eqpoll *) (* ********************************************************************** *) lemma Sigma_fun_space_not0: "\0\A; B \ A\ \ (nat->\(A)) * B \ 0" by (blast dest!: Sigma_empty_iff [THEN iffD1] Union_empty_iff [THEN iffD1]) lemma inj_lemma: "C \ A \ (\g \ (nat->\(A))*C. (\n \ nat. if(n=0, snd(g), fst(g)`(n #- 1)))) \ inj((nat->\(A))*C, (nat->\(A)) ) " unfolding inj_def apply (rule CollectI) apply (fast intro!: lam_type if_type apply_type fst_type snd_type, auto) apply (rule fun_extension, assumption+) apply (drule lam_eqE [OF _ nat_succI], assumption, simp) apply (drule lam_eqE [OF _ nat_0I], simp) done lemma Sigma_fun_space_eqpoll: "\C \ A; 0\A\ \ (nat->\(A)) * C \ (nat->\(A))" apply (rule eqpollI) apply (simp add: lepoll_def) apply (fast intro!: inj_lemma) apply (fast elim!: prod_lepoll_self not_sym [THEN not_emptyE] subst_elem elim: swap) done (* ********************************************************************** *) (* AC6 \ AC7 *) (* ********************************************************************** *) lemma AC6_AC7: "AC6 \ AC7" by (unfold AC6_def AC7_def, blast) (* ********************************************************************** *) (* AC7 \ AC6, Rubin \ Rubin p. 12, Theorem 2.8 *) (* The case of the empty family of sets added in order to complete *) (* the proof. *) (* ********************************************************************** *) lemma lemma1_1: "y \ (\B \ A. Y*B) \ (\B \ A. snd(y`B)) \ (\B \ A. B)" by (fast intro!: lam_type snd_type apply_type) lemma lemma1_2: "y \ (\B \ {Y*C. C \ A}. B) \ (\B \ A. y`(Y*B)) \ (\B \ A. Y*B)" apply (fast intro!: lam_type apply_type) done lemma AC7_AC6_lemma1: "(\B \ {(nat->\(A))*C. C \ A}. B) \ 0 \ (\B \ A. B) \ 0" by (fast intro!: equals0I lemma1_1 lemma1_2) lemma AC7_AC6_lemma2: "0 \ A \ 0 \ {(nat -> \(A)) * C. C \ A}" by (blast dest: Sigma_fun_space_not0) lemma AC7_AC6: "AC7 \ AC6" -apply (unfold AC6_def AC7_def) + unfolding AC6_def AC7_def apply (rule allI) apply (rule impI) apply (case_tac "A=0", simp) apply (rule AC7_AC6_lemma1) apply (erule allE) apply (blast del: notI intro!: AC7_AC6_lemma2 intro: eqpoll_sym eqpoll_trans Sigma_fun_space_eqpoll) done (* ********************************************************************** *) (* AC1 \ AC8 *) (* ********************************************************************** *) lemma AC1_AC8_lemma1: "\B \ A. \B1 B2. B=\B1,B2\ \ B1 \ B2 \ 0 \ { bij(fst(B),snd(B)). B \ A }" apply (unfold eqpoll_def, auto) done lemma AC1_AC8_lemma2: "\f \ (\X \ RepFun(A,p). X); D \ A\ \ (\x \ A. f`p(x))`D \ p(D)" apply (simp, fast elim!: apply_type) done lemma AC1_AC8: "AC1 \ AC8" -apply (unfold AC1_def AC8_def) + unfolding AC1_def AC8_def apply (fast dest: AC1_AC8_lemma1 AC1_AC8_lemma2) done (* ********************************************************************** *) (* AC8 \ AC9 *) (* - this proof replaces the following two from Rubin \ Rubin: *) (* AC8 \ AC1 and AC1 \ AC9 *) (* ********************************************************************** *) lemma AC8_AC9_lemma: "\B1 \ A. \B2 \ A. B1 \ B2 \ \B \ A*A. \B1 B2. B=\B1,B2\ \ B1 \ B2" by fast lemma AC8_AC9: "AC8 \ AC9" -apply (unfold AC8_def AC9_def) + unfolding AC8_def AC9_def apply (intro allI impI) apply (erule allE) apply (erule impE, erule AC8_AC9_lemma, force) done (* ********************************************************************** *) (* AC9 \ AC1 *) (* The idea of this proof comes from "Equivalents of the Axiom of Choice" *) (* by Rubin \ Rubin. But (x * y) is not necessarily equipollent to *) (* (x * y) \ {0} when y is a set of total functions acting from nat to *) (* \(A) -- therefore we have used the set (y * nat) instead of y. *) (* ********************************************************************** *) lemma snd_lepoll_SigmaI: "b \ B \ X \ B \ X" by (blast intro: lepoll_trans prod_lepoll_self eqpoll_imp_lepoll prod_commute_eqpoll) lemma nat_lepoll_lemma: "\0 \ A; B \ A\ \ nat \ ((nat \ \(A)) \ B) \ nat" by (blast dest: Sigma_fun_space_not0 intro: snd_lepoll_SigmaI) lemma AC9_AC1_lemma1: "\0\A; A\0; C = {((nat->\(A))*B)*nat. B \ A} \ {cons(0,((nat->\(A))*B)*nat). B \ A}; B1 \ C; B2 \ C\ \ B1 \ B2" by (blast intro!: nat_lepoll_lemma Sigma_fun_space_eqpoll nat_cons_eqpoll [THEN eqpoll_trans] prod_eqpoll_cong [OF _ eqpoll_refl] intro: eqpoll_trans eqpoll_sym ) lemma AC9_AC1_lemma2: "\B1 \ {(F*B)*N. B \ A} \ {cons(0,(F*B)*N). B \ A}. \B2 \ {(F*B)*N. B \ A} \ {cons(0,(F*B)*N). B \ A}. f`\B1,B2\ \ bij(B1, B2) \ (\B \ A. snd(fst((f`)`0))) \ (\X \ A. X)" apply (intro lam_type snd_type fst_type) apply (rule apply_type [OF _ consI1]) apply (fast intro!: fun_weaken_type bij_is_fun) done lemma AC9_AC1: "AC9 \ AC1" -apply (unfold AC1_def AC9_def) + unfolding AC1_def AC9_def apply (intro allI impI) apply (erule allE) apply (case_tac "A\0") apply (blast dest: AC9_AC1_lemma1 AC9_AC1_lemma2, force) done end diff --git a/src/ZF/AC/DC.thy b/src/ZF/AC/DC.thy --- a/src/ZF/AC/DC.thy +++ b/src/ZF/AC/DC.thy @@ -1,585 +1,585 @@ (* Title: ZF/AC/DC.thy Author: Krzysztof Grabczewski The proofs concerning the Axiom of Dependent Choice. *) theory DC imports AC_Equiv Hartog Cardinal_aux begin lemma RepFun_lepoll: "Ord(a) \ {P(b). b \ a} \ a" unfolding lepoll_def apply (rule_tac x = "\z \ RepFun (a,P) . \ i. z=P (i) " in exI) apply (rule_tac d="\z. P (z)" in lam_injective) apply (fast intro!: Least_in_Ord) apply (rule sym) apply (fast intro: LeastI Ord_in_Ord) done text\Trivial in the presence of AC, but here we need a wellordering of X\ lemma image_Ord_lepoll: "\f \ X->Y; Ord(X)\ \ f``X \ X" unfolding lepoll_def apply (rule_tac x = "\x \ f``X. \ y. f`y = x" in exI) apply (rule_tac d = "\z. f`z" in lam_injective) apply (fast intro!: Least_in_Ord apply_equality, clarify) apply (rule LeastI) apply (erule apply_equality, assumption+) apply (blast intro: Ord_in_Ord) done lemma range_subset_domain: "\R \ X*X; \g. g \ X \ \u. \g,u\ \ R\ \ range(R) \ domain(R)" by blast lemma cons_fun_type: "g \ n->X \ cons(\n,x\, g) \ succ(n) -> cons(x, X)" unfolding succ_def apply (fast intro!: fun_extend elim!: mem_irrefl) done lemma cons_fun_type2: "\g \ n->X; x \ X\ \ cons(\n,x\, g) \ succ(n) -> X" by (erule cons_absorb [THEN subst], erule cons_fun_type) lemma cons_image_n: "n \ nat \ cons(\n,x\, g)``n = g``n" by (fast elim!: mem_irrefl) lemma cons_val_n: "g \ n->X \ cons(\n,x\, g)`n = x" by (fast intro!: apply_equality elim!: cons_fun_type) lemma cons_image_k: "k \ n \ cons(\n,x\, g)``k = g``k" by (fast elim: mem_asym) lemma cons_val_k: "\k \ n; g \ n->X\ \ cons(\n,x\, g)`k = g`k" by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair) lemma domain_cons_eq_succ: "domain(f)=x \ domain(cons(\x,y\, f)) = succ(x)" by (simp add: domain_cons succ_def) lemma restrict_cons_eq: "g \ n->X \ restrict(cons(\n,x\, g), n) = g" apply (simp add: restrict_def Pi_iff) apply (blast intro: elim: mem_irrefl) done lemma succ_in_succ: "\Ord(k); i \ k\ \ succ(i) \ succ(k)" apply (rule Ord_linear [of "succ(i)" "succ(k)", THEN disjE]) apply (fast elim: Ord_in_Ord mem_irrefl mem_asym)+ done lemma restrict_eq_imp_val_eq: "\restrict(f, domain(g)) = g; x \ domain(g)\ \ f`x = g`x" by (erule subst, simp add: restrict) lemma domain_eq_imp_fun_type: "\domain(f)=A; f \ B->C\ \ f \ A->C" by (frule domain_of_fun, fast) lemma ex_in_domain: "\R \ A * B; R \ 0\ \ \x. x \ domain(R)" by (fast elim!: not_emptyE) definition DC :: "i \ o" where "DC(a) \ \X R. R \ Pow(X)*X \ (\Y \ Pow(X). Y \ a \ (\x \ X. \Y,x\ \ R)) \ (\f \ a->X. \b \ R)" definition DC0 :: o where "DC0 \ \A B R. R \ A*B \ R\0 \ range(R) \ domain(R) \ (\f \ nat->domain(R). \n \ nat. :R)" definition ff :: "[i, i, i, i] \ i" where "ff(b, X, Q, R) \ transrec(b, \c r. THE x. first(x, {x \ X. \ R}, Q))" locale DC0_imp = fixes XX and RR and X and R assumes all_ex: "\Y \ Pow(X). Y \ nat \ (\x \ X. \Y, x\ \ R)" defines XX_def: "XX \ (\n \ nat. {f \ n->X. \k \ n. \ R})" and RR_def: "RR \ {\z1,z2\:XX*XX. domain(z2)=succ(domain(z1)) \ restrict(z2, domain(z1)) = z1}" begin (* ********************************************************************** *) (* DC \ DC(omega) *) (* *) (* The scheme of the proof: *) (* *) (* Assume DC. Let R and X satisfy the premise of DC(omega). *) (* *) (* Define XX and RR as follows: *) (* *) (* XX = (\n \ nat. {f \ n->X. \k \ n. \ R}) *) (* f RR g iff domain(g)=succ(domain(f)) \ *) (* restrict(g, domain(f)) = f *) (* *) (* Then RR satisfies the hypotheses of DC. *) (* So applying DC: *) (* *) (* \f \ nat->XX. \n \ nat. f`n RR f`succ(n) *) (* *) (* Thence *) (* *) (* ff = {. n \ nat} *) (* *) (* is the desired function. *) (* *) (* ********************************************************************** *) lemma lemma1_1: "RR \ XX*XX" by (unfold RR_def, fast) lemma lemma1_2: "RR \ 0" -apply (unfold RR_def XX_def) + unfolding RR_def XX_def apply (rule all_ex [THEN ballE]) apply (erule_tac [2] notE [OF _ empty_subsetI [THEN PowI]]) apply (erule_tac impE [OF _ nat_0I [THEN n_lesspoll_nat]]) apply (erule bexE) apply (rule_tac a = "<0, {\0, x\}>" in not_emptyI) apply (rule CollectI) apply (rule SigmaI) apply (rule nat_0I [THEN UN_I]) apply (simp (no_asm_simp) add: nat_0I [THEN UN_I]) apply (rule nat_1I [THEN UN_I]) apply (force intro!: singleton_fun [THEN Pi_type] simp add: singleton_0 [symmetric]) apply (simp add: singleton_0) done lemma lemma1_3: "range(RR) \ domain(RR)" -apply (unfold RR_def XX_def) + unfolding RR_def XX_def apply (rule range_subset_domain, blast, clarify) apply (frule fun_is_rel [THEN image_subset, THEN PowI, THEN all_ex [THEN bspec]]) apply (erule impE[OF _ lesspoll_trans1[OF image_Ord_lepoll [OF _ nat_into_Ord] n_lesspoll_nat]], assumption+) apply (erule bexE) apply (rule_tac x = "cons (\n,x\, g) " in exI) apply (rule CollectI) apply (force elim!: cons_fun_type2 simp add: cons_image_n cons_val_n cons_image_k cons_val_k) apply (simp add: domain_of_fun succ_def restrict_cons_eq) done lemma lemma2: "\\n \ nat. \ RR; f \ nat -> XX; n \ nat\ \ \k \ nat. f`succ(n) \ k -> X \ n \ k \ \ R" apply (induct_tac "n") apply (drule apply_type [OF _ nat_1I]) apply (drule bspec [OF _ nat_0I]) apply (simp add: XX_def, safe) apply (rule rev_bexI, assumption) apply (subgoal_tac "0 \ y", force) apply (force simp add: RR_def intro: ltD elim!: nat_0_le [THEN leE]) (** LEVEL 7, other subgoal **) apply (drule bspec [OF _ nat_succI], assumption) apply (subgoal_tac "f ` succ (succ (x)) \ succ (k) ->X") apply (drule apply_type [OF _ nat_succI [THEN nat_succI]], assumption) apply (simp (no_asm_use) add: XX_def RR_def) apply safe apply (frule_tac a="succ(k)" in domain_of_fun [symmetric, THEN trans], assumption) apply (frule_tac a=y in domain_of_fun [symmetric, THEN trans], assumption) apply (fast elim!: nat_into_Ord [THEN succ_in_succ] dest!: bspec [OF _ nat_into_Ord [THEN succ_in_succ]]) apply (drule domain_of_fun) apply (simp add: XX_def RR_def, clarify) apply (blast dest: domain_of_fun [symmetric, THEN trans] ) done lemma lemma3_1: "\\n \ nat. \ RR; f \ nat -> XX; m \ nat\ \ {f`succ(x)`x. x \ m} = {f`succ(m)`x. x \ m}" apply (subgoal_tac "\x \ m. f`succ (m) `x = f`succ (x) `x") apply simp apply (induct_tac "m", blast) apply (rule ballI) apply (erule succE) apply (rule restrict_eq_imp_val_eq) apply (drule bspec [OF _ nat_succI], assumption) apply (simp add: RR_def) apply (drule lemma2, assumption+) apply (fast dest!: domain_of_fun) apply (drule_tac x = xa in bspec, assumption) apply (erule sym [THEN trans, symmetric]) apply (rule restrict_eq_imp_val_eq [symmetric]) apply (drule bspec [OF _ nat_succI], assumption) apply (simp add: RR_def) apply (drule lemma2, assumption+) apply (blast dest!: domain_of_fun intro: nat_into_Ord OrdmemD [THEN subsetD]) done lemma lemma3: "\\n \ nat. \ RR; f \ nat -> XX; m \ nat\ \ (\x \ nat. f`succ(x)`x) `` m = f`succ(m)``m" apply (erule natE, simp) apply (subst image_lam) apply (fast elim!: OrdmemD [OF nat_succI Ord_nat]) apply (subst lemma3_1, assumption+) apply fast apply (fast dest!: lemma2 elim!: image_fun [symmetric, OF _ OrdmemD [OF _ nat_into_Ord]]) done end theorem DC0_imp_DC_nat: "DC0 \ DC(nat)" apply (unfold DC_def DC0_def, clarify) apply (elim allE) apply (erule impE) (*these three results comprise Lemma 1*) apply (blast intro!: DC0_imp.lemma1_1 [OF DC0_imp.intro] DC0_imp.lemma1_2 [OF DC0_imp.intro] DC0_imp.lemma1_3 [OF DC0_imp.intro]) apply (erule bexE) apply (rule_tac x = "\n \ nat. f`succ (n) `n" in rev_bexI) apply (rule lam_type, blast dest!: DC0_imp.lemma2 [OF DC0_imp.intro] intro: fun_weaken_type) apply (rule oallI) apply (frule DC0_imp.lemma2 [OF DC0_imp.intro], assumption) apply (blast intro: fun_weaken_type) apply (erule ltD) (** LEVEL 11: last subgoal **) apply (subst DC0_imp.lemma3 [OF DC0_imp.intro], assumption+) apply (fast elim!: fun_weaken_type) apply (erule ltD) apply (force simp add: lt_def) done (* ************************************************************************ DC(omega) \ DC The scheme of the proof: Assume DC(omega). Let R and x satisfy the premise of DC. Define XX and RR as follows: XX = (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R}) RR = {\z1,z2\:Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(z2, domain(f)) = f)) | (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(g, domain(f)) = f)) \ z2={\0,x\})} Then XX and RR satisfy the hypotheses of DC(omega). So applying DC: \f \ nat->XX. \n \ nat. f``n RR f`n Thence ff = {. n \ nat} is the desired function. ************************************************************************* *) lemma singleton_in_funs: "x \ X \ {\0,x\} \ (\n \ nat. {f \ succ(n)->X. \k \ n. \ R})" apply (rule nat_0I [THEN UN_I]) apply (force simp add: singleton_0 [symmetric] intro!: singleton_fun [THEN Pi_type]) done locale imp_DC0 = fixes XX and RR and x and R and f and allRR defines XX_def: "XX \ (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R})" and RR_def: "RR \ {\z1,z2\:Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(z2, domain(f)) = f)) | (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(g, domain(f)) = f)) \ z2={\0,x\})}" and allRR_def: "allRR \ \b \ {\z1,z2\\Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. domain(f)) = b \ (\f \ z1. restrict(z2,domain(f)) = f))}" begin lemma lemma4: "\range(R) \ domain(R); x \ domain(R)\ \ RR \ Pow(XX)*XX \ (\Y \ Pow(XX). Y \ nat \ (\x \ XX. \Y,x\:RR))" apply (rule conjI) apply (force dest!: FinD [THEN PowI] simp add: RR_def) apply (rule impI [THEN ballI]) apply (drule Finite_Fin [OF lesspoll_nat_is_Finite PowD], assumption) apply (case_tac "\g \ XX. domain (g) = succ(\f \ Y. domain(f)) \ (\f\Y. restrict(g, domain(f)) = f)") apply (simp add: RR_def, blast) apply (safe del: domainE) -apply (unfold XX_def RR_def) + unfolding XX_def RR_def apply (rule rev_bexI, erule singleton_in_funs) apply (simp add: nat_0I [THEN rev_bexI] cons_fun_type2) done lemma UN_image_succ_eq: "\f \ nat->X; n \ nat\ \ (\x \ f``succ(n). P(x)) = P(f`n) \ (\x \ f``n. P(x))" by (simp add: image_fun OrdmemD) lemma UN_image_succ_eq_succ: "\(\x \ f``n. P(x)) = y; P(f`n) = succ(y); f \ nat -> X; n \ nat\ \ (\x \ f``succ(n). P(x)) = succ(y)" by (simp add: UN_image_succ_eq, blast) lemma apply_domain_type: "\h \ succ(n) -> D; n \ nat; domain(h)=succ(y)\ \ h`y \ D" by (fast elim: apply_type dest!: trans [OF sym domain_of_fun]) lemma image_fun_succ: "\h \ nat -> X; n \ nat\ \ h``succ(n) = cons(h`n, h``n)" by (simp add: image_fun OrdmemD) lemma f_n_type: "\domain(f`n) = succ(k); f \ nat -> XX; n \ nat\ \ f`n \ succ(k) -> domain(R)" unfolding XX_def apply (drule apply_type, assumption) apply (fast elim: domain_eq_imp_fun_type) done lemma f_n_pairs_in_R [rule_format]: "\h \ nat -> XX; domain(h`n) = succ(k); n \ nat\ \ \i \ k. \ R" unfolding XX_def apply (drule apply_type, assumption) apply (elim UN_E CollectE) apply (drule domain_of_fun [symmetric, THEN trans], assumption, simp) done lemma restrict_cons_eq_restrict: "\restrict(h, domain(u))=u; h \ n->X; domain(u) \ n\ \ restrict(cons(\n, y\, h), domain(u)) = u" unfolding restrict_def apply (simp add: restrict_def Pi_iff) apply (erule sym [THEN trans, symmetric]) apply (blast elim: mem_irrefl) done lemma all_in_image_restrict_eq: "\\x \ f``n. restrict(f`n, domain(x))=x; f \ nat -> XX; n \ nat; domain(f`n) = succ(n); (\x \ f``n. domain(x)) \ n\ \ \x \ f``succ(n). restrict(cons(, f`n), domain(x)) = x" apply (rule ballI) apply (simp add: image_fun_succ) apply (drule f_n_type, assumption+) apply (erule disjE) apply (simp add: domain_of_fun restrict_cons_eq) apply (blast intro!: restrict_cons_eq_restrict) done lemma simplify_recursion: "\\b \ RR; f \ nat -> XX; range(R) \ domain(R); x \ domain(R)\ \ allRR" -apply (unfold RR_def allRR_def) + unfolding RR_def allRR_def apply (rule oallI, drule ltD) apply (erule nat_induct) apply (drule_tac x=0 in ospec, blast intro: Limit_has_0) apply (force simp add: singleton_fun [THEN domain_of_fun] singleton_in_funs) (*induction step*) (** LEVEL 5 **) (*prevent simplification of \\ to \\ *) apply (simp only: separation split) apply (drule_tac x="succ(xa)" in ospec, blast intro: ltI) apply (elim conjE disjE) apply (force elim!: trans subst_context intro!: UN_image_succ_eq_succ) apply (erule notE) apply (simp add: XX_def UN_image_succ_eq_succ) apply (elim conjE bexE) apply (drule apply_domain_type, assumption+) apply (erule domainE)+ apply (frule f_n_type) apply (simp add: XX_def, assumption+) apply (rule rev_bexI, erule nat_succI) apply (rename_tac m i j y z) apply (rule_tac x = "cons(, f`m)" in bexI) prefer 2 apply (blast intro: cons_fun_type2) apply (rule conjI) prefer 2 apply (fast del: ballI subsetI elim: trans [OF _ subst_context, THEN domain_cons_eq_succ] subst_context all_in_image_restrict_eq [simplified XX_def] trans equalityD1) (*one remaining subgoal*) apply (rule ballI) apply (erule succE) (** LEVEL 25 **) apply (simp add: cons_val_n cons_val_k) (*assumption+ will not perform the required backtracking!*) apply (drule f_n_pairs_in_R [simplified XX_def, OF _ domain_of_fun], assumption, assumption, assumption) apply (simp add: nat_into_Ord [THEN succ_in_succ] succI2 cons_val_k) done lemma lemma2: "\allRR; f \ nat->XX; range(R) \ domain(R); x \ domain(R); n \ nat\ \ f`n \ succ(n) -> domain(R) \ (\i \ n. :R)" unfolding allRR_def apply (drule ospec) apply (erule ltI [OF _ Ord_nat]) apply (erule CollectE, simp) apply (rule conjI) prefer 2 apply (fast elim!: f_n_pairs_in_R trans subst_context) unfolding XX_def apply (fast elim!: trans [THEN domain_eq_imp_fun_type] subst_context) done lemma lemma3: "\allRR; f \ nat->XX; n\nat; range(R) \ domain(R); x \ domain(R)\ \ f`n`n = f`succ(n)`n" apply (frule lemma2 [THEN conjunct1, THEN domain_of_fun], assumption+) unfolding allRR_def apply (drule ospec) apply (drule ltI [OF nat_succI Ord_nat], assumption, simp) apply (elim conjE ballE) apply (erule restrict_eq_imp_val_eq [symmetric], force) apply (simp add: image_fun OrdmemD) done end theorem DC_nat_imp_DC0: "DC(nat) \ DC0" -apply (unfold DC_def DC0_def) + unfolding DC_def DC0_def apply (intro allI impI) apply (erule asm_rl conjE ex_in_domain [THEN exE] allE)+ apply (erule impE [OF _ imp_DC0.lemma4], assumption+) apply (erule bexE) apply (drule imp_DC0.simplify_recursion, assumption+) apply (rule_tac x = "\n \ nat. f`n`n" in bexI) apply (rule_tac [2] lam_type) apply (erule_tac [2] apply_type [OF imp_DC0.lemma2 [THEN conjunct1] succI1]) apply (rule ballI) apply (frule_tac n="succ(n)" in imp_DC0.lemma2, (assumption|erule nat_succI)+) apply (drule imp_DC0.lemma3, auto) done (* ********************************************************************** *) (* \K. Card(K) \ DC(K) \ WO3 *) (* ********************************************************************** *) lemma fun_Ord_inj: "\f \ a->X; Ord(a); \b c. \b a\ \ f`b\f`c\ \ f \ inj(a, X)" apply (unfold inj_def, simp) apply (intro ballI impI) apply (rule_tac j=x in Ord_in_Ord [THEN Ord_linear_lt], assumption+) apply (blast intro: Ord_in_Ord, auto) apply (atomize, blast dest: not_sym) done lemma value_in_image: "\f \ X->Y; A \ X; a \ A\ \ f`a \ f``A" by (fast elim!: image_fun [THEN ssubst]) lemma lesspoll_lemma: "\\ A \ B; C \ B\ \ A - C \ 0" unfolding lesspoll_def apply (fast dest!: Diff_eq_0_iff [THEN iffD1, THEN subset_imp_lepoll] intro!: eqpollI elim: notE elim!: eqpollE lepoll_trans) done theorem DC_WO3: "(\K. Card(K) \ DC(K)) \ WO3" -apply (unfold DC_def WO3_def) + unfolding DC_def WO3_def apply (rule allI) apply (case_tac "A \ Hartog (A)") apply (fast dest!: lesspoll_imp_ex_lt_eqpoll intro!: Ord_Hartog leI [THEN le_imp_subset]) apply (erule allE impE)+ apply (rule Card_Hartog) apply (erule_tac x = A in allE) apply (erule_tac x = "{\z1,z2\ \ Pow (A) *A . z1 \ Hartog (A) \ z2 \ z1}" in allE) apply simp apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE]) apply (erule bexE) apply (rule Hartog_lepoll_selfE) apply (rule lepoll_def [THEN def_imp_iff, THEN iffD2]) apply (rule exI, rule fun_Ord_inj, assumption, rule Ord_Hartog) apply (drule value_in_image) apply (drule OrdmemD, rule Ord_Hartog, assumption+, erule ltD) apply (drule ospec) apply (blast intro: ltI Ord_Hartog, force) done (* ********************************************************************** *) (* WO1 \ \K. Card(K) \ DC(K) *) (* ********************************************************************** *) lemma images_eq: "\\x \ A. f`x=g`x; f \ Df->Cf; g \ Dg->Cg; A \ Df; A \ Dg\ \ f``A = g``A" apply (simp (no_asm_simp) add: image_fun) done lemma lam_images_eq: "\Ord(a); b \ a\ \ (\x \ a. h(x))``b = (\x \ b. h(x))``b" apply (rule images_eq) apply (rule ballI) apply (drule OrdmemD [THEN subsetD], assumption+) apply simp apply (fast elim!: RepFunI OrdmemD intro!: lam_type)+ done lemma lam_type_RepFun: "(\b \ a. h(b)) \ a -> {h(b). b \ a}" by (fast intro!: lam_type RepFunI) lemma lemmaX: "\\Y \ Pow(X). Y \ K \ (\x \ X. \Y, x\ \ R); b \ K; Z \ Pow(X); Z \ K\ \ {x \ X. \Z,x\ \ R} \ 0" by blast lemma WO1_DC_lemma: "\Card(K); well_ord(X,Q); \Y \ Pow(X). Y \ K \ (\x \ X. \Y, x\ \ R); b \ K\ \ ff(b, X, Q, R) \ {x \ X. <(\c \ b. ff(c, X, Q, R))``b, x> \ R}" apply (rule_tac P = "b \ K" in impE, (erule_tac [2] asm_rl)+) apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct], assumption+) apply (rule impI) apply (rule ff_def [THEN def_transrec, THEN ssubst]) apply (erule the_first_in, fast) apply (simp add: image_fun [OF lam_type_RepFun subset_refl]) apply (erule lemmaX, assumption) apply (blast intro: Card_is_Ord OrdmemD [THEN subsetD]) apply (blast intro: lesspoll_trans1 in_Card_imp_lesspoll RepFun_lepoll) done theorem WO1_DC_Card: "WO1 \ \K. Card(K) \ DC(K)" -apply (unfold DC_def WO1_def) + unfolding DC_def WO1_def apply (rule allI impI)+ apply (erule allE exE conjE)+ apply (rule_tac x = "\b \ K. ff (b, X, Ra, R) " in bexI) apply (simp add: lam_images_eq [OF Card_is_Ord ltD]) apply (fast elim!: ltE WO1_DC_lemma [THEN CollectD2]) apply (rule_tac lam_type) apply (rule WO1_DC_lemma [THEN CollectD1], assumption+) done end diff --git a/src/ZF/AC/HH.thy b/src/ZF/AC/HH.thy --- a/src/ZF/AC/HH.thy +++ b/src/ZF/AC/HH.thy @@ -1,245 +1,245 @@ (* Title: ZF/AC/HH.thy Author: Krzysztof Grabczewski Some properties of the recursive definition of HH used in the proofs of AC17 \ AC1 AC1 \ WO2 AC15 \ WO6 *) theory HH imports AC_Equiv Hartog begin definition HH :: "[i, i, i] \ i" where "HH(f,x,a) \ transrec(a, \b r. let z = x - (\c \ b. r`c) in if f`z \ Pow(z)-{0} then f`z else {x})" subsection\Lemmas useful in each of the three proofs\ lemma HH_def_satisfies_eq: "HH(f,x,a) = (let z = x - (\b \ a. HH(f,x,b)) in if f`z \ Pow(z)-{0} then f`z else {x})" by (rule HH_def [THEN def_transrec, THEN trans], simp) lemma HH_values: "HH(f,x,a) \ Pow(x)-{0} | HH(f,x,a)={x}" apply (rule HH_def_satisfies_eq [THEN ssubst]) apply (simp add: Let_def Diff_subset [THEN PowI], fast) done lemma subset_imp_Diff_eq: "B \ A \ X-(\a \ A. P(a)) = X-(\a \ A-B. P(a))-(\b \ B. P(b))" by fast lemma Ord_DiffE: "\c \ a-b; b \ c=b | b cy. y\A \ P(y) = {x}) \ x - (\y \ A. P(y)) = x" by (simp, fast elim!: mem_irrefl) lemma HH_eq: "x - (\b \ a. HH(f,x,b)) = x - (\b \ a1. HH(f,x,b)) \ HH(f,x,a) = HH(f,x,a1)" apply (subst HH_def_satisfies_eq [of _ _ a1]) apply (rule HH_def_satisfies_eq [THEN trans], simp) done lemma HH_is_x_gt_too: "\HH(f,x,b)={x}; b \ HH(f,x,a)={x}" apply (rule_tac P = "bHH(f,x,a) \ Pow(x)-{0}; b \ HH(f,x,b) \ Pow(x)-{0}" apply (rule HH_values [THEN disjE], assumption) apply (drule HH_is_x_gt_too, assumption) apply (drule subst, assumption) apply (fast elim!: mem_irrefl) done lemma HH_subset_x_imp_subset_Diff_UN: "HH(f,x,a) \ Pow(x)-{0} \ HH(f,x,a) \ Pow(x - (\b \ a. HH(f,x,b)))-{0}" apply (drule HH_def_satisfies_eq [THEN subst]) apply (rule HH_def_satisfies_eq [THEN ssubst]) apply (simp add: Let_def Diff_subset [THEN PowI]) apply (drule split_if [THEN iffD1]) apply (fast elim!: mem_irrefl) done lemma HH_eq_arg_lt: "\HH(f,x,v)=HH(f,x,w); HH(f,x,v) \ Pow(x)-{0}; v \ w\ \ P" apply (frule_tac P = "\y. y \ Pow (x) -{0}" in subst, assumption) apply (drule_tac a = w in HH_subset_x_imp_subset_Diff_UN) apply (drule subst_elem, assumption) apply (fast intro!: singleton_iff [THEN iffD2] equals0I) done lemma HH_eq_imp_arg_eq: "\HH(f,x,v)=HH(f,x,w); HH(f,x,w) \ Pow(x)-{0}; Ord(v); Ord(w)\ \ v=w" apply (rule_tac j = w in Ord_linear_lt) apply (simp_all (no_asm_simp)) apply (drule subst_elem, assumption) apply (blast dest: ltD HH_eq_arg_lt) apply (blast dest: HH_eq_arg_lt [OF sym] ltD) done lemma HH_subset_x_imp_lepoll: "\HH(f, x, i) \ Pow(x)-{0}; Ord(i)\ \ i \ Pow(x)-{0}" -apply (unfold lepoll_def inj_def) + unfolding lepoll_def inj_def apply (rule_tac x = "\j \ i. HH (f, x, j) " in exI) apply (simp (no_asm_simp)) apply (fast del: DiffE elim!: HH_eq_imp_arg_eq Ord_in_Ord HH_subset_x_lt_too intro!: lam_type ballI ltI intro: bexI) done lemma HH_Hartog_is_x: "HH(f, x, Hartog(Pow(x)-{0})) = {x}" apply (rule HH_values [THEN disjE]) prefer 2 apply assumption apply (fast del: DiffE intro!: Ord_Hartog dest!: HH_subset_x_imp_lepoll elim!: Hartog_lepoll_selfE) done lemma HH_Least_eq_x: "HH(f, x, \ i. HH(f, x, i) = {x}) = {x}" by (fast intro!: Ord_Hartog HH_Hartog_is_x LeastI) lemma less_Least_subset_x: "a \ (\ i. HH(f,x,i)={x}) \ HH(f,x,a) \ Pow(x)-{0}" apply (rule HH_values [THEN disjE], assumption) apply (rule less_LeastE) apply (erule_tac [2] ltI [OF _ Ord_Least], assumption) done subsection\Lemmas used in the proofs of AC1 \ WO2 and AC17 \ AC1\ lemma lam_Least_HH_inj_Pow: "(\a \ (\ i. HH(f,x,i)={x}). HH(f,x,a)) \ inj(\ i. HH(f,x,i)={x}, Pow(x)-{0})" apply (unfold inj_def, simp) apply (fast intro!: lam_type dest: less_Least_subset_x elim!: HH_eq_imp_arg_eq Ord_Least [THEN Ord_in_Ord]) done lemma lam_Least_HH_inj: "\a \ (\ i. HH(f,x,i)={x}). \z \ x. HH(f,x,a) = {z} \ (\a \ (\ i. HH(f,x,i)={x}). HH(f,x,a)) \ inj(\ i. HH(f,x,i)={x}, {{y}. y \ x})" by (rule lam_Least_HH_inj_Pow [THEN inj_strengthen_type], simp) lemma lam_surj_sing: "\x - (\a \ A. F(a)) = 0; \a \ A. \z \ x. F(a) = {z}\ \ (\a \ A. F(a)) \ surj(A, {{y}. y \ x})" apply (simp add: surj_def lam_type Diff_eq_0_iff) apply (blast elim: equalityE) done lemma not_emptyI2: "y \ Pow(x)-{0} \ x \ 0" by auto lemma f_subset_imp_HH_subset: "f`(x - (\j \ i. HH(f,x,j))) \ Pow(x - (\j \ i. HH(f,x,j)))-{0} \ HH(f, x, i) \ Pow(x) - {0}" apply (rule HH_def_satisfies_eq [THEN ssubst]) apply (simp add: Let_def Diff_subset [THEN PowI] not_emptyI2 [THEN if_P], fast) done lemma f_subsets_imp_UN_HH_eq_x: "\z \ Pow(x)-{0}. f`z \ Pow(z)-{0} \ x - (\j \ (\ i. HH(f,x,i)={x}). HH(f,x,j)) = 0" apply (case_tac "P \ {0}" for P, fast) apply (drule Diff_subset [THEN PowI, THEN DiffI]) apply (drule bspec, assumption) apply (drule f_subset_imp_HH_subset) apply (blast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]] elim!: mem_irrefl) done lemma HH_values2: "HH(f,x,i) = f`(x - (\j \ i. HH(f,x,j))) | HH(f,x,i)={x}" apply (rule HH_def_satisfies_eq [THEN ssubst]) apply (simp add: Let_def Diff_subset [THEN PowI]) done lemma HH_subset_imp_eq: "HH(f,x,i): Pow(x)-{0} \ HH(f,x,i)=f`(x - (\j \ i. HH(f,x,j)))" apply (rule HH_values2 [THEN disjE], assumption) apply (fast elim!: equalityE mem_irrefl dest!: singleton_subsetD) done lemma f_sing_imp_HH_sing: "\f \ (Pow(x)-{0}) -> {{z}. z \ x}; a \ (\ i. HH(f,x,i)={x})\ \ \z \ x. HH(f,x,a) = {z}" apply (drule less_Least_subset_x) apply (frule HH_subset_imp_eq) apply (drule apply_type) apply (rule Diff_subset [THEN PowI, THEN DiffI]) apply (fast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2], force) done lemma f_sing_lam_bij: "\x - (\j \ (\ i. HH(f,x,i)={x}). HH(f,x,j)) = 0; f \ (Pow(x)-{0}) -> {{z}. z \ x}\ \ (\a \ (\ i. HH(f,x,i)={x}). HH(f,x,a)) \ bij(\ i. HH(f,x,i)={x}, {{y}. y \ x})" unfolding bij_def apply (fast intro!: lam_Least_HH_inj lam_surj_sing f_sing_imp_HH_sing) done lemma lam_singI: "f \ (\X \ Pow(x)-{0}. F(X)) \ (\X \ Pow(x)-{0}. {f`X}) \ (\X \ Pow(x)-{0}. {{z}. z \ F(X)})" by (fast del: DiffI DiffE intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type) (*FIXME: both uses have the form ...[THEN bij_converse_bij], so simplification is needed!*) lemmas bij_Least_HH_x = comp_bij [OF f_sing_lam_bij [OF _ lam_singI] lam_sing_bij [THEN bij_converse_bij]] subsection\The proof of AC1 \ WO2\ (*Establishing the existence of a bijection, namely converse (converse(\x\x. {x}) O Lambda (\ i. HH(\X\Pow(x) - {0}. {f ` X}, x, i) = {x}, HH(\X\Pow(x) - {0}. {f ` X}, x))) Perhaps it could be simplified. *) lemma bijection: "f \ (\X \ Pow(x) - {0}. X) \ \g. g \ bij(x, \ i. HH(\X \ Pow(x)-{0}. {f`X}, x, i) = {x})" apply (rule exI) apply (rule bij_Least_HH_x [THEN bij_converse_bij]) apply (rule f_subsets_imp_UN_HH_eq_x) apply (intro ballI apply_type) apply (fast intro: lam_type apply_type del: DiffE, assumption) apply (fast intro: Pi_weaken_type) done lemma AC1_WO2: "AC1 \ WO2" -apply (unfold AC1_def WO2_def eqpoll_def) + unfolding AC1_def WO2_def eqpoll_def apply (intro allI) apply (drule_tac x = "Pow(A) - {0}" in spec) apply (blast dest: bijection) done end diff --git a/src/ZF/AC/WO1_AC.thy b/src/ZF/AC/WO1_AC.thy --- a/src/ZF/AC/WO1_AC.thy +++ b/src/ZF/AC/WO1_AC.thy @@ -1,106 +1,106 @@ (* Title: ZF/AC/WO1_AC.thy Author: Krzysztof Grabczewski The proofs of WO1 \ AC1 and WO1 \ AC10(n) for n >= 1 The latter proof is referred to as clear by the Rubins. However it seems to be quite complicated. The formal proof presented below is a mechanisation of the proof by Lawrence C. Paulson which is the following: Assume WO1. Let s be a set of infinite sets. Suppose x \ s. Then x is equipollent to |x| (by WO1), an infinite cardinal call it K. Since K = K \ K = |K+K| (by InfCard_cdouble_eq) there is an isomorphism h \ bij(K+K, x). (Here + means disjoint sum.) So there is a partition of x into 2-element sets, namely {{h(Inl(i)), h(Inr(i))} . i \ K} So for all x \ s the desired partition exists. By AC1 (which follows from WO1) there exists a function f that chooses a partition for each x \ s. Therefore we have AC10(2). *) theory WO1_AC imports AC_Equiv begin (* ********************************************************************** *) (* WO1 \ AC1 *) (* ********************************************************************** *) theorem WO1_AC1: "WO1 \ AC1" by (unfold AC1_def WO1_def, fast elim!: ex_choice_fun) (* ********************************************************************** *) (* WO1 \ AC10(n) (n >= 1) *) (* ********************************************************************** *) lemma lemma1: "\WO1; \B \ A. \C \ D(B). P(C,B)\ \ \f. \B \ A. P(f`B,B)" unfolding WO1_def apply (erule_tac x = "\({{C \ D (B) . P (C,B) }. B \ A}) " in allE) apply (erule exE, drule ex_choice_fun, fast) apply (erule exE) apply (rule_tac x = "\x \ A. f`{C \ D (x) . P (C,x) }" in exI) apply (simp, blast dest!: apply_type [OF _ RepFunI]) done lemma lemma2_1: "\\Finite(B); WO1\ \ |B| + |B| \ B" unfolding WO1_def apply (rule eqpoll_trans) prefer 2 apply (fast elim!: well_ord_cardinal_eqpoll) apply (rule eqpoll_sym [THEN eqpoll_trans]) apply (fast elim!: well_ord_cardinal_eqpoll) apply (drule spec [of _ B]) apply (clarify dest!: eqpoll_imp_Finite_iff [OF well_ord_cardinal_eqpoll]) apply (simp add: cadd_def [symmetric] eqpoll_refl InfCard_cdouble_eq Card_cardinal Inf_Card_is_InfCard) done lemma lemma2_2: "f \ bij(D+D, B) \ {{f`Inl(i), f`Inr(i)}. i \ D} \ Pow(Pow(B))" by (fast elim!: bij_is_fun [THEN apply_type]) lemma lemma2_3: "f \ bij(D+D, B) \ pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \ D})" unfolding pairwise_disjoint_def apply (blast dest: bij_is_inj [THEN inj_apply_equality]) done lemma lemma2_4: "\f \ bij(D+D, B); 1\n\ \ sets_of_size_between({{f`Inl(i), f`Inr(i)}. i \ D}, 2, succ(n))" apply (simp (no_asm_simp) add: sets_of_size_between_def succ_def) apply (blast intro!: cons_lepoll_cong intro: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] le_imp_subset [THEN subset_imp_lepoll] lepoll_trans dest: bij_is_inj [THEN inj_apply_equality] elim!: mem_irrefl) done lemma lemma2_5: "f \ bij(D+D, B) \ \({{f`Inl(i), f`Inr(i)}. i \ D})=B" -apply (unfold bij_def surj_def) + unfolding bij_def surj_def apply (fast elim!: inj_is_fun [THEN apply_type]) done lemma lemma2: "\WO1; \Finite(B); 1\n\ \ \C \ Pow(Pow(B)). pairwise_disjoint(C) \ sets_of_size_between(C, 2, succ(n)) \ \(C)=B" apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], assumption) apply (blast intro!: lemma2_2 lemma2_3 lemma2_4 lemma2_5) done theorem WO1_AC10: "\WO1; 1\n\ \ AC10(n)" unfolding AC10_def apply (fast intro!: lemma1 elim!: lemma2) done end diff --git a/src/ZF/AC/WO1_WO7.thy b/src/ZF/AC/WO1_WO7.thy --- a/src/ZF/AC/WO1_WO7.thy +++ b/src/ZF/AC/WO1_WO7.thy @@ -1,114 +1,114 @@ (* Title: ZF/AC/WO1_WO7.thy Author: Lawrence C Paulson, CU Computer Laboratory Copyright 1998 University of Cambridge WO7 \ LEMMA \ WO1 (Rubin \ Rubin p. 5) LEMMA is the sentence denoted by (**) Also, WO1 \ WO8 *) theory WO1_WO7 imports AC_Equiv begin definition "LEMMA \ \X. \Finite(X) \ (\R. well_ord(X,R) \ \well_ord(X,converse(R)))" (* ********************************************************************** *) (* It is easy to see that WO7 is equivalent to (**) *) (* ********************************************************************** *) lemma WO7_iff_LEMMA: "WO7 \ LEMMA" -apply (unfold WO7_def LEMMA_def) + unfolding WO7_def LEMMA_def apply (blast intro: Finite_well_ord_converse) done (* ********************************************************************** *) (* It is also easy to show that LEMMA implies WO1. *) (* ********************************************************************** *) lemma LEMMA_imp_WO1: "LEMMA \ WO1" -apply (unfold WO1_def LEMMA_def Finite_def eqpoll_def) + unfolding WO1_def LEMMA_def Finite_def eqpoll_def apply (blast intro!: well_ord_rvimage [OF bij_is_inj nat_implies_well_ord]) done (* ********************************************************************** *) (* The Rubins' proof of the other implication is contained within the *) (* following sentence \ *) (* "... each infinite ordinal is well ordered by < but not by >." *) (* This statement can be proved by the following two theorems. *) (* But moreover we need to show similar property for any well ordered *) (* infinite set. It is not very difficult thanks to Isabelle order types *) (* We show that if a set is well ordered by some relation and by its *) (* converse, then apropriate order type is well ordered by the converse *) (* of it's membership relation, which in connection with the previous *) (* gives the conclusion. *) (* ********************************************************************** *) lemma converse_Memrel_not_wf_on: "\Ord(a); \Finite(a)\ \ \wf[a](converse(Memrel(a)))" -apply (unfold wf_on_def wf_def) + unfolding wf_on_def wf_def apply (drule nat_le_infinite_Ord [THEN le_imp_subset], assumption) apply (rule notI) apply (erule_tac x = nat in allE, blast) done lemma converse_Memrel_not_well_ord: "\Ord(a); \Finite(a)\ \ \well_ord(a,converse(Memrel(a)))" unfolding well_ord_def apply (blast dest: converse_Memrel_not_wf_on) done lemma well_ord_rvimage_ordertype: "well_ord(A,r) \ rvimage (ordertype(A,r), converse(ordermap(A,r)),r) = Memrel(ordertype(A,r))" by (blast intro: ordertype_ord_iso [THEN ord_iso_sym] ord_iso_rvimage_eq Memrel_type [THEN subset_Int_iff [THEN iffD1]] trans) lemma well_ord_converse_Memrel: "\well_ord(A,r); well_ord(A,converse(r))\ \ well_ord(ordertype(A,r), converse(Memrel(ordertype(A,r))))" apply (subst well_ord_rvimage_ordertype [symmetric], assumption) apply (rule rvimage_converse [THEN subst]) apply (blast intro: ordertype_ord_iso ord_iso_sym ord_iso_is_bij bij_is_inj well_ord_rvimage) done lemma WO1_imp_LEMMA: "WO1 \ LEMMA" apply (unfold WO1_def LEMMA_def, clarify) apply (blast dest: well_ord_converse_Memrel Ord_ordertype [THEN converse_Memrel_not_well_ord] intro: ordertype_ord_iso ord_iso_is_bij bij_is_inj lepoll_Finite lepoll_def [THEN def_imp_iff, THEN iffD2] ) done lemma WO1_iff_WO7: "WO1 \ WO7" apply (simp add: WO7_iff_LEMMA) apply (blast intro: LEMMA_imp_WO1 WO1_imp_LEMMA) done (* ********************************************************************** *) (* The proof of WO8 \ WO1 (Rubin \ Rubin p. 6) *) (* ********************************************************************** *) lemma WO1_WO8: "WO1 \ WO8" by (unfold WO1_def WO8_def, fast) (* The implication "WO8 \ WO1": a faithful image of Rubin \ Rubin's proof*) lemma WO8_WO1: "WO8 \ WO1" -apply (unfold WO1_def WO8_def) + unfolding WO1_def WO8_def apply (rule allI) apply (erule_tac x = "{{x}. x \ A}" in allE) apply (erule impE) apply (rule_tac x = "\a \ {{x}. x \ A}. THE x. a={x}" in exI) apply (force intro!: lam_type simp add: singleton_eq_iff the_equality) apply (blast intro: lam_sing_bij bij_is_inj well_ord_rvimage) done end diff --git a/src/ZF/AC/WO2_AC16.thy b/src/ZF/AC/WO2_AC16.thy --- a/src/ZF/AC/WO2_AC16.thy +++ b/src/ZF/AC/WO2_AC16.thy @@ -1,579 +1,579 @@ (* Title: ZF/AC/WO2_AC16.thy Author: Krzysztof Grabczewski The proof of WO2 \ AC16(k #+ m, k) The main part of the proof is the inductive reasoning concerning properties of constructed family T_gamma. The proof deals with three cases for ordinals: 0, succ and limit ordinal. The first instance is trivial, the third not difficult, but the second is very complicated requiring many lemmas. We also need to prove that at any stage gamma the set (s - \(...) - k_gamma) (Rubin \ Rubin page 15) contains m distinct elements (in fact is equipollent to s) *) theory WO2_AC16 imports AC_Equiv AC16_lemmas Cardinal_aux begin (**** A recursive definition used in the proof of WO2 \ AC16 ****) definition recfunAC16 :: "[i,i,i,i] \ i" where "recfunAC16(f,h,i,a) \ transrec2(i, 0, \g r. if (\y \ r. h`g \ y) then r else r \ {f`(\ i. h`g \ f`i \ (\b f`i \ (\t \ r. \ h`b \ t))))})" (* ********************************************************************** *) (* Basic properties of recfunAC16 *) (* ********************************************************************** *) lemma recfunAC16_0: "recfunAC16(f,h,0,a) = 0" by (simp add: recfunAC16_def) lemma recfunAC16_succ: "recfunAC16(f,h,succ(i),a) = (if (\y \ recfunAC16(f,h,i,a). h ` i \ y) then recfunAC16(f,h,i,a) else recfunAC16(f,h,i,a) \ {f ` (\ j. h ` i \ f ` j \ (\b f`j \ (\t \ recfunAC16(f,h,i,a). \ h`b \ t))))})" apply (simp add: recfunAC16_def) done lemma recfunAC16_Limit: "Limit(i) \ recfunAC16(f,h,i,a) = (\j\g r. r \ B(g,r); Ord(i)\ \ j transrec2(j, 0, B) \ transrec2(i, 0, B)" apply (erule trans_induct) apply (rule Ord_cases, assumption+, fast) apply (simp (no_asm_simp)) apply (blast elim!: leE) apply (simp add: transrec2_Limit) apply (blast intro: OUN_I ltI Ord_in_Ord [THEN le_refl] elim!: Limit_has_succ [THEN ltE]) done lemma transrec2_mono: "\\g r. r \ B(g,r); j\i\ \ transrec2(j, 0, B) \ transrec2(i, 0, B)" apply (erule leE) apply (rule transrec2_mono_lemma) apply (auto intro: lt_Ord2 ) done (* ********************************************************************** *) (* Monotonicity of recfunAC16 *) (* ********************************************************************** *) lemma recfunAC16_mono: "i\j \ recfunAC16(f, g, i, a) \ recfunAC16(f, g, j, a)" unfolding recfunAC16_def apply (rule transrec2_mono, auto) done (* ********************************************************************** *) (* case of limit ordinal *) (* ********************************************************************** *) lemma lemma3_1: "\\yzY \ F(y). f(z)<=Y) \ (\! Y. Y \ F(y) \ f(z)<=Y); \i j. i\j \ F(i) \ F(j); j\i; i F(i); f(z)<=V; W \ F(j); f(z)<=W\ \ V = W" apply (erule asm_rl allE impE)+ apply (drule subsetD, assumption, blast) done lemma lemma3: "\\yzY \ F(y). f(z)<=Y) \ (\! Y. Y \ F(y) \ f(z)<=Y); \i j. i\j \ F(i) \ F(j); i F(i); f(z)<=V; W \ F(j); f(z)<=W\ \ V = W" apply (rule_tac j=j in Ord_linear_le [OF lt_Ord lt_Ord], assumption+) apply (erule lemma3_1 [symmetric], assumption+) apply (erule lemma3_1, assumption+) done lemma lemma4: "\\y X \ (\xY \ F(y). h(x) \ Y) \ (\! Y. Y \ F(y) \ h(x) \ Y)); x < a\ \ \yzY \ F(y). h(z) \ Y) \ (\! Y. Y \ F(y) \ h(z) \ Y)" apply (intro oallI impI) apply (drule ospec, assumption, clarify) apply (blast elim!: oallE ) done lemma lemma5: "\\y X \ (\xY \ F(y). h(x) \ Y) \ (\! Y. Y \ F(y) \ h(x) \ Y)); x < a; Limit(x); \i j. i\j \ F(i) \ F(j)\ \ (\x X \ (\xax \ \x x) \ (\! Y. Y \ (\x h(xa) \ Y))" apply (rule conjI) apply (rule subsetI) apply (erule OUN_E) apply (drule ospec, assumption, fast) apply (drule lemma4, assumption) apply (rule oallI) apply (rule impI) apply (erule disjE) apply (frule ospec, erule Limit_has_succ, assumption) apply (drule_tac A = a and x = xa in ospec, assumption) apply (erule impE, rule le_refl [THEN disjI1], erule lt_Ord) apply (blast intro: lemma3 Limit_has_succ) apply (blast intro: lemma3) done (* ********************************************************************** *) (* case of successor ordinal *) (* ********************************************************************** *) (* First quite complicated proof of the fact used in the recursive construction of the family T_gamma (WO2 \ AC16(k #+ m, k)) - the fact that at any stage gamma the set (s - \(...) - k_gamma) is equipollent to s (Rubin \ Rubin page 15). *) (* ********************************************************************** *) (* dbl_Diff_eqpoll_Card *) (* ********************************************************************** *) lemma dbl_Diff_eqpoll_Card: "\A\a; Card(a); \Finite(a); B\a; C\a\ \ A - B - C\a" by (blast intro: Diff_lesspoll_eqpoll_Card) (* ********************************************************************** *) (* Case of finite ordinals *) (* ********************************************************************** *) lemma Finite_lesspoll_infinite_Ord: "\Finite(X); \Finite(a); Ord(a)\ \ X\a" unfolding lesspoll_def apply (rule conjI) apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption) unfolding Finite_def apply (blast intro: leI [THEN le_imp_subset, THEN subset_imp_lepoll] ltI eqpoll_imp_lepoll lepoll_trans) apply (blast intro: eqpoll_sym [THEN eqpoll_trans]) done lemma Union_lesspoll: "\\x \ X. x \ n \ x \ T; well_ord(T, R); X \ b; bFinite(a); Card(a); n \ nat\ \ \(X)\a" apply (case_tac "Finite (X)") apply (blast intro: Card_is_Ord Finite_lesspoll_infinite_Ord lepoll_nat_imp_Finite Finite_Union) apply (drule lepoll_imp_ex_le_eqpoll) apply (erule lt_Ord) apply (elim exE conjE) apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) apply (erule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], THEN exE]) apply (frule bij_is_surj [THEN surj_image_eq]) apply (drule image_fun [OF bij_is_fun subset_refl]) apply (drule sym [THEN trans], assumption) apply (blast intro: lt_Ord UN_lepoll lt_Card_imp_lesspoll lt_trans1 lesspoll_trans1) done (* ********************************************************************** *) (* recfunAC16_lepoll_index *) (* ********************************************************************** *) lemma Un_sing_eq_cons: "A \ {a} = cons(a, A)" by fast lemma Un_lepoll_succ: "A \ B \ A \ {a} \ succ(B)" apply (simp add: Un_sing_eq_cons succ_def) apply (blast elim!: mem_irrefl intro: cons_lepoll_cong) done lemma Diff_UN_succ_empty: "Ord(a) \ F(a) - (\b F(a) \ X - (\b X" by blast lemma recfunAC16_Diff_lepoll_1: "Ord(x) \ recfunAC16(f, g, x, a) - (\i 1" apply (erule Ord_cases) apply (simp add: recfunAC16_0 empty_subsetI [THEN subset_imp_lepoll]) (*Limit case*) prefer 2 apply (simp add: recfunAC16_Limit Diff_cancel empty_subsetI [THEN subset_imp_lepoll]) (*succ case*) apply (simp add: recfunAC16_succ Diff_UN_succ_empty [of _ "\j. recfunAC16(f,g,j,a)"] empty_subsetI [THEN subset_imp_lepoll]) apply (best intro: Diff_UN_succ_subset [THEN subset_imp_lepoll] singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] lepoll_trans) done lemma in_Least_Diff: "\z \ F(x); Ord(x)\ \ z \ F(\ i. z \ F(i)) - (\j<(\ i. z \ F(i)). F(j))" by (fast elim: less_LeastE elim!: LeastI) lemma Least_eq_imp_ex: "\(\ i. w \ F(i)) = (\ i. z \ F(i)); w \ (\i (\i \ \b (F(b) - (\c z \ (F(b) - (\cA \ 1; a \ A; b \ A\ \ a=b" by (fast dest!: lepoll_1_is_sing) lemma UN_lepoll_index: "\\ij 1; Limit(a)\ \ (\x a" apply (rule lepoll_def [THEN def_imp_iff [THEN iffD2]]) apply (rule_tac x = "\z \ (\x i. z \ F (i) " in exI) unfolding inj_def apply (rule CollectI) apply (rule lam_type) apply (erule OUN_E) apply (erule Least_in_Ord) apply (erule ltD) apply (erule lt_Ord2) apply (intro ballI) apply (simp (no_asm_simp)) apply (rule impI) apply (drule Least_eq_imp_ex, assumption+) apply (fast elim!: two_in_lepoll_1) done lemma recfunAC16_lepoll_index: "Ord(y) \ recfunAC16(f, h, y, a) \ y" apply (erule trans_induct3) (*0 case*) apply (simp (no_asm_simp) add: recfunAC16_0 lepoll_refl) (*succ case*) apply (simp (no_asm_simp) add: recfunAC16_succ) apply (blast dest!: succI1 [THEN rev_bspec] intro: subset_succI [THEN subset_imp_lepoll] Un_lepoll_succ lepoll_trans) apply (simp (no_asm_simp) add: recfunAC16_Limit) apply (blast intro: lt_Ord [THEN recfunAC16_Diff_lepoll_1] UN_lepoll_index) done lemma Union_recfunAC16_lesspoll: "\recfunAC16(f,g,y,a) \ {X \ Pow(A). X\n}; A\a; yFinite(a); Card(a); n \ nat\ \ \(recfunAC16(f,g,y,a))\a" apply (erule eqpoll_def [THEN def_imp_iff, THEN iffD1, THEN exE]) apply (rule_tac T=A in Union_lesspoll, simp_all) apply (blast intro!: eqpoll_imp_lepoll) apply (blast intro: bij_is_inj Card_is_Ord [THEN well_ord_Memrel] well_ord_rvimage) apply (erule lt_Ord [THEN recfunAC16_lepoll_index]) done lemma dbl_Diff_eqpoll: "\recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; Card(a); \ Finite(a); A\a; k \ nat; y bij(a, {Y \ Pow(A). Y\succ(k)})\ \ A - \(recfunAC16(f, h, y, a)) - h`y\a" apply (rule dbl_Diff_eqpoll_Card, simp_all) apply (simp add: Union_recfunAC16_lesspoll) apply (rule Finite_lesspoll_infinite_Ord) apply (rule Finite_def [THEN def_imp_iff, THEN iffD2]) apply (blast dest: ltD bij_is_fun [THEN apply_type], assumption) apply (blast intro: Card_is_Ord) done (* back to the proof *) lemmas disj_Un_eqpoll_nat_sum = eqpoll_trans [THEN eqpoll_trans, OF disj_Un_eqpoll_sum sum_eqpoll_cong nat_sum_eqpoll_sum] lemma Un_in_Collect: "\x \ Pow(A - B - h`i); x\m; h \ bij(a, {x \ Pow(A) . x\k}); i nat; m \ nat\ \ h ` i \ x \ {x \ Pow(A) . x\k #+ m}" by (blast intro: disj_Un_eqpoll_nat_sum dest: ltD bij_is_fun [THEN apply_type]) (* ********************************************************************** *) (* Lemmas simplifying assumptions *) (* ********************************************************************** *) lemma lemma6: "\\y (\x Q(x,y)); succ(j) \ F(j)<=X \ (\x Q(x,j))" by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]) lemma lemma7: "\\x Q(x,j); succ(j) \ P(j,j) \ (\xj | P(x,j) \ Q(x,j))" by (fast elim!: leE) (* ********************************************************************** *) (* Lemmas needed to prove ex_next_set, which means that for any successor *) (* ordinal there is a set satisfying certain properties *) (* ********************************************************************** *) lemma ex_subset_eqpoll: "\A\a; \ Finite(a); Ord(a); m \ nat\ \ \X \ Pow(A). X\m" apply (rule lepoll_imp_eqpoll_subset [of m A, THEN exE]) apply (rule lepoll_trans, rule leI [THEN le_imp_lepoll]) apply (blast intro: lt_trans2 [OF ltI nat_le_infinite_Ord] Ord_nat) apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) apply (fast elim!: eqpoll_sym) done lemma subset_Un_disjoint: "\A \ B \ C; A \ C = 0\ \ A \ B" by blast lemma Int_empty: "\X \ Pow(A - \(B) -C); T \ B; F \ T\ \ F \ X = 0" by blast (* ********************************************************************** *) (* equipollent subset (and finite) is the whole set *) (* ********************************************************************** *) lemma subset_imp_eq_lemma: "m \ nat \ \A B. A \ B \ m \ A \ B \ m \ A=B" apply (induct_tac "m") apply (fast dest!: lepoll_0_is_0) apply (intro allI impI) apply (elim conjE) apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], assumption) apply (frule subsetD [THEN Diff_sing_lepoll], assumption+) apply (frule lepoll_Diff_sing) apply (erule allE impE)+ apply (rule conjI) prefer 2 apply fast apply fast apply (blast elim: equalityE) done lemma subset_imp_eq: "\A \ B; m \ A; B \ m; m \ nat\ \ A=B" by (blast dest!: subset_imp_eq_lemma) lemma bij_imp_arg_eq: "\f \ bij(a, {Y \ X. Y\succ(k)}); k \ nat; f`b \ f`y; b \ b=y" apply (drule subset_imp_eq) apply (erule_tac [3] nat_succI) -apply (unfold bij_def inj_def) + unfolding bij_def inj_def apply (blast intro: eqpoll_sym eqpoll_imp_lepoll dest: ltD apply_type)+ done lemma ex_next_set: "\recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; Card(a); \ Finite(a); A\a; k \ nat; m \ nat; y bij(a, {Y \ Pow(A). Y\succ(k)}); \ (\Y \ recfunAC16(f, h, y, a). h`y \ Y)\ \ \X \ {Y \ Pow(A). Y\succ(k #+ m)}. h`y \ X \ (\b X \ (\T \ recfunAC16(f, h, y, a). \ h`b \ T))" apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], assumption+) apply (erule Card_is_Ord, assumption) apply (frule Un_in_Collect, (erule asm_rl nat_succI)+) apply (erule CollectE) apply (rule rev_bexI, simp) apply (rule conjI, blast) apply (intro ballI impI oallI notI) apply (drule subset_Un_disjoint, rule Int_empty, assumption+) apply (blast dest: bij_imp_arg_eq) done (* ********************************************************************** *) (* Lemma ex_next_Ord states that for any successor *) (* ordinal there is a number of the set satisfying certain properties *) (* ********************************************************************** *) lemma ex_next_Ord: "\recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; Card(a); \ Finite(a); A\a; k \ nat; m \ nat; y bij(a, {Y \ Pow(A). Y\succ(k)}); f \ bij(a, {Y \ Pow(A). Y\succ(k #+ m)}); \ (\Y \ recfunAC16(f, h, y, a). h`y \ Y)\ \ \c f`c \ (\b f`c \ (\T \ recfunAC16(f, h, y, a). \ h`b \ T))" apply (drule ex_next_set, assumption+) apply (erule bexE) apply (rule_tac x="converse(f)`X" in oexI) apply (simp add: right_inverse_bij) apply (blast intro: bij_converse_bij bij_is_fun [THEN apply_type] ltI Card_is_Ord) done (* ********************************************************************** *) (* Lemma simplifying assumptions *) (* ********************************************************************** *) lemma lemma8: "\\xxa \ F(j). P(x, xa)) \ (\! Y. Y \ F(j) \ P(x, Y)); F(j) \ X; L \ X; P(j, L) \ (\x (\xa \ F(j). \P(x, xa)))\ \ F(j) \ {L} \ X \ (\xj | (\xa \ (F(j) \ {L}). P(x, xa)) \ (\! Y. Y \ (F(j) \ {L}) \ P(x, Y)))" apply (rule conjI) apply (fast intro!: singleton_subsetI) apply (rule oallI) apply (blast elim!: leE oallE) done (* ********************************************************************** *) (* The main part of the proof: inductive proof of the property of T_gamma *) (* lemma main_induct *) (* ********************************************************************** *) lemma main_induct: "\b < a; f \ bij(a, {Y \ Pow(A) . Y\succ(k #+ m)}); h \ bij(a, {Y \ Pow(A) . Y\succ(k)}); \Finite(a); Card(a); A\a; k \ nat; m \ nat\ \ recfunAC16(f, h, b, a) \ {X \ Pow(A) . X\succ(k #+ m)} \ (\xY \ recfunAC16(f, h, b, a). h ` x \ Y) \ (\! Y. Y \ recfunAC16(f, h, b, a) \ h ` x \ Y))" apply (erule lt_induct) apply (frule lt_Ord) apply (erule Ord_cases) (* case 0 *) apply (simp add: recfunAC16_0) (* case Limit *) prefer 2 apply (simp add: recfunAC16_Limit) apply (rule lemma5, assumption+) apply (blast dest!: recfunAC16_mono) (* case succ *) apply clarify apply (erule lemma6 [THEN conjE], assumption) apply (simp (no_asm_simp) split del: split_if add: recfunAC16_succ) apply (rule conjI [THEN split_if [THEN iffD2]]) apply (simp, erule lemma7, assumption) apply (rule impI) apply (rule ex_next_Ord [THEN oexE], assumption+, rule le_refl [THEN lt_trans], assumption+) apply (erule lemma8, assumption) apply (rule bij_is_fun [THEN apply_type], assumption) apply (erule Least_le [THEN lt_trans2, THEN ltD]) apply (erule lt_Ord) apply (erule succ_leI) apply (erule LeastI) apply (erule lt_Ord) done (* ********************************************************************** *) (* Lemma to simplify the inductive proof *) (* - the desired property is a consequence of the inductive assumption *) (* ********************************************************************** *) lemma lemma_simp_induct: "\\b. b F(b) \ S \ (\xY \ F(b). f`x \ Y)) \ (\! Y. Y \ F(b) \ f`x \ Y)); f \ a->f``(a); Limit(a); \i j. i\j \ F(i) \ F(j)\ \ (\j S \ (\x \ f``a. \! Y. Y \ (\j x \ Y)" apply (rule conjI) apply (rule subsetI) apply (erule OUN_E, blast) apply (rule ballI) apply (erule imageE) apply (drule ltI, erule Limit_is_Ord) apply (drule Limit_has_succ, assumption) apply (frule_tac x1="succ(xa)" in spec [THEN mp], assumption) apply (erule conjE) apply (drule ospec) (** LEVEL 10 **) apply (erule leI [THEN succ_leE]) apply (erule impE) apply (fast elim!: leI [THEN succ_leE, THEN lt_Ord, THEN le_refl]) apply (drule apply_equality, assumption) apply (elim conjE ex1E) (** LEVEL 15 **) apply (rule ex1I, blast) apply (elim conjE OUN_E) apply (erule_tac i="succ(xa)" and j=aa in Ord_linear_le [OF lt_Ord lt_Ord], assumption) prefer 2 apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) (** LEVEL 20 **) apply (drule_tac x1=aa in spec [THEN mp], assumption) apply (frule succ_leE) apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) done (* ********************************************************************** *) (* The target theorem *) (* ********************************************************************** *) theorem WO2_AC16: "\WO2; 0 nat; m \ nat\ \ AC16(k #+ m,k)" unfolding AC16_def apply (rule allI) apply (rule impI) apply (frule WO2_infinite_subsets_eqpoll_X, assumption+) apply (frule_tac n="k #+ m" in WO2_infinite_subsets_eqpoll_X, simp, simp) apply (frule WO2_imp_ex_Card) apply (elim exE conjE) apply (drule eqpoll_trans [THEN eqpoll_sym, THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], assumption) apply (drule eqpoll_trans [THEN eqpoll_sym, THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], assumption+) apply (elim exE) apply (rename_tac h) apply (rule_tac x = "\j WO1. Every proof (except WO6 \ WO1 and WO1 \ WO2) are described as "clear" by Rubin \ Rubin (page 2). They refer reader to a book by Gödel to see the proof WO1 \ WO2. Fortunately order types made this proof also very easy. *) theory WO6_WO1 imports Cardinal_aux begin (* Auxiliary definitions used in proof *) definition NN :: "i \ i" where "NN(y) \ {m \ nat. \a. \f. Ord(a) \ domain(f)=a \ (\b (\b m)}" definition uu :: "[i, i, i, i] \ i" where "uu(f, beta, gamma, delta) \ (f`beta * f`gamma) \ f`delta" (** Definitions for case 1 **) definition vv1 :: "[i, i, i] \ i" where "vv1(f,m,b) \ let g = \ g. (\d. Ord(d) \ (domain(uu(f,b,g,d)) \ 0 \ domain(uu(f,b,g,d)) \ m)); d = \ d. domain(uu(f,b,g,d)) \ 0 \ domain(uu(f,b,g,d)) \ m in if f`b \ 0 then domain(uu(f,b,g,d)) else 0" definition ww1 :: "[i, i, i] \ i" where "ww1(f,m,b) \ f`b - vv1(f,m,b)" definition gg1 :: "[i, i, i] \ i" where "gg1(f,a,m) \ \b \ a++a. if b i" where "vv2(f,b,g,s) \ if f`g \ 0 then {uu(f, b, g, \ d. uu(f,b,g,d) \ 0)`s} else 0" definition ww2 :: "[i, i, i, i] \ i" where "ww2(f,b,g,s) \ f`g - vv2(f,b,g,s)" definition gg2 :: "[i, i, i, i] \ i" where "gg2(f,a,b,s) \ \g \ a++a. if g WO3" by (unfold WO2_def WO3_def, fast) (* ********************************************************************** *) lemma WO3_WO1: "WO3 \ WO1" -apply (unfold eqpoll_def WO1_def WO3_def) + unfolding eqpoll_def WO1_def WO3_def apply (intro allI) apply (drule_tac x=A in spec) apply (blast intro: bij_is_inj well_ord_rvimage well_ord_Memrel [THEN well_ord_subset]) done (* ********************************************************************** *) lemma WO1_WO2: "WO1 \ WO2" -apply (unfold eqpoll_def WO1_def WO2_def) + unfolding eqpoll_def WO1_def WO2_def apply (blast intro!: Ord_ordertype ordermap_bij) done (* ********************************************************************** *) lemma lam_sets: "f \ A->B \ (\x \ A. {f`x}): A -> {{b}. b \ B}" by (fast intro!: lam_type apply_type) lemma surj_imp_eq': "f \ surj(A,B) \ (\a \ A. {f`a}) = B" unfolding surj_def apply (fast elim!: apply_type) done lemma surj_imp_eq: "\f \ surj(A,B); Ord(A)\ \ (\a WO4(1)" -apply (unfold WO1_def WO4_def) + unfolding WO1_def WO4_def apply (rule allI) apply (erule_tac x = A in allE) apply (erule exE) apply (intro exI conjI) apply (erule Ord_ordertype) apply (erule ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN lam_sets, THEN domain_of_fun]) apply (simp_all add: singleton_eqpoll_1 eqpoll_imp_lepoll Ord_ordertype ordermap_bij [THEN bij_converse_bij, THEN bij_is_surj, THEN surj_imp_eq] ltD) done (* ********************************************************************** *) lemma WO4_mono: "\m\n; WO4(m)\ \ WO4(n)" unfolding WO4_def apply (blast dest!: spec intro: lepoll_trans [OF _ le_imp_lepoll]) done (* ********************************************************************** *) lemma WO4_WO5: "\m \ nat; 1\m; WO4(m)\ \ WO5" by (unfold WO4_def WO5_def, blast) (* ********************************************************************** *) lemma WO5_WO6: "WO5 \ WO6" by (unfold WO4_def WO5_def WO6_def, blast) (* ********************************************************************** The proof of "WO6 \ WO1". Simplified by L C Paulson. From the book "Equivalents of the Axiom of Choice" by Rubin \ Rubin, pages 2-5 ************************************************************************* *) lemma lt_oadd_odiff_disj: "\k < i++j; Ord(i); Ord(j)\ \ k < i | (\ k k = i ++ (k--i) \ (k--i) f`b" by (unfold uu_def, blast) lemma quant_domain_uu_lepoll_m: "\b m \ \bgd m" by (blast intro: domain_uu_subset [THEN subset_imp_lepoll] lepoll_trans) lemma uu_subset1: "uu(f,b,g,d) \ f`b * f`g" by (unfold uu_def, blast) lemma uu_subset2: "uu(f,b,g,d) \ f`d" by (unfold uu_def, blast) lemma uu_lepoll_m: "\\b m; d \ uu(f,b,g,d) \ m" by (blast intro: uu_subset2 [THEN subset_imp_lepoll] lepoll_trans) (* ********************************************************************** *) (* Two cases for lemma ii *) (* ********************************************************************** *) lemma cases: "\bgd m \ (\b 0 \ (\gd 0 \ u(f,b,g,d) \ m)) | (\b 0 \ (\gd 0 \ u(f,b,g,d) \ m))" unfolding lesspoll_def apply (blast del: equalityI) done (* ********************************************************************** *) (* Lemmas used in both cases *) (* ********************************************************************** *) lemma UN_oadd: "Ord(a) \ (\bb C(a++b))" by (blast intro: ltI lt_oadd1 oadd_lt_mono2 dest!: lt_oadd_disj) (* ********************************************************************** *) (* Case 1: lemmas *) (* ********************************************************************** *) lemma vv1_subset: "vv1(f,m,b) \ f`b" by (simp add: vv1_def Let_def domain_uu_subset) (* ********************************************************************** *) (* Case 1: Union of images is the whole "y" *) (* ********************************************************************** *) lemma UN_gg1_eq: "\Ord(a); m \ nat\ \ (\bbP(a, b); Ord(a); Ord(b); Least_a = (\ a. \x. Ord(x) \ P(a, x))\ \ P(Least_a, \ b. P(Least_a, b))" apply (erule ssubst) apply (rule_tac Q = "\z. P (z, \ b. P (z, b))" in LeastI2) apply (fast elim!: LeastI)+ done lemmas nested_Least_instance = nested_LeastI [of "\g d. domain(uu(f,b,g,d)) \ 0 \ domain(uu(f,b,g,d)) \ m"] for f b m lemma gg1_lepoll_m: "\Ord(a); m \ nat; \b0 \ (\gd 0 \ domain(uu(f,b,g,d)) \ m); \b succ(m); b \ gg1(f,a,m)`b \ m" apply (simp add: gg1_def empty_lepollI) apply (safe dest!: lt_oadd_odiff_disj) (*Case b show vv1(f,m,b) \ m *) apply (simp add: vv1_def Let_def empty_lepollI) apply (fast intro: nested_Least_instance [THEN conjunct2] elim!: lt_Ord) (*Case a\b \ show ww1(f,m,b--a) \ m *) apply (simp add: ww1_def empty_lepollI) apply (case_tac "f` (b--a) = 0", simp add: empty_lepollI) apply (rule Diff_lepoll, blast) apply (rule vv1_subset) apply (drule ospec [THEN mp], assumption+) apply (elim oexE conjE) apply (simp add: vv1_def Let_def lt_Ord nested_Least_instance [THEN conjunct1]) done (* ********************************************************************** *) (* Case 2: lemmas *) (* ********************************************************************** *) (* ********************************************************************** *) (* Case 2: vv2_subset *) (* ********************************************************************** *) lemma ex_d_uu_not_empty: "\b0; f`g\0; y*y \ y; (\b \ \d 0" by (unfold uu_def, blast) lemma uu_not_empty: "\b0; f`g\0; y*y \ y; (\b \ uu(f,b,g,\ d. (uu(f,b,g,d) \ 0)) \ 0" apply (drule ex_d_uu_not_empty, assumption+) apply (fast elim!: LeastI lt_Ord) done lemma not_empty_rel_imp_domain: "\r \ A*B; r\0\ \ domain(r)\0" by blast lemma Least_uu_not_empty_lt_a: "\b0; f`g\0; y*y \ y; (\b \ (\ d. uu(f,b,g,d) \ 0) < a" apply (erule ex_d_uu_not_empty [THEN oexE], assumption+) apply (blast intro: Least_le [THEN lt_trans1] lt_Ord) done lemma subset_Diff_sing: "\B \ A; a\B\ \ B \ A-{a}" by blast (*Could this be proved more directly?*) lemma supset_lepoll_imp_eq: "\A \ m; m \ B; B \ A; m \ nat\ \ A=B" apply (erule natE) apply (fast dest!: lepoll_0_is_0 intro!: equalityI) apply (safe intro!: equalityI) apply (rule ccontr) apply (rule succ_lepoll_natE) apply (erule lepoll_trans) apply (rule lepoll_trans) apply (erule subset_Diff_sing [THEN subset_imp_lepoll], assumption) apply (rule Diff_sing_lepoll, assumption+) done lemma uu_Least_is_fun: "\\gd0 \ domain(uu(f, b, g, d)) \ succ(m); \b succ(m); y*y \ y; (\b0; f`g\0; m \ nat; s \ f`b\ \ uu(f, b, g, \ d. uu(f,b,g,d)\0) \ f`b -> f`g" apply (drule_tac x2=g in ospec [THEN ospec, THEN mp]) apply (rule_tac [3] not_empty_rel_imp_domain [OF uu_subset1 uu_not_empty]) apply (rule_tac [2] Least_uu_not_empty_lt_a, assumption+) apply (rule rel_is_fun) apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) apply (erule uu_lepoll_m) apply (rule Least_uu_not_empty_lt_a, assumption+) apply (rule uu_subset1) apply (rule supset_lepoll_imp_eq [OF _ eqpoll_sym [THEN eqpoll_imp_lepoll]]) apply (fast intro!: domain_uu_subset)+ done lemma vv2_subset: "\\gd0 \ domain(uu(f, b, g, d)) \ succ(m); \b succ(m); y*y \ y; (\b nat; s \ f`b\ \ vv2(f,b,g,s) \ f`g" apply (simp add: vv2_def) apply (blast intro: uu_Least_is_fun [THEN apply_type]) done (* ********************************************************************** *) (* Case 2: Union of images is the whole "y" *) (* ********************************************************************** *) lemma UN_gg2_eq: "\\gd 0 \ domain(uu(f,b,g,d)) \ succ(m); \b succ(m); y*y \ y; (\b nat; s \ f`b; b \ (\gm \ nat; m\0\ \ vv2(f,b,g,s) \ m" unfolding vv2_def apply (simp add: empty_lepollI) apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0] intro!: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans] not_lt_imp_le [THEN le_imp_subset, THEN subset_imp_lepoll] nat_into_Ord nat_1I) done lemma ww2_lepoll: "\\b succ(m); g nat; vv2(f,b,g,d) \ f`g\ \ ww2(f,b,g,d) \ m" unfolding ww2_def apply (case_tac "f`g = 0") apply (simp add: empty_lepollI) apply (drule ospec, assumption) apply (rule Diff_lepoll, assumption+) apply (simp add: vv2_def not_emptyI) done lemma gg2_lepoll_m: "\\gd 0 \ domain(uu(f,b,g,d)) \ succ(m); \b succ(m); y*y \ y; (\b f`b; m \ nat; m\ 0; g \ gg2(f,a,b,s) ` g \ m" apply (simp add: gg2_def empty_lepollI) apply (safe elim!: lt_Ord2 dest!: lt_oadd_odiff_disj) apply (simp add: vv2_lepoll) apply (simp add: ww2_lepoll vv2_subset) done (* ********************************************************************** *) (* lemma ii *) (* ********************************************************************** *) lemma lemma_ii: "\succ(m) \ NN(y); y*y \ y; m \ nat; m\0\ \ m \ NN(y)" unfolding NN_def apply (elim CollectE exE conjE) apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption) (* case 1 *) apply (simp add: lesspoll_succ_iff) apply (rule_tac x = "a++a" in exI) apply (fast intro!: Ord_oadd domain_gg1 UN_gg1_eq gg1_lepoll_m) (* case 2 *) apply (elim oexE conjE) apply (rule_tac A = "f`B" for B in not_emptyE, assumption) apply (rule CollectI) apply (erule succ_natD) apply (rule_tac x = "a++a" in exI) apply (rule_tac x = "gg2 (f,a,b,x) " in exI) apply (simp add: Ord_oadd domain_gg2 UN_gg2_eq gg2_lepoll_m) done (* ********************************************************************** *) (* lemma iv - p. 4: *) (* For every set x there is a set y such that x \ (y * y) \ y *) (* ********************************************************************** *) (* The leading \-quantifier looks odd but makes the proofs shorter (used only in the following two lemmas) *) lemma z_n_subset_z_succ_n: "\n \ nat. rec(n, x, \k r. r \ r*r) \ rec(succ(n), x, \k r. r \ r*r)" by (fast intro: rec_succ [THEN ssubst]) lemma le_subsets: "\\n \ nat. f(n)<=f(succ(n)); n\m; n \ nat; m \ nat\ \ f(n)<=f(m)" apply (erule_tac P = "n\m" in rev_mp) apply (rule_tac P = "\z. n\z \ f (n) \ f (z) " in nat_induct) apply (auto simp add: le_iff) done lemma le_imp_rec_subset: "\n\m; m \ nat\ \ rec(n, x, \k r. r \ r*r) \ rec(m, x, \k r. r \ r*r)" apply (rule z_n_subset_z_succ_n [THEN le_subsets]) apply (blast intro: lt_nat_in_nat)+ done lemma lemma_iv: "\y. x \ y*y \ y" apply (rule_tac x = "\n \ nat. rec (n, x, \k r. r \ r*r) " in exI) apply safe apply (rule nat_0I [THEN UN_I], simp) apply (rule_tac a = "succ (n \ na) " in UN_I) apply (erule Un_nat_type [THEN nat_succI], assumption) apply (auto intro: le_imp_rec_subset [THEN subsetD] intro!: Un_upper1_le Un_upper2_le Un_nat_type elim!: nat_into_Ord) done (* ********************************************************************** *) (* Rubin \ Rubin wrote, *) (* "It follows from (ii) and mathematical induction that if y*y \ y then *) (* y can be well-ordered" *) (* In fact we have to prove *) (* * WO6 \ NN(y) \ 0 *) (* * reverse induction which lets us infer that 1 \ NN(y) *) (* * 1 \ NN(y) \ y can be well-ordered *) (* ********************************************************************** *) (* ********************************************************************** *) (* WO6 \ NN(y) \ 0 *) (* ********************************************************************** *) lemma WO6_imp_NN_not_empty: "WO6 \ NN(y) \ 0" by (unfold WO6_def NN_def, clarify, blast) (* ********************************************************************** *) (* 1 \ NN(y) \ y can be well-ordered *) (* ********************************************************************** *) lemma lemma1: "\(\b y; \b 1; Ord(a)\ \ \c(\b y; \b 1; Ord(a)\ \ f` (\ i. f`i = {x}) = {x}" apply (drule lemma1, assumption+) apply (fast elim!: lt_Ord intro: LeastI) done lemma NN_imp_ex_inj: "1 \ NN(y) \ \a f. Ord(a) \ f \ inj(y, a)" unfolding NN_def apply (elim CollectE exE conjE) apply (rule_tac x = a in exI) apply (rule_tac x = "\x \ y. \ i. f`i = {x}" in exI) apply (rule conjI, assumption) apply (rule_tac d = "\i. THE x. x \ f`i" in lam_injective) apply (drule lemma1, assumption+) apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord) apply (rule lemma2 [THEN ssubst], assumption+, blast) done lemma y_well_ord: "\y*y \ y; 1 \ NN(y)\ \ \r. well_ord(y, r)" apply (drule NN_imp_ex_inj) apply (fast elim!: well_ord_rvimage [OF _ well_ord_Memrel]) done (* ********************************************************************** *) (* reverse induction which lets us infer that 1 \ NN(y) *) (* ********************************************************************** *) lemma rev_induct_lemma [rule_format]: "\n \ nat; \m. \m \ nat; m\0; P(succ(m))\ \ P(m)\ \ n\0 \ P(n) \ P(1)" by (erule nat_induct, blast+) lemma rev_induct: "\n \ nat; P(n); n\0; \m. \m \ nat; m\0; P(succ(m))\ \ P(m)\ \ P(1)" by (rule rev_induct_lemma, blast+) lemma NN_into_nat: "n \ NN(y) \ n \ nat" by (simp add: NN_def) lemma lemma3: "\n \ NN(y); y*y \ y; n\0\ \ 1 \ NN(y)" apply (rule rev_induct [OF NN_into_nat], assumption+) apply (rule lemma_ii, assumption+) done (* ********************************************************************** *) (* Main theorem "WO6 \ WO1" *) (* ********************************************************************** *) (* another helpful lemma *) lemma NN_y_0: "0 \ NN(y) \ y=0" unfolding NN_def apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst) done lemma WO6_imp_WO1: "WO6 \ WO1" unfolding WO1_def apply (rule allI) apply (case_tac "A=0") apply (fast intro!: well_ord_Memrel nat_0I [THEN nat_into_Ord]) apply (rule_tac x = A in lemma_iv [elim_format]) apply (erule exE) apply (drule WO6_imp_NN_not_empty) apply (erule Un_subset_iff [THEN iffD1, THEN conjE]) apply (erule_tac A = "NN (y) " in not_emptyE) apply (frule y_well_ord) apply (fast intro!: lemma3 dest!: NN_y_0 elim!: not_emptyE) apply (fast elim: well_ord_subset) done end diff --git a/src/ZF/Bin.thy b/src/ZF/Bin.thy --- a/src/ZF/Bin.thy +++ b/src/ZF/Bin.thy @@ -1,763 +1,763 @@ (* Title: ZF/Bin.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge The sign Pls stands for an infinite string of leading 0's. The sign Min stands for an infinite string of leading 1's. A number can have multiple representations, namely leading 0's with sign Pls and leading 1's with sign Min. See twos-compl.ML/int_of_binary for the numerical interpretation. The representation expects that (m mod 2) is 0 or 1, even if m is negative; For instance, \5 div 2 = \3 and \5 mod 2 = 1; thus \5 = (\3)*2 + 1 *) section\Arithmetic on Binary Integers\ theory Bin imports Int Datatype begin consts bin :: i datatype "bin" = Pls | Min | Bit ("w \ bin", "b \ bool") (infixl \BIT\ 90) consts integ_of :: "i\i" NCons :: "[i,i]\i" bin_succ :: "i\i" bin_pred :: "i\i" bin_minus :: "i\i" bin_adder :: "i\i" bin_mult :: "[i,i]\i" primrec integ_of_Pls: "integ_of (Pls) = $# 0" integ_of_Min: "integ_of (Min) = $-($#1)" integ_of_BIT: "integ_of (w BIT b) = $#b $+ integ_of(w) $+ integ_of(w)" (** recall that cond(1,b,c)=b and cond(0,b,c)=0 **) primrec (*NCons adds a bit, suppressing leading 0s and 1s*) NCons_Pls: "NCons (Pls,b) = cond(b,Pls BIT b,Pls)" NCons_Min: "NCons (Min,b) = cond(b,Min,Min BIT b)" NCons_BIT: "NCons (w BIT c,b) = w BIT c BIT b" primrec (*successor. If a BIT, can change a 0 to a 1 without recursion.*) bin_succ_Pls: "bin_succ (Pls) = Pls BIT 1" bin_succ_Min: "bin_succ (Min) = Pls" bin_succ_BIT: "bin_succ (w BIT b) = cond(b, bin_succ(w) BIT 0, NCons(w,1))" primrec (*predecessor*) bin_pred_Pls: "bin_pred (Pls) = Min" bin_pred_Min: "bin_pred (Min) = Min BIT 0" bin_pred_BIT: "bin_pred (w BIT b) = cond(b, NCons(w,0), bin_pred(w) BIT 1)" primrec (*unary negation*) bin_minus_Pls: "bin_minus (Pls) = Pls" bin_minus_Min: "bin_minus (Min) = Pls BIT 1" bin_minus_BIT: "bin_minus (w BIT b) = cond(b, bin_pred(NCons(bin_minus(w),0)), bin_minus(w) BIT 0)" primrec (*sum*) bin_adder_Pls: "bin_adder (Pls) = (\w\bin. w)" bin_adder_Min: "bin_adder (Min) = (\w\bin. bin_pred(w))" bin_adder_BIT: "bin_adder (v BIT x) = (\w\bin. bin_case (v BIT x, bin_pred(v BIT x), \w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w), x xor y), w))" (*The bin_case above replaces the following mutually recursive function: primrec "adding (v,x,Pls) = v BIT x" "adding (v,x,Min) = bin_pred(v BIT x)" "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)), x xor y)" *) definition bin_add :: "[i,i]\i" where "bin_add(v,w) \ bin_adder(v)`w" primrec bin_mult_Pls: "bin_mult (Pls,w) = Pls" bin_mult_Min: "bin_mult (Min,w) = bin_minus(w)" bin_mult_BIT: "bin_mult (v BIT b,w) = cond(b, bin_add(NCons(bin_mult(v,w),0),w), NCons(bin_mult(v,w),0))" syntax "_Int0" :: i (\#' 0\) "_Int1" :: i (\#' 1\) "_Int2" :: i (\#' 2\) "_Neg_Int1" :: i (\#-' 1\) "_Neg_Int2" :: i (\#-' 2\) translations "#0" \ "CONST integ_of(CONST Pls)" "#1" \ "CONST integ_of(CONST Pls BIT 1)" "#2" \ "CONST integ_of(CONST Pls BIT 1 BIT 0)" "#-1" \ "CONST integ_of(CONST Min)" "#-2" \ "CONST integ_of(CONST Min BIT 0)" syntax "_Int" :: "num_token \ i" (\#_\ 1000) "_Neg_Int" :: "num_token \ i" (\#-_\ 1000) ML_file \Tools/numeral_syntax.ML\ declare bin.intros [simp,TC] lemma NCons_Pls_0: "NCons(Pls,0) = Pls" by simp lemma NCons_Pls_1: "NCons(Pls,1) = Pls BIT 1" by simp lemma NCons_Min_0: "NCons(Min,0) = Min BIT 0" by simp lemma NCons_Min_1: "NCons(Min,1) = Min" by simp lemma NCons_BIT: "NCons(w BIT x,b) = w BIT x BIT b" by (simp add: bin.case_eqns) lemmas NCons_simps [simp] = NCons_Pls_0 NCons_Pls_1 NCons_Min_0 NCons_Min_1 NCons_BIT (** Type checking **) lemma integ_of_type [TC]: "w \ bin \ integ_of(w) \ int" apply (induct_tac "w") apply (simp_all add: bool_into_nat) done lemma NCons_type [TC]: "\w \ bin; b \ bool\ \ NCons(w,b) \ bin" by (induct_tac "w", auto) lemma bin_succ_type [TC]: "w \ bin \ bin_succ(w) \ bin" by (induct_tac "w", auto) lemma bin_pred_type [TC]: "w \ bin \ bin_pred(w) \ bin" by (induct_tac "w", auto) lemma bin_minus_type [TC]: "w \ bin \ bin_minus(w) \ bin" by (induct_tac "w", auto) (*This proof is complicated by the mutual recursion*) lemma bin_add_type [rule_format]: "v \ bin \ \w\bin. bin_add(v,w) \ bin" unfolding bin_add_def apply (induct_tac "v") apply (rule_tac [3] ballI) apply (rename_tac [3] "w'") apply (induct_tac [3] "w'") apply (simp_all add: NCons_type) done declare bin_add_type [TC] lemma bin_mult_type [TC]: "\v \ bin; w \ bin\ \ bin_mult(v,w) \ bin" by (induct_tac "v", auto) subsubsection\The Carry and Borrow Functions, \<^term>\bin_succ\ and \<^term>\bin_pred\\ (*NCons preserves the integer value of its argument*) lemma integ_of_NCons [simp]: "\w \ bin; b \ bool\ \ integ_of(NCons(w,b)) = integ_of(w BIT b)" apply (erule bin.cases) apply (auto elim!: boolE) done lemma integ_of_succ [simp]: "w \ bin \ integ_of(bin_succ(w)) = $#1 $+ integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac elim!: boolE) done lemma integ_of_pred [simp]: "w \ bin \ integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac elim!: boolE) done subsubsection\\<^term>\bin_minus\: Unary Negation of Binary Integers\ lemma integ_of_minus: "w \ bin \ integ_of(bin_minus(w)) = $- integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac zminus_zadd_distrib elim!: boolE) done subsubsection\\<^term>\bin_add\: Binary Addition\ lemma bin_add_Pls [simp]: "w \ bin \ bin_add(Pls,w) = w" by (unfold bin_add_def, simp) lemma bin_add_Pls_right: "w \ bin \ bin_add(w,Pls) = w" unfolding bin_add_def apply (erule bin.induct, auto) done lemma bin_add_Min [simp]: "w \ bin \ bin_add(Min,w) = bin_pred(w)" by (unfold bin_add_def, simp) lemma bin_add_Min_right: "w \ bin \ bin_add(w,Min) = bin_pred(w)" unfolding bin_add_def apply (erule bin.induct, auto) done lemma bin_add_BIT_Pls [simp]: "bin_add(v BIT x,Pls) = v BIT x" by (unfold bin_add_def, simp) lemma bin_add_BIT_Min [simp]: "bin_add(v BIT x,Min) = bin_pred(v BIT x)" by (unfold bin_add_def, simp) lemma bin_add_BIT_BIT [simp]: "\w \ bin; y \ bool\ \ bin_add(v BIT x, w BIT y) = NCons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)" by (unfold bin_add_def, simp) lemma integ_of_add [rule_format]: "v \ bin \ \w\bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)" apply (erule bin.induct, simp, simp) apply (rule ballI) apply (induct_tac "wa") apply (auto simp add: zadd_ac elim!: boolE) done (*Subtraction*) lemma diff_integ_of_eq: "\v \ bin; w \ bin\ \ integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))" unfolding zdiff_def apply (simp add: integ_of_add integ_of_minus) done subsubsection\\<^term>\bin_mult\: Binary Multiplication\ lemma integ_of_mult: "\v \ bin; w \ bin\ \ integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)" apply (induct_tac "v", simp) apply (simp add: integ_of_minus) apply (auto simp add: zadd_ac integ_of_add zadd_zmult_distrib elim!: boolE) done subsection\Computations\ (** extra rules for bin_succ, bin_pred **) lemma bin_succ_1: "bin_succ(w BIT 1) = bin_succ(w) BIT 0" by simp lemma bin_succ_0: "bin_succ(w BIT 0) = NCons(w,1)" by simp lemma bin_pred_1: "bin_pred(w BIT 1) = NCons(w,0)" by simp lemma bin_pred_0: "bin_pred(w BIT 0) = bin_pred(w) BIT 1" by simp (** extra rules for bin_minus **) lemma bin_minus_1: "bin_minus(w BIT 1) = bin_pred(NCons(bin_minus(w), 0))" by simp lemma bin_minus_0: "bin_minus(w BIT 0) = bin_minus(w) BIT 0" by simp (** extra rules for bin_add **) lemma bin_add_BIT_11: "w \ bin \ bin_add(v BIT 1, w BIT 1) = NCons(bin_add(v, bin_succ(w)), 0)" by simp lemma bin_add_BIT_10: "w \ bin \ bin_add(v BIT 1, w BIT 0) = NCons(bin_add(v,w), 1)" by simp lemma bin_add_BIT_0: "\w \ bin; y \ bool\ \ bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)" by simp (** extra rules for bin_mult **) lemma bin_mult_1: "bin_mult(v BIT 1, w) = bin_add(NCons(bin_mult(v,w),0), w)" by simp lemma bin_mult_0: "bin_mult(v BIT 0, w) = NCons(bin_mult(v,w),0)" by simp (** Simplification rules with integer constants **) lemma int_of_0: "$#0 = #0" by simp lemma int_of_succ: "$# succ(n) = #1 $+ $#n" by (simp add: int_of_add [symmetric] natify_succ) lemma zminus_0 [simp]: "$- #0 = #0" by simp lemma zadd_0_intify [simp]: "#0 $+ z = intify(z)" by simp lemma zadd_0_right_intify [simp]: "z $+ #0 = intify(z)" by simp lemma zmult_1_intify [simp]: "#1 $* z = intify(z)" by simp lemma zmult_1_right_intify [simp]: "z $* #1 = intify(z)" by (subst zmult_commute, simp) lemma zmult_0 [simp]: "#0 $* z = #0" by simp lemma zmult_0_right [simp]: "z $* #0 = #0" by (subst zmult_commute, simp) lemma zmult_minus1 [simp]: "#-1 $* z = $-z" by (simp add: zcompare_rls) lemma zmult_minus1_right [simp]: "z $* #-1 = $-z" apply (subst zmult_commute) apply (rule zmult_minus1) done subsection\Simplification Rules for Comparison of Binary Numbers\ text\Thanks to Norbert Voelker\ (** Equals (=) **) lemma eq_integ_of_eq: "\v \ bin; w \ bin\ \ ((integ_of(v)) = integ_of(w)) \ iszero (integ_of (bin_add (v, bin_minus(w))))" unfolding iszero_def apply (simp add: zcompare_rls integ_of_add integ_of_minus) done lemma iszero_integ_of_Pls: "iszero (integ_of(Pls))" by (unfold iszero_def, simp) lemma nonzero_integ_of_Min: "\ iszero (integ_of(Min))" unfolding iszero_def apply (simp add: zminus_equation) done lemma iszero_integ_of_BIT: "\w \ bin; x \ bool\ \ iszero (integ_of (w BIT x)) \ (x=0 \ iszero (integ_of(w)))" apply (unfold iszero_def, simp) apply (subgoal_tac "integ_of (w) \ int") apply typecheck apply (drule int_cases) apply (safe elim!: boolE) apply (simp_all (asm_lr) add: zcompare_rls zminus_zadd_distrib [symmetric] int_of_add [symmetric]) done lemma iszero_integ_of_0: "w \ bin \ iszero (integ_of (w BIT 0)) \ iszero (integ_of(w))" by (simp only: iszero_integ_of_BIT, blast) lemma iszero_integ_of_1: "w \ bin \ \ iszero (integ_of (w BIT 1))" by (simp only: iszero_integ_of_BIT, blast) (** Less-than (<) **) lemma less_integ_of_eq_neg: "\v \ bin; w \ bin\ \ integ_of(v) $< integ_of(w) \ znegative (integ_of (bin_add (v, bin_minus(w))))" -apply (unfold zless_def zdiff_def) + unfolding zless_def zdiff_def apply (simp add: integ_of_minus integ_of_add) done lemma not_neg_integ_of_Pls: "\ znegative (integ_of(Pls))" by simp lemma neg_integ_of_Min: "znegative (integ_of(Min))" by simp lemma neg_integ_of_BIT: "\w \ bin; x \ bool\ \ znegative (integ_of (w BIT x)) \ znegative (integ_of(w))" apply simp apply (subgoal_tac "integ_of (w) \ int") apply typecheck apply (drule int_cases) apply (auto elim!: boolE simp add: int_of_add [symmetric] zcompare_rls) apply (simp_all add: zminus_zadd_distrib [symmetric] zdiff_def int_of_add [symmetric]) apply (subgoal_tac "$#1 $- $# succ (succ (n #+ n)) = $- $# succ (n #+ n) ") apply (simp add: zdiff_def) apply (simp add: equation_zminus int_of_diff [symmetric]) done (** Less-than-or-equals (<=) **) lemma le_integ_of_eq_not_less: "(integ_of(x) $\ (integ_of(w))) \ \ (integ_of(w) $< (integ_of(x)))" by (simp add: not_zless_iff_zle [THEN iff_sym]) (*Delete the original rewrites, with their clumsy conditional expressions*) declare bin_succ_BIT [simp del] bin_pred_BIT [simp del] bin_minus_BIT [simp del] NCons_Pls [simp del] NCons_Min [simp del] bin_adder_BIT [simp del] bin_mult_BIT [simp del] (*Hide the binary representation of integer constants*) declare integ_of_Pls [simp del] integ_of_Min [simp del] integ_of_BIT [simp del] lemmas bin_arith_extra_simps = integ_of_add [symmetric] integ_of_minus [symmetric] integ_of_mult [symmetric] bin_succ_1 bin_succ_0 bin_pred_1 bin_pred_0 bin_minus_1 bin_minus_0 bin_add_Pls_right bin_add_Min_right bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11 diff_integ_of_eq bin_mult_1 bin_mult_0 NCons_simps (*For making a minimal simpset, one must include these default simprules of thy. Also include simp_thms, or at least (\False)=True*) lemmas bin_arith_simps = bin_pred_Pls bin_pred_Min bin_succ_Pls bin_succ_Min bin_add_Pls bin_add_Min bin_minus_Pls bin_minus_Min bin_mult_Pls bin_mult_Min bin_arith_extra_simps (*Simplification of relational operations*) lemmas bin_rel_simps = eq_integ_of_eq iszero_integ_of_Pls nonzero_integ_of_Min iszero_integ_of_0 iszero_integ_of_1 less_integ_of_eq_neg not_neg_integ_of_Pls neg_integ_of_Min neg_integ_of_BIT le_integ_of_eq_not_less declare bin_arith_simps [simp] declare bin_rel_simps [simp] (** Simplification of arithmetic when nested to the right **) lemma add_integ_of_left [simp]: "\v \ bin; w \ bin\ \ integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)" by (simp add: zadd_assoc [symmetric]) lemma mult_integ_of_left [simp]: "\v \ bin; w \ bin\ \ integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)" by (simp add: zmult_assoc [symmetric]) lemma add_integ_of_diff1 [simp]: "\v \ bin; w \ bin\ \ integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)" unfolding zdiff_def apply (rule add_integ_of_left, auto) done lemma add_integ_of_diff2 [simp]: "\v \ bin; w \ bin\ \ integ_of(v) $+ (c $- integ_of(w)) = integ_of (bin_add (v, bin_minus(w))) $+ (c)" apply (subst diff_integ_of_eq [symmetric]) apply (simp_all add: zdiff_def zadd_ac) done (** More for integer constants **) declare int_of_0 [simp] int_of_succ [simp] lemma zdiff0 [simp]: "#0 $- x = $-x" by (simp add: zdiff_def) lemma zdiff0_right [simp]: "x $- #0 = intify(x)" by (simp add: zdiff_def) lemma zdiff_self [simp]: "x $- x = #0" by (simp add: zdiff_def) lemma znegative_iff_zless_0: "k \ int \ znegative(k) \ k $< #0" by (simp add: zless_def) lemma zero_zless_imp_znegative_zminus: "\#0 $< k; k \ int\ \ znegative($-k)" by (simp add: zless_def) lemma zero_zle_int_of [simp]: "#0 $\ $# n" by (simp add: not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym]) lemma nat_of_0 [simp]: "nat_of(#0) = 0" by (simp only: natify_0 int_of_0 [symmetric] nat_of_int_of) lemma nat_le_int0_lemma: "\z $\ $#0; z \ int\ \ nat_of(z) = 0" by (auto simp add: znegative_iff_zless_0 [THEN iff_sym] zle_def zneg_nat_of) lemma nat_le_int0: "z $\ $#0 \ nat_of(z) = 0" apply (subgoal_tac "nat_of (intify (z)) = 0") apply (rule_tac [2] nat_le_int0_lemma, auto) done lemma int_of_eq_0_imp_natify_eq_0: "$# n = #0 \ natify(n) = 0" by (rule not_znegative_imp_zero, auto) lemma nat_of_zminus_int_of: "nat_of($- $# n) = 0" by (simp add: nat_of_def int_of_def raw_nat_of zminus image_intrel_int) lemma int_of_nat_of: "#0 $\ z \ $# nat_of(z) = intify(z)" apply (rule not_zneg_nat_of_intify) apply (simp add: znegative_iff_zless_0 not_zless_iff_zle) done declare int_of_nat_of [simp] nat_of_zminus_int_of [simp] lemma int_of_nat_of_if: "$# nat_of(z) = (if #0 $\ z then intify(z) else #0)" by (simp add: int_of_nat_of znegative_iff_zless_0 not_zle_iff_zless) lemma zless_nat_iff_int_zless: "\m \ nat; z \ int\ \ (m < nat_of(z)) \ ($#m $< z)" apply (case_tac "znegative (z) ") apply (erule_tac [2] not_zneg_nat_of [THEN subst]) apply (auto dest: zless_trans dest!: zero_zle_int_of [THEN zle_zless_trans] simp add: znegative_iff_zless_0) done (** nat_of and zless **) (*An alternative condition is @{term"$#0 \ w"} *) lemma zless_nat_conj_lemma: "$#0 $< z \ (nat_of(w) < nat_of(z)) \ (w $< z)" apply (rule iff_trans) apply (rule zless_int_of [THEN iff_sym]) apply (auto simp add: int_of_nat_of_if simp del: zless_int_of) apply (auto elim: zless_asym simp add: not_zle_iff_zless) apply (blast intro: zless_zle_trans) done lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) \ ($#0 $< z \ w $< z)" apply (case_tac "$#0 $< z") apply (auto simp add: zless_nat_conj_lemma nat_le_int0 not_zless_iff_zle) done (*This simprule cannot be added unless we can find a way to make eq_integ_of_eq unconditional! [The condition "True" is a hack to prevent looping. Conditional rewrite rules are tried after unconditional ones, so a rule like eq_nat_number_of will be tried first to eliminate #mm=#nn.] lemma integ_of_reorient [simp]: "True \ (integ_of(w) = x) \ (x = integ_of(w))" by auto *) lemma integ_of_minus_reorient [simp]: "(integ_of(w) = $- x) \ ($- x = integ_of(w))" by auto lemma integ_of_add_reorient [simp]: "(integ_of(w) = x $+ y) \ (x $+ y = integ_of(w))" by auto lemma integ_of_diff_reorient [simp]: "(integ_of(w) = x $- y) \ (x $- y = integ_of(w))" by auto lemma integ_of_mult_reorient [simp]: "(integ_of(w) = x $* y) \ (x $* y = integ_of(w))" by auto (** To simplify inequalities involving integer negation and literals, such as -x = #3 **) lemmas [simp] = zminus_equation [where y = "integ_of(w)"] equation_zminus [where x = "integ_of(w)"] for w lemmas [iff] = zminus_zless [where y = "integ_of(w)"] zless_zminus [where x = "integ_of(w)"] for w lemmas [iff] = zminus_zle [where y = "integ_of(w)"] zle_zminus [where x = "integ_of(w)"] for w lemmas [simp] = Let_def [where s = "integ_of(w)"] for w (*** Simprocs for numeric literals ***) (** Combining of literal coefficients in sums of products **) lemma zless_iff_zdiff_zless_0: "(x $< y) \ (x$-y $< #0)" by (simp add: zcompare_rls) lemma eq_iff_zdiff_eq_0: "\x \ int; y \ int\ \ (x = y) \ (x$-y = #0)" by (simp add: zcompare_rls) lemma zle_iff_zdiff_zle_0: "(x $\ y) \ (x$-y $\ #0)" by (simp add: zcompare_rls) (** For combine_numerals **) lemma left_zadd_zmult_distrib: "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k" by (simp add: zadd_zmult_distrib zadd_ac) (** For cancel_numerals **) lemma eq_add_iff1: "(i$*u $+ m = j$*u $+ n) \ ((i$-j)$*u $+ m = intify(n))" apply (simp add: zdiff_def zadd_zmult_distrib) apply (simp add: zcompare_rls) apply (simp add: zadd_ac) done lemma eq_add_iff2: "(i$*u $+ m = j$*u $+ n) \ (intify(m) = (j$-i)$*u $+ n)" apply (simp add: zdiff_def zadd_zmult_distrib) apply (simp add: zcompare_rls) apply (simp add: zadd_ac) done context fixes n :: i begin lemmas rel_iff_rel_0_rls = zless_iff_zdiff_zless_0 [where y = "u $+ v"] eq_iff_zdiff_eq_0 [where y = "u $+ v"] zle_iff_zdiff_zle_0 [where y = "u $+ v"] zless_iff_zdiff_zless_0 [where y = n] eq_iff_zdiff_eq_0 [where y = n] zle_iff_zdiff_zle_0 [where y = n] for u v lemma less_add_iff1: "(i$*u $+ m $< j$*u $+ n) \ ((i$-j)$*u $+ m $< n)" apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls) done lemma less_add_iff2: "(i$*u $+ m $< j$*u $+ n) \ (m $< (j$-i)$*u $+ n)" apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls) done end lemma le_add_iff1: "(i$*u $+ m $\ j$*u $+ n) \ ((i$-j)$*u $+ m $\ n)" apply (simp add: zdiff_def zadd_zmult_distrib) apply (simp add: zcompare_rls) apply (simp add: zadd_ac) done lemma le_add_iff2: "(i$*u $+ m $\ j$*u $+ n) \ (m $\ (j$-i)$*u $+ n)" apply (simp add: zdiff_def zadd_zmult_distrib) apply (simp add: zcompare_rls) apply (simp add: zadd_ac) done ML_file \int_arith.ML\ subsection \examples:\ text \\combine_numerals_prod\ (products of separate literals)\ lemma "#5 $* x $* #3 = y" apply simp oops schematic_goal "y2 $+ ?x42 = y $+ y2" apply simp oops lemma "oo : int \ l $+ (l $+ #2) $+ oo = oo" apply simp oops lemma "#9$*x $+ y = x$*#23 $+ z" apply simp oops lemma "y $+ x = x $+ z" apply simp oops lemma "x : int \ x $+ y $+ z = x $+ z" apply simp oops lemma "x : int \ y $+ (z $+ x) = z $+ x" apply simp oops lemma "z : int \ x $+ y $+ z = (z $+ y) $+ (x $+ w)" apply simp oops lemma "z : int \ x$*y $+ z = (z $+ y) $+ (y$*x $+ w)" apply simp oops lemma "#-3 $* x $+ y $\ x $* #2 $+ z" apply simp oops lemma "y $+ x $\ x $+ z" apply simp oops lemma "x $+ y $+ z $\ x $+ z" apply simp oops lemma "y $+ (z $+ x) $< z $+ x" apply simp oops lemma "x $+ y $+ z $< (z $+ y) $+ (x $+ w)" apply simp oops lemma "x$*y $+ z $< (z $+ y) $+ (y$*x $+ w)" apply simp oops lemma "l $+ #2 $+ #2 $+ #2 $+ (l $+ #2) $+ (oo $+ #2) = uu" apply simp oops lemma "u : int \ #2 $* u = u" apply simp oops lemma "(i $+ j $+ #12 $+ k) $- #15 = y" apply simp oops lemma "(i $+ j $+ #12 $+ k) $- #5 = y" apply simp oops lemma "y $- b $< b" apply simp oops lemma "y $- (#3 $* b $+ c) $< b $- #2 $* c" apply simp oops lemma "(#2 $* x $- (u $* v) $+ y) $- v $* #3 $* u = w" apply simp oops lemma "(#2 $* x $* u $* v $+ (u $* v) $* #4 $+ y) $- v $* u $* #4 = w" apply simp oops lemma "(#2 $* x $* u $* v $+ (u $* v) $* #4 $+ y) $- v $* u = w" apply simp oops lemma "u $* v $- (x $* u $* v $+ (u $* v) $* #4 $+ y) = w" apply simp oops lemma "(i $+ j $+ #12 $+ k) = u $+ #15 $+ y" apply simp oops lemma "(i $+ j $* #2 $+ #12 $+ k) = j $+ #5 $+ y" apply simp oops lemma "#2 $* y $+ #3 $* z $+ #6 $* w $+ #2 $* y $+ #3 $* z $+ #2 $* u = #2 $* y' $+ #3 $* z' $+ #6 $* w' $+ #2 $* y' $+ #3 $* z' $+ u $+ vv" apply simp oops lemma "a $+ $-(b$+c) $+ b = d" apply simp oops lemma "a $+ $-(b$+c) $- b = d" apply simp oops text \negative numerals\ lemma "(i $+ j $+ #-2 $+ k) $- (u $+ #5 $+ y) = zz" apply simp oops lemma "(i $+ j $+ #-3 $+ k) $< u $+ #5 $+ y" apply simp oops lemma "(i $+ j $+ #3 $+ k) $< u $+ #-6 $+ y" apply simp oops lemma "(i $+ j $+ #-12 $+ k) $- #15 = y" apply simp oops lemma "(i $+ j $+ #12 $+ k) $- #-15 = y" apply simp oops lemma "(i $+ j $+ #-12 $+ k) $- #-15 = y" apply simp oops text \Multiplying separated numerals\ lemma "#6 $* ($# x $* #2) = uu" apply simp oops lemma "#4 $* ($# x $* $# x) $* (#2 $* $# x) = uu" apply simp oops end diff --git a/src/ZF/Cardinal.thy b/src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy +++ b/src/ZF/Cardinal.thy @@ -1,1180 +1,1180 @@ (* Title: ZF/Cardinal.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Cardinal Numbers Without the Axiom of Choice\ theory Cardinal imports OrderType Finite Nat Sum begin definition (*least ordinal operator*) Least :: "(i\o) \ i" (binder \\ \ 10) where "Least(P) \ THE i. Ord(i) \ P(i) \ (\j. j \P(j))" definition eqpoll :: "[i,i] \ o" (infixl \\\ 50) where "A \ B \ \f. f \ bij(A,B)" definition lepoll :: "[i,i] \ o" (infixl \\\ 50) where "A \ B \ \f. f \ inj(A,B)" definition lesspoll :: "[i,i] \ o" (infixl \\\ 50) where "A \ B \ A \ B \ \(A \ B)" definition cardinal :: "i\i" (\|_|\) where "|A| \ (\ i. i \ A)" definition Finite :: "i\o" where "Finite(A) \ \n\nat. A \ n" definition Card :: "i\o" where "Card(i) \ (i = |i|)" subsection\The Schroeder-Bernstein Theorem\ text\See Davey and Priestly, page 106\ (** Lemma: Banach's Decomposition Theorem **) lemma decomp_bnd_mono: "bnd_mono(X, \W. X - g``(Y - f``W))" by (rule bnd_monoI, blast+) lemma Banach_last_equation: "g \ Y->X \ g``(Y - f`` lfp(X, \W. X - g``(Y - f``W))) = X - lfp(X, \W. X - g``(Y - f``W))" apply (rule_tac P = "\u. v = X-u" for v in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst]) apply (simp add: double_complement fun_is_rel [THEN image_subset]) done lemma decomposition: "\f \ X->Y; g \ Y->X\ \ \XA XB YA YB. (XA \ XB = 0) \ (XA \ XB = X) \ (YA \ YB = 0) \ (YA \ YB = Y) \ f``XA=YA \ g``YB=XB" apply (intro exI conjI) apply (rule_tac [6] Banach_last_equation) apply (rule_tac [5] refl) apply (assumption | rule Diff_disjoint Diff_partition fun_is_rel image_subset lfp_subset)+ done lemma schroeder_bernstein: "\f \ inj(X,Y); g \ inj(Y,X)\ \ \h. h \ bij(X,Y)" apply (insert decomposition [of f X Y g]) apply (simp add: inj_is_fun) apply (blast intro!: restrict_bij bij_disjoint_Un intro: bij_converse_bij) (* The instantiation of exI to @{term"restrict(f,XA) \ converse(restrict(g,YB))"} is forced by the context\*) done (** Equipollence is an equivalence relation **) lemma bij_imp_eqpoll: "f \ bij(A,B) \ A \ B" unfolding eqpoll_def apply (erule exI) done (*A \ A*) lemmas eqpoll_refl = id_bij [THEN bij_imp_eqpoll, simp] lemma eqpoll_sym: "X \ Y \ Y \ X" unfolding eqpoll_def apply (blast intro: bij_converse_bij) done lemma eqpoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" unfolding eqpoll_def apply (blast intro: comp_bij) done (** Le-pollence is a partial ordering **) lemma subset_imp_lepoll: "X<=Y \ X \ Y" unfolding lepoll_def apply (rule exI) apply (erule id_subset_inj) done lemmas lepoll_refl = subset_refl [THEN subset_imp_lepoll, simp] lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll] lemma eqpoll_imp_lepoll: "X \ Y \ X \ Y" by (unfold eqpoll_def bij_def lepoll_def, blast) lemma lepoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" unfolding lepoll_def apply (blast intro: comp_inj) done lemma eq_lepoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (blast intro: eqpoll_imp_lepoll lepoll_trans) lemma lepoll_eq_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (blast intro: eqpoll_imp_lepoll lepoll_trans) (*Asymmetry law*) lemma eqpollI: "\X \ Y; Y \ X\ \ X \ Y" -apply (unfold lepoll_def eqpoll_def) + unfolding lepoll_def eqpoll_def apply (elim exE) apply (rule schroeder_bernstein, assumption+) done lemma eqpollE: "\X \ Y; \X \ Y; Y \ X\ \ P\ \ P" by (blast intro: eqpoll_imp_lepoll eqpoll_sym) lemma eqpoll_iff: "X \ Y \ X \ Y \ Y \ X" by (blast intro: eqpollI elim!: eqpollE) lemma lepoll_0_is_0: "A \ 0 \ A = 0" -apply (unfold lepoll_def inj_def) + unfolding lepoll_def inj_def apply (blast dest: apply_type) done (*@{term"0 \ Y"}*) lemmas empty_lepollI = empty_subsetI [THEN subset_imp_lepoll] lemma lepoll_0_iff: "A \ 0 \ A=0" by (blast intro: lepoll_0_is_0 lepoll_refl) lemma Un_lepoll_Un: "\A \ B; C \ D; B \ D = 0\ \ A \ C \ B \ D" unfolding lepoll_def apply (blast intro: inj_disjoint_Un) done (*A \ 0 \ A=0*) lemmas eqpoll_0_is_0 = eqpoll_imp_lepoll [THEN lepoll_0_is_0] lemma eqpoll_0_iff: "A \ 0 \ A=0" by (blast intro: eqpoll_0_is_0 eqpoll_refl) lemma eqpoll_disjoint_Un: "\A \ B; C \ D; A \ C = 0; B \ D = 0\ \ A \ C \ B \ D" unfolding eqpoll_def apply (blast intro: bij_disjoint_Un) done subsection\lesspoll: contributions by Krzysztof Grabczewski\ lemma lesspoll_not_refl: "\ (i \ i)" by (simp add: lesspoll_def) lemma lesspoll_irrefl [elim!]: "i \ i \ P" by (simp add: lesspoll_def) lemma lesspoll_imp_lepoll: "A \ B \ A \ B" by (unfold lesspoll_def, blast) lemma lepoll_well_ord: "\A \ B; well_ord(B,r)\ \ \s. well_ord(A,s)" unfolding lepoll_def apply (blast intro: well_ord_rvimage) done lemma lepoll_iff_leqpoll: "A \ B \ A \ B | A \ B" unfolding lesspoll_def apply (blast intro!: eqpollI elim!: eqpollE) done lemma inj_not_surj_succ: assumes fi: "f \ inj(A, succ(m))" and fns: "f \ surj(A, succ(m))" shows "\f. f \ inj(A,m)" proof - from fi [THEN inj_is_fun] fns obtain y where y: "y \ succ(m)" "\x. x\A \ f ` x \ y" by (auto simp add: surj_def) show ?thesis proof show "(\z\A. if f`z = m then y else f`z) \ inj(A, m)" using y fi by (simp add: inj_def) (auto intro!: if_type [THEN lam_type] intro: Pi_type dest: apply_funtype) qed qed (** Variations on transitivity **) lemma lesspoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" unfolding lesspoll_def apply (blast elim!: eqpollE intro: eqpollI lepoll_trans) done lemma lesspoll_trans1 [trans]: "\X \ Y; Y \ Z\ \ X \ Z" unfolding lesspoll_def apply (blast elim!: eqpollE intro: eqpollI lepoll_trans) done lemma lesspoll_trans2 [trans]: "\X \ Y; Y \ Z\ \ X \ Z" unfolding lesspoll_def apply (blast elim!: eqpollE intro: eqpollI lepoll_trans) done lemma eq_lesspoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (blast intro: eqpoll_imp_lepoll lesspoll_trans1) lemma lesspoll_eq_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (blast intro: eqpoll_imp_lepoll lesspoll_trans2) (** \ -- the least number operator [from HOL/Univ.ML] **) lemma Least_equality: "\P(i); Ord(i); \x. x \P(x)\ \ (\ x. P(x)) = i" unfolding Least_def apply (rule the_equality, blast) apply (elim conjE) apply (erule Ord_linear_lt, assumption, blast+) done lemma LeastI: assumes P: "P(i)" and i: "Ord(i)" shows "P(\ x. P(x))" proof - { from i have "P(i) \ P(\ x. P(x))" proof (induct i rule: trans_induct) case (step i) show ?case proof (cases "P(\ a. P(a))") case True thus ?thesis . next case False hence "\x. x \ i \ \P(x)" using step by blast hence "(\ a. P(a)) = i" using step by (blast intro: Least_equality ltD) thus ?thesis using step.prems by simp qed qed } thus ?thesis using P . qed text\The proof is almost identical to the one above!\ lemma Least_le: assumes P: "P(i)" and i: "Ord(i)" shows "(\ x. P(x)) \ i" proof - { from i have "P(i) \ (\ x. P(x)) \ i" proof (induct i rule: trans_induct) case (step i) show ?case proof (cases "(\ a. P(a)) \ i") case True thus ?thesis . next case False hence "\x. x \ i \ \ (\ a. P(a)) \ i" using step by blast hence "(\ a. P(a)) = i" using step by (blast elim: ltE intro: ltI Least_equality lt_trans1) thus ?thesis using step by simp qed qed } thus ?thesis using P . qed (*\ really is the smallest*) lemma less_LeastE: "\P(i); i < (\ x. P(x))\ \ Q" apply (rule Least_le [THEN [2] lt_trans2, THEN lt_irrefl], assumption+) apply (simp add: lt_Ord) done (*Easier to apply than LeastI: conclusion has only one occurrence of P*) lemma LeastI2: "\P(i); Ord(i); \j. P(j) \ Q(j)\ \ Q(\ j. P(j))" by (blast intro: LeastI ) (*If there is no such P then \ is vacuously 0*) lemma Least_0: "\\ (\i. Ord(i) \ P(i))\ \ (\ x. P(x)) = 0" unfolding Least_def apply (rule the_0, blast) done lemma Ord_Least [intro,simp,TC]: "Ord(\ x. P(x))" proof (cases "\i. Ord(i) \ P(i)") case True then obtain i where "P(i)" "Ord(i)" by auto hence " (\ x. P(x)) \ i" by (rule Least_le) thus ?thesis by (elim ltE) next case False hence "(\ x. P(x)) = 0" by (rule Least_0) thus ?thesis by auto qed subsection\Basic Properties of Cardinals\ (*Not needed for simplification, but helpful below*) lemma Least_cong: "(\y. P(y) \ Q(y)) \ (\ x. P(x)) = (\ x. Q(x))" by simp (*Need AC to get @{term"X \ Y \ |X| \ |Y|"}; see well_ord_lepoll_imp_cardinal_le Converse also requires AC, but see well_ord_cardinal_eqE*) lemma cardinal_cong: "X \ Y \ |X| = |Y|" -apply (unfold eqpoll_def cardinal_def) + unfolding eqpoll_def cardinal_def apply (rule Least_cong) apply (blast intro: comp_bij bij_converse_bij) done (*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*) lemma well_ord_cardinal_eqpoll: assumes r: "well_ord(A,r)" shows "|A| \ A" proof (unfold cardinal_def) show "(\ i. i \ A) \ A" by (best intro: LeastI Ord_ordertype ordermap_bij bij_converse_bij bij_imp_eqpoll r) qed (* @{term"Ord(A) \ |A| \ A"} *) lemmas Ord_cardinal_eqpoll = well_ord_Memrel [THEN well_ord_cardinal_eqpoll] lemma Ord_cardinal_idem: "Ord(A) \ ||A|| = |A|" by (rule Ord_cardinal_eqpoll [THEN cardinal_cong]) lemma well_ord_cardinal_eqE: assumes woX: "well_ord(X,r)" and woY: "well_ord(Y,s)" and eq: "|X| = |Y|" shows "X \ Y" proof - have "X \ |X|" by (blast intro: well_ord_cardinal_eqpoll [OF woX] eqpoll_sym) also have "... = |Y|" by (rule eq) also have "... \ Y" by (rule well_ord_cardinal_eqpoll [OF woY]) finally show ?thesis . qed lemma well_ord_cardinal_eqpoll_iff: "\well_ord(X,r); well_ord(Y,s)\ \ |X| = |Y| \ X \ Y" by (blast intro: cardinal_cong well_ord_cardinal_eqE) (** Observations from Kunen, page 28 **) lemma Ord_cardinal_le: "Ord(i) \ |i| \ i" unfolding cardinal_def apply (erule eqpoll_refl [THEN Least_le]) done lemma Card_cardinal_eq: "Card(K) \ |K| = K" unfolding Card_def apply (erule sym) done (* Could replace the @{term"\(j \ i)"} by @{term"\(i \ j)"}. *) lemma CardI: "\Ord(i); \j. j \(j \ i)\ \ Card(i)" -apply (unfold Card_def cardinal_def) + unfolding Card_def cardinal_def apply (subst Least_equality) apply (blast intro: eqpoll_refl)+ done lemma Card_is_Ord: "Card(i) \ Ord(i)" -apply (unfold Card_def cardinal_def) + unfolding Card_def cardinal_def apply (erule ssubst) apply (rule Ord_Least) done lemma Card_cardinal_le: "Card(K) \ K \ |K|" apply (simp (no_asm_simp) add: Card_is_Ord Card_cardinal_eq) done lemma Ord_cardinal [simp,intro!]: "Ord(|A|)" unfolding cardinal_def apply (rule Ord_Least) done text\The cardinals are the initial ordinals.\ lemma Card_iff_initial: "Card(K) \ Ord(K) \ (\j. j \ j \ K)" proof - { fix j assume K: "Card(K)" "j \ K" assume "j < K" also have "... = (\ i. i \ K)" using K by (simp add: Card_def cardinal_def) finally have "j < (\ i. i \ K)" . hence "False" using K by (best dest: less_LeastE) } then show ?thesis by (blast intro: CardI Card_is_Ord) qed lemma lt_Card_imp_lesspoll: "\Card(a); i \ i \ a" unfolding lesspoll_def apply (drule Card_iff_initial [THEN iffD1]) apply (blast intro!: leI [THEN le_imp_lepoll]) done lemma Card_0: "Card(0)" apply (rule Ord_0 [THEN CardI]) apply (blast elim!: ltE) done lemma Card_Un: "\Card(K); Card(L)\ \ Card(K \ L)" apply (rule Ord_linear_le [of K L]) apply (simp_all add: subset_Un_iff [THEN iffD1] Card_is_Ord le_imp_subset subset_Un_iff2 [THEN iffD1]) done (*Infinite unions of cardinals? See Devlin, Lemma 6.7, page 98*) lemma Card_cardinal [iff]: "Card(|A|)" proof (unfold cardinal_def) show "Card(\ i. i \ A)" proof (cases "\i. Ord (i) \ i \ A") case False thus ?thesis \ \degenerate case\ by (simp add: Least_0 Card_0) next case True \ \real case: \<^term>\A\ is isomorphic to some ordinal\ then obtain i where i: "Ord(i)" "i \ A" by blast show ?thesis proof (rule CardI [OF Ord_Least], rule notI) fix j assume j: "j < (\ i. i \ A)" assume "j \ (\ i. i \ A)" also have "... \ A" using i by (auto intro: LeastI) finally have "j \ A" . thus False by (rule less_LeastE [OF _ j]) qed qed qed (*Kunen's Lemma 10.5*) lemma cardinal_eq_lemma: assumes i:"|i| \ j" and j: "j \ i" shows "|j| = |i|" proof (rule eqpollI [THEN cardinal_cong]) show "j \ i" by (rule le_imp_lepoll [OF j]) next have Oi: "Ord(i)" using j by (rule le_Ord2) hence "i \ |i|" by (blast intro: Ord_cardinal_eqpoll eqpoll_sym) also have "... \ j" by (blast intro: le_imp_lepoll i) finally show "i \ j" . qed lemma cardinal_mono: assumes ij: "i \ j" shows "|i| \ |j|" using Ord_cardinal [of i] Ord_cardinal [of j] proof (cases rule: Ord_linear_le) case le thus ?thesis . next case ge have i: "Ord(i)" using ij by (simp add: lt_Ord) have ci: "|i| \ j" by (blast intro: Ord_cardinal_le ij le_trans i) have "|i| = ||i||" by (auto simp add: Ord_cardinal_idem i) also have "... = |j|" by (rule cardinal_eq_lemma [OF ge ci]) finally have "|i| = |j|" . thus ?thesis by simp qed text\Since we have \<^term>\|succ(nat)| \ |nat|\, the converse of \cardinal_mono\ fails!\ lemma cardinal_lt_imp_lt: "\|i| < |j|; Ord(i); Ord(j)\ \ i < j" apply (rule Ord_linear2 [of i j], assumption+) apply (erule lt_trans2 [THEN lt_irrefl]) apply (erule cardinal_mono) done lemma Card_lt_imp_lt: "\|i| < K; Ord(i); Card(K)\ \ i < K" by (simp (no_asm_simp) add: cardinal_lt_imp_lt Card_is_Ord Card_cardinal_eq) lemma Card_lt_iff: "\Ord(i); Card(K)\ \ (|i| < K) \ (i < K)" by (blast intro: Card_lt_imp_lt Ord_cardinal_le [THEN lt_trans1]) lemma Card_le_iff: "\Ord(i); Card(K)\ \ (K \ |i|) \ (K \ i)" by (simp add: Card_lt_iff Card_is_Ord Ord_cardinal not_lt_iff_le [THEN iff_sym]) (*Can use AC or finiteness to discharge first premise*) lemma well_ord_lepoll_imp_cardinal_le: assumes wB: "well_ord(B,r)" and AB: "A \ B" shows "|A| \ |B|" using Ord_cardinal [of A] Ord_cardinal [of B] proof (cases rule: Ord_linear_le) case le thus ?thesis . next case ge from lepoll_well_ord [OF AB wB] obtain s where s: "well_ord(A, s)" by blast have "B \ |B|" by (blast intro: wB eqpoll_sym well_ord_cardinal_eqpoll) also have "... \ |A|" by (rule le_imp_lepoll [OF ge]) also have "... \ A" by (rule well_ord_cardinal_eqpoll [OF s]) finally have "B \ A" . hence "A \ B" by (blast intro: eqpollI AB) hence "|A| = |B|" by (rule cardinal_cong) thus ?thesis by simp qed lemma lepoll_cardinal_le: "\A \ i; Ord(i)\ \ |A| \ i" apply (rule le_trans) apply (erule well_ord_Memrel [THEN well_ord_lepoll_imp_cardinal_le], assumption) apply (erule Ord_cardinal_le) done lemma lepoll_Ord_imp_eqpoll: "\A \ i; Ord(i)\ \ |A| \ A" by (blast intro: lepoll_cardinal_le well_ord_Memrel well_ord_cardinal_eqpoll dest!: lepoll_well_ord) lemma lesspoll_imp_eqpoll: "\A \ i; Ord(i)\ \ |A| \ A" unfolding lesspoll_def apply (blast intro: lepoll_Ord_imp_eqpoll) done lemma cardinal_subset_Ord: "\A<=i; Ord(i)\ \ |A| \ i" apply (drule subset_imp_lepoll [THEN lepoll_cardinal_le]) apply (auto simp add: lt_def) apply (blast intro: Ord_trans) done subsection\The finite cardinals\ lemma cons_lepoll_consD: "\cons(u,A) \ cons(v,B); u\A; v\B\ \ A \ B" apply (unfold lepoll_def inj_def, safe) apply (rule_tac x = "\x\A. if f`x=v then f`u else f`x" in exI) apply (rule CollectI) (*Proving it's in the function space A->B*) apply (rule if_type [THEN lam_type]) apply (blast dest: apply_funtype) apply (blast elim!: mem_irrefl dest: apply_funtype) (*Proving it's injective*) apply (simp (no_asm_simp)) apply blast done lemma cons_eqpoll_consD: "\cons(u,A) \ cons(v,B); u\A; v\B\ \ A \ B" apply (simp add: eqpoll_iff) apply (blast intro: cons_lepoll_consD) done (*Lemma suggested by Mike Fourman*) lemma succ_lepoll_succD: "succ(m) \ succ(n) \ m \ n" unfolding succ_def apply (erule cons_lepoll_consD) apply (rule mem_not_refl)+ done lemma nat_lepoll_imp_le: "m \ nat \ n \ nat \ m \ n \ m \ n" proof (induct m arbitrary: n rule: nat_induct) case 0 thus ?case by (blast intro!: nat_0_le) next case (succ m) show ?case using \n \ nat\ proof (cases rule: natE) case 0 thus ?thesis using succ by (simp add: lepoll_def inj_def) next case (succ n') thus ?thesis using succ.hyps \ succ(m) \ n\ by (blast intro!: succ_leI dest!: succ_lepoll_succD) qed qed lemma nat_eqpoll_iff: "\m \ nat; n \ nat\ \ m \ n \ m = n" apply (rule iffI) apply (blast intro: nat_lepoll_imp_le le_anti_sym elim!: eqpollE) apply (simp add: eqpoll_refl) done (*The object of all this work: every natural number is a (finite) cardinal*) lemma nat_into_Card: assumes n: "n \ nat" shows "Card(n)" proof (unfold Card_def cardinal_def, rule sym) have "Ord(n)" using n by auto moreover { fix i assume "i < n" "i \ n" hence False using n by (auto simp add: lt_nat_in_nat [THEN nat_eqpoll_iff]) } ultimately show "(\ i. i \ n) = n" by (auto intro!: Least_equality) qed lemmas cardinal_0 = nat_0I [THEN nat_into_Card, THEN Card_cardinal_eq, iff] lemmas cardinal_1 = nat_1I [THEN nat_into_Card, THEN Card_cardinal_eq, iff] (*Part of Kunen's Lemma 10.6*) lemma succ_lepoll_natE: "\succ(n) \ n; n \ nat\ \ P" by (rule nat_lepoll_imp_le [THEN lt_irrefl], auto) lemma nat_lepoll_imp_ex_eqpoll_n: "\n \ nat; nat \ X\ \ \Y. Y \ X \ n \ Y" -apply (unfold lepoll_def eqpoll_def) + unfolding lepoll_def eqpoll_def apply (fast del: subsetI subsetCE intro!: subset_SIs dest!: Ord_nat [THEN [2] OrdmemD, THEN [2] restrict_inj] elim!: restrict_bij inj_is_fun [THEN fun_is_rel, THEN image_subset]) done (** \, \ and natural numbers **) lemma lepoll_succ: "i \ succ(i)" by (blast intro: subset_imp_lepoll) lemma lepoll_imp_lesspoll_succ: assumes A: "A \ m" and m: "m \ nat" shows "A \ succ(m)" proof - { assume "A \ succ(m)" hence "succ(m) \ A" by (rule eqpoll_sym) also have "... \ m" by (rule A) finally have "succ(m) \ m" . hence False by (rule succ_lepoll_natE) (rule m) } moreover have "A \ succ(m)" by (blast intro: lepoll_trans A lepoll_succ) ultimately show ?thesis by (auto simp add: lesspoll_def) qed lemma lesspoll_succ_imp_lepoll: "\A \ succ(m); m \ nat\ \ A \ m" -apply (unfold lesspoll_def lepoll_def eqpoll_def bij_def) + unfolding lesspoll_def lepoll_def eqpoll_def bij_def apply (auto dest: inj_not_surj_succ) done lemma lesspoll_succ_iff: "m \ nat \ A \ succ(m) \ A \ m" by (blast intro!: lepoll_imp_lesspoll_succ lesspoll_succ_imp_lepoll) lemma lepoll_succ_disj: "\A \ succ(m); m \ nat\ \ A \ m | A \ succ(m)" apply (rule disjCI) apply (rule lesspoll_succ_imp_lepoll) prefer 2 apply assumption apply (simp (no_asm_simp) add: lesspoll_def) done lemma lesspoll_cardinal_lt: "\A \ i; Ord(i)\ \ |A| < i" apply (unfold lesspoll_def, clarify) apply (frule lepoll_cardinal_le, assumption) apply (blast intro: well_ord_Memrel well_ord_cardinal_eqpoll [THEN eqpoll_sym] dest: lepoll_well_ord elim!: leE) done subsection\The first infinite cardinal: Omega, or nat\ (*This implies Kunen's Lemma 10.6*) lemma lt_not_lepoll: assumes n: "n nat" shows "\ i \ n" proof - { assume i: "i \ n" have "succ(n) \ i" using n by (elim ltE, blast intro: Ord_succ_subsetI [THEN subset_imp_lepoll]) also have "... \ n" by (rule i) finally have "succ(n) \ n" . hence False by (rule succ_lepoll_natE) (rule n) } thus ?thesis by auto qed text\A slightly weaker version of \nat_eqpoll_iff\\ lemma Ord_nat_eqpoll_iff: assumes i: "Ord(i)" and n: "n \ nat" shows "i \ n \ i=n" using i nat_into_Ord [OF n] proof (cases rule: Ord_linear_lt) case lt hence "i \ nat" by (rule lt_nat_in_nat) (rule n) thus ?thesis by (simp add: nat_eqpoll_iff n) next case eq thus ?thesis by (simp add: eqpoll_refl) next case gt hence "\ i \ n" using n by (rule lt_not_lepoll) hence "\ i \ n" using n by (blast intro: eqpoll_imp_lepoll) moreover have "i \ n" using \n by auto ultimately show ?thesis by blast qed lemma Card_nat: "Card(nat)" proof - { fix i assume i: "i < nat" "i \ nat" hence "\ nat \ i" by (simp add: lt_def lt_not_lepoll) hence False using i by (simp add: eqpoll_iff) } hence "(\ i. i \ nat) = nat" by (blast intro: Least_equality eqpoll_refl) thus ?thesis by (auto simp add: Card_def cardinal_def) qed (*Allows showing that |i| is a limit cardinal*) lemma nat_le_cardinal: "nat \ i \ nat \ |i|" apply (rule Card_nat [THEN Card_cardinal_eq, THEN subst]) apply (erule cardinal_mono) done lemma n_lesspoll_nat: "n \ nat \ n \ nat" by (blast intro: Ord_nat Card_nat ltI lt_Card_imp_lesspoll) subsection\Towards Cardinal Arithmetic\ (** Congruence laws for successor, cardinal addition and multiplication **) (*Congruence law for cons under equipollence*) lemma cons_lepoll_cong: "\A \ B; b \ B\ \ cons(a,A) \ cons(b,B)" apply (unfold lepoll_def, safe) apply (rule_tac x = "\y\cons (a,A) . if y=a then b else f`y" in exI) apply (rule_tac d = "\z. if z \ B then converse (f) `z else a" in lam_injective) apply (safe elim!: consE') apply simp_all apply (blast intro: inj_is_fun [THEN apply_type])+ done lemma cons_eqpoll_cong: "\A \ B; a \ A; b \ B\ \ cons(a,A) \ cons(b,B)" by (simp add: eqpoll_iff cons_lepoll_cong) lemma cons_lepoll_cons_iff: "\a \ A; b \ B\ \ cons(a,A) \ cons(b,B) \ A \ B" by (blast intro: cons_lepoll_cong cons_lepoll_consD) lemma cons_eqpoll_cons_iff: "\a \ A; b \ B\ \ cons(a,A) \ cons(b,B) \ A \ B" by (blast intro: cons_eqpoll_cong cons_eqpoll_consD) lemma singleton_eqpoll_1: "{a} \ 1" unfolding succ_def apply (blast intro!: eqpoll_refl [THEN cons_eqpoll_cong]) done lemma cardinal_singleton: "|{a}| = 1" apply (rule singleton_eqpoll_1 [THEN cardinal_cong, THEN trans]) apply (simp (no_asm) add: nat_into_Card [THEN Card_cardinal_eq]) done lemma not_0_is_lepoll_1: "A \ 0 \ 1 \ A" apply (erule not_emptyE) apply (rule_tac a = "cons (x, A-{x}) " in subst) apply (rule_tac [2] a = "cons(0,0)" and P= "\y. y \ cons (x, A-{x})" in subst) prefer 3 apply (blast intro: cons_lepoll_cong subset_imp_lepoll, auto) done (*Congruence law for succ under equipollence*) lemma succ_eqpoll_cong: "A \ B \ succ(A) \ succ(B)" unfolding succ_def apply (simp add: cons_eqpoll_cong mem_not_refl) done (*Congruence law for + under equipollence*) lemma sum_eqpoll_cong: "\A \ C; B \ D\ \ A+B \ C+D" unfolding eqpoll_def apply (blast intro!: sum_bij) done (*Congruence law for * under equipollence*) lemma prod_eqpoll_cong: "\A \ C; B \ D\ \ A*B \ C*D" unfolding eqpoll_def apply (blast intro!: prod_bij) done lemma inj_disjoint_eqpoll: "\f \ inj(A,B); A \ B = 0\ \ A \ (B - range(f)) \ B" unfolding eqpoll_def apply (rule exI) apply (rule_tac c = "\x. if x \ A then f`x else x" and d = "\y. if y \ range (f) then converse (f) `y else y" in lam_bijective) apply (blast intro!: if_type inj_is_fun [THEN apply_type]) apply (simp (no_asm_simp) add: inj_converse_fun [THEN apply_funtype]) apply (safe elim!: UnE') apply (simp_all add: inj_is_fun [THEN apply_rangeI]) apply (blast intro: inj_converse_fun [THEN apply_type])+ done subsection\Lemmas by Krzysztof Grabczewski\ (*New proofs using cons_lepoll_cons. Could generalise from succ to cons.*) text\If \<^term>\A\ has at most \<^term>\n+1\ elements and \<^term>\a \ A\ then \<^term>\A-{a}\ has at most \<^term>\n\.\ lemma Diff_sing_lepoll: "\a \ A; A \ succ(n)\ \ A - {a} \ n" unfolding succ_def apply (rule cons_lepoll_consD) apply (rule_tac [3] mem_not_refl) apply (erule cons_Diff [THEN ssubst], safe) done text\If \<^term>\A\ has at least \<^term>\n+1\ elements then \<^term>\A-{a}\ has at least \<^term>\n\.\ lemma lepoll_Diff_sing: assumes A: "succ(n) \ A" shows "n \ A - {a}" proof - have "cons(n,n) \ A" using A by (unfold succ_def) also have "... \ cons(a, A-{a})" by (blast intro: subset_imp_lepoll) finally have "cons(n,n) \ cons(a, A-{a})" . thus ?thesis by (blast intro: cons_lepoll_consD mem_irrefl) qed lemma Diff_sing_eqpoll: "\a \ A; A \ succ(n)\ \ A - {a} \ n" by (blast intro!: eqpollI elim!: eqpollE intro: Diff_sing_lepoll lepoll_Diff_sing) lemma lepoll_1_is_sing: "\A \ 1; a \ A\ \ A = {a}" apply (frule Diff_sing_lepoll, assumption) apply (drule lepoll_0_is_0) apply (blast elim: equalityE) done lemma Un_lepoll_sum: "A \ B \ A+B" unfolding lepoll_def apply (rule_tac x = "\x\A \ B. if x\A then Inl (x) else Inr (x)" in exI) apply (rule_tac d = "\z. snd (z)" in lam_injective) apply force apply (simp add: Inl_def Inr_def) done lemma well_ord_Un: "\well_ord(X,R); well_ord(Y,S)\ \ \T. well_ord(X \ Y, T)" by (erule well_ord_radd [THEN Un_lepoll_sum [THEN lepoll_well_ord]], assumption) (*Krzysztof Grabczewski*) lemma disj_Un_eqpoll_sum: "A \ B = 0 \ A \ B \ A + B" unfolding eqpoll_def apply (rule_tac x = "\a\A \ B. if a \ A then Inl (a) else Inr (a)" in exI) apply (rule_tac d = "\z. case (\x. x, \x. x, z)" in lam_bijective) apply auto done subsection \Finite and infinite sets\ lemma eqpoll_imp_Finite_iff: "A \ B \ Finite(A) \ Finite(B)" unfolding Finite_def apply (blast intro: eqpoll_trans eqpoll_sym) done lemma Finite_0 [simp]: "Finite(0)" unfolding Finite_def apply (blast intro!: eqpoll_refl nat_0I) done lemma Finite_cons: "Finite(x) \ Finite(cons(y,x))" unfolding Finite_def apply (case_tac "y \ x") apply (simp add: cons_absorb) apply (erule bexE) apply (rule bexI) apply (erule_tac [2] nat_succI) apply (simp (no_asm_simp) add: succ_def cons_eqpoll_cong mem_not_refl) done lemma Finite_succ: "Finite(x) \ Finite(succ(x))" unfolding succ_def apply (erule Finite_cons) done lemma lepoll_nat_imp_Finite: assumes A: "A \ n" and n: "n \ nat" shows "Finite(A)" proof - have "A \ n \ Finite(A)" using n proof (induct n) case 0 hence "A = 0" by (rule lepoll_0_is_0) thus ?case by simp next case (succ n) hence "A \ n \ A \ succ(n)" by (blast dest: lepoll_succ_disj) thus ?case using succ by (auto simp add: Finite_def) qed thus ?thesis using A . qed lemma lesspoll_nat_is_Finite: "A \ nat \ Finite(A)" unfolding Finite_def apply (blast dest: ltD lesspoll_cardinal_lt lesspoll_imp_eqpoll [THEN eqpoll_sym]) done lemma lepoll_Finite: assumes Y: "Y \ X" and X: "Finite(X)" shows "Finite(Y)" proof - obtain n where n: "n \ nat" "X \ n" using X by (auto simp add: Finite_def) have "Y \ X" by (rule Y) also have "... \ n" by (rule n) finally have "Y \ n" . thus ?thesis using n by (simp add: lepoll_nat_imp_Finite) qed lemmas subset_Finite = subset_imp_lepoll [THEN lepoll_Finite] lemma Finite_cons_iff [iff]: "Finite(cons(y,x)) \ Finite(x)" by (blast intro: Finite_cons subset_Finite) lemma Finite_succ_iff [iff]: "Finite(succ(x)) \ Finite(x)" by (simp add: succ_def) lemma Finite_Int: "Finite(A) | Finite(B) \ Finite(A \ B)" by (blast intro: subset_Finite) lemmas Finite_Diff = Diff_subset [THEN subset_Finite] lemma nat_le_infinite_Ord: "\Ord(i); \ Finite(i)\ \ nat \ i" unfolding Finite_def apply (erule Ord_nat [THEN [2] Ord_linear2]) prefer 2 apply assumption apply (blast intro!: eqpoll_refl elim!: ltE) done lemma Finite_imp_well_ord: "Finite(A) \ \r. well_ord(A,r)" -apply (unfold Finite_def eqpoll_def) + unfolding Finite_def eqpoll_def apply (blast intro: well_ord_rvimage bij_is_inj well_ord_Memrel nat_into_Ord) done lemma succ_lepoll_imp_not_empty: "succ(x) \ y \ y \ 0" by (fast dest!: lepoll_0_is_0) lemma eqpoll_succ_imp_not_empty: "x \ succ(n) \ x \ 0" by (fast elim!: eqpoll_sym [THEN eqpoll_0_is_0, THEN succ_neq_0]) lemma Finite_Fin_lemma [rule_format]: "n \ nat \ \A. (A\n \ A \ X) \ A \ Fin(X)" apply (induct_tac n) apply (rule allI) apply (fast intro!: Fin.emptyI dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0]) apply (rule allI) apply (rule impI) apply (erule conjE) apply (rule eqpoll_succ_imp_not_empty [THEN not_emptyE], assumption) apply (frule Diff_sing_eqpoll, assumption) apply (erule allE) apply (erule impE, fast) apply (drule subsetD, assumption) apply (drule Fin.consI, assumption) apply (simp add: cons_Diff) done lemma Finite_Fin: "\Finite(A); A \ X\ \ A \ Fin(X)" by (unfold Finite_def, blast intro: Finite_Fin_lemma) lemma Fin_lemma [rule_format]: "n \ nat \ \A. A \ n \ A \ Fin(A)" apply (induct_tac n) apply (simp add: eqpoll_0_iff, clarify) apply (subgoal_tac "\u. u \ A") apply (erule exE) apply (rule Diff_sing_eqpoll [elim_format]) prefer 2 apply assumption apply assumption apply (rule_tac b = A in cons_Diff [THEN subst], assumption) apply (rule Fin.consI, blast) apply (blast intro: subset_consI [THEN Fin_mono, THEN subsetD]) (*Now for the lemma assumed above*) unfolding eqpoll_def apply (blast intro: bij_converse_bij [THEN bij_is_fun, THEN apply_type]) done lemma Finite_into_Fin: "Finite(A) \ A \ Fin(A)" unfolding Finite_def apply (blast intro: Fin_lemma) done lemma Fin_into_Finite: "A \ Fin(U) \ Finite(A)" by (fast intro!: Finite_0 Finite_cons elim: Fin_induct) lemma Finite_Fin_iff: "Finite(A) \ A \ Fin(A)" by (blast intro: Finite_into_Fin Fin_into_Finite) lemma Finite_Un: "\Finite(A); Finite(B)\ \ Finite(A \ B)" by (blast intro!: Fin_into_Finite Fin_UnI dest!: Finite_into_Fin intro: Un_upper1 [THEN Fin_mono, THEN subsetD] Un_upper2 [THEN Fin_mono, THEN subsetD]) lemma Finite_Un_iff [simp]: "Finite(A \ B) \ (Finite(A) \ Finite(B))" by (blast intro: subset_Finite Finite_Un) text\The converse must hold too.\ lemma Finite_Union: "\\y\X. Finite(y); Finite(X)\ \ Finite(\(X))" apply (simp add: Finite_Fin_iff) apply (rule Fin_UnionI) apply (erule Fin_induct, simp) apply (blast intro: Fin.consI Fin_mono [THEN [2] rev_subsetD]) done (* Induction principle for Finite(A), by Sidi Ehmety *) lemma Finite_induct [case_names 0 cons, induct set: Finite]: "\Finite(A); P(0); \x B. \Finite(B); x \ B; P(B)\ \ P(cons(x, B))\ \ P(A)" apply (erule Finite_into_Fin [THEN Fin_induct]) apply (blast intro: Fin_into_Finite)+ done (*Sidi Ehmety. The contrapositive says \Finite(A) \ \Finite(A-{a}) *) lemma Diff_sing_Finite: "Finite(A - {a}) \ Finite(A)" unfolding Finite_def apply (case_tac "a \ A") apply (subgoal_tac [2] "A-{a}=A", auto) apply (rule_tac x = "succ (n) " in bexI) apply (subgoal_tac "cons (a, A - {a}) = A \ cons (n, n) = succ (n) ") apply (drule_tac a = a and b = n in cons_eqpoll_cong) apply (auto dest: mem_irrefl) done (*Sidi Ehmety. And the contrapositive of this says \\Finite(A); Finite(B)\ \ \Finite(A-B) *) lemma Diff_Finite [rule_format]: "Finite(B) \ Finite(A-B) \ Finite(A)" apply (erule Finite_induct, auto) apply (case_tac "x \ A") apply (subgoal_tac [2] "A-cons (x, B) = A - B") apply (subgoal_tac "A - cons (x, B) = (A - B) - {x}", simp) apply (drule Diff_sing_Finite, auto) done lemma Finite_RepFun: "Finite(A) \ Finite(RepFun(A,f))" by (erule Finite_induct, simp_all) lemma Finite_RepFun_iff_lemma [rule_format]: "\Finite(x); \x y. f(x)=f(y) \ x=y\ \ \A. x = RepFun(A,f) \ Finite(A)" apply (erule Finite_induct) apply clarify apply (case_tac "A=0", simp) apply (blast del: allE, clarify) apply (subgoal_tac "\z\A. x = f(z)") prefer 2 apply (blast del: allE elim: equalityE, clarify) apply (subgoal_tac "B = {f(u) . u \ A - {z}}") apply (blast intro: Diff_sing_Finite) apply (thin_tac "\A. P(A) \ Finite(A)" for P) apply (rule equalityI) apply (blast intro: elim: equalityE) apply (blast intro: elim: equalityCE) done text\I don't know why, but if the premise is expressed using meta-connectives then the simplifier cannot prove it automatically in conditional rewriting.\ lemma Finite_RepFun_iff: "(\x y. f(x)=f(y) \ x=y) \ Finite(RepFun(A,f)) \ Finite(A)" by (blast intro: Finite_RepFun Finite_RepFun_iff_lemma [of _ f]) lemma Finite_Pow: "Finite(A) \ Finite(Pow(A))" apply (erule Finite_induct) apply (simp_all add: Pow_insert Finite_Un Finite_RepFun) done lemma Finite_Pow_imp_Finite: "Finite(Pow(A)) \ Finite(A)" apply (subgoal_tac "Finite({{x} . x \ A})") apply (simp add: Finite_RepFun_iff ) apply (blast intro: subset_Finite) done lemma Finite_Pow_iff [iff]: "Finite(Pow(A)) \ Finite(A)" by (blast intro: Finite_Pow Finite_Pow_imp_Finite) lemma Finite_cardinal_iff: assumes i: "Ord(i)" shows "Finite(|i|) \ Finite(i)" by (auto simp add: Finite_def) (blast intro: eqpoll_trans eqpoll_sym Ord_cardinal_eqpoll [OF i])+ (*Krzysztof Grabczewski's proof that the converse of a finite, well-ordered set is well-ordered. Proofs simplified by lcp. *) lemma nat_wf_on_converse_Memrel: "n \ nat \ wf[n](converse(Memrel(n)))" proof (induct n rule: nat_induct) case 0 thus ?case by (blast intro: wf_onI) next case (succ x) hence wfx: "\Z. Z = 0 \ (\z\Z. \y. z \ y \ z \ x \ y \ x \ z \ x \ y \ Z)" by (simp add: wf_on_def wf_def) \ \not easy to erase the duplicate \<^term>\z \ x\!\ show ?case proof (rule wf_onI) fix Z u assume Z: "u \ Z" "\z\Z. \y\Z. \y, z\ \ converse(Memrel(succ(x)))" show False proof (cases "x \ Z") case True thus False using Z by (blast elim: mem_irrefl mem_asym) next case False thus False using wfx [of Z] Z by blast qed qed qed lemma nat_well_ord_converse_Memrel: "n \ nat \ well_ord(n,converse(Memrel(n)))" apply (frule Ord_nat [THEN Ord_in_Ord, THEN well_ord_Memrel]) apply (simp add: well_ord_def tot_ord_converse nat_wf_on_converse_Memrel) done lemma well_ord_converse: "\well_ord(A,r); well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))\ \ well_ord(A,converse(r))" apply (rule well_ord_Int_iff [THEN iffD1]) apply (frule ordermap_bij [THEN bij_is_inj, THEN well_ord_rvimage], assumption) apply (simp add: rvimage_converse converse_Int converse_prod ordertype_ord_iso [THEN ord_iso_rvimage_eq]) done lemma ordertype_eq_n: assumes r: "well_ord(A,r)" and A: "A \ n" and n: "n \ nat" shows "ordertype(A,r) = n" proof - have "ordertype(A,r) \ A" by (blast intro: bij_imp_eqpoll bij_converse_bij ordermap_bij r) also have "... \ n" by (rule A) finally have "ordertype(A,r) \ n" . thus ?thesis by (simp add: Ord_nat_eqpoll_iff Ord_ordertype n r) qed lemma Finite_well_ord_converse: "\Finite(A); well_ord(A,r)\ \ well_ord(A,converse(r))" unfolding Finite_def apply (rule well_ord_converse, assumption) apply (blast dest: ordertype_eq_n intro!: nat_well_ord_converse_Memrel) done lemma nat_into_Finite: "n \ nat \ Finite(n)" by (auto simp add: Finite_def intro: eqpoll_refl) lemma nat_not_Finite: "\ Finite(nat)" proof - { fix n assume n: "n \ nat" "nat \ n" have "n \ nat" by (rule n) also have "... = n" using n by (simp add: Ord_nat_eqpoll_iff Ord_nat) finally have "n \ n" . hence False by (blast elim: mem_irrefl) } thus ?thesis by (auto simp add: Finite_def) qed end diff --git a/src/ZF/CardinalArith.thy b/src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy +++ b/src/ZF/CardinalArith.thy @@ -1,940 +1,940 @@ (* Title: ZF/CardinalArith.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Cardinal Arithmetic Without the Axiom of Choice\ theory CardinalArith imports Cardinal OrderArith ArithSimp Finite begin definition InfCard :: "i\o" where "InfCard(i) \ Card(i) \ nat \ i" definition cmult :: "[i,i]\i" (infixl \\\ 70) where "i \ j \ |i*j|" definition cadd :: "[i,i]\i" (infixl \\\ 65) where "i \ j \ |i+j|" definition csquare_rel :: "i\i" where "csquare_rel(K) \ rvimage(K*K, lam \x,y\:K*K. y, x, y>, rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))" definition jump_cardinal :: "i\i" where \ \This definition is more complex than Kunen's but it more easily proved to be a cardinal\ "jump_cardinal(K) \ \X\Pow(K). {z. r \ Pow(K*K), well_ord(X,r) \ z = ordertype(X,r)}" definition csucc :: "i\i" where \ \needed because \<^term>\jump_cardinal(K)\ might not be the successor of \<^term>\K\\ "csucc(K) \ \ L. Card(L) \ Kx. x\A \ Card(x)" shows "Card(\(A))" proof (rule CardI) show "Ord(\A)" using A by (simp add: Card_is_Ord) next fix j assume j: "j < \A" hence "\c\A. j < c \ Card(c)" using A by (auto simp add: lt_def intro: Card_is_Ord) then obtain c where c: "c\A" "j < c" "Card(c)" by blast hence jls: "j \ c" by (simp add: lt_Card_imp_lesspoll) { assume eqp: "j \ \A" have "c \ \A" using c by (blast intro: subset_imp_lepoll) also have "... \ j" by (rule eqpoll_sym [OF eqp]) also have "... \ c" by (rule jls) finally have "c \ c" . hence False by auto } thus "\ j \ \A" by blast qed lemma Card_UN: "(\x. x \ A \ Card(K(x))) \ Card(\x\A. K(x))" by blast lemma Card_OUN [simp,intro,TC]: "(\x. x \ A \ Card(K(x))) \ Card(\xCard(K); b \ K\ \ b \ K" unfolding lesspoll_def apply (simp add: Card_iff_initial) apply (fast intro!: le_imp_lepoll ltI leI) done subsection\Cardinal addition\ text\Note: Could omit proving the algebraic laws for cardinal addition and multiplication. On finite cardinals these operations coincide with addition and multiplication of natural numbers; on infinite cardinals they coincide with union (maximum). Either way we get most laws for free.\ subsubsection\Cardinal addition is commutative\ lemma sum_commute_eqpoll: "A+B \ B+A" proof (unfold eqpoll_def, rule exI) show "(\z\A+B. case(Inr,Inl,z)) \ bij(A+B, B+A)" by (auto intro: lam_bijective [where d = "case(Inr,Inl)"]) qed lemma cadd_commute: "i \ j = j \ i" unfolding cadd_def apply (rule sum_commute_eqpoll [THEN cardinal_cong]) done subsubsection\Cardinal addition is associative\ lemma sum_assoc_eqpoll: "(A+B)+C \ A+(B+C)" unfolding eqpoll_def apply (rule exI) apply (rule sum_assoc_bij) done text\Unconditional version requires AC\ lemma well_ord_cadd_assoc: assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)" shows "(i \ j) \ k = i \ (j \ k)" proof (unfold cadd_def, rule cardinal_cong) have "|i + j| + k \ (i + j) + k" by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd i j) also have "... \ i + (j + k)" by (rule sum_assoc_eqpoll) also have "... \ i + |j + k|" by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd j k eqpoll_sym) finally show "|i + j| + k \ i + |j + k|" . qed subsubsection\0 is the identity for addition\ lemma sum_0_eqpoll: "0+A \ A" unfolding eqpoll_def apply (rule exI) apply (rule bij_0_sum) done lemma cadd_0 [simp]: "Card(K) \ 0 \ K = K" unfolding cadd_def apply (simp add: sum_0_eqpoll [THEN cardinal_cong] Card_cardinal_eq) done subsubsection\Addition by another cardinal\ lemma sum_lepoll_self: "A \ A+B" proof (unfold lepoll_def, rule exI) show "(\x\A. Inl (x)) \ inj(A, A + B)" by (simp add: inj_def) qed (*Could probably weaken the premises to well_ord(K,r), or removing using AC*) lemma cadd_le_self: assumes K: "Card(K)" and L: "Ord(L)" shows "K \ (K \ L)" proof (unfold cadd_def) have "K \ |K|" by (rule Card_cardinal_le [OF K]) moreover have "|K| \ |K + L|" using K L by (blast intro: well_ord_lepoll_imp_cardinal_le sum_lepoll_self well_ord_radd well_ord_Memrel Card_is_Ord) ultimately show "K \ |K + L|" by (blast intro: le_trans) qed subsubsection\Monotonicity of addition\ lemma sum_lepoll_mono: "\A \ C; B \ D\ \ A + B \ C + D" unfolding lepoll_def apply (elim exE) apply (rule_tac x = "\z\A+B. case (\w. Inl(f`w), \y. Inr(fa`y), z)" in exI) apply (rule_tac d = "case (\w. Inl(converse(f) `w), \y. Inr(converse(fa) ` y))" in lam_injective) apply (typecheck add: inj_is_fun, auto) done lemma cadd_le_mono: "\K' \ K; L' \ L\ \ (K' \ L') \ (K \ L)" unfolding cadd_def apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule well_ord_lepoll_imp_cardinal_le) apply (blast intro: well_ord_radd well_ord_Memrel) apply (blast intro: sum_lepoll_mono subset_imp_lepoll) done subsubsection\Addition of finite cardinals is "ordinary" addition\ lemma sum_succ_eqpoll: "succ(A)+B \ succ(A+B)" unfolding eqpoll_def apply (rule exI) apply (rule_tac c = "\z. if z=Inl (A) then A+B else z" and d = "\z. if z=A+B then Inl (A) else z" in lam_bijective) apply simp_all apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+ done (*Pulling the succ(...) outside the |...| requires m, n \ nat *) (*Unconditional version requires AC*) lemma cadd_succ_lemma: assumes "Ord(m)" "Ord(n)" shows "succ(m) \ n = |succ(m \ n)|" proof (unfold cadd_def) have [intro]: "m + n \ |m + n|" using assms by (blast intro: eqpoll_sym well_ord_cardinal_eqpoll well_ord_radd well_ord_Memrel) have "|succ(m) + n| = |succ(m + n)|" by (rule sum_succ_eqpoll [THEN cardinal_cong]) also have "... = |succ(|m + n|)|" by (blast intro: succ_eqpoll_cong cardinal_cong) finally show "|succ(m) + n| = |succ(|m + n|)|" . qed lemma nat_cadd_eq_add: assumes m: "m \ nat" and [simp]: "n \ nat" shows"m \ n = m #+ n" using m proof (induct m) case 0 thus ?case by (simp add: nat_into_Card cadd_0) next case (succ m) thus ?case by (simp add: cadd_succ_lemma nat_into_Card Card_cardinal_eq) qed subsection\Cardinal multiplication\ subsubsection\Cardinal multiplication is commutative\ lemma prod_commute_eqpoll: "A*B \ B*A" unfolding eqpoll_def apply (rule exI) apply (rule_tac c = "\\x,y\.\y,x\" and d = "\\x,y\.\y,x\" in lam_bijective, auto) done lemma cmult_commute: "i \ j = j \ i" unfolding cmult_def apply (rule prod_commute_eqpoll [THEN cardinal_cong]) done subsubsection\Cardinal multiplication is associative\ lemma prod_assoc_eqpoll: "(A*B)*C \ A*(B*C)" unfolding eqpoll_def apply (rule exI) apply (rule prod_assoc_bij) done text\Unconditional version requires AC\ lemma well_ord_cmult_assoc: assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)" shows "(i \ j) \ k = i \ (j \ k)" proof (unfold cmult_def, rule cardinal_cong) have "|i * j| * k \ (i * j) * k" by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_rmult i j) also have "... \ i * (j * k)" by (rule prod_assoc_eqpoll) also have "... \ i * |j * k|" by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_rmult j k eqpoll_sym) finally show "|i * j| * k \ i * |j * k|" . qed subsubsection\Cardinal multiplication distributes over addition\ lemma sum_prod_distrib_eqpoll: "(A+B)*C \ (A*C)+(B*C)" unfolding eqpoll_def apply (rule exI) apply (rule sum_prod_distrib_bij) done lemma well_ord_cadd_cmult_distrib: assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)" shows "(i \ j) \ k = (i \ k) \ (j \ k)" proof (unfold cadd_def cmult_def, rule cardinal_cong) have "|i + j| * k \ (i + j) * k" by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_refl well_ord_radd i j) also have "... \ i * k + j * k" by (rule sum_prod_distrib_eqpoll) also have "... \ |i * k| + |j * k|" by (blast intro: sum_eqpoll_cong well_ord_cardinal_eqpoll well_ord_rmult i j k eqpoll_sym) finally show "|i + j| * k \ |i * k| + |j * k|" . qed subsubsection\Multiplication by 0 yields 0\ lemma prod_0_eqpoll: "0*A \ 0" unfolding eqpoll_def apply (rule exI) apply (rule lam_bijective, safe) done lemma cmult_0 [simp]: "0 \ i = 0" by (simp add: cmult_def prod_0_eqpoll [THEN cardinal_cong]) subsubsection\1 is the identity for multiplication\ lemma prod_singleton_eqpoll: "{x}*A \ A" unfolding eqpoll_def apply (rule exI) apply (rule singleton_prod_bij [THEN bij_converse_bij]) done lemma cmult_1 [simp]: "Card(K) \ 1 \ K = K" -apply (unfold cmult_def succ_def) + unfolding cmult_def succ_def apply (simp add: prod_singleton_eqpoll [THEN cardinal_cong] Card_cardinal_eq) done subsection\Some inequalities for multiplication\ lemma prod_square_lepoll: "A \ A*A" -apply (unfold lepoll_def inj_def) + unfolding lepoll_def inj_def apply (rule_tac x = "\x\A. \x,x\" in exI, simp) done (*Could probably weaken the premise to well_ord(K,r), or remove using AC*) lemma cmult_square_le: "Card(K) \ K \ K \ K" unfolding cmult_def apply (rule le_trans) apply (rule_tac [2] well_ord_lepoll_imp_cardinal_le) apply (rule_tac [3] prod_square_lepoll) apply (simp add: le_refl Card_is_Ord Card_cardinal_eq) apply (blast intro: well_ord_rmult well_ord_Memrel Card_is_Ord) done subsubsection\Multiplication by a non-zero cardinal\ lemma prod_lepoll_self: "b \ B \ A \ A*B" -apply (unfold lepoll_def inj_def) + unfolding lepoll_def inj_def apply (rule_tac x = "\x\A. \x,b\" in exI, simp) done (*Could probably weaken the premises to well_ord(K,r), or removing using AC*) lemma cmult_le_self: "\Card(K); Ord(L); 0 \ K \ (K \ L)" unfolding cmult_def apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_cardinal_le]) apply assumption apply (blast intro: well_ord_rmult well_ord_Memrel Card_is_Ord) apply (blast intro: prod_lepoll_self ltD) done subsubsection\Monotonicity of multiplication\ lemma prod_lepoll_mono: "\A \ C; B \ D\ \ A * B \ C * D" unfolding lepoll_def apply (elim exE) apply (rule_tac x = "lam \w,y\:A*B. " in exI) apply (rule_tac d = "\\w,y\. " in lam_injective) apply (typecheck add: inj_is_fun, auto) done lemma cmult_le_mono: "\K' \ K; L' \ L\ \ (K' \ L') \ (K \ L)" unfolding cmult_def apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule well_ord_lepoll_imp_cardinal_le) apply (blast intro: well_ord_rmult well_ord_Memrel) apply (blast intro: prod_lepoll_mono subset_imp_lepoll) done subsection\Multiplication of finite cardinals is "ordinary" multiplication\ lemma prod_succ_eqpoll: "succ(A)*B \ B + A*B" unfolding eqpoll_def apply (rule exI) apply (rule_tac c = "\\x,y\. if x=A then Inl (y) else Inr (\x,y\)" and d = "case (\y. \A,y\, \z. z)" in lam_bijective) apply safe apply (simp_all add: succI2 if_type mem_imp_not_eq) done (*Unconditional version requires AC*) lemma cmult_succ_lemma: "\Ord(m); Ord(n)\ \ succ(m) \ n = n \ (m \ n)" -apply (unfold cmult_def cadd_def) + unfolding cmult_def cadd_def apply (rule prod_succ_eqpoll [THEN cardinal_cong, THEN trans]) apply (rule cardinal_cong [symmetric]) apply (rule sum_eqpoll_cong [OF eqpoll_refl well_ord_cardinal_eqpoll]) apply (blast intro: well_ord_rmult well_ord_Memrel) done lemma nat_cmult_eq_mult: "\m \ nat; n \ nat\ \ m \ n = m#*n" apply (induct_tac m) apply (simp_all add: cmult_succ_lemma nat_cadd_eq_add) done lemma cmult_2: "Card(n) \ 2 \ n = n \ n" by (simp add: cmult_succ_lemma Card_is_Ord cadd_commute [of _ 0]) lemma sum_lepoll_prod: assumes C: "2 \ C" shows "B+B \ C*B" proof - have "B+B \ 2*B" by (simp add: sum_eq_2_times) also have "... \ C*B" by (blast intro: prod_lepoll_mono lepoll_refl C) finally show "B+B \ C*B" . qed lemma lepoll_imp_sum_lepoll_prod: "\A \ B; 2 \ A\ \ A+B \ A*B" by (blast intro: sum_lepoll_mono sum_lepoll_prod lepoll_trans lepoll_refl) subsection\Infinite Cardinals are Limit Ordinals\ (*This proof is modelled upon one assuming nat<=A, with injection \z\cons(u,A). if z=u then 0 else if z \ nat then succ(z) else z and inverse \y. if y \ nat then nat_case(u, \z. z, y) else y. \ If f \ inj(nat,A) then range(f) behaves like the natural numbers.*) lemma nat_cons_lepoll: "nat \ A \ cons(u,A) \ A" unfolding lepoll_def apply (erule exE) apply (rule_tac x = "\z\cons (u,A). if z=u then f`0 else if z \ range (f) then f`succ (converse (f) `z) else z" in exI) apply (rule_tac d = "\y. if y \ range(f) then nat_case (u, \z. f`z, converse(f) `y) else y" in lam_injective) apply (fast intro!: if_type apply_type intro: inj_is_fun inj_converse_fun) apply (simp add: inj_is_fun [THEN apply_rangeI] inj_converse_fun [THEN apply_rangeI] inj_converse_fun [THEN apply_funtype]) done lemma nat_cons_eqpoll: "nat \ A \ cons(u,A) \ A" apply (erule nat_cons_lepoll [THEN eqpollI]) apply (rule subset_consI [THEN subset_imp_lepoll]) done (*Specialized version required below*) lemma nat_succ_eqpoll: "nat \ A \ succ(A) \ A" unfolding succ_def apply (erule subset_imp_lepoll [THEN nat_cons_eqpoll]) done lemma InfCard_nat: "InfCard(nat)" unfolding InfCard_def apply (blast intro: Card_nat le_refl Card_is_Ord) done lemma InfCard_is_Card: "InfCard(K) \ Card(K)" unfolding InfCard_def apply (erule conjunct1) done lemma InfCard_Un: "\InfCard(K); Card(L)\ \ InfCard(K \ L)" unfolding InfCard_def apply (simp add: Card_Un Un_upper1_le [THEN [2] le_trans] Card_is_Ord) done (*Kunen's Lemma 10.11*) lemma InfCard_is_Limit: "InfCard(K) \ Limit(K)" unfolding InfCard_def apply (erule conjE) apply (frule Card_is_Ord) apply (rule ltI [THEN non_succ_LimitI]) apply (erule le_imp_subset [THEN subsetD]) apply (safe dest!: Limit_nat [THEN Limit_le_succD]) unfolding Card_def apply (drule trans) apply (erule le_imp_subset [THEN nat_succ_eqpoll, THEN cardinal_cong]) apply (erule Ord_cardinal_le [THEN lt_trans2, THEN lt_irrefl]) apply (rule le_eqI, assumption) apply (rule Ord_cardinal) done (*** An infinite cardinal equals its square (Kunen, Thm 10.12, page 29) ***) (*A general fact about ordermap*) lemma ordermap_eqpoll_pred: "\well_ord(A,r); x \ A\ \ ordermap(A,r)`x \ Order.pred(A,x,r)" unfolding eqpoll_def apply (rule exI) apply (simp add: ordermap_eq_image well_ord_is_wf) apply (erule ordermap_bij [THEN bij_is_inj, THEN restrict_bij, THEN bij_converse_bij]) apply (rule pred_subset) done subsubsection\Establishing the well-ordering\ lemma well_ord_csquare: assumes K: "Ord(K)" shows "well_ord(K*K, csquare_rel(K))" proof (unfold csquare_rel_def, rule well_ord_rvimage) show "(\\x,y\\K \ K. \x \ y, x, y\) \ inj(K \ K, K \ K \ K)" using K by (force simp add: inj_def intro: lam_type Un_least_lt [THEN ltD] ltI) next show "well_ord(K \ K \ K, rmult(K, Memrel(K), K \ K, rmult(K, Memrel(K), K, Memrel(K))))" using K by (blast intro: well_ord_rmult well_ord_Memrel) qed subsubsection\Characterising initial segments of the well-ordering\ lemma csquareD: "\<\x,y\, \z,z\> \ csquare_rel(K); x \ x \ z \ y \ z" unfolding csquare_rel_def apply (erule rev_mp) apply (elim ltE) apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD) apply (safe elim!: mem_irrefl intro!: Un_upper1_le Un_upper2_le) apply (simp_all add: lt_def succI2) done lemma pred_csquare_subset: "z Order.pred(K*K, \z,z\, csquare_rel(K)) \ succ(z)*succ(z)" -apply (unfold Order.pred_def) + unfolding Order.pred_def apply (safe del: SigmaI dest!: csquareD) apply (unfold lt_def, auto) done lemma csquare_ltI: "\x \ <\x,y\, \z,z\> \ csquare_rel(K)" unfolding csquare_rel_def apply (subgoal_tac "x y apply *) lemma csquare_or_eqI: "\x \ z; y \ z; z \ <\x,y\, \z,z\> \ csquare_rel(K) | x=z \ y=z" unfolding csquare_rel_def apply (subgoal_tac "x yThe cardinality of initial segments\ lemma ordermap_z_lt: "\Limit(K); x y)\ \ ordermap(K*K, csquare_rel(K)) ` \x,y\ < ordermap(K*K, csquare_rel(K)) ` \z,z\" apply (subgoal_tac "z well_ord (K*K, csquare_rel (K))") prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ Limit_is_Ord [THEN well_ord_csquare], clarify) apply (rule csquare_ltI [THEN ordermap_mono, THEN ltI]) apply (erule_tac [4] well_ord_is_wf) apply (blast intro!: Un_upper1_le Un_upper2_le Ord_ordermap elim!: ltE)+ done text\Kunen: "each \<^term>\\x,y\ \ K \ K\ has no more than \<^term>\z \ z\ predecessors..." (page 29)\ lemma ordermap_csquare_le: assumes K: "Limit(K)" and x: "x succ(x \ y)" shows "|ordermap(K \ K, csquare_rel(K)) ` \x,y\| \ |succ(z)| \ |succ(z)|" proof (unfold cmult_def, rule well_ord_lepoll_imp_cardinal_le) show "well_ord(|succ(z)| \ |succ(z)|, rmult(|succ(z)|, Memrel(|succ(z)|), |succ(z)|, Memrel(|succ(z)|)))" by (blast intro: Ord_cardinal well_ord_Memrel well_ord_rmult) next have zK: "z K, csquare_rel(K)) ` \x,y\ \ ordermap(K \ K, csquare_rel(K)) ` \z,z\" using z_def by (blast intro: ordermap_z_lt leI le_imp_lepoll K x y) also have "... \ Order.pred(K \ K, \z,z\, csquare_rel(K))" proof (rule ordermap_eqpoll_pred) show "well_ord(K \ K, csquare_rel(K))" using K by (rule Limit_is_Ord [THEN well_ord_csquare]) next show "\z, z\ \ K \ K" using zK by (blast intro: ltD) qed also have "... \ succ(z) \ succ(z)" using zK by (rule pred_csquare_subset [THEN subset_imp_lepoll]) also have "... \ |succ(z)| \ |succ(z)|" using oz by (blast intro: prod_eqpoll_cong Ord_succ Ord_cardinal_eqpoll eqpoll_sym) finally show "ordermap(K \ K, csquare_rel(K)) ` \x,y\ \ |succ(z)| \ |succ(z)|" . qed text\Kunen: "... so the order type is \\\ K"\ lemma ordertype_csquare_le: assumes IK: "InfCard(K)" and eq: "\y. y\K \ InfCard(y) \ y \ y = y" shows "ordertype(K*K, csquare_rel(K)) \ K" proof - have CK: "Card(K)" using IK by (rule InfCard_is_Card) hence OK: "Ord(K)" by (rule Card_is_Ord) moreover have "Ord(ordertype(K \ K, csquare_rel(K)))" using OK by (rule well_ord_csquare [THEN Ord_ordertype]) ultimately show ?thesis proof (rule all_lt_imp_le) fix i assume i: "i < ordertype(K \ K, csquare_rel(K))" hence Oi: "Ord(i)" by (elim ltE) obtain x y where x: "x \ K" and y: "y \ K" and ieq: "i = ordermap(K \ K, csquare_rel(K)) ` \x,y\" using i by (auto simp add: ordertype_unfold elim: ltE) hence xy: "Ord(x)" "Ord(y)" "x < K" "y < K" using OK by (blast intro: Ord_in_Ord ltI)+ hence ou: "Ord(x \ y)" by (simp add: Ord_Un) show "i < K" proof (rule Card_lt_imp_lt [OF _ Oi CK]) have "|i| \ |succ(succ(x \ y))| \ |succ(succ(x \ y))|" using IK xy by (auto simp add: ieq intro: InfCard_is_Limit [THEN ordermap_csquare_le]) moreover have "|succ(succ(x \ y))| \ |succ(succ(x \ y))| < K" proof (cases rule: Ord_linear2 [OF ou Ord_nat]) assume "x \ y < nat" hence "|succ(succ(x \ y))| \ |succ(succ(x \ y))| \ nat" by (simp add: lt_def nat_cmult_eq_mult nat_succI mult_type nat_into_Card [THEN Card_cardinal_eq] Ord_nat) also have "... \ K" using IK by (simp add: InfCard_def le_imp_subset) finally show "|succ(succ(x \ y))| \ |succ(succ(x \ y))| < K" by (simp add: ltI OK) next assume natxy: "nat \ x \ y" hence seq: "|succ(succ(x \ y))| = |x \ y|" using xy by (simp add: le_imp_subset nat_succ_eqpoll [THEN cardinal_cong] le_succ_iff) also have "... < K" using xy by (simp add: Un_least_lt Ord_cardinal_le [THEN lt_trans1]) finally have "|succ(succ(x \ y))| < K" . moreover have "InfCard(|succ(succ(x \ y))|)" using xy natxy by (simp add: seq InfCard_def Card_cardinal nat_le_cardinal) ultimately show ?thesis by (simp add: eq ltD) qed ultimately show "|i| < K" by (blast intro: lt_trans1) qed qed qed (*Main result: Kunen's Theorem 10.12*) lemma InfCard_csquare_eq: assumes IK: "InfCard(K)" shows "K \ K = K" proof - have OK: "Ord(K)" using IK by (simp add: Card_is_Ord InfCard_is_Card) show "K \ K = K" using OK IK proof (induct rule: trans_induct) case (step i) show "i \ i = i" proof (rule le_anti_sym) have "|i \ i| = |ordertype(i \ i, csquare_rel(i))|" by (rule cardinal_cong, simp add: step.hyps well_ord_csquare [THEN ordermap_bij, THEN bij_imp_eqpoll]) hence "i \ i \ ordertype(i \ i, csquare_rel(i))" by (simp add: step.hyps cmult_def Ord_cardinal_le well_ord_csquare [THEN Ord_ordertype]) moreover have "ordertype(i \ i, csquare_rel(i)) \ i" using step by (simp add: ordertype_csquare_le) ultimately show "i \ i \ i" by (rule le_trans) next show "i \ i \ i" using step by (blast intro: cmult_square_le InfCard_is_Card) qed qed qed (*Corollary for arbitrary well-ordered sets (all sets, assuming AC)*) lemma well_ord_InfCard_square_eq: assumes r: "well_ord(A,r)" and I: "InfCard(|A|)" shows "A \ A \ A" proof - have "A \ A \ |A| \ |A|" by (blast intro: prod_eqpoll_cong well_ord_cardinal_eqpoll eqpoll_sym r) also have "... \ A" proof (rule well_ord_cardinal_eqE [OF _ r]) show "well_ord(|A| \ |A|, rmult(|A|, Memrel(|A|), |A|, Memrel(|A|)))" by (blast intro: Ord_cardinal well_ord_rmult well_ord_Memrel r) next show "||A| \ |A|| = |A|" using InfCard_csquare_eq I by (simp add: cmult_def) qed finally show ?thesis . qed lemma InfCard_square_eqpoll: "InfCard(K) \ K \ K \ K" apply (rule well_ord_InfCard_square_eq) apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel]) apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]) done lemma Inf_Card_is_InfCard: "\Card(i); \ Finite(i)\ \ InfCard(i)" by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord]) subsubsection\Toward's Kunen's Corollary 10.13 (1)\ lemma InfCard_le_cmult_eq: "\InfCard(K); L \ K; 0 \ K \ L = K" apply (rule le_anti_sym) prefer 2 apply (erule ltE, blast intro: cmult_le_self InfCard_is_Card) apply (frule InfCard_is_Card [THEN Card_is_Ord, THEN le_refl]) apply (rule cmult_le_mono [THEN le_trans], assumption+) apply (simp add: InfCard_csquare_eq) done (*Corollary 10.13 (1), for cardinal multiplication*) lemma InfCard_cmult_eq: "\InfCard(K); InfCard(L)\ \ K \ L = K \ L" apply (rule_tac i = K and j = L in Ord_linear_le) apply (typecheck add: InfCard_is_Card Card_is_Ord) apply (rule cmult_commute [THEN ssubst]) apply (rule Un_commute [THEN ssubst]) apply (simp_all add: InfCard_is_Limit [THEN Limit_has_0] InfCard_le_cmult_eq subset_Un_iff2 [THEN iffD1] le_imp_subset) done lemma InfCard_cdouble_eq: "InfCard(K) \ K \ K = K" apply (simp add: cmult_2 [symmetric] InfCard_is_Card cmult_commute) apply (simp add: InfCard_le_cmult_eq InfCard_is_Limit Limit_has_0 Limit_has_succ) done (*Corollary 10.13 (1), for cardinal addition*) lemma InfCard_le_cadd_eq: "\InfCard(K); L \ K\ \ K \ L = K" apply (rule le_anti_sym) prefer 2 apply (erule ltE, blast intro: cadd_le_self InfCard_is_Card) apply (frule InfCard_is_Card [THEN Card_is_Ord, THEN le_refl]) apply (rule cadd_le_mono [THEN le_trans], assumption+) apply (simp add: InfCard_cdouble_eq) done lemma InfCard_cadd_eq: "\InfCard(K); InfCard(L)\ \ K \ L = K \ L" apply (rule_tac i = K and j = L in Ord_linear_le) apply (typecheck add: InfCard_is_Card Card_is_Ord) apply (rule cadd_commute [THEN ssubst]) apply (rule Un_commute [THEN ssubst]) apply (simp_all add: InfCard_le_cadd_eq subset_Un_iff2 [THEN iffD1] le_imp_subset) done (*The other part, Corollary 10.13 (2), refers to the cardinality of the set of all n-tuples of elements of K. A better version for the Isabelle theory might be InfCard(K) \ |list(K)| = K. *) subsection\For Every Cardinal Number There Exists A Greater One\ text\This result is Kunen's Theorem 10.16, which would be trivial using AC\ lemma Ord_jump_cardinal: "Ord(jump_cardinal(K))" unfolding jump_cardinal_def apply (rule Ord_is_Transset [THEN [2] OrdI]) prefer 2 apply (blast intro!: Ord_ordertype) unfolding Transset_def apply (safe del: subsetI) apply (simp add: ordertype_pred_unfold, safe) apply (rule UN_I) apply (rule_tac [2] ReplaceI) prefer 4 apply (blast intro: well_ord_subset elim!: predE)+ done (*Allows selective unfolding. Less work than deriving intro/elim rules*) lemma jump_cardinal_iff: "i \ jump_cardinal(K) \ (\r X. r \ K*K \ X \ K \ well_ord(X,r) \ i = ordertype(X,r))" unfolding jump_cardinal_def apply (blast del: subsetI) done (*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*) lemma K_lt_jump_cardinal: "Ord(K) \ K < jump_cardinal(K)" apply (rule Ord_jump_cardinal [THEN [2] ltI]) apply (rule jump_cardinal_iff [THEN iffD2]) apply (rule_tac x="Memrel(K)" in exI) apply (rule_tac x=K in exI) apply (simp add: ordertype_Memrel well_ord_Memrel) apply (simp add: Memrel_def subset_iff) done (*The proof by contradiction: the bijection f yields a wellordering of X whose ordertype is jump_cardinal(K). *) lemma Card_jump_cardinal_lemma: "\well_ord(X,r); r \ K * K; X \ K; f \ bij(ordertype(X,r), jump_cardinal(K))\ \ jump_cardinal(K) \ jump_cardinal(K)" apply (subgoal_tac "f O ordermap (X,r) \ bij (X, jump_cardinal (K))") prefer 2 apply (blast intro: comp_bij ordermap_bij) apply (rule jump_cardinal_iff [THEN iffD2]) apply (intro exI conjI) apply (rule subset_trans [OF rvimage_type Sigma_mono], assumption+) apply (erule bij_is_inj [THEN well_ord_rvimage]) apply (rule Ord_jump_cardinal [THEN well_ord_Memrel]) apply (simp add: well_ord_Memrel [THEN [2] bij_ordertype_vimage] ordertype_Memrel Ord_jump_cardinal) done (*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*) lemma Card_jump_cardinal: "Card(jump_cardinal(K))" apply (rule Ord_jump_cardinal [THEN CardI]) unfolding eqpoll_def apply (safe dest!: ltD jump_cardinal_iff [THEN iffD1]) apply (blast intro: Card_jump_cardinal_lemma [THEN mem_irrefl]) done subsection\Basic Properties of Successor Cardinals\ lemma csucc_basic: "Ord(K) \ Card(csucc(K)) \ K < csucc(K)" unfolding csucc_def apply (rule LeastI) apply (blast intro: Card_jump_cardinal K_lt_jump_cardinal Ord_jump_cardinal)+ done lemmas Card_csucc = csucc_basic [THEN conjunct1] lemmas lt_csucc = csucc_basic [THEN conjunct2] lemma Ord_0_lt_csucc: "Ord(K) \ 0 < csucc(K)" by (blast intro: Ord_0_le lt_csucc lt_trans1) lemma csucc_le: "\Card(L); K \ csucc(K) \ L" unfolding csucc_def apply (rule Least_le) apply (blast intro: Card_is_Ord)+ done lemma lt_csucc_iff: "\Ord(i); Card(K)\ \ i < csucc(K) \ |i| \ K" apply (rule iffI) apply (rule_tac [2] Card_lt_imp_lt) apply (erule_tac [2] lt_trans1) apply (simp_all add: lt_csucc Card_csucc Card_is_Ord) apply (rule notI [THEN not_lt_imp_le]) apply (rule Card_cardinal [THEN csucc_le, THEN lt_trans1, THEN lt_irrefl], assumption) apply (rule Ord_cardinal_le [THEN lt_trans1]) apply (simp_all add: Ord_cardinal Card_is_Ord) done lemma Card_lt_csucc_iff: "\Card(K'); Card(K)\ \ K' < csucc(K) \ K' \ K" by (simp add: lt_csucc_iff Card_cardinal_eq Card_is_Ord) lemma InfCard_csucc: "InfCard(K) \ InfCard(csucc(K))" by (simp add: InfCard_def Card_csucc Card_is_Ord lt_csucc [THEN leI, THEN [2] le_trans]) subsubsection\Removing elements from a finite set decreases its cardinality\ lemma Finite_imp_cardinal_cons [simp]: assumes FA: "Finite(A)" and a: "a\A" shows "|cons(a,A)| = succ(|A|)" proof - { fix X have "Finite(X) \ a \ X \ cons(a,X) \ X \ False" proof (induct X rule: Finite_induct) case 0 thus False by (simp add: lepoll_0_iff) next case (cons x Y) hence "cons(x, cons(a, Y)) \ cons(x, Y)" by (simp add: cons_commute) hence "cons(a, Y) \ Y" using cons by (blast dest: cons_lepoll_consD) thus False using cons by auto qed } hence [simp]: "\ cons(a,A) \ A" using a FA by auto have [simp]: "|A| \ A" using Finite_imp_well_ord [OF FA] by (blast intro: well_ord_cardinal_eqpoll) have "(\ i. i \ cons(a, A)) = succ(|A|)" proof (rule Least_equality [OF _ _ notI]) show "succ(|A|) \ cons(a, A)" by (simp add: succ_def cons_eqpoll_cong mem_not_refl a) next show "Ord(succ(|A|))" by simp next fix i assume i: "i \ |A|" "i \ cons(a, A)" have "cons(a, A) \ i" by (rule eqpoll_sym) (rule i) also have "... \ |A|" by (rule le_imp_lepoll) (rule i) also have "... \ A" by simp finally have "cons(a, A) \ A" . thus False by simp qed thus ?thesis by (simp add: cardinal_def) qed lemma Finite_imp_succ_cardinal_Diff: "\Finite(A); a \ A\ \ succ(|A-{a}|) = |A|" apply (rule_tac b = A in cons_Diff [THEN subst], assumption) apply (simp add: Finite_imp_cardinal_cons Diff_subset [THEN subset_Finite]) apply (simp add: cons_Diff) done lemma Finite_imp_cardinal_Diff: "\Finite(A); a \ A\ \ |A-{a}| < |A|" apply (rule succ_leE) apply (simp add: Finite_imp_succ_cardinal_Diff) done lemma Finite_cardinal_in_nat [simp]: "Finite(A) \ |A| \ nat" proof (induct rule: Finite_induct) case 0 thus ?case by (simp add: cardinal_0) next case (cons x A) thus ?case by (simp add: Finite_imp_cardinal_cons) qed lemma card_Un_Int: "\Finite(A); Finite(B)\ \ |A| #+ |B| = |A \ B| #+ |A \ B|" apply (erule Finite_induct, simp) apply (simp add: Finite_Int cons_absorb Un_cons Int_cons_left) done lemma card_Un_disjoint: "\Finite(A); Finite(B); A \ B = 0\ \ |A \ B| = |A| #+ |B|" by (simp add: Finite_Un card_Un_Int) lemma card_partition: assumes FC: "Finite(C)" shows "Finite (\ C) \ (\c\C. |c| = k) \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = 0) \ k #* |C| = |\ C|" using FC proof (induct rule: Finite_induct) case 0 thus ?case by simp next case (cons x B) hence "x \ \B = 0" by auto thus ?case using cons by (auto simp add: card_Un_disjoint) qed subsubsection\Theorems by Krzysztof Grabczewski, proofs by lcp\ lemmas nat_implies_well_ord = nat_into_Ord [THEN well_ord_Memrel] lemma nat_sum_eqpoll_sum: assumes m: "m \ nat" and n: "n \ nat" shows "m + n \ m #+ n" proof - have "m + n \ |m+n|" using m n by (blast intro: nat_implies_well_ord well_ord_radd well_ord_cardinal_eqpoll eqpoll_sym) also have "... = m #+ n" using m n by (simp add: nat_cadd_eq_add [symmetric] cadd_def) finally show ?thesis . qed lemma Ord_subset_natD [rule_format]: "Ord(i) \ i \ nat \ i \ nat | i=nat" proof (induct i rule: trans_induct3) case 0 thus ?case by auto next case (succ i) thus ?case by auto next case (limit l) thus ?case by (blast dest: nat_le_Limit le_imp_subset) qed lemma Ord_nat_subset_into_Card: "\Ord(i); i \ nat\ \ Card(i)" by (blast dest: Ord_subset_natD intro: Card_nat nat_into_Card) end diff --git a/src/ZF/Coind/ECR.thy b/src/ZF/Coind/ECR.thy --- a/src/ZF/Coind/ECR.thy +++ b/src/ZF/Coind/ECR.thy @@ -1,167 +1,167 @@ (* Title: ZF/Coind/ECR.thy Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) theory ECR imports Static Dynamic begin (* The extended correspondence relation *) consts HasTyRel :: i coinductive domains "HasTyRel" \ "Val * Ty" intros htr_constI [intro!]: "\c \ Const; t \ Ty; isof(c,t)\ \ \ HasTyRel" htr_closI [intro]: "\x \ ExVar; e \ Exp; t \ Ty; ve \ ValEnv; te \ TyEnv; \ ElabRel; ve_dom(ve) = te_dom(te); {.y \ ve_dom(ve)} \ Pow(HasTyRel)\ \ \ HasTyRel" monos Pow_mono type_intros Val_ValEnv.intros (* Pointwise extension to environments *) definition hastyenv :: "[i,i] \ o" where "hastyenv(ve,te) \ ve_dom(ve) = te_dom(te) \ (\x \ ve_dom(ve). \ HasTyRel)" (* Specialised co-induction rule *) lemma htr_closCI [intro]: "\x \ ExVar; e \ Exp; t \ Ty; ve \ ValEnv; te \ TyEnv; \ ElabRel; ve_dom(ve) = te_dom(te); {.y \ ve_dom(ve)} \ Pow({} \ HasTyRel)\ \ \ HasTyRel" apply (rule singletonI [THEN HasTyRel.coinduct], auto) done (* Specialised elimination rules *) inductive_cases htr_constE [elim!]: " \ HasTyRel" and htr_closE [elim]: " \ HasTyRel" (* Properties of the pointwise extension to environments *) lemmas HasTyRel_non_zero = HasTyRel.dom_subset [THEN subsetD, THEN SigmaD1, THEN ValNEE] lemma hastyenv_owr: "\ve \ ValEnv; te \ TyEnv; hastyenv(ve,te); \v,t\ \ HasTyRel\ \ hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))" by (auto simp add: hastyenv_def ve_app_owr HasTyRel_non_zero) lemma basic_consistency_lem: "\ve \ ValEnv; te \ TyEnv; isofenv(ve,te)\ \ hastyenv(ve,te)" -apply (unfold isofenv_def hastyenv_def) + unfolding isofenv_def hastyenv_def apply (force intro: te_appI ve_domI) done (* ############################################################ *) (* The Consistency theorem *) (* ############################################################ *) lemma consistency_const: "\c \ Const; hastyenv(ve,te); \ ElabRel\ \ \ HasTyRel" by blast lemma consistency_var: "\x \ ve_dom(ve); hastyenv(ve,te); \ ElabRel\ \ \ HasTyRel" by (unfold hastyenv_def, blast) lemma consistency_fn: "\ve \ ValEnv; x \ ExVar; e \ Exp; hastyenv(ve,te); \ ElabRel \ \ \ HasTyRel" by (unfold hastyenv_def, blast) declare ElabRel.dom_subset [THEN subsetD, dest] declare Ty.intros [simp, intro!] declare TyEnv.intros [simp, intro!] declare Val_ValEnv.intros [simp, intro!] lemma consistency_fix: "\ve \ ValEnv; x \ ExVar; e \ Exp; f \ ExVar; cl \ Val; v_clos(x,e,ve_owr(ve,f,cl)) = cl; hastyenv(ve,te); \ ElabRel\ \ \cl,t\ \ HasTyRel" unfolding hastyenv_def apply (erule elab_fixE, safe) apply hypsubst_thin apply (rule subst, assumption) apply (rule_tac te="te_owr(te, f, t_fun(t1, t2))" in htr_closCI) apply simp_all apply (blast intro: ve_owrI) apply (rule ElabRel.fnI) apply (simp_all add: ValNEE, force) done lemma consistency_app1: "\ve \ ValEnv; e1 \ Exp; e2 \ Exp; c1 \ Const; c2 \ Const; \ EvalRel; \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; \ EvalRel; \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; hastyenv(ve, te); \ ElabRel\ \ \ HasTyRel" by (blast intro!: c_appI intro: isof_app) lemma consistency_app2: "\ve \ ValEnv; vem \ ValEnv; e1 \ Exp; e2 \ Exp; em \ Exp; xm \ ExVar; v \ Val; \ EvalRel; \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; \ EvalRel; \t te. hastyenv(ve,te) \ \ ElabRel \ \v2,t\ \ HasTyRel; \ EvalRel; \t te. hastyenv(ve_owr(vem,xm,v2),te) \ \ ElabRel \ \v,t\ \ HasTyRel; hastyenv(ve,te); \ ElabRel\ \ \v,t\ \ HasTyRel" apply (erule elab_appE) apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+) apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+) apply (erule htr_closE) apply (erule elab_fnE, simp) apply clarify apply (drule spec [THEN spec, THEN mp, THEN mp]) prefer 2 apply assumption+ apply (rule hastyenv_owr, assumption) apply assumption apply (simp add: hastyenv_def, blast+) done lemma consistency [rule_format]: " \ EvalRel \ (\t te. hastyenv(ve,te) \ \ ElabRel \ \v,t\ \ HasTyRel)" apply (erule EvalRel.induct) apply (simp_all add: consistency_const consistency_var consistency_fn consistency_fix consistency_app1) apply (blast intro: consistency_app2) done lemma basic_consistency: "\ve \ ValEnv; te \ TyEnv; isofenv(ve,te); \ EvalRel; \ ElabRel\ \ isof(c,t)" by (blast dest: consistency intro!: basic_consistency_lem) end diff --git a/src/ZF/Coind/Map.thy b/src/ZF/Coind/Map.thy --- a/src/ZF/Coind/Map.thy +++ b/src/ZF/Coind/Map.thy @@ -1,184 +1,184 @@ (* Title: ZF/Coind/Map.thy Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge Some sample proofs of inclusions for the final coalgebra "U" (by lcp). *) theory Map imports ZF begin definition TMap :: "[i,i] \ i" where "TMap(A,B) \ {m \ Pow(A*\(B)).\a \ A. m``{a} \ B}" definition PMap :: "[i,i] \ i" where "PMap(A,B) \ TMap(A,cons(0,B))" (* Note: 0 \ B \ TMap(A,B) = PMap(A,B) *) definition map_emp :: i where "map_emp \ 0" definition map_owr :: "[i,i,i]\i" where "map_owr(m,a,b) \ \x \ {a} \ domain(m). if x=a then b else m``{x}" definition map_app :: "[i,i]\i" where "map_app(m,a) \ m``{a}" lemma "{0,1} \ {1} \ TMap(I, {0,1})" by (unfold TMap_def, blast) lemma "{0} \ TMap(I,1) \ {1} \ TMap(I, {0} \ TMap(I,1))" by (unfold TMap_def, blast) lemma "{0,1} \ TMap(I,2) \ {1} \ TMap(I, {0,1} \ TMap(I,2))" by (unfold TMap_def, blast) (*A bit too slow. lemma "{0,1} \ TMap(I,TMap(I,2)) \ TMap(I,2) \ {1} \ TMap(I, {0,1} \ TMap(I,TMap(I,2)) \ TMap(I,2))" by (unfold TMap_def, blast) *) (* ############################################################ *) (* Lemmas *) (* ############################################################ *) lemma qbeta_if: "Sigma(A,B)``{a} = (if a \ A then B(a) else 0)" by auto lemma qbeta: "a \ A \ Sigma(A,B)``{a} = B(a)" by fast lemma qbeta_emp: "a\A \ Sigma(A,B)``{a} = 0" by fast lemma image_Sigma1: "a \ A \ Sigma(A,B)``{a}=0" by fast (* ############################################################ *) (* Inclusion in Quine Universes *) (* ############################################################ *) (* Lemmas *) lemma MapQU_lemma: "A \ univ(X) \ Pow(A * \(quniv(X))) \ quniv(X)" unfolding quniv_def apply (rule Pow_mono) apply (rule subset_trans [OF Sigma_mono product_univ]) apply (erule subset_trans) apply (rule arg_subset_eclose [THEN univ_mono]) apply (simp add: Union_Pow_eq) done (* Theorems *) lemma mapQU: "\m \ PMap(A,quniv(B)); \x. x \ A \ x \ univ(B)\ \ m \ quniv(B)" -apply (unfold PMap_def TMap_def) + unfolding PMap_def TMap_def apply (blast intro!: MapQU_lemma [THEN subsetD]) done (* ############################################################ *) (* Monotonicity *) (* ############################################################ *) lemma PMap_mono: "B \ C \ PMap(A,B)<=PMap(A,C)" by (unfold PMap_def TMap_def, blast) (* ############################################################ *) (* Introduction Rules *) (* ############################################################ *) (** map_emp **) lemma pmap_empI: "map_emp \ PMap(A,B)" by (unfold map_emp_def PMap_def TMap_def, auto) (** map_owr **) lemma pmap_owrI: "\m \ PMap(A,B); a \ A; b \ B\ \ map_owr(m,a,b):PMap(A,B)" apply (unfold map_owr_def PMap_def TMap_def, safe) apply (simp_all add: if_iff, auto) (*Remaining subgoal*) apply (rule excluded_middle [THEN disjE]) apply (erule image_Sigma1) apply (drule_tac psi = "uu \ B" for uu in asm_rl) apply (auto simp add: qbeta) done (** map_app **) lemma tmap_app_notempty: "\m \ TMap(A,B); a \ domain(m)\ \ map_app(m,a) \0" by (unfold TMap_def map_app_def, blast) lemma tmap_appI: "\m \ TMap(A,B); a \ domain(m)\ \ map_app(m,a):B" by (unfold TMap_def map_app_def domain_def, blast) lemma pmap_appI: "\m \ PMap(A,B); a \ domain(m)\ \ map_app(m,a):B" unfolding PMap_def apply (frule tmap_app_notempty, assumption) apply (drule tmap_appI, auto) done (** domain **) lemma tmap_domainD: "\m \ TMap(A,B); a \ domain(m)\ \ a \ A" by (unfold TMap_def, blast) lemma pmap_domainD: "\m \ PMap(A,B); a \ domain(m)\ \ a \ A" by (unfold PMap_def TMap_def, blast) (* ############################################################ *) (* Equalities *) (* ############################################################ *) (** Domain **) (* Lemmas *) lemma domain_UN: "domain(\x \ A. B(x)) = (\x \ A. domain(B(x)))" by fast lemma domain_Sigma: "domain(Sigma(A,B)) = {x \ A. \y. y \ B(x)}" by blast (* Theorems *) lemma map_domain_emp: "domain(map_emp) = 0" by (unfold map_emp_def, blast) lemma map_domain_owr: "b \ 0 \ domain(map_owr(f,a,b)) = {a} \ domain(f)" unfolding map_owr_def apply (auto simp add: domain_Sigma) done (** Application **) lemma map_app_owr: "map_app(map_owr(f,a,b),c) = (if c=a then b else map_app(f,c))" by (simp add: qbeta_if map_app_def map_owr_def, blast) lemma map_app_owr1: "map_app(map_owr(f,a,b),a) = b" by (simp add: map_app_owr) lemma map_app_owr2: "c \ a \ map_app(map_owr(f,a,b),c)= map_app(f,c)" by (simp add: map_app_owr) end diff --git a/src/ZF/Coind/Values.thy b/src/ZF/Coind/Values.thy --- a/src/ZF/Coind/Values.thy +++ b/src/ZF/Coind/Values.thy @@ -1,115 +1,115 @@ (* Title: ZF/Coind/Values.thy Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) theory Values imports Language Map begin (* Values, values environments and associated operators *) consts Val :: i ValEnv :: i Val_ValEnv :: i codatatype "Val" = v_const ("c \ Const") | v_clos ("x \ ExVar","e \ Exp","ve \ ValEnv") and "ValEnv" = ve_mk ("m \ PMap(ExVar,Val)") monos PMap_mono type_intros A_into_univ mapQU consts ve_owr :: "[i,i,i] \ i" ve_dom :: "i\i" ve_app :: "[i,i] \ i" primrec "ve_owr(ve_mk(m), x, v) = ve_mk(map_owr(m,x,v))" primrec "ve_dom(ve_mk(m)) = domain(m)" primrec "ve_app(ve_mk(m), a) = map_app(m,a)" definition ve_emp :: i where "ve_emp \ ve_mk(map_emp)" (* Elimination rules *) lemma ValEnvE: "\ve \ ValEnv; \m.\ve=ve_mk(m); m \ PMap(ExVar,Val)\ \ Q\ \ Q" apply (unfold Part_def Val_def ValEnv_def, clarify) apply (erule Val_ValEnv.cases) apply (auto simp add: Val_def Part_def Val_ValEnv.con_defs) done lemma ValE: "\v \ Val; \c. \v = v_const(c); c \ Const\ \ Q; \e ve x. \v = v_clos(x,e,ve); x \ ExVar; e \ Exp; ve \ ValEnv\ \ Q \ \ Q" apply (unfold Part_def Val_def ValEnv_def, clarify) apply (erule Val_ValEnv.cases) apply (auto simp add: ValEnv_def Part_def Val_ValEnv.con_defs) done (* Nonempty sets *) lemma v_closNE [simp]: "v_clos(x,e,ve) \ 0" by (unfold QPair_def QInl_def QInr_def Val_ValEnv.con_defs, blast) declare v_closNE [THEN notE, elim!] lemma v_constNE [simp]: "c \ Const \ v_const(c) \ 0" -apply (unfold QPair_def QInl_def QInr_def Val_ValEnv.con_defs) + unfolding QPair_def QInl_def QInr_def Val_ValEnv.con_defs apply (drule constNEE, auto) done (* Proving that the empty set is not a value *) lemma ValNEE: "v \ Val \ v \ 0" by (erule ValE, auto) (* Equalities for value environments *) lemma ve_dom_owr [simp]: "\ve \ ValEnv; v \0\ \ ve_dom(ve_owr(ve,x,v)) = ve_dom(ve) \ {x}" apply (erule ValEnvE) apply (auto simp add: map_domain_owr) done lemma ve_app_owr [simp]: "ve \ ValEnv \ ve_app(ve_owr(ve,y,v),x) = (if x=y then v else ve_app(ve,x))" by (erule ValEnvE, simp add: map_app_owr) (* Introduction rules for operators on value environments *) lemma ve_appI: "\ve \ ValEnv; x \ ve_dom(ve)\ \ ve_app(ve,x):Val" by (erule ValEnvE, simp add: pmap_appI) lemma ve_domI: "\ve \ ValEnv; x \ ve_dom(ve)\ \ x \ ExVar" apply (erule ValEnvE, simp) apply (blast dest: pmap_domainD) done lemma ve_empI: "ve_emp \ ValEnv" unfolding ve_emp_def apply (rule Val_ValEnv.intros) apply (rule pmap_empI) done lemma ve_owrI: "\ve \ ValEnv; x \ ExVar; v \ Val\ \ ve_owr(ve,x,v):ValEnv" apply (erule ValEnvE, simp) apply (blast intro: pmap_owrI Val_ValEnv.intros) done end diff --git a/src/ZF/Constructible/L_axioms.thy b/src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy +++ b/src/ZF/Constructible/L_axioms.thy @@ -1,1403 +1,1403 @@ (* Title: ZF/Constructible/L_axioms.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) section \The ZF Axioms (Except Separation) in L\ theory L_axioms imports Formula Relative Reflection MetaExists begin text \The class L satisfies the premises of locale \M_trivial\\ lemma transL: "\y\x; L(x)\ \ L(y)" apply (insert Transset_Lset) apply (simp add: Transset_def L_def, blast) done lemma nonempty: "L(0)" apply (simp add: L_def) apply (blast intro: zero_in_Lset) done theorem upair_ax: "upair_ax(L)" apply (simp add: upair_ax_def upair_def, clarify) apply (rule_tac x="{x,y}" in rexI) apply (simp_all add: doubleton_in_L) done theorem Union_ax: "Union_ax(L)" apply (simp add: Union_ax_def big_union_def, clarify) apply (rule_tac x="\(x)" in rexI) apply (simp_all add: Union_in_L, auto) apply (blast intro: transL) done theorem power_ax: "power_ax(L)" apply (simp add: power_ax_def powerset_def Relative.subset_def, clarify) apply (rule_tac x="{y \ Pow(x). L(y)}" in rexI) apply (simp_all add: LPow_in_L, auto) apply (blast intro: transL) done text\We don't actually need \<^term>\L\ to satisfy the foundation axiom.\ theorem foundation_ax: "foundation_ax(L)" apply (simp add: foundation_ax_def) apply (rule rallI) apply (cut_tac A=x in foundation) apply (blast intro: transL) done subsection\For L to satisfy Replacement\ (*Can't move these to Formula unless the definition of univalent is moved there too!*) lemma LReplace_in_Lset: "\X \ Lset(i); univalent(L,X,Q); Ord(i)\ \ \j. Ord(j) \ Replace(X, \x y. Q(x,y) \ L(y)) \ Lset(j)" apply (rule_tac x="\y \ Replace(X, \x y. Q(x,y) \ L(y)). succ(lrank(y))" in exI) apply simp apply clarify apply (rule_tac a=x in UN_I) apply (simp_all add: Replace_iff univalent_def) apply (blast dest: transL L_I) done lemma LReplace_in_L: "\L(X); univalent(L,X,Q)\ \ \Y. L(Y) \ Replace(X, \x y. Q(x,y) \ L(y)) \ Y" apply (drule L_D, clarify) apply (drule LReplace_in_Lset, assumption+) apply (blast intro: L_I Lset_in_Lset_succ) done theorem replacement: "replacement(L,P)" apply (simp add: replacement_def, clarify) apply (frule LReplace_in_L, assumption+, clarify) apply (rule_tac x=Y in rexI) apply (simp_all add: Replace_iff univalent_def, blast) done lemma strong_replacementI [rule_format]: "\\B[L]. separation(L, \u. \x[L]. x\B \ P(x,u))\ \ strong_replacement(L,P)" apply (simp add: strong_replacement_def, clarify) apply (frule replacementD [OF replacement], assumption, clarify) apply (drule_tac x=A in rspec, clarify) apply (drule_tac z=Y in separationD, assumption, clarify) apply (rule_tac x=y in rexI, force, assumption) done subsection\Instantiating the locale \M_trivial\\ text\No instances of Separation yet.\ lemma Lset_mono_le: "mono_le_subset(Lset)" by (simp add: mono_le_subset_def le_imp_subset Lset_mono) lemma Lset_cont: "cont_Ord(Lset)" by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) lemmas L_nat = Ord_in_L [OF Ord_nat] theorem M_trivial_L: "M_trivial(L)" apply (rule M_trivial.intro) apply (rule M_trans.intro) apply (erule (1) transL) apply(rule exI,rule nonempty) apply (rule M_trivial_axioms.intro) apply (rule upair_ax) apply (rule Union_ax) done interpretation L: M_trivial L by (rule M_trivial_L) (* Replaces the following declarations... lemmas rall_abs = M_trivial.rall_abs [OF M_trivial_L] and rex_abs = M_trivial.rex_abs [OF M_trivial_L] ... declare rall_abs [simp] declare rex_abs [simp] ...and dozens of similar ones. *) subsection\Instantiation of the locale \reflection\\ text\instances of locale constants\ definition L_F0 :: "[i\o,i] \ i" where "L_F0(P,y) \ \ b. (\z. L(z) \ P(\y,z\)) \ (\z\Lset(b). P(\y,z\))" definition L_FF :: "[i\o,i] \ i" where "L_FF(P) \ \a. \y\Lset(a). L_F0(P,y)" definition L_ClEx :: "[i\o,i] \ o" where "L_ClEx(P) \ \a. Limit(a) \ normalize(L_FF(P),a) = a" text\We must use the meta-existential quantifier; otherwise the reflection terms become enormous!\ definition L_Reflects :: "[i\o,[i,i]\o] \ prop" (\(3REFLECTS/ [_,/ _])\) where "REFLECTS[P,Q] \ (\Cl. Closed_Unbounded(Cl) \ (\a. Cl(a) \ (\x \ Lset(a). P(x) \ Q(a,x))))" theorem Triv_reflection: "REFLECTS[P, \a x. P(x)]" apply (simp add: L_Reflects_def) apply (rule meta_exI) apply (rule Closed_Unbounded_Ord) done theorem Not_reflection: "REFLECTS[P,Q] \ REFLECTS[\x. \P(x), \a x. \Q(a,x)]" unfolding L_Reflects_def apply (erule meta_exE) apply (rule_tac x=Cl in meta_exI, simp) done theorem And_reflection: "\REFLECTS[P,Q]; REFLECTS[P',Q']\ \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" unfolding L_Reflects_def apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done theorem Or_reflection: "\REFLECTS[P,Q]; REFLECTS[P',Q']\ \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" unfolding L_Reflects_def apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done theorem Imp_reflection: "\REFLECTS[P,Q]; REFLECTS[P',Q']\ \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" unfolding L_Reflects_def apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done theorem Iff_reflection: "\REFLECTS[P,Q]; REFLECTS[P',Q']\ \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" unfolding L_Reflects_def apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done lemma reflection_Lset: "reflection(Lset)" by (blast intro: reflection.intro Lset_mono_le Lset_cont Formula.Pair_in_LLimit)+ theorem Ex_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] \ REFLECTS[\x. \z. L(z) \ P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" -apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) + unfolding L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def apply (elim meta_exE) apply (rule meta_exI) apply (erule reflection.Ex_reflection [OF reflection_Lset]) done theorem All_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] \ REFLECTS[\x. \z. L(z) \ P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" -apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) + unfolding L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def apply (elim meta_exE) apply (rule meta_exI) apply (erule reflection.All_reflection [OF reflection_Lset]) done theorem Rex_reflection: "REFLECTS[ \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" unfolding rex_def apply (intro And_reflection Ex_reflection, assumption) done theorem Rall_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" unfolding rall_def apply (intro Imp_reflection All_reflection, assumption) done text\This version handles an alternative form of the bounded quantifier in the second argument of \REFLECTS\.\ theorem Rex_reflection': "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z[##Lset(a)]. Q(a,x,z)]" -apply (unfold setclass_def rex_def) + unfolding setclass_def rex_def apply (erule Rex_reflection [unfolded rex_def Bex_def]) done text\As above.\ theorem Rall_reflection': "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z[##Lset(a)]. Q(a,x,z)]" -apply (unfold setclass_def rall_def) + unfolding setclass_def rall_def apply (erule Rall_reflection [unfolded rall_def Ball_def]) done lemmas FOL_reflections = Triv_reflection Not_reflection And_reflection Or_reflection Imp_reflection Iff_reflection Ex_reflection All_reflection Rex_reflection Rall_reflection Rex_reflection' Rall_reflection' lemma ReflectsD: "\REFLECTS[P,Q]; Ord(i)\ \ \j. i (\x \ Lset(j). P(x) \ Q(j,x))" -apply (unfold L_Reflects_def Closed_Unbounded_def) + unfolding L_Reflects_def Closed_Unbounded_def apply (elim meta_exE, clarify) apply (blast dest!: UnboundedD) done lemma ReflectsE: "\REFLECTS[P,Q]; Ord(i); \j. \ix \ Lset(j). P(x) \ Q(j,x)\ \ R\ \ R" by (drule ReflectsD, assumption, blast) lemma Collect_mem_eq: "{x\A. x\B} = A \ B" by blast subsection\Internalized Formulas for some Set-Theoretic Concepts\ subsubsection\Some numbers to help write de Bruijn indices\ abbreviation digit3 :: i (\3\) where "3 \ succ(2)" abbreviation digit4 :: i (\4\) where "4 \ succ(3)" abbreviation digit5 :: i (\5\) where "5 \ succ(4)" abbreviation digit6 :: i (\6\) where "6 \ succ(5)" abbreviation digit7 :: i (\7\) where "7 \ succ(6)" abbreviation digit8 :: i (\8\) where "8 \ succ(7)" abbreviation digit9 :: i (\9\) where "9 \ succ(8)" subsubsection\The Empty Set, Internalized\ definition empty_fm :: "i\i" where "empty_fm(x) \ Forall(Neg(Member(0,succ(x))))" lemma empty_type [TC]: "x \ nat \ empty_fm(x) \ formula" by (simp add: empty_fm_def) lemma sats_empty_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, empty_fm(x), env) \ empty(##A, nth(x,env))" by (simp add: empty_fm_def empty_def) lemma empty_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ empty(##A, x) \ sats(A, empty_fm(i), env)" by simp theorem empty_reflection: "REFLECTS[\x. empty(L,f(x)), \i x. empty(##Lset(i),f(x))]" apply (simp only: empty_def) apply (intro FOL_reflections) done text\Not used. But maybe useful?\ lemma Transset_sats_empty_fm_eq_0: "\n \ nat; env \ list(A); Transset(A)\ \ sats(A, empty_fm(n), env) \ nth(n,env) = 0" apply (simp add: empty_fm_def empty_def Transset_def, auto) apply (case_tac "n < length(env)") apply (frule nth_type, assumption+, blast) apply (simp_all add: not_lt_iff_le nth_eq_0) done subsubsection\Unordered Pairs, Internalized\ definition upair_fm :: "[i,i,i]\i" where "upair_fm(x,y,z) \ And(Member(x,z), And(Member(y,z), Forall(Implies(Member(0,succ(z)), Or(Equal(0,succ(x)), Equal(0,succ(y)))))))" lemma upair_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ upair_fm(x,y,z) \ formula" by (simp add: upair_fm_def) lemma sats_upair_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, upair_fm(x,y,z), env) \ upair(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: upair_fm_def upair_def) lemma upair_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ upair(##A, x, y, z) \ sats(A, upair_fm(i,j,k), env)" by (simp) text\Useful? At least it refers to "real" unordered pairs\ lemma sats_upair_fm2 [simp]: "\x \ nat; y \ nat; z < length(env); env \ list(A); Transset(A)\ \ sats(A, upair_fm(x,y,z), env) \ nth(z,env) = {nth(x,env), nth(y,env)}" apply (frule lt_length_in_nat, assumption) apply (simp add: upair_fm_def Transset_def, auto) apply (blast intro: nth_type) done theorem upair_reflection: "REFLECTS[\x. upair(L,f(x),g(x),h(x)), \i x. upair(##Lset(i),f(x),g(x),h(x))]" apply (simp add: upair_def) apply (intro FOL_reflections) done subsubsection\Ordered pairs, Internalized\ definition pair_fm :: "[i,i,i]\i" where "pair_fm(x,y,z) \ Exists(And(upair_fm(succ(x),succ(x),0), Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0), upair_fm(1,0,succ(succ(z)))))))" lemma pair_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ pair_fm(x,y,z) \ formula" by (simp add: pair_fm_def) lemma sats_pair_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, pair_fm(x,y,z), env) \ pair(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: pair_fm_def pair_def) lemma pair_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ pair(##A, x, y, z) \ sats(A, pair_fm(i,j,k), env)" by (simp) theorem pair_reflection: "REFLECTS[\x. pair(L,f(x),g(x),h(x)), \i x. pair(##Lset(i),f(x),g(x),h(x))]" apply (simp only: pair_def) apply (intro FOL_reflections upair_reflection) done subsubsection\Binary Unions, Internalized\ definition union_fm :: "[i,i,i]\i" where "union_fm(x,y,z) \ Forall(Iff(Member(0,succ(z)), Or(Member(0,succ(x)),Member(0,succ(y)))))" lemma union_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ union_fm(x,y,z) \ formula" by (simp add: union_fm_def) lemma sats_union_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, union_fm(x,y,z), env) \ union(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: union_fm_def union_def) lemma union_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ union(##A, x, y, z) \ sats(A, union_fm(i,j,k), env)" by (simp) theorem union_reflection: "REFLECTS[\x. union(L,f(x),g(x),h(x)), \i x. union(##Lset(i),f(x),g(x),h(x))]" apply (simp only: union_def) apply (intro FOL_reflections) done subsubsection\Set ``Cons,'' Internalized\ definition cons_fm :: "[i,i,i]\i" where "cons_fm(x,y,z) \ Exists(And(upair_fm(succ(x),succ(x),0), union_fm(0,succ(y),succ(z))))" lemma cons_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ cons_fm(x,y,z) \ formula" by (simp add: cons_fm_def) lemma sats_cons_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, cons_fm(x,y,z), env) \ is_cons(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: cons_fm_def is_cons_def) lemma cons_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ is_cons(##A, x, y, z) \ sats(A, cons_fm(i,j,k), env)" by simp theorem cons_reflection: "REFLECTS[\x. is_cons(L,f(x),g(x),h(x)), \i x. is_cons(##Lset(i),f(x),g(x),h(x))]" apply (simp only: is_cons_def) apply (intro FOL_reflections upair_reflection union_reflection) done subsubsection\Successor Function, Internalized\ definition succ_fm :: "[i,i]\i" where "succ_fm(x,y) \ cons_fm(x,x,y)" lemma succ_type [TC]: "\x \ nat; y \ nat\ \ succ_fm(x,y) \ formula" by (simp add: succ_fm_def) lemma sats_succ_fm [simp]: "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, succ_fm(x,y), env) \ successor(##A, nth(x,env), nth(y,env))" by (simp add: succ_fm_def successor_def) lemma successor_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)\ \ successor(##A, x, y) \ sats(A, succ_fm(i,j), env)" by simp theorem successor_reflection: "REFLECTS[\x. successor(L,f(x),g(x)), \i x. successor(##Lset(i),f(x),g(x))]" apply (simp only: successor_def) apply (intro cons_reflection) done subsubsection\The Number 1, Internalized\ (* "number1(M,a) \ (\x[M]. empty(M,x) \ successor(M,x,a))" *) definition number1_fm :: "i\i" where "number1_fm(a) \ Exists(And(empty_fm(0), succ_fm(0,succ(a))))" lemma number1_type [TC]: "x \ nat \ number1_fm(x) \ formula" by (simp add: number1_fm_def) lemma sats_number1_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, number1_fm(x), env) \ number1(##A, nth(x,env))" by (simp add: number1_fm_def number1_def) lemma number1_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ number1(##A, x) \ sats(A, number1_fm(i), env)" by simp theorem number1_reflection: "REFLECTS[\x. number1(L,f(x)), \i x. number1(##Lset(i),f(x))]" apply (simp only: number1_def) apply (intro FOL_reflections empty_reflection successor_reflection) done subsubsection\Big Union, Internalized\ (* "big_union(M,A,z) \ \x[M]. x \ z \ (\y[M]. y\A \ x \ y)" *) definition big_union_fm :: "[i,i]\i" where "big_union_fm(A,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), Member(1,0)))))" lemma big_union_type [TC]: "\x \ nat; y \ nat\ \ big_union_fm(x,y) \ formula" by (simp add: big_union_fm_def) lemma sats_big_union_fm [simp]: "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, big_union_fm(x,y), env) \ big_union(##A, nth(x,env), nth(y,env))" by (simp add: big_union_fm_def big_union_def) lemma big_union_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)\ \ big_union(##A, x, y) \ sats(A, big_union_fm(i,j), env)" by simp theorem big_union_reflection: "REFLECTS[\x. big_union(L,f(x),g(x)), \i x. big_union(##Lset(i),f(x),g(x))]" apply (simp only: big_union_def) apply (intro FOL_reflections) done subsubsection\Variants of Satisfaction Definitions for Ordinals, etc.\ text\The \sats\ theorems below are standard versions of the ones proved in theory \Formula\. They relate elements of type \<^term>\formula\ to relativized concepts such as \<^term>\subset\ or \<^term>\ordinal\ rather than to real concepts such as \<^term>\Ord\. Now that we have instantiated the locale \M_trivial\, we no longer require the earlier versions.\ lemma sats_subset_fm': "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, subset_fm(x,y), env) \ subset(##A, nth(x,env), nth(y,env))" by (simp add: subset_fm_def Relative.subset_def) theorem subset_reflection: "REFLECTS[\x. subset(L,f(x),g(x)), \i x. subset(##Lset(i),f(x),g(x))]" apply (simp only: Relative.subset_def) apply (intro FOL_reflections) done lemma sats_transset_fm': "\x \ nat; env \ list(A)\ \ sats(A, transset_fm(x), env) \ transitive_set(##A, nth(x,env))" by (simp add: sats_subset_fm' transset_fm_def transitive_set_def) theorem transitive_set_reflection: "REFLECTS[\x. transitive_set(L,f(x)), \i x. transitive_set(##Lset(i),f(x))]" apply (simp only: transitive_set_def) apply (intro FOL_reflections subset_reflection) done lemma sats_ordinal_fm': "\x \ nat; env \ list(A)\ \ sats(A, ordinal_fm(x), env) \ ordinal(##A,nth(x,env))" by (simp add: sats_transset_fm' ordinal_fm_def ordinal_def) lemma ordinal_iff_sats: "\nth(i,env) = x; i \ nat; env \ list(A)\ \ ordinal(##A, x) \ sats(A, ordinal_fm(i), env)" by (simp add: sats_ordinal_fm') theorem ordinal_reflection: "REFLECTS[\x. ordinal(L,f(x)), \i x. ordinal(##Lset(i),f(x))]" apply (simp only: ordinal_def) apply (intro FOL_reflections transitive_set_reflection) done subsubsection\Membership Relation, Internalized\ definition Memrel_fm :: "[i,i]\i" where "Memrel_fm(A,r) \ Forall(Iff(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), Exists(And(Member(0,succ(succ(succ(A)))), And(Member(1,0), pair_fm(1,0,2))))))))" lemma Memrel_type [TC]: "\x \ nat; y \ nat\ \ Memrel_fm(x,y) \ formula" by (simp add: Memrel_fm_def) lemma sats_Memrel_fm [simp]: "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, Memrel_fm(x,y), env) \ membership(##A, nth(x,env), nth(y,env))" by (simp add: Memrel_fm_def membership_def) lemma Memrel_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)\ \ membership(##A, x, y) \ sats(A, Memrel_fm(i,j), env)" by simp theorem membership_reflection: "REFLECTS[\x. membership(L,f(x),g(x)), \i x. membership(##Lset(i),f(x),g(x))]" apply (simp only: membership_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Predecessor Set, Internalized\ definition pred_set_fm :: "[i,i,i,i]\i" where "pred_set_fm(A,x,r,B) \ Forall(Iff(Member(0,succ(B)), Exists(And(Member(0,succ(succ(r))), And(Member(1,succ(succ(A))), pair_fm(1,succ(succ(x)),0))))))" lemma pred_set_type [TC]: "\A \ nat; x \ nat; r \ nat; B \ nat\ \ pred_set_fm(A,x,r,B) \ formula" by (simp add: pred_set_fm_def) lemma sats_pred_set_fm [simp]: "\U \ nat; x \ nat; r \ nat; B \ nat; env \ list(A)\ \ sats(A, pred_set_fm(U,x,r,B), env) \ pred_set(##A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))" by (simp add: pred_set_fm_def pred_set_def) lemma pred_set_iff_sats: "\nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B; i \ nat; j \ nat; k \ nat; l \ nat; env \ list(A)\ \ pred_set(##A,U,x,r,B) \ sats(A, pred_set_fm(i,j,k,l), env)" by (simp) theorem pred_set_reflection: "REFLECTS[\x. pred_set(L,f(x),g(x),h(x),b(x)), \i x. pred_set(##Lset(i),f(x),g(x),h(x),b(x))]" apply (simp only: pred_set_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Domain of a Relation, Internalized\ (* "is_domain(M,r,z) \ \x[M]. (x \ z \ (\w[M]. w\r \ (\y[M]. pair(M,x,y,w))))" *) definition domain_fm :: "[i,i]\i" where "domain_fm(r,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(pair_fm(2,0,1))))))" lemma domain_type [TC]: "\x \ nat; y \ nat\ \ domain_fm(x,y) \ formula" by (simp add: domain_fm_def) lemma sats_domain_fm [simp]: "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, domain_fm(x,y), env) \ is_domain(##A, nth(x,env), nth(y,env))" by (simp add: domain_fm_def is_domain_def) lemma domain_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)\ \ is_domain(##A, x, y) \ sats(A, domain_fm(i,j), env)" by simp theorem domain_reflection: "REFLECTS[\x. is_domain(L,f(x),g(x)), \i x. is_domain(##Lset(i),f(x),g(x))]" apply (simp only: is_domain_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Range of a Relation, Internalized\ (* "is_range(M,r,z) \ \y[M]. (y \ z \ (\w[M]. w\r \ (\x[M]. pair(M,x,y,w))))" *) definition range_fm :: "[i,i]\i" where "range_fm(r,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(pair_fm(0,2,1))))))" lemma range_type [TC]: "\x \ nat; y \ nat\ \ range_fm(x,y) \ formula" by (simp add: range_fm_def) lemma sats_range_fm [simp]: "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, range_fm(x,y), env) \ is_range(##A, nth(x,env), nth(y,env))" by (simp add: range_fm_def is_range_def) lemma range_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)\ \ is_range(##A, x, y) \ sats(A, range_fm(i,j), env)" by simp theorem range_reflection: "REFLECTS[\x. is_range(L,f(x),g(x)), \i x. is_range(##Lset(i),f(x),g(x))]" apply (simp only: is_range_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Field of a Relation, Internalized\ (* "is_field(M,r,z) \ \dr[M]. is_domain(M,r,dr) \ (\rr[M]. is_range(M,r,rr) \ union(M,dr,rr,z))" *) definition field_fm :: "[i,i]\i" where "field_fm(r,z) \ Exists(And(domain_fm(succ(r),0), Exists(And(range_fm(succ(succ(r)),0), union_fm(1,0,succ(succ(z)))))))" lemma field_type [TC]: "\x \ nat; y \ nat\ \ field_fm(x,y) \ formula" by (simp add: field_fm_def) lemma sats_field_fm [simp]: "\x \ nat; y \ nat; env \ list(A)\ \ sats(A, field_fm(x,y), env) \ is_field(##A, nth(x,env), nth(y,env))" by (simp add: field_fm_def is_field_def) lemma field_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)\ \ is_field(##A, x, y) \ sats(A, field_fm(i,j), env)" by simp theorem field_reflection: "REFLECTS[\x. is_field(L,f(x),g(x)), \i x. is_field(##Lset(i),f(x),g(x))]" apply (simp only: is_field_def) apply (intro FOL_reflections domain_reflection range_reflection union_reflection) done subsubsection\Image under a Relation, Internalized\ (* "image(M,r,A,z) \ \y[M]. (y \ z \ (\w[M]. w\r \ (\x[M]. x\A \ pair(M,x,y,w))))" *) definition image_fm :: "[i,i,i]\i" where "image_fm(r,A,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(And(Member(0,succ(succ(succ(A)))), pair_fm(0,2,1)))))))" lemma image_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ image_fm(x,y,z) \ formula" by (simp add: image_fm_def) lemma sats_image_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, image_fm(x,y,z), env) \ image(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: image_fm_def Relative.image_def) lemma image_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ image(##A, x, y, z) \ sats(A, image_fm(i,j,k), env)" by (simp) theorem image_reflection: "REFLECTS[\x. image(L,f(x),g(x),h(x)), \i x. image(##Lset(i),f(x),g(x),h(x))]" apply (simp only: Relative.image_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Pre-Image under a Relation, Internalized\ (* "pre_image(M,r,A,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. y\A \ pair(M,x,y,w)))" *) definition pre_image_fm :: "[i,i,i]\i" where "pre_image_fm(r,A,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(And(Member(0,succ(succ(succ(A)))), pair_fm(2,0,1)))))))" lemma pre_image_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ pre_image_fm(x,y,z) \ formula" by (simp add: pre_image_fm_def) lemma sats_pre_image_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, pre_image_fm(x,y,z), env) \ pre_image(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: pre_image_fm_def Relative.pre_image_def) lemma pre_image_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ pre_image(##A, x, y, z) \ sats(A, pre_image_fm(i,j,k), env)" by (simp) theorem pre_image_reflection: "REFLECTS[\x. pre_image(L,f(x),g(x),h(x)), \i x. pre_image(##Lset(i),f(x),g(x),h(x))]" apply (simp only: Relative.pre_image_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Function Application, Internalized\ (* "fun_apply(M,f,x,y) \ (\xs[M]. \fxs[M]. upair(M,x,x,xs) \ image(M,f,xs,fxs) \ big_union(M,fxs,y))" *) definition fun_apply_fm :: "[i,i,i]\i" where "fun_apply_fm(f,x,y) \ Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1), And(image_fm(succ(succ(f)), 1, 0), big_union_fm(0,succ(succ(y)))))))" lemma fun_apply_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ fun_apply_fm(x,y,z) \ formula" by (simp add: fun_apply_fm_def) lemma sats_fun_apply_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, fun_apply_fm(x,y,z), env) \ fun_apply(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: fun_apply_fm_def fun_apply_def) lemma fun_apply_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ fun_apply(##A, x, y, z) \ sats(A, fun_apply_fm(i,j,k), env)" by simp theorem fun_apply_reflection: "REFLECTS[\x. fun_apply(L,f(x),g(x),h(x)), \i x. fun_apply(##Lset(i),f(x),g(x),h(x))]" apply (simp only: fun_apply_def) apply (intro FOL_reflections upair_reflection image_reflection big_union_reflection) done subsubsection\The Concept of Relation, Internalized\ (* "is_relation(M,r) \ (\z[M]. z\r \ (\x[M]. \y[M]. pair(M,x,y,z)))" *) definition relation_fm :: "i\i" where "relation_fm(r) \ Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))" lemma relation_type [TC]: "\x \ nat\ \ relation_fm(x) \ formula" by (simp add: relation_fm_def) lemma sats_relation_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, relation_fm(x), env) \ is_relation(##A, nth(x,env))" by (simp add: relation_fm_def is_relation_def) lemma relation_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ is_relation(##A, x) \ sats(A, relation_fm(i), env)" by simp theorem is_relation_reflection: "REFLECTS[\x. is_relation(L,f(x)), \i x. is_relation(##Lset(i),f(x))]" apply (simp only: is_relation_def) apply (intro FOL_reflections pair_reflection) done subsubsection\The Concept of Function, Internalized\ (* "is_function(M,r) \ \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x,y',p') \ p\r \ p'\r \ y=y'" *) definition function_fm :: "i\i" where "function_fm(r) \ Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,3,1), Implies(pair_fm(4,2,0), Implies(Member(1,r#+5), Implies(Member(0,r#+5), Equal(3,2))))))))))" lemma function_type [TC]: "\x \ nat\ \ function_fm(x) \ formula" by (simp add: function_fm_def) lemma sats_function_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, function_fm(x), env) \ is_function(##A, nth(x,env))" by (simp add: function_fm_def is_function_def) lemma is_function_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ is_function(##A, x) \ sats(A, function_fm(i), env)" by simp theorem is_function_reflection: "REFLECTS[\x. is_function(L,f(x)), \i x. is_function(##Lset(i),f(x))]" apply (simp only: is_function_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Typed Functions, Internalized\ (* "typed_function(M,A,B,r) \ is_function(M,r) \ is_relation(M,r) \ is_domain(M,r,A) \ (\u[M]. u\r \ (\x[M]. \y[M]. pair(M,x,y,u) \ y\B))" *) definition typed_function_fm :: "[i,i,i]\i" where "typed_function_fm(A,B,r) \ And(function_fm(r), And(relation_fm(r), And(domain_fm(r,A), Forall(Implies(Member(0,succ(r)), Forall(Forall(Implies(pair_fm(1,0,2),Member(0,B#+3)))))))))" lemma typed_function_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ typed_function_fm(x,y,z) \ formula" by (simp add: typed_function_fm_def) lemma sats_typed_function_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, typed_function_fm(x,y,z), env) \ typed_function(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: typed_function_fm_def typed_function_def) lemma typed_function_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ typed_function(##A, x, y, z) \ sats(A, typed_function_fm(i,j,k), env)" by simp lemmas function_reflections = empty_reflection number1_reflection upair_reflection pair_reflection union_reflection big_union_reflection cons_reflection successor_reflection fun_apply_reflection subset_reflection transitive_set_reflection membership_reflection pred_set_reflection domain_reflection range_reflection field_reflection image_reflection pre_image_reflection is_relation_reflection is_function_reflection lemmas function_iff_sats = empty_iff_sats number1_iff_sats upair_iff_sats pair_iff_sats union_iff_sats big_union_iff_sats cons_iff_sats successor_iff_sats fun_apply_iff_sats Memrel_iff_sats pred_set_iff_sats domain_iff_sats range_iff_sats field_iff_sats image_iff_sats pre_image_iff_sats relation_iff_sats is_function_iff_sats theorem typed_function_reflection: "REFLECTS[\x. typed_function(L,f(x),g(x),h(x)), \i x. typed_function(##Lset(i),f(x),g(x),h(x))]" apply (simp only: typed_function_def) apply (intro FOL_reflections function_reflections) done subsubsection\Composition of Relations, Internalized\ (* "composition(M,r,s,t) \ \p[M]. p \ t \ (\x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. pair(M,x,z,p) \ pair(M,x,y,xy) \ pair(M,y,z,yz) \ xy \ s \ yz \ r)" *) definition composition_fm :: "[i,i,i]\i" where "composition_fm(r,s,t) \ Forall(Iff(Member(0,succ(t)), Exists(Exists(Exists(Exists(Exists( And(pair_fm(4,2,5), And(pair_fm(4,3,1), And(pair_fm(3,2,0), And(Member(1,s#+6), Member(0,r#+6))))))))))))" lemma composition_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ composition_fm(x,y,z) \ formula" by (simp add: composition_fm_def) lemma sats_composition_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, composition_fm(x,y,z), env) \ composition(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: composition_fm_def composition_def) lemma composition_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ composition(##A, x, y, z) \ sats(A, composition_fm(i,j,k), env)" by simp theorem composition_reflection: "REFLECTS[\x. composition(L,f(x),g(x),h(x)), \i x. composition(##Lset(i),f(x),g(x),h(x))]" apply (simp only: composition_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Injections, Internalized\ (* "injection(M,A,B,f) \ typed_function(M,A,B,f) \ (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x',y,p') \ p\f \ p'\f \ x=x')" *) definition injection_fm :: "[i,i,i]\i" where "injection_fm(A,B,f) \ And(typed_function_fm(A,B,f), Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,2,1), Implies(pair_fm(3,2,0), Implies(Member(1,f#+5), Implies(Member(0,f#+5), Equal(4,3)))))))))))" lemma injection_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ injection_fm(x,y,z) \ formula" by (simp add: injection_fm_def) lemma sats_injection_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, injection_fm(x,y,z), env) \ injection(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: injection_fm_def injection_def) lemma injection_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ injection(##A, x, y, z) \ sats(A, injection_fm(i,j,k), env)" by simp theorem injection_reflection: "REFLECTS[\x. injection(L,f(x),g(x),h(x)), \i x. injection(##Lset(i),f(x),g(x),h(x))]" apply (simp only: injection_def) apply (intro FOL_reflections function_reflections typed_function_reflection) done subsubsection\Surjections, Internalized\ (* surjection :: "[i\o,i,i,i] \ o" "surjection(M,A,B,f) \ typed_function(M,A,B,f) \ (\y[M]. y\B \ (\x[M]. x\A \ fun_apply(M,f,x,y)))" *) definition surjection_fm :: "[i,i,i]\i" where "surjection_fm(A,B,f) \ And(typed_function_fm(A,B,f), Forall(Implies(Member(0,succ(B)), Exists(And(Member(0,succ(succ(A))), fun_apply_fm(succ(succ(f)),0,1))))))" lemma surjection_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ surjection_fm(x,y,z) \ formula" by (simp add: surjection_fm_def) lemma sats_surjection_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, surjection_fm(x,y,z), env) \ surjection(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: surjection_fm_def surjection_def) lemma surjection_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ surjection(##A, x, y, z) \ sats(A, surjection_fm(i,j,k), env)" by simp theorem surjection_reflection: "REFLECTS[\x. surjection(L,f(x),g(x),h(x)), \i x. surjection(##Lset(i),f(x),g(x),h(x))]" apply (simp only: surjection_def) apply (intro FOL_reflections function_reflections typed_function_reflection) done subsubsection\Bijections, Internalized\ (* bijection :: "[i\o,i,i,i] \ o" "bijection(M,A,B,f) \ injection(M,A,B,f) \ surjection(M,A,B,f)" *) definition bijection_fm :: "[i,i,i]\i" where "bijection_fm(A,B,f) \ And(injection_fm(A,B,f), surjection_fm(A,B,f))" lemma bijection_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ bijection_fm(x,y,z) \ formula" by (simp add: bijection_fm_def) lemma sats_bijection_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, bijection_fm(x,y,z), env) \ bijection(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: bijection_fm_def bijection_def) lemma bijection_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ bijection(##A, x, y, z) \ sats(A, bijection_fm(i,j,k), env)" by simp theorem bijection_reflection: "REFLECTS[\x. bijection(L,f(x),g(x),h(x)), \i x. bijection(##Lset(i),f(x),g(x),h(x))]" apply (simp only: bijection_def) apply (intro And_reflection injection_reflection surjection_reflection) done subsubsection\Restriction of a Relation, Internalized\ (* "restriction(M,r,A,z) \ \x[M]. x \ z \ (x \ r \ (\u[M]. u\A \ (\v[M]. pair(M,u,v,x))))" *) definition restriction_fm :: "[i,i,i]\i" where "restriction_fm(r,A,z) \ Forall(Iff(Member(0,succ(z)), And(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), Exists(pair_fm(1,0,2)))))))" lemma restriction_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ restriction_fm(x,y,z) \ formula" by (simp add: restriction_fm_def) lemma sats_restriction_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, restriction_fm(x,y,z), env) \ restriction(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: restriction_fm_def restriction_def) lemma restriction_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ restriction(##A, x, y, z) \ sats(A, restriction_fm(i,j,k), env)" by simp theorem restriction_reflection: "REFLECTS[\x. restriction(L,f(x),g(x),h(x)), \i x. restriction(##Lset(i),f(x),g(x),h(x))]" apply (simp only: restriction_def) apply (intro FOL_reflections pair_reflection) done subsubsection\Order-Isomorphisms, Internalized\ (* order_isomorphism :: "[i\o,i,i,i,i,i] \ o" "order_isomorphism(M,A,r,B,s,f) \ bijection(M,A,B,f) \ (\x[M]. x\A \ (\y[M]. y\A \ (\p[M]. \fx[M]. \fy[M]. \q[M]. pair(M,x,y,p) \ fun_apply(M,f,x,fx) \ fun_apply(M,f,y,fy) \ pair(M,fx,fy,q) \ (p\r \ q\s))))" *) definition order_isomorphism_fm :: "[i,i,i,i,i]\i" where "order_isomorphism_fm(A,r,B,s,f) \ And(bijection_fm(A,B,f), Forall(Implies(Member(0,succ(A)), Forall(Implies(Member(0,succ(succ(A))), Forall(Forall(Forall(Forall( Implies(pair_fm(5,4,3), Implies(fun_apply_fm(f#+6,5,2), Implies(fun_apply_fm(f#+6,4,1), Implies(pair_fm(2,1,0), Iff(Member(3,r#+6), Member(0,s#+6)))))))))))))))" lemma order_isomorphism_type [TC]: "\A \ nat; r \ nat; B \ nat; s \ nat; f \ nat\ \ order_isomorphism_fm(A,r,B,s,f) \ formula" by (simp add: order_isomorphism_fm_def) lemma sats_order_isomorphism_fm [simp]: "\U \ nat; r \ nat; B \ nat; s \ nat; f \ nat; env \ list(A)\ \ sats(A, order_isomorphism_fm(U,r,B,s,f), env) \ order_isomorphism(##A, nth(U,env), nth(r,env), nth(B,env), nth(s,env), nth(f,env))" by (simp add: order_isomorphism_fm_def order_isomorphism_def) lemma order_isomorphism_iff_sats: "\nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s; nth(k',env) = f; i \ nat; j \ nat; k \ nat; j' \ nat; k' \ nat; env \ list(A)\ \ order_isomorphism(##A,U,r,B,s,f) \ sats(A, order_isomorphism_fm(i,j,k,j',k'), env)" by simp theorem order_isomorphism_reflection: "REFLECTS[\x. order_isomorphism(L,f(x),g(x),h(x),g'(x),h'(x)), \i x. order_isomorphism(##Lset(i),f(x),g(x),h(x),g'(x),h'(x))]" apply (simp only: order_isomorphism_def) apply (intro FOL_reflections function_reflections bijection_reflection) done subsubsection\Limit Ordinals, Internalized\ text\A limit ordinal is a non-empty, successor-closed ordinal\ (* "limit_ordinal(M,a) \ ordinal(M,a) \ \ empty(M,a) \ (\x[M]. x\a \ (\y[M]. y\a \ successor(M,x,y)))" *) definition limit_ordinal_fm :: "i\i" where "limit_ordinal_fm(x) \ And(ordinal_fm(x), And(Neg(empty_fm(x)), Forall(Implies(Member(0,succ(x)), Exists(And(Member(0,succ(succ(x))), succ_fm(1,0)))))))" lemma limit_ordinal_type [TC]: "x \ nat \ limit_ordinal_fm(x) \ formula" by (simp add: limit_ordinal_fm_def) lemma sats_limit_ordinal_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, limit_ordinal_fm(x), env) \ limit_ordinal(##A, nth(x,env))" by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm') lemma limit_ordinal_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ limit_ordinal(##A, x) \ sats(A, limit_ordinal_fm(i), env)" by simp theorem limit_ordinal_reflection: "REFLECTS[\x. limit_ordinal(L,f(x)), \i x. limit_ordinal(##Lset(i),f(x))]" apply (simp only: limit_ordinal_def) apply (intro FOL_reflections ordinal_reflection empty_reflection successor_reflection) done subsubsection\Finite Ordinals: The Predicate ``Is A Natural Number''\ (* "finite_ordinal(M,a) \ ordinal(M,a) \ \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x))" *) definition finite_ordinal_fm :: "i\i" where "finite_ordinal_fm(x) \ And(ordinal_fm(x), And(Neg(limit_ordinal_fm(x)), Forall(Implies(Member(0,succ(x)), Neg(limit_ordinal_fm(0))))))" lemma finite_ordinal_type [TC]: "x \ nat \ finite_ordinal_fm(x) \ formula" by (simp add: finite_ordinal_fm_def) lemma sats_finite_ordinal_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, finite_ordinal_fm(x), env) \ finite_ordinal(##A, nth(x,env))" by (simp add: finite_ordinal_fm_def sats_ordinal_fm' finite_ordinal_def) lemma finite_ordinal_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ finite_ordinal(##A, x) \ sats(A, finite_ordinal_fm(i), env)" by simp theorem finite_ordinal_reflection: "REFLECTS[\x. finite_ordinal(L,f(x)), \i x. finite_ordinal(##Lset(i),f(x))]" apply (simp only: finite_ordinal_def) apply (intro FOL_reflections ordinal_reflection limit_ordinal_reflection) done subsubsection\Omega: The Set of Natural Numbers\ (* omega(M,a) \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x)) *) definition omega_fm :: "i\i" where "omega_fm(x) \ And(limit_ordinal_fm(x), Forall(Implies(Member(0,succ(x)), Neg(limit_ordinal_fm(0)))))" lemma omega_type [TC]: "x \ nat \ omega_fm(x) \ formula" by (simp add: omega_fm_def) lemma sats_omega_fm [simp]: "\x \ nat; env \ list(A)\ \ sats(A, omega_fm(x), env) \ omega(##A, nth(x,env))" by (simp add: omega_fm_def omega_def) lemma omega_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)\ \ omega(##A, x) \ sats(A, omega_fm(i), env)" by simp theorem omega_reflection: "REFLECTS[\x. omega(L,f(x)), \i x. omega(##Lset(i),f(x))]" apply (simp only: omega_def) apply (intro FOL_reflections limit_ordinal_reflection) done lemmas fun_plus_reflections = typed_function_reflection composition_reflection injection_reflection surjection_reflection bijection_reflection restriction_reflection order_isomorphism_reflection finite_ordinal_reflection ordinal_reflection limit_ordinal_reflection omega_reflection lemmas fun_plus_iff_sats = typed_function_iff_sats composition_iff_sats injection_iff_sats surjection_iff_sats bijection_iff_sats restriction_iff_sats order_isomorphism_iff_sats finite_ordinal_iff_sats ordinal_iff_sats limit_ordinal_iff_sats omega_iff_sats end diff --git a/src/ZF/Constructible/Rank.thy b/src/ZF/Constructible/Rank.thy --- a/src/ZF/Constructible/Rank.thy +++ b/src/ZF/Constructible/Rank.thy @@ -1,942 +1,942 @@ (* Title: ZF/Constructible/Rank.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) section \Absoluteness for Order Types, Rank Functions and Well-Founded Relations\ theory Rank imports WF_absolute begin subsection \Order Types: A Direct Construction by Replacement\ locale M_ordertype = M_basic + assumes well_ord_iso_separation: "\M(A); M(f); M(r)\ \ separation (M, \x. x\A \ (\y[M]. (\p[M]. fun_apply(M,f,x,y) \ pair(M,y,x,p) \ p \ r)))" and obase_separation: \ \part of the order type formalization\ "\M(A); M(r)\ \ separation(M, \a. \x[M]. \g[M]. \mx[M]. \par[M]. ordinal(M,x) \ membership(M,x,mx) \ pred_set(M,A,a,r,par) \ order_isomorphism(M,par,r,x,mx,g))" and obase_equals_separation: "\M(A); M(r)\ \ separation (M, \x. x\A \ \(\y[M]. \g[M]. ordinal(M,y) \ (\my[M]. \pxr[M]. membership(M,y,my) \ pred_set(M,A,x,r,pxr) \ order_isomorphism(M,pxr,r,y,my,g))))" and omap_replacement: "\M(A); M(r)\ \ strong_replacement(M, \a z. \x[M]. \g[M]. \mx[M]. \par[M]. ordinal(M,x) \ pair(M,a,x,z) \ membership(M,x,mx) \ pred_set(M,A,a,r,par) \ order_isomorphism(M,par,r,x,mx,g))" text\Inductive argument for Kunen's Lemma I 6.1, etc. Simple proof from Halmos, page 72\ lemma (in M_ordertype) wellordered_iso_subset_lemma: "\wellordered(M,A,r); f \ ord_iso(A,r, A',r); A'<= A; y \ A; M(A); M(f); M(r)\ \ \ \ r" -apply (unfold wellordered_def ord_iso_def) + unfolding wellordered_def ord_iso_def apply (elim conjE CollectE) apply (erule wellfounded_on_induct, assumption+) apply (insert well_ord_iso_separation [of A f r]) apply (simp, clarify) apply (drule_tac a = x in bij_is_fun [THEN apply_type], assumption, blast) done text\Kunen's Lemma I 6.1, page 14: there's no order-isomorphism to an initial segment of a well-ordering\ lemma (in M_ordertype) wellordered_iso_predD: "\wellordered(M,A,r); f \ ord_iso(A, r, Order.pred(A,x,r), r); M(A); M(f); M(r)\ \ x \ A" apply (rule notI) apply (frule wellordered_iso_subset_lemma, assumption) apply (auto elim: predE) (*Now we know \ (f`x < x) *) apply (drule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption) (*Now we also know f`x \ pred(A,x,r); contradiction! *) apply (simp add: Order.pred_def) done lemma (in M_ordertype) wellordered_iso_pred_eq_lemma: "\f \ \Order.pred(A,y,r), r\ \ \Order.pred(A,x,r), r\; wellordered(M,A,r); x\A; y\A; M(A); M(f); M(r)\ \ \x,y\ \ r" apply (frule wellordered_is_trans_on, assumption) apply (rule notI) apply (drule_tac x2=y and x=x and r2=r in wellordered_subset [OF _ pred_subset, THEN wellordered_iso_predD]) apply (simp add: trans_pred_pred_eq) apply (blast intro: predI dest: transM)+ done text\Simple consequence of Lemma 6.1\ lemma (in M_ordertype) wellordered_iso_pred_eq: "\wellordered(M,A,r); f \ ord_iso(Order.pred(A,a,r), r, Order.pred(A,c,r), r); M(A); M(f); M(r); a\A; c\A\ \ a=c" apply (frule wellordered_is_trans_on, assumption) apply (frule wellordered_is_linear, assumption) apply (erule_tac x=a and y=c in linearE, auto) apply (drule ord_iso_sym) (*two symmetric cases*) apply (blast dest: wellordered_iso_pred_eq_lemma)+ done text\Following Kunen's Theorem I 7.6, page 17. Note that this material is not required elsewhere.\ text\Can't use \well_ord_iso_preserving\ because it needs the strong premise \<^term>\well_ord(A,r)\\ lemma (in M_ordertype) ord_iso_pred_imp_lt: "\f \ ord_iso(Order.pred(A,x,r), r, i, Memrel(i)); g \ ord_iso(Order.pred(A,y,r), r, j, Memrel(j)); wellordered(M,A,r); x \ A; y \ A; M(A); M(r); M(f); M(g); M(j); Ord(i); Ord(j); \x,y\ \ r\ \ i < j" apply (frule wellordered_is_trans_on, assumption) apply (frule_tac y=y in transM, assumption) apply (rule_tac i=i and j=j in Ord_linear_lt, auto) txt\case \<^term>\i=j\ yields a contradiction\ apply (rule_tac x1=x and A1="Order.pred(A,y,r)" in wellordered_iso_predD [THEN notE]) apply (blast intro: wellordered_subset [OF _ pred_subset]) apply (simp add: trans_pred_pred_eq) apply (blast intro: Ord_iso_implies_eq ord_iso_sym ord_iso_trans) apply (simp_all add: pred_iff) txt\case \<^term>\j also yields a contradiction\ apply (frule restrict_ord_iso2, assumption+) apply (frule ord_iso_sym [THEN ord_iso_is_bij, THEN bij_is_fun]) apply (frule apply_type, blast intro: ltD) \ \thus \<^term>\converse(f)`j \ Order.pred(A,x,r)\\ apply (simp add: pred_iff) apply (subgoal_tac "\h[M]. h \ ord_iso(Order.pred(A,y,r), r, Order.pred(A, converse(f)`j, r), r)") apply (clarify, frule wellordered_iso_pred_eq, assumption+) apply (blast dest: wellordered_asym) apply (intro rexI) apply (blast intro: Ord_iso_implies_eq ord_iso_sym ord_iso_trans)+ done lemma ord_iso_converse1: "\f: ord_iso(A,r,B,s); : s; a:A; b:B\ \ \ r" apply (frule ord_iso_converse, assumption+) apply (blast intro: ord_iso_is_bij [THEN bij_is_fun, THEN apply_funtype]) apply (simp add: left_inverse_bij [OF ord_iso_is_bij]) done definition obase :: "[i\o,i,i] \ i" where \ \the domain of \om\, eventually shown to equal \A\\ "obase(M,A,r) \ {a\A. \x[M]. \g[M]. Ord(x) \ g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x))}" definition omap :: "[i\o,i,i,i] \ o" where \ \the function that maps wosets to order types\ "omap(M,A,r,f) \ \z[M]. z \ f \ (\a\A. \x[M]. \g[M]. z = \a,x\ \ Ord(x) \ g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))" definition otype :: "[i\o,i,i,i] \ o" where \ \the order types themselves\ "otype(M,A,r,i) \ \f[M]. omap(M,A,r,f) \ is_range(M,f,i)" text\Can also be proved with the premise \<^term>\M(z)\ instead of \<^term>\M(f)\, but that version is less useful. This lemma is also more useful than the definition, \omap_def\.\ lemma (in M_ordertype) omap_iff: "\omap(M,A,r,f); M(A); M(f)\ \ z \ f \ (\a\A. \x[M]. \g[M]. z = \a,x\ \ Ord(x) \ g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))" apply (simp add: omap_def) apply (rule iffI) apply (drule_tac [2] x=z in rspec) apply (drule_tac x=z in rspec) apply (blast dest: transM)+ done lemma (in M_ordertype) omap_unique: "\omap(M,A,r,f); omap(M,A,r,f'); M(A); M(r); M(f); M(f')\ \ f' = f" apply (rule equality_iffI) apply (simp add: omap_iff) done lemma (in M_ordertype) omap_yields_Ord: "\omap(M,A,r,f); \a,x\ \ f; M(a); M(x)\ \ Ord(x)" by (simp add: omap_def) lemma (in M_ordertype) otype_iff: "\otype(M,A,r,i); M(A); M(r); M(i)\ \ x \ i \ (M(x) \ Ord(x) \ (\a\A. \g[M]. g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x))))" apply (auto simp add: omap_iff otype_def) apply (blast intro: transM) apply (rule rangeI) apply (frule transM, assumption) apply (simp add: omap_iff, blast) done lemma (in M_ordertype) otype_eq_range: "\omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\ \ i = range(f)" apply (auto simp add: otype_def omap_iff) apply (blast dest: omap_unique) done lemma (in M_ordertype) Ord_otype: "\otype(M,A,r,i); trans[A](r); M(A); M(r); M(i)\ \ Ord(i)" apply (rule OrdI) prefer 2 apply (simp add: Ord_def otype_def omap_def) apply clarify apply (frule pair_components_in_M, assumption) apply blast apply (auto simp add: Transset_def otype_iff) apply (blast intro: transM) apply (blast intro: Ord_in_Ord) apply (rename_tac y a g) apply (frule ord_iso_sym [THEN ord_iso_is_bij, THEN bij_is_fun, THEN apply_funtype], assumption) apply (rule_tac x="converse(g)`y" in bexI) apply (frule_tac a="converse(g) ` y" in ord_iso_restrict_pred, assumption) apply (safe elim!: predE) apply (blast intro: restrict_ord_iso ord_iso_sym ltI dest: transM) done lemma (in M_ordertype) domain_omap: "\omap(M,A,r,f); M(A); M(r); M(B); M(f)\ \ domain(f) = obase(M,A,r)" apply (simp add: obase_def) apply (rule equality_iffI) apply (simp add: domain_iff omap_iff, blast) done lemma (in M_ordertype) omap_subset: "\omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(B); M(i)\ \ f \ obase(M,A,r) * i" apply clarify apply (simp add: omap_iff obase_def) apply (force simp add: otype_iff) done lemma (in M_ordertype) omap_funtype: "\omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\ \ f \ obase(M,A,r) -> i" apply (simp add: domain_omap omap_subset Pi_iff function_def omap_iff) apply (blast intro: Ord_iso_implies_eq ord_iso_sym ord_iso_trans) done lemma (in M_ordertype) wellordered_omap_bij: "\wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\ \ f \ bij(obase(M,A,r),i)" apply (insert omap_funtype [of A r f i]) apply (auto simp add: bij_def inj_def) prefer 2 apply (blast intro: fun_is_surj dest: otype_eq_range) apply (frule_tac a=w in apply_Pair, assumption) apply (frule_tac a=x in apply_Pair, assumption) apply (simp add: omap_iff) apply (blast intro: wellordered_iso_pred_eq ord_iso_sym ord_iso_trans) done text\This is not the final result: we must show \<^term>\oB(A,r) = A\\ lemma (in M_ordertype) omap_ord_iso: "\wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\ \ f \ ord_iso(obase(M,A,r),r,i,Memrel(i))" apply (rule ord_isoI) apply (erule wellordered_omap_bij, assumption+) apply (insert omap_funtype [of A r f i], simp) apply (frule_tac a=x in apply_Pair, assumption) apply (frule_tac a=y in apply_Pair, assumption) apply (auto simp add: omap_iff) txt\direction 1: assuming \<^term>\\x,y\ \ r\\ apply (blast intro: ltD ord_iso_pred_imp_lt) txt\direction 2: proving \<^term>\\x,y\ \ r\ using linearity of \<^term>\r\\ apply (rename_tac x y g ga) apply (frule wellordered_is_linear, assumption, erule_tac x=x and y=y in linearE, assumption+) txt\the case \<^term>\x=y\ leads to immediate contradiction\ apply (blast elim: mem_irrefl) txt\the case \<^term>\\y,x\ \ r\: handle like the opposite direction\ apply (blast dest: ord_iso_pred_imp_lt ltD elim: mem_asym) done lemma (in M_ordertype) Ord_omap_image_pred: "\wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i); b \ A\ \ Ord(f `` Order.pred(A,b,r))" apply (frule wellordered_is_trans_on, assumption) apply (rule OrdI) prefer 2 apply (simp add: image_iff omap_iff Ord_def, blast) txt\Hard part is to show that the image is a transitive set.\ apply (simp add: Transset_def, clarify) apply (simp add: image_iff pred_iff apply_iff [OF omap_funtype [of A r f i]]) apply (rename_tac c j, clarify) apply (frule omap_funtype [of A r f, THEN apply_funtype], assumption+) apply (subgoal_tac "j \ i") prefer 2 apply (blast intro: Ord_trans Ord_otype) apply (subgoal_tac "converse(f) ` j \ obase(M,A,r)") prefer 2 apply (blast dest: wellordered_omap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN apply_funtype]) apply (rule_tac x="converse(f) ` j" in bexI) apply (simp add: right_inverse_bij [OF wellordered_omap_bij]) apply (intro predI conjI) apply (erule_tac b=c in trans_onD) apply (rule ord_iso_converse1 [OF omap_ord_iso [of A r f i]]) apply (auto simp add: obase_def) done lemma (in M_ordertype) restrict_omap_ord_iso: "\wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i); D \ obase(M,A,r); M(A); M(r); M(f); M(i)\ \ restrict(f,D) \ (\D,r\ \ \f``D, Memrel(f``D)\)" apply (frule ord_iso_restrict_image [OF omap_ord_iso [of A r f i]], assumption+) apply (drule ord_iso_sym [THEN subset_ord_iso_Memrel]) apply (blast dest: subsetD [OF omap_subset]) apply (drule ord_iso_sym, simp) done lemma (in M_ordertype) obase_equals: "\wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\ \ obase(M,A,r) = A" apply (rule equalityI, force simp add: obase_def, clarify) apply (unfold obase_def, simp) apply (frule wellordered_is_wellfounded_on, assumption) apply (erule wellfounded_on_induct, assumption+) apply (frule obase_equals_separation [of A r], assumption) apply (simp, clarify) apply (rename_tac b) apply (subgoal_tac "Order.pred(A,b,r) \ obase(M,A,r)") apply (blast intro!: restrict_omap_ord_iso Ord_omap_image_pred) apply (force simp add: pred_iff obase_def) done text\Main result: \<^term>\om\ gives the order-isomorphism \<^term>\\A,r\ \ \i, Memrel(i)\\\ theorem (in M_ordertype) omap_ord_iso_otype: "\wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i); M(A); M(r); M(f); M(i)\ \ f \ ord_iso(A, r, i, Memrel(i))" apply (frule omap_ord_iso, assumption+) apply (simp add: obase_equals) done lemma (in M_ordertype) obase_exists: "\M(A); M(r)\ \ M(obase(M,A,r))" apply (simp add: obase_def) apply (insert obase_separation [of A r]) apply (simp add: separation_def) done lemma (in M_ordertype) omap_exists: "\M(A); M(r)\ \ \z[M]. omap(M,A,r,z)" apply (simp add: omap_def) apply (insert omap_replacement [of A r]) apply (simp add: strong_replacement_def) apply (drule_tac x="obase(M,A,r)" in rspec) apply (simp add: obase_exists) apply (simp add: obase_def) apply (erule impE) apply (clarsimp simp add: univalent_def) apply (blast intro: Ord_iso_implies_eq ord_iso_sym ord_iso_trans, clarify) apply (rule_tac x=Y in rexI) apply (simp add: obase_def, blast, assumption) done lemma (in M_ordertype) otype_exists: "\wellordered(M,A,r); M(A); M(r)\ \ \i[M]. otype(M,A,r,i)" apply (insert omap_exists [of A r]) apply (simp add: otype_def, safe) apply (rule_tac x="range(x)" in rexI) apply blast+ done lemma (in M_ordertype) ordertype_exists: "\wellordered(M,A,r); M(A); M(r)\ \ \f[M]. (\i[M]. Ord(i) \ f \ ord_iso(A, r, i, Memrel(i)))" apply (insert obase_exists [of A r] omap_exists [of A r] otype_exists [of A r], simp, clarify) apply (rename_tac i) apply (subgoal_tac "Ord(i)", blast intro: omap_ord_iso_otype) apply (rule Ord_otype) apply (force simp add: otype_def) apply (simp_all add: wellordered_is_trans_on) done lemma (in M_ordertype) relativized_imp_well_ord: "\wellordered(M,A,r); M(A); M(r)\ \ well_ord(A,r)" apply (insert ordertype_exists [of A r], simp) apply (blast intro: well_ord_ord_iso well_ord_Memrel) done subsection \Kunen's theorem 5.4, page 127\ text\(a) The notion of Wellordering is absolute\ theorem (in M_ordertype) well_ord_abs [simp]: "\M(A); M(r)\ \ wellordered(M,A,r) \ well_ord(A,r)" by (blast intro: well_ord_imp_relativized relativized_imp_well_ord) text\(b) Order types are absolute\ theorem (in M_ordertype) ordertypes_are_absolute: "\wellordered(M,A,r); f \ ord_iso(A, r, i, Memrel(i)); M(A); M(r); M(f); M(i); Ord(i)\ \ i = ordertype(A,r)" by (blast intro: Ord_ordertype relativized_imp_well_ord ordertype_ord_iso Ord_iso_implies_eq ord_iso_sym ord_iso_trans) subsection\Ordinal Arithmetic: Two Examples of Recursion\ text\Note: the remainder of this theory is not needed elsewhere.\ subsubsection\Ordinal Addition\ (*FIXME: update to use new techniques\*) (*This expresses ordinal addition in the language of ZF. It also provides an abbreviation that can be used in the instance of strong replacement below. Here j is used to define the relation, namely Memrel(succ(j)), while x determines the domain of f.*) definition is_oadd_fun :: "[i\o,i,i,i,i] \ o" where "is_oadd_fun(M,i,j,x,f) \ (\sj msj. M(sj) \ M(msj) \ successor(M,j,sj) \ membership(M,sj,msj) \ M_is_recfun(M, \x g y. \gx[M]. image(M,g,x,gx) \ union(M,i,gx,y), msj, x, f))" definition is_oadd :: "[i\o,i,i,i] \ o" where "is_oadd(M,i,j,k) \ (\ ordinal(M,i) \ \ ordinal(M,j) \ k=0) | (\ ordinal(M,i) \ ordinal(M,j) \ k=j) | (ordinal(M,i) \ \ ordinal(M,j) \ k=i) | (ordinal(M,i) \ ordinal(M,j) \ (\f fj sj. M(f) \ M(fj) \ M(sj) \ successor(M,j,sj) \ is_oadd_fun(M,i,sj,sj,f) \ fun_apply(M,f,j,fj) \ fj = k))" definition (*NEEDS RELATIVIZATION*) omult_eqns :: "[i,i,i,i] \ o" where "omult_eqns(i,x,g,z) \ Ord(x) \ (x=0 \ z=0) \ (\j. x = succ(j) \ z = g`j ++ i) \ (Limit(x) \ z = \(g``x))" definition is_omult_fun :: "[i\o,i,i,i] \ o" where "is_omult_fun(M,i,j,f) \ (\df. M(df) \ is_function(M,f) \ is_domain(M,f,df) \ subset(M, j, df)) \ (\x\j. omult_eqns(i,x,f,f`x))" definition is_omult :: "[i\o,i,i,i] \ o" where "is_omult(M,i,j,k) \ \f fj sj. M(f) \ M(fj) \ M(sj) \ successor(M,j,sj) \ is_omult_fun(M,i,sj,f) \ fun_apply(M,f,j,fj) \ fj = k" locale M_ord_arith = M_ordertype + assumes oadd_strong_replacement: "\M(i); M(j)\ \ strong_replacement(M, \x z. \y[M]. pair(M,x,y,z) \ (\f[M]. \fx[M]. is_oadd_fun(M,i,j,x,f) \ image(M,f,x,fx) \ y = i \ fx))" and omult_strong_replacement': "\M(i); M(j)\ \ strong_replacement(M, \x z. \y[M]. z = \x,y\ \ (\g[M]. is_recfun(Memrel(succ(j)),x,\x g. THE z. omult_eqns(i,x,g,z),g) \ y = (THE z. omult_eqns(i, x, g, z))))" text\\is_oadd_fun\: Relating the pure "language of set theory" to Isabelle/ZF\ lemma (in M_ord_arith) is_oadd_fun_iff: "\a\j; M(i); M(j); M(a); M(f)\ \ is_oadd_fun(M,i,j,a,f) \ f \ a \ range(f) \ (\x. M(x) \ x < a \ f`x = i \ f``x)" apply (frule lt_Ord) apply (simp add: is_oadd_fun_def relation2_def is_recfun_abs [of "\x g. i \ g``x"] is_recfun_iff_equation Ball_def lt_trans [OF ltI, of _ a] lt_Memrel) apply (simp add: lt_def) apply (blast dest: transM) done lemma (in M_ord_arith) oadd_strong_replacement': "\M(i); M(j)\ \ strong_replacement(M, \x z. \y[M]. z = \x,y\ \ (\g[M]. is_recfun(Memrel(succ(j)),x,\x g. i \ g``x,g) \ y = i \ g``x))" apply (insert oadd_strong_replacement [of i j]) apply (simp add: is_oadd_fun_def relation2_def is_recfun_abs [of "\x g. i \ g``x"]) done lemma (in M_ord_arith) exists_oadd: "\Ord(j); M(i); M(j)\ \ \f[M]. is_recfun(Memrel(succ(j)), j, \x g. i \ g``x, f)" apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel]) apply (simp_all add: Memrel_type oadd_strong_replacement') done lemma (in M_ord_arith) exists_oadd_fun: "\Ord(j); M(i); M(j)\ \ \f[M]. is_oadd_fun(M,i,succ(j),succ(j),f)" apply (rule exists_oadd [THEN rexE]) apply (erule Ord_succ, assumption, simp) apply (rename_tac f) apply (frule is_recfun_type) apply (rule_tac x=f in rexI) apply (simp add: fun_is_function domain_of_fun lt_Memrel apply_recfun lt_def is_oadd_fun_iff Ord_trans [OF _ succI1], assumption) done lemma (in M_ord_arith) is_oadd_fun_apply: "\x < j; M(i); M(j); M(f); is_oadd_fun(M,i,j,j,f)\ \ f`x = i \ (\k\x. {f ` k})" apply (simp add: is_oadd_fun_iff lt_Ord2, clarify) apply (frule lt_closed, simp) apply (frule leI [THEN le_imp_subset]) apply (simp add: image_fun, blast) done lemma (in M_ord_arith) is_oadd_fun_iff_oadd [rule_format]: "\is_oadd_fun(M,i,J,J,f); M(i); M(J); M(f); Ord(i); Ord(j)\ \ j f`j = i++j" apply (erule_tac i=j in trans_induct, clarify) apply (subgoal_tac "\k\x. kM(i); M(j); M(k); Ord(i); Ord(j)\ \ is_oadd(M,i,j,k) \ k = i++j" apply (simp add: is_oadd_def is_oadd_fun_iff_oadd) apply (frule exists_oadd_fun [of j i], blast+) done lemma (in M_ord_arith) oadd_abs: "\M(i); M(j); M(k)\ \ is_oadd(M,i,j,k) \ k = i++j" apply (case_tac "Ord(i) \ Ord(j)") apply (simp add: Ord_oadd_abs) apply (auto simp add: is_oadd_def oadd_eq_if_raw_oadd) done lemma (in M_ord_arith) oadd_closed [intro,simp]: "\M(i); M(j)\ \ M(i++j)" apply (simp add: oadd_eq_if_raw_oadd, clarify) apply (simp add: raw_oadd_eq_oadd) apply (frule exists_oadd_fun [of j i], auto) apply (simp add: is_oadd_fun_iff_oadd [symmetric]) done subsubsection\Ordinal Multiplication\ lemma omult_eqns_unique: "\omult_eqns(i,x,g,z); omult_eqns(i,x,g,z')\ \ z=z'" apply (simp add: omult_eqns_def, clarify) apply (erule Ord_cases, simp_all) done lemma omult_eqns_0: "omult_eqns(i,0,g,z) \ z=0" by (simp add: omult_eqns_def) lemma the_omult_eqns_0: "(THE z. omult_eqns(i,0,g,z)) = 0" by (simp add: omult_eqns_0) lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) \ Ord(j) \ z = g`j ++ i" by (simp add: omult_eqns_def) lemma the_omult_eqns_succ: "Ord(j) \ (THE z. omult_eqns(i,succ(j),g,z)) = g`j ++ i" by (simp add: omult_eqns_succ) lemma omult_eqns_Limit: "Limit(x) \ omult_eqns(i,x,g,z) \ z = \(g``x)" apply (simp add: omult_eqns_def) apply (blast intro: Limit_is_Ord) done lemma the_omult_eqns_Limit: "Limit(x) \ (THE z. omult_eqns(i,x,g,z)) = \(g``x)" by (simp add: omult_eqns_Limit) lemma omult_eqns_Not: "\ Ord(x) \ \ omult_eqns(i,x,g,z)" by (simp add: omult_eqns_def) lemma (in M_ord_arith) the_omult_eqns_closed: "\M(i); M(x); M(g); function(g)\ \ M(THE z. omult_eqns(i, x, g, z))" apply (case_tac "Ord(x)") prefer 2 apply (simp add: omult_eqns_Not) \ \trivial, non-Ord case\ apply (erule Ord_cases) apply (simp add: omult_eqns_0) apply (simp add: omult_eqns_succ) apply (simp add: omult_eqns_Limit) done lemma (in M_ord_arith) exists_omult: "\Ord(j); M(i); M(j)\ \ \f[M]. is_recfun(Memrel(succ(j)), j, \x g. THE z. omult_eqns(i,x,g,z), f)" apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel]) apply (simp_all add: Memrel_type omult_strong_replacement') apply (blast intro: the_omult_eqns_closed) done lemma (in M_ord_arith) exists_omult_fun: "\Ord(j); M(i); M(j)\ \ \f[M]. is_omult_fun(M,i,succ(j),f)" apply (rule exists_omult [THEN rexE]) apply (erule Ord_succ, assumption, simp) apply (rename_tac f) apply (frule is_recfun_type) apply (rule_tac x=f in rexI) apply (simp add: fun_is_function domain_of_fun lt_Memrel apply_recfun lt_def is_omult_fun_def Ord_trans [OF _ succI1]) apply (force dest: Ord_in_Ord' simp add: omult_eqns_def the_omult_eqns_0 the_omult_eqns_succ the_omult_eqns_Limit, assumption) done lemma (in M_ord_arith) is_omult_fun_apply_0: "\0 < j; is_omult_fun(M,i,j,f)\ \ f`0 = 0" by (simp add: is_omult_fun_def omult_eqns_def lt_def ball_conj_distrib) lemma (in M_ord_arith) is_omult_fun_apply_succ: "\succ(x) < j; is_omult_fun(M,i,j,f)\ \ f`succ(x) = f`x ++ i" by (simp add: is_omult_fun_def omult_eqns_def lt_def, blast) lemma (in M_ord_arith) is_omult_fun_apply_Limit: "\x < j; Limit(x); M(j); M(f); is_omult_fun(M,i,j,f)\ \ f ` x = (\y\x. f`y)" apply (simp add: is_omult_fun_def omult_eqns_def lt_def, clarify) apply (drule subset_trans [OF OrdmemD], assumption+) apply (simp add: ball_conj_distrib omult_Limit image_function) done lemma (in M_ord_arith) is_omult_fun_eq_omult: "\is_omult_fun(M,i,J,f); M(J); M(f); Ord(i); Ord(j)\ \ j f`j = i**j" apply (erule_tac i=j in trans_induct3) apply (safe del: impCE) apply (simp add: is_omult_fun_apply_0) apply (subgoal_tac "xk\x. kM(i); M(j); M(k); Ord(i); Ord(j)\ \ is_omult(M,i,j,k) \ k = i**j" apply (simp add: is_omult_def is_omult_fun_eq_omult) apply (frule exists_omult_fun [of j i], blast+) done subsection \Absoluteness of Well-Founded Relations\ text\Relativized to \<^term>\M\: Every well-founded relation is a subset of some inverse image of an ordinal. Key step is the construction (in \<^term>\M\) of a rank function.\ locale M_wfrank = M_trancl + assumes wfrank_separation: "M(r) \ separation (M, \x. \rplus[M]. tran_closure(M,r,rplus) \ \ (\f[M]. M_is_recfun(M, \x f y. is_range(M,f,y), rplus, x, f)))" and wfrank_strong_replacement: "M(r) \ strong_replacement(M, \x z. \rplus[M]. tran_closure(M,r,rplus) \ (\y[M]. \f[M]. pair(M,x,y,z) \ M_is_recfun(M, \x f y. is_range(M,f,y), rplus, x, f) \ is_range(M,f,y)))" and Ord_wfrank_separation: "M(r) \ separation (M, \x. \rplus[M]. tran_closure(M,r,rplus) \ \ (\f[M]. \rangef[M]. is_range(M,f,rangef) \ M_is_recfun(M, \x f y. is_range(M,f,y), rplus, x, f) \ ordinal(M,rangef)))" text\Proving that the relativized instances of Separation or Replacement agree with the "real" ones.\ lemma (in M_wfrank) wfrank_separation': "M(r) \ separation (M, \x. \ (\f[M]. is_recfun(r^+, x, \x f. range(f), f)))" apply (insert wfrank_separation [of r]) apply (simp add: relation2_def is_recfun_abs [of "\x. range"]) done lemma (in M_wfrank) wfrank_strong_replacement': "M(r) \ strong_replacement(M, \x z. \y[M]. \f[M]. pair(M,x,y,z) \ is_recfun(r^+, x, \x f. range(f), f) \ y = range(f))" apply (insert wfrank_strong_replacement [of r]) apply (simp add: relation2_def is_recfun_abs [of "\x. range"]) done lemma (in M_wfrank) Ord_wfrank_separation': "M(r) \ separation (M, \x. \ (\f[M]. is_recfun(r^+, x, \x. range, f) \ Ord(range(f))))" apply (insert Ord_wfrank_separation [of r]) apply (simp add: relation2_def is_recfun_abs [of "\x. range"]) done text\This function, defined using replacement, is a rank function for well-founded relations within the class M.\ definition wellfoundedrank :: "[i\o,i,i] \ i" where "wellfoundedrank(M,r,A) \ {p. x\A, \y[M]. \f[M]. p = \x,y\ \ is_recfun(r^+, x, \x f. range(f), f) \ y = range(f)}" lemma (in M_wfrank) exists_wfrank: "\wellfounded(M,r); M(a); M(r)\ \ \f[M]. is_recfun(r^+, a, \x f. range(f), f)" apply (rule wellfounded_exists_is_recfun) apply (blast intro: wellfounded_trancl) apply (rule trans_trancl) apply (erule wfrank_separation') apply (erule wfrank_strong_replacement') apply (simp_all add: trancl_subset_times) done lemma (in M_wfrank) M_wellfoundedrank: "\wellfounded(M,r); M(r); M(A)\ \ M(wellfoundedrank(M,r,A))" apply (insert wfrank_strong_replacement' [of r]) apply (simp add: wellfoundedrank_def) apply (rule strong_replacement_closed) apply assumption+ apply (rule univalent_is_recfun) apply (blast intro: wellfounded_trancl) apply (rule trans_trancl) apply (simp add: trancl_subset_times) apply (blast dest: transM) done lemma (in M_wfrank) Ord_wfrank_range [rule_format]: "\wellfounded(M,r); a\A; M(r); M(A)\ \ \f[M]. is_recfun(r^+, a, \x f. range(f), f) \ Ord(range(f))" apply (drule wellfounded_trancl, assumption) apply (rule wellfounded_induct, assumption, erule (1) transM) apply simp apply (blast intro: Ord_wfrank_separation', clarify) txt\The reasoning in both cases is that we get \<^term>\y\ such that \<^term>\\y, x\ \ r^+\. We find that \<^term>\f`y = restrict(f, r^+ -`` {y})\.\ apply (rule OrdI [OF _ Ord_is_Transset]) txt\An ordinal is a transitive set...\ apply (simp add: Transset_def) apply clarify apply (frule apply_recfun2, assumption) apply (force simp add: restrict_iff) txt\...of ordinals. This second case requires the induction hyp.\ apply clarify apply (rename_tac i y) apply (frule apply_recfun2, assumption) apply (frule is_recfun_imp_in_r, assumption) apply (frule is_recfun_restrict) (*simp_all won't work*) apply (simp add: trans_trancl trancl_subset_times)+ apply (drule spec [THEN mp], assumption) apply (subgoal_tac "M(restrict(f, r^+ -`` {y}))") apply (drule_tac x="restrict(f, r^+ -`` {y})" in rspec) apply assumption apply (simp add: function_apply_equality [OF _ is_recfun_imp_function]) apply (blast dest: pair_components_in_M) done lemma (in M_wfrank) Ord_range_wellfoundedrank: "\wellfounded(M,r); r \ A*A; M(r); M(A)\ \ Ord (range(wellfoundedrank(M,r,A)))" apply (frule wellfounded_trancl, assumption) apply (frule trancl_subset_times) apply (simp add: wellfoundedrank_def) apply (rule OrdI [OF _ Ord_is_Transset]) prefer 2 txt\by our previous result the range consists of ordinals.\ apply (blast intro: Ord_wfrank_range) txt\We still must show that the range is a transitive set.\ apply (simp add: Transset_def, clarify, simp) apply (rename_tac x i f u) apply (frule is_recfun_imp_in_r, assumption) apply (subgoal_tac "M(u) \ M(i) \ M(x)") prefer 2 apply (blast dest: transM, clarify) apply (rule_tac a=u in rangeI) apply (rule_tac x=u in ReplaceI) apply simp apply (rule_tac x="restrict(f, r^+ -`` {u})" in rexI) apply (blast intro: is_recfun_restrict trans_trancl dest: apply_recfun2) apply simp apply blast txt\Unicity requirement of Replacement\ apply clarify apply (frule apply_recfun2, assumption) apply (simp add: trans_trancl is_recfun_cut) done lemma (in M_wfrank) function_wellfoundedrank: "\wellfounded(M,r); M(r); M(A)\ \ function(wellfoundedrank(M,r,A))" apply (simp add: wellfoundedrank_def function_def, clarify) txt\Uniqueness: repeated below!\ apply (drule is_recfun_functional, assumption) apply (blast intro: wellfounded_trancl) apply (simp_all add: trancl_subset_times trans_trancl) done lemma (in M_wfrank) domain_wellfoundedrank: "\wellfounded(M,r); M(r); M(A)\ \ domain(wellfoundedrank(M,r,A)) = A" apply (simp add: wellfoundedrank_def function_def) apply (rule equalityI, auto) apply (frule transM, assumption) apply (frule_tac a=x in exists_wfrank, assumption+, clarify) apply (rule_tac b="range(f)" in domainI) apply (rule_tac x=x in ReplaceI) apply simp apply (rule_tac x=f in rexI, blast, simp_all) txt\Uniqueness (for Replacement): repeated above!\ apply clarify apply (drule is_recfun_functional, assumption) apply (blast intro: wellfounded_trancl) apply (simp_all add: trancl_subset_times trans_trancl) done lemma (in M_wfrank) wellfoundedrank_type: "\wellfounded(M,r); M(r); M(A)\ \ wellfoundedrank(M,r,A) \ A -> range(wellfoundedrank(M,r,A))" apply (frule function_wellfoundedrank [of r A], assumption+) apply (frule function_imp_Pi) apply (simp add: wellfoundedrank_def relation_def) apply blast apply (simp add: domain_wellfoundedrank) done lemma (in M_wfrank) Ord_wellfoundedrank: "\wellfounded(M,r); a \ A; r \ A*A; M(r); M(A)\ \ Ord(wellfoundedrank(M,r,A) ` a)" by (blast intro: apply_funtype [OF wellfoundedrank_type] Ord_in_Ord [OF Ord_range_wellfoundedrank]) lemma (in M_wfrank) wellfoundedrank_eq: "\is_recfun(r^+, a, \x. range, f); wellfounded(M,r); a \ A; M(f); M(r); M(A)\ \ wellfoundedrank(M,r,A) ` a = range(f)" apply (rule apply_equality) prefer 2 apply (blast intro: wellfoundedrank_type) apply (simp add: wellfoundedrank_def) apply (rule ReplaceI) apply (rule_tac x="range(f)" in rexI) apply blast apply simp_all txt\Unicity requirement of Replacement\ apply clarify apply (drule is_recfun_functional, assumption) apply (blast intro: wellfounded_trancl) apply (simp_all add: trancl_subset_times trans_trancl) done lemma (in M_wfrank) wellfoundedrank_lt: "\\a,b\ \ r; wellfounded(M,r); r \ A*A; M(r); M(A)\ \ wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b" apply (frule wellfounded_trancl, assumption) apply (subgoal_tac "a\A \ b\A") prefer 2 apply blast apply (simp add: lt_def Ord_wellfoundedrank, clarify) apply (frule exists_wfrank [of concl: _ b], erule (1) transM, assumption) apply clarify apply (rename_tac fb) apply (frule is_recfun_restrict [of concl: "r^+" a]) apply (rule trans_trancl, assumption) apply (simp_all add: r_into_trancl trancl_subset_times) txt\Still the same goal, but with new \is_recfun\ assumptions.\ apply (simp add: wellfoundedrank_eq) apply (frule_tac a=a in wellfoundedrank_eq, assumption+) apply (simp_all add: transM [of a]) txt\We have used equations for wellfoundedrank and now must use some for \is_recfun\.\ apply (rule_tac a=a in rangeI) apply (simp add: is_recfun_type [THEN apply_iff] vimage_singleton_iff r_into_trancl apply_recfun) done lemma (in M_wfrank) wellfounded_imp_subset_rvimage: "\wellfounded(M,r); r \ A*A; M(r); M(A)\ \ \i f. Ord(i) \ r \ rvimage(A, f, Memrel(i))" apply (rule_tac x="range(wellfoundedrank(M,r,A))" in exI) apply (rule_tac x="wellfoundedrank(M,r,A)" in exI) apply (simp add: Ord_range_wellfoundedrank, clarify) apply (frule subsetD, assumption, clarify) apply (simp add: rvimage_iff wellfoundedrank_lt [THEN ltD]) apply (blast intro: apply_rangeI wellfoundedrank_type) done lemma (in M_wfrank) wellfounded_imp_wf: "\wellfounded(M,r); relation(r); M(r)\ \ wf(r)" by (blast dest!: relation_field_times_field wellfounded_imp_subset_rvimage intro: wf_rvimage_Ord [THEN wf_subset]) lemma (in M_wfrank) wellfounded_on_imp_wf_on: "\wellfounded_on(M,A,r); relation(r); M(r); M(A)\ \ wf[A](r)" apply (simp add: wellfounded_on_iff_wellfounded wf_on_def) apply (rule wellfounded_imp_wf) apply (simp_all add: relation_def) done theorem (in M_wfrank) wf_abs: "\relation(r); M(r)\ \ wellfounded(M,r) \ wf(r)" by (blast intro: wellfounded_imp_wf wf_imp_relativized) theorem (in M_wfrank) wf_on_abs: "\relation(r); M(r); M(A)\ \ wellfounded_on(M,A,r) \ wf[A](r)" by (blast intro: wellfounded_on_imp_wf_on wf_on_imp_relativized) end diff --git a/src/ZF/Constructible/Reflection.thy b/src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy +++ b/src/ZF/Constructible/Reflection.thy @@ -1,370 +1,370 @@ (* Title: ZF/Constructible/Reflection.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) section \The Reflection Theorem\ theory Reflection imports Normal begin lemma all_iff_not_ex_not: "(\x. P(x)) \ (\ (\x. \ P(x)))" by blast lemma ball_iff_not_bex_not: "(\x\A. P(x)) \ (\ (\x\A. \ P(x)))" by blast text\From the notes of A. S. Kechris, page 6, and from Andrzej Mostowski, \emph{Constructible Sets with Applications}, North-Holland, 1969, page 23.\ subsection\Basic Definitions\ text\First part: the cumulative hierarchy defining the class \M\. To avoid handling multiple arguments, we assume that \Mset(l)\ is closed under ordered pairing provided \l\ is limit. Possibly this could be avoided: the induction hypothesis \<^term>\Cl_reflects\ (in locale \ex_reflection\) could be weakened to \<^term>\\y\Mset(a). \z\Mset(a). P(\y,z\) \ Q(a,\y,z\)\, removing most uses of \<^term>\Pair_in_Mset\. But there isn't much point in doing so, since ultimately the \ex_reflection\ proof is packaged up using the predicate \Reflects\. \ locale reflection = fixes Mset and M and Reflects assumes Mset_mono_le : "mono_le_subset(Mset)" and Mset_cont : "cont_Ord(Mset)" and Pair_in_Mset : "\x \ Mset(a); y \ Mset(a); Limit(a)\ \ \x,y\ \ Mset(a)" defines "M(x) \ \a. Ord(a) \ x \ Mset(a)" and "Reflects(Cl,P,Q) \ Closed_Unbounded(Cl) \ (\a. Cl(a) \ (\x\Mset(a). P(x) \ Q(a,x)))" fixes F0 \ \ordinal for a specific value \<^term>\y\\ fixes FF \ \sup over the whole level, \<^term>\y\Mset(a)\\ fixes ClEx \ \Reflecting ordinals for the formula \<^term>\\z. P\\ defines "F0(P,y) \ \ b. (\z. M(z) \ P(\y,z\)) \ (\z\Mset(b). P(\y,z\))" and "FF(P) \ \a. \y\Mset(a). F0(P,y)" and "ClEx(P,a) \ Limit(a) \ normalize(FF(P),a) = a" begin lemma Mset_mono: "i\j \ Mset(i) \ Mset(j)" using Mset_mono_le by (simp add: mono_le_subset_def leI) text\Awkward: we need a version of \ClEx_def\ as an equality at the level of classes, which do not really exist\ lemma ClEx_eq: "ClEx(P) \ \a. Limit(a) \ normalize(FF(P),a) = a" by (simp add: ClEx_def [symmetric]) subsection\Easy Cases of the Reflection Theorem\ theorem Triv_reflection [intro]: "Reflects(Ord, P, \a x. P(x))" by (simp add: Reflects_def) theorem Not_reflection [intro]: "Reflects(Cl,P,Q) \ Reflects(Cl, \x. \P(x), \a x. \Q(a,x))" by (simp add: Reflects_def) theorem And_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem Or_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem Imp_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem Iff_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) subsection\Reflection for Existential Quantifiers\ lemma F0_works: "\y\Mset(a); Ord(a); M(z); P(\y,z\)\ \ \z\Mset(F0(P,y)). P(\y,z\)" unfolding F0_def M_def apply clarify apply (rule LeastI2) apply (blast intro: Mset_mono [THEN subsetD]) apply (blast intro: lt_Ord2, blast) done lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))" by (simp add: F0_def) lemma Ord_FF [intro,simp]: "Ord(FF(P,y))" by (simp add: FF_def) lemma cont_Ord_FF: "cont_Ord(FF(P))" using Mset_cont by (simp add: cont_Ord_def FF_def, blast) text\Recall that \<^term>\F0\ depends upon \<^term>\y\Mset(a)\, while \<^term>\FF\ depends only upon \<^term>\a\.\ lemma FF_works: "\M(z); y\Mset(a); P(\y,z\); Ord(a)\ \ \z\Mset(FF(P,a)). P(\y,z\)" apply (simp add: FF_def) apply (simp_all add: cont_Ord_Union [of concl: Mset] Mset_cont Mset_mono_le not_emptyI) apply (blast intro: F0_works) done lemma FFN_works: "\M(z); y\Mset(a); P(\y,z\); Ord(a)\ \ \z\Mset(normalize(FF(P),a)). P(\y,z\)" apply (drule FF_works [of concl: P], assumption+) apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD]) done end text\Locale for the induction hypothesis\ locale ex_reflection = reflection + fixes P \ \the original formula\ fixes Q \ \the reflected formula\ fixes Cl \ \the class of reflecting ordinals\ assumes Cl_reflects: "\Cl(a); Ord(a)\ \ \x\Mset(a). P(x) \ Q(a,x)" begin lemma ClEx_downward: "\M(z); y\Mset(a); P(\y,z\); Cl(a); ClEx(P,a)\ \ \z\Mset(a). Q(a,\y,z\)" apply (simp add: ClEx_def, clarify) apply (frule Limit_is_Ord) apply (frule FFN_works [of concl: P], assumption+) apply (drule Cl_reflects, assumption+) apply (auto simp add: Limit_is_Ord Pair_in_Mset) done lemma ClEx_upward: "\z\Mset(a); y\Mset(a); Q(a,\y,z\); Cl(a); ClEx(P,a)\ \ \z. M(z) \ P(\y,z\)" apply (simp add: ClEx_def M_def) apply (blast dest: Cl_reflects intro: Limit_is_Ord Pair_in_Mset) done text\Class \ClEx\ indeed consists of reflecting ordinals...\ lemma ZF_ClEx_iff: "\y\Mset(a); Cl(a); ClEx(P,a)\ \ (\z. M(z) \ P(\y,z\)) \ (\z\Mset(a). Q(a,\y,z\))" by (blast intro: dest: ClEx_downward ClEx_upward) text\...and it is closed and unbounded\ lemma ZF_Closed_Unbounded_ClEx: "Closed_Unbounded(ClEx(P))" apply (simp add: ClEx_eq) apply (fast intro: Closed_Unbounded_Int Normal_imp_fp_Closed_Unbounded Closed_Unbounded_Limit Normal_normalize) done end text\The same two theorems, exported to locale \reflection\.\ context reflection begin text\Class \ClEx\ indeed consists of reflecting ordinals...\ lemma ClEx_iff: "\y\Mset(a); Cl(a); ClEx(P,a); \a. \Cl(a); Ord(a)\ \ \x\Mset(a). P(x) \ Q(a,x)\ \ (\z. M(z) \ P(\y,z\)) \ (\z\Mset(a). Q(a,\y,z\))" -apply (unfold ClEx_def FF_def F0_def M_def) + unfolding ClEx_def FF_def F0_def M_def apply (rule ex_reflection.ZF_ClEx_iff [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro, of Mset Cl]) apply (simp_all add: Mset_mono_le Mset_cont Pair_in_Mset) done (*Alternative proof, less unfolding: apply (rule Reflection.ZF_ClEx_iff [of Mset _ _ Cl, folded M_def]) apply (fold ClEx_def FF_def F0_def) apply (rule ex_reflection.intro, assumption) apply (simp add: ex_reflection_axioms.intro, assumption+) *) lemma Closed_Unbounded_ClEx: "(\a. \Cl(a); Ord(a)\ \ \x\Mset(a). P(x) \ Q(a,x)) \ Closed_Unbounded(ClEx(P))" -apply (unfold ClEx_eq FF_def F0_def M_def) + unfolding ClEx_eq FF_def F0_def M_def apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl]) apply (rule ex_reflection.intro, rule reflection_axioms) apply (blast intro: ex_reflection_axioms.intro) done subsection\Packaging the Quantifier Reflection Rules\ lemma Ex_reflection_0: "Reflects(Cl,P0,Q0) \ Reflects(\a. Cl(a) \ ClEx(P0,a), \x. \z. M(z) \ P0(\x,z\), \a x. \z\Mset(a). Q0(a,\x,z\))" apply (simp add: Reflects_def) apply (intro conjI Closed_Unbounded_Int) apply blast apply (rule Closed_Unbounded_ClEx [of Cl P0 Q0], blast, clarify) apply (rule_tac Cl=Cl in ClEx_iff, assumption+, blast) done lemma All_reflection_0: "Reflects(Cl,P0,Q0) \ Reflects(\a. Cl(a) \ ClEx(\x.\P0(x), a), \x. \z. M(z) \ P0(\x,z\), \a x. \z\Mset(a). Q0(a,\x,z\))" apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) apply (rule Not_reflection, drule Not_reflection, simp) apply (erule Ex_reflection_0) done theorem Ex_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) \ Reflects(\a. Cl(a) \ ClEx(\x. P(fst(x),snd(x)), a), \x. \z. M(z) \ P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (rule Ex_reflection_0 [of _ " \x. P(fst(x),snd(x))" "\a x. Q(a,fst(x),snd(x))", simplified]) theorem All_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) \ Reflects(\a. Cl(a) \ ClEx(\x. \P(fst(x),snd(x)), a), \x. \z. M(z) \ P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (rule All_reflection_0 [of _ "\x. P(fst(x),snd(x))" "\a x. Q(a,fst(x),snd(x))", simplified]) text\And again, this time using class-bounded quantifiers\ theorem Rex_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) \ Reflects(\a. Cl(a) \ ClEx(\x. P(fst(x),snd(x)), a), \x. \z[M]. P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (unfold rex_def, blast) theorem Rall_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) \ Reflects(\a. Cl(a) \ ClEx(\x. \P(fst(x),snd(x)), a), \x. \z[M]. P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (unfold rall_def, blast) text\No point considering bounded quantifiers, where reflection is trivial.\ subsection\Simple Examples of Reflection\ text\Example 1: reflecting a simple formula. The reflecting class is first given as the variable \?Cl\ and later retrieved from the final proof state.\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ x \ y, \a x. \y\Mset(a). x \ y)" by fast text\Problem here: there needs to be a conjunction (class intersection) in the class of reflecting ordinals. The \<^term>\Ord(a)\ is redundant, though harmless.\ lemma "Reflects(\a. Ord(a) \ ClEx(\x. fst(x) \ snd(x), a), \x. \y. M(y) \ x \ y, \a x. \y\Mset(a). x \ y)" by fast text\Example 2\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ (\z. M(z) \ z \ x \ z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x \ z \ y)" by fast text\Example 2'. We give the reflecting class explicitly.\ lemma "Reflects (\a. (Ord(a) \ ClEx(\x. \ (snd(x) \ fst(fst(x)) \ snd(x) \ snd(fst(x))), a)) \ ClEx(\x. \z. M(z) \ z \ fst(x) \ z \ snd(x), a), \x. \y. M(y) \ (\z. M(z) \ z \ x \ z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x \ z \ y)" by fast text\Example 2''. We expand the subset relation.\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ (\z. M(z) \ (\w. M(w) \ w\z \ w\x) \ z\y), \a x. \y\Mset(a). \z\Mset(a). (\w\Mset(a). w\z \ w\x) \ z\y)" by fast text\Example 2'''. Single-step version, to reveal the reflecting class.\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ (\z. M(z) \ z \ x \ z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x \ z \ y)" apply (rule Ex_reflection) txt\ @{goals[display,indent=0,margin=60]} \ apply (rule All_reflection) txt\ @{goals[display,indent=0,margin=60]} \ apply (rule Triv_reflection) txt\ @{goals[display,indent=0,margin=60]} \ done text\Example 3. Warning: the following examples make sense only if \<^term>\P\ is quantifier-free, since it is not being relativized.\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ (\z. M(z) \ z \ y \ z \ x \ P(z)), \a x. \y\Mset(a). \z\Mset(a). z \ y \ z \ x \ P(z))" by fast text\Example 3'\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ y = Collect(x,P), \a x. \y\Mset(a). y = Collect(x,P))" by fast text\Example 3''\ schematic_goal "Reflects(?Cl, \x. \y. M(y) \ y = Replace(x,P), \a x. \y\Mset(a). y = Replace(x,P))" by fast text\Example 4: Axiom of Choice. Possibly wrong, since \\\ needs to be relativized.\ schematic_goal "Reflects(?Cl, \A. 0\A \ (\f. M(f) \ f \ (\X \ A. X)), \a A. 0\A \ (\f\Mset(a). f \ (\X \ A. X)))" by fast end end diff --git a/src/ZF/Constructible/Relative.thy b/src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy +++ b/src/ZF/Constructible/Relative.thy @@ -1,1671 +1,1671 @@ (* Title: ZF/Constructible/Relative.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory With modifications by E. Gunther, M. Pagano, and P. Sánchez Terraf *) section \Relativization and Absoluteness\ theory Relative imports ZF begin subsection\Relativized versions of standard set-theoretic concepts\ definition empty :: "[i\o,i] \ o" where "empty(M,z) \ \x[M]. x \ z" definition subset :: "[i\o,i,i] \ o" where "subset(M,A,B) \ \x[M]. x\A \ x \ B" definition upair :: "[i\o,i,i,i] \ o" where "upair(M,a,b,z) \ a \ z \ b \ z \ (\x[M]. x\z \ x = a \ x = b)" definition pair :: "[i\o,i,i,i] \ o" where "pair(M,a,b,z) \ \x[M]. upair(M,a,a,x) \ (\y[M]. upair(M,a,b,y) \ upair(M,x,y,z))" definition union :: "[i\o,i,i,i] \ o" where "union(M,a,b,z) \ \x[M]. x \ z \ x \ a \ x \ b" definition is_cons :: "[i\o,i,i,i] \ o" where "is_cons(M,a,b,z) \ \x[M]. upair(M,a,a,x) \ union(M,x,b,z)" definition successor :: "[i\o,i,i] \ o" where "successor(M,a,z) \ is_cons(M,a,a,z)" definition number1 :: "[i\o,i] \ o" where "number1(M,a) \ \x[M]. empty(M,x) \ successor(M,x,a)" definition number2 :: "[i\o,i] \ o" where "number2(M,a) \ \x[M]. number1(M,x) \ successor(M,x,a)" definition number3 :: "[i\o,i] \ o" where "number3(M,a) \ \x[M]. number2(M,x) \ successor(M,x,a)" definition powerset :: "[i\o,i,i] \ o" where "powerset(M,A,z) \ \x[M]. x \ z \ subset(M,x,A)" definition is_Collect :: "[i\o,i,i\o,i] \ o" where "is_Collect(M,A,P,z) \ \x[M]. x \ z \ x \ A \ P(x)" definition is_Replace :: "[i\o,i,[i,i]\o,i] \ o" where "is_Replace(M,A,P,z) \ \u[M]. u \ z \ (\x[M]. x\A \ P(x,u))" definition inter :: "[i\o,i,i,i] \ o" where "inter(M,a,b,z) \ \x[M]. x \ z \ x \ a \ x \ b" definition setdiff :: "[i\o,i,i,i] \ o" where "setdiff(M,a,b,z) \ \x[M]. x \ z \ x \ a \ x \ b" definition big_union :: "[i\o,i,i] \ o" where "big_union(M,A,z) \ \x[M]. x \ z \ (\y[M]. y\A \ x \ y)" definition big_inter :: "[i\o,i,i] \ o" where "big_inter(M,A,z) \ (A=0 \ z=0) \ (A\0 \ (\x[M]. x \ z \ (\y[M]. y\A \ x \ y)))" definition cartprod :: "[i\o,i,i,i] \ o" where "cartprod(M,A,B,z) \ \u[M]. u \ z \ (\x[M]. x\A \ (\y[M]. y\B \ pair(M,x,y,u)))" definition is_sum :: "[i\o,i,i,i] \ o" where "is_sum(M,A,B,Z) \ \A0[M]. \n1[M]. \s1[M]. \B1[M]. number1(M,n1) \ cartprod(M,n1,A,A0) \ upair(M,n1,n1,s1) \ cartprod(M,s1,B,B1) \ union(M,A0,B1,Z)" definition is_Inl :: "[i\o,i,i] \ o" where "is_Inl(M,a,z) \ \zero[M]. empty(M,zero) \ pair(M,zero,a,z)" definition is_Inr :: "[i\o,i,i] \ o" where "is_Inr(M,a,z) \ \n1[M]. number1(M,n1) \ pair(M,n1,a,z)" definition is_converse :: "[i\o,i,i] \ o" where "is_converse(M,r,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\u[M]. \v[M]. pair(M,u,v,w) \ pair(M,v,u,x)))" definition pre_image :: "[i\o,i,i,i] \ o" where "pre_image(M,r,A,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. y\A \ pair(M,x,y,w)))" definition is_domain :: "[i\o,i,i] \ o" where "is_domain(M,r,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. pair(M,x,y,w)))" definition image :: "[i\o,i,i,i] \ o" where "image(M,r,A,z) \ \y[M]. y \ z \ (\w[M]. w\r \ (\x[M]. x\A \ pair(M,x,y,w)))" definition is_range :: "[i\o,i,i] \ o" where \ \the cleaner \<^term>\\r'[M]. is_converse(M,r,r') \ is_domain(M,r',z)\ unfortunately needs an instance of separation in order to prove \<^term>\M(converse(r))\.\ "is_range(M,r,z) \ \y[M]. y \ z \ (\w[M]. w\r \ (\x[M]. pair(M,x,y,w)))" definition is_field :: "[i\o,i,i] \ o" where "is_field(M,r,z) \ \dr[M]. \rr[M]. is_domain(M,r,dr) \ is_range(M,r,rr) \ union(M,dr,rr,z)" definition is_relation :: "[i\o,i] \ o" where "is_relation(M,r) \ (\z[M]. z\r \ (\x[M]. \y[M]. pair(M,x,y,z)))" definition is_function :: "[i\o,i] \ o" where "is_function(M,r) \ \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x,y',p') \ p\r \ p'\r \ y=y'" definition fun_apply :: "[i\o,i,i,i] \ o" where "fun_apply(M,f,x,y) \ (\xs[M]. \fxs[M]. upair(M,x,x,xs) \ image(M,f,xs,fxs) \ big_union(M,fxs,y))" definition typed_function :: "[i\o,i,i,i] \ o" where "typed_function(M,A,B,r) \ is_function(M,r) \ is_relation(M,r) \ is_domain(M,r,A) \ (\u[M]. u\r \ (\x[M]. \y[M]. pair(M,x,y,u) \ y\B))" definition is_funspace :: "[i\o,i,i,i] \ o" where "is_funspace(M,A,B,F) \ \f[M]. f \ F \ typed_function(M,A,B,f)" definition composition :: "[i\o,i,i,i] \ o" where "composition(M,r,s,t) \ \p[M]. p \ t \ (\x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. pair(M,x,z,p) \ pair(M,x,y,xy) \ pair(M,y,z,yz) \ xy \ s \ yz \ r)" definition injection :: "[i\o,i,i,i] \ o" where "injection(M,A,B,f) \ typed_function(M,A,B,f) \ (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x',y,p') \ p\f \ p'\f \ x=x')" definition surjection :: "[i\o,i,i,i] \ o" where "surjection(M,A,B,f) \ typed_function(M,A,B,f) \ (\y[M]. y\B \ (\x[M]. x\A \ fun_apply(M,f,x,y)))" definition bijection :: "[i\o,i,i,i] \ o" where "bijection(M,A,B,f) \ injection(M,A,B,f) \ surjection(M,A,B,f)" definition restriction :: "[i\o,i,i,i] \ o" where "restriction(M,r,A,z) \ \x[M]. x \ z \ (x \ r \ (\u[M]. u\A \ (\v[M]. pair(M,u,v,x))))" definition transitive_set :: "[i\o,i] \ o" where "transitive_set(M,a) \ \x[M]. x\a \ subset(M,x,a)" definition ordinal :: "[i\o,i] \ o" where \ \an ordinal is a transitive set of transitive sets\ "ordinal(M,a) \ transitive_set(M,a) \ (\x[M]. x\a \ transitive_set(M,x))" definition limit_ordinal :: "[i\o,i] \ o" where \ \a limit ordinal is a non-empty, successor-closed ordinal\ "limit_ordinal(M,a) \ ordinal(M,a) \ \ empty(M,a) \ (\x[M]. x\a \ (\y[M]. y\a \ successor(M,x,y)))" definition successor_ordinal :: "[i\o,i] \ o" where \ \a successor ordinal is any ordinal that is neither empty nor limit\ "successor_ordinal(M,a) \ ordinal(M,a) \ \ empty(M,a) \ \ limit_ordinal(M,a)" definition finite_ordinal :: "[i\o,i] \ o" where \ \an ordinal is finite if neither it nor any of its elements are limit\ "finite_ordinal(M,a) \ ordinal(M,a) \ \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x))" definition omega :: "[i\o,i] \ o" where \ \omega is a limit ordinal none of whose elements are limit\ "omega(M,a) \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x))" definition is_quasinat :: "[i\o,i] \ o" where "is_quasinat(M,z) \ empty(M,z) \ (\m[M]. successor(M,m,z))" definition is_nat_case :: "[i\o, i, [i,i]\o, i, i] \ o" where "is_nat_case(M, a, is_b, k, z) \ (empty(M,k) \ z=a) \ (\m[M]. successor(M,m,k) \ is_b(m,z)) \ (is_quasinat(M,k) \ empty(M,z))" definition relation1 :: "[i\o, [i,i]\o, i\i] \ o" where "relation1(M,is_f,f) \ \x[M]. \y[M]. is_f(x,y) \ y = f(x)" definition Relation1 :: "[i\o, i, [i,i]\o, i\i] \ o" where \ \as above, but typed\ "Relation1(M,A,is_f,f) \ \x[M]. \y[M]. x\A \ is_f(x,y) \ y = f(x)" definition relation2 :: "[i\o, [i,i,i]\o, [i,i]\i] \ o" where "relation2(M,is_f,f) \ \x[M]. \y[M]. \z[M]. is_f(x,y,z) \ z = f(x,y)" definition Relation2 :: "[i\o, i, i, [i,i,i]\o, [i,i]\i] \ o" where "Relation2(M,A,B,is_f,f) \ \x[M]. \y[M]. \z[M]. x\A \ y\B \ is_f(x,y,z) \ z = f(x,y)" definition relation3 :: "[i\o, [i,i,i,i]\o, [i,i,i]\i] \ o" where "relation3(M,is_f,f) \ \x[M]. \y[M]. \z[M]. \u[M]. is_f(x,y,z,u) \ u = f(x,y,z)" definition Relation3 :: "[i\o, i, i, i, [i,i,i,i]\o, [i,i,i]\i] \ o" where "Relation3(M,A,B,C,is_f,f) \ \x[M]. \y[M]. \z[M]. \u[M]. x\A \ y\B \ z\C \ is_f(x,y,z,u) \ u = f(x,y,z)" definition relation4 :: "[i\o, [i,i,i,i,i]\o, [i,i,i,i]\i] \ o" where "relation4(M,is_f,f) \ \u[M]. \x[M]. \y[M]. \z[M]. \a[M]. is_f(u,x,y,z,a) \ a = f(u,x,y,z)" text\Useful when absoluteness reasoning has replaced the predicates by terms\ lemma triv_Relation1: "Relation1(M, A, \x y. y = f(x), f)" by (simp add: Relation1_def) lemma triv_Relation2: "Relation2(M, A, B, \x y a. a = f(x,y), f)" by (simp add: Relation2_def) subsection \The relativized ZF axioms\ definition extensionality :: "(i\o) \ o" where "extensionality(M) \ \x[M]. \y[M]. (\z[M]. z \ x \ z \ y) \ x=y" definition separation :: "[i\o, i\o] \ o" where \ \The formula \P\ should only involve parameters belonging to \M\ and all its quantifiers must be relativized to \M\. We do not have separation as a scheme; every instance that we need must be assumed (and later proved) separately.\ "separation(M,P) \ \z[M]. \y[M]. \x[M]. x \ y \ x \ z \ P(x)" definition upair_ax :: "(i\o) \ o" where "upair_ax(M) \ \x[M]. \y[M]. \z[M]. upair(M,x,y,z)" definition Union_ax :: "(i\o) \ o" where "Union_ax(M) \ \x[M]. \z[M]. big_union(M,x,z)" definition power_ax :: "(i\o) \ o" where "power_ax(M) \ \x[M]. \z[M]. powerset(M,x,z)" definition univalent :: "[i\o, i, [i,i]\o] \ o" where "univalent(M,A,P) \ \x[M]. x\A \ (\y[M]. \z[M]. P(x,y) \ P(x,z) \ y=z)" definition replacement :: "[i\o, [i,i]\o] \ o" where "replacement(M,P) \ \A[M]. univalent(M,A,P) \ (\Y[M]. \b[M]. (\x[M]. x\A \ P(x,b)) \ b \ Y)" definition strong_replacement :: "[i\o, [i,i]\o] \ o" where "strong_replacement(M,P) \ \A[M]. univalent(M,A,P) \ (\Y[M]. \b[M]. b \ Y \ (\x[M]. x\A \ P(x,b)))" definition foundation_ax :: "(i\o) \ o" where "foundation_ax(M) \ \x[M]. (\y[M]. y\x) \ (\y[M]. y\x \ \(\z[M]. z\x \ z \ y))" subsection\A trivial consistency proof for $V_\omega$\ text\We prove that $V_\omega$ (or \univ\ in Isabelle) satisfies some ZF axioms. Kunen, Theorem IV 3.13, page 123.\ lemma univ0_downwards_mem: "\y \ x; x \ univ(0)\ \ y \ univ(0)" apply (insert Transset_univ [OF Transset_0]) apply (simp add: Transset_def, blast) done lemma univ0_Ball_abs [simp]: "A \ univ(0) \ (\x\A. x \ univ(0) \ P(x)) \ (\x\A. P(x))" by (blast intro: univ0_downwards_mem) lemma univ0_Bex_abs [simp]: "A \ univ(0) \ (\x\A. x \ univ(0) \ P(x)) \ (\x\A. P(x))" by (blast intro: univ0_downwards_mem) text\Congruence rule for separation: can assume the variable is in \M\\ lemma separation_cong [cong]: "(\x. M(x) \ P(x) \ P'(x)) \ separation(M, \x. P(x)) \ separation(M, \x. P'(x))" by (simp add: separation_def) lemma univalent_cong [cong]: "\A=A'; \x y. \x\A; M(x); M(y)\ \ P(x,y) \ P'(x,y)\ \ univalent(M, A, \x y. P(x,y)) \ univalent(M, A', \x y. P'(x,y))" by (simp add: univalent_def) lemma univalent_triv [intro,simp]: "univalent(M, A, \x y. y = f(x))" by (simp add: univalent_def) lemma univalent_conjI2 [intro,simp]: "univalent(M,A,Q) \ univalent(M, A, \x y. P(x,y) \ Q(x,y))" by (simp add: univalent_def, blast) text\Congruence rule for replacement\ lemma strong_replacement_cong [cong]: "\\x y. \M(x); M(y)\ \ P(x,y) \ P'(x,y)\ \ strong_replacement(M, \x y. P(x,y)) \ strong_replacement(M, \x y. P'(x,y))" by (simp add: strong_replacement_def) text\The extensionality axiom\ lemma "extensionality(\x. x \ univ(0))" apply (simp add: extensionality_def) apply (blast intro: univ0_downwards_mem) done text\The separation axiom requires some lemmas\ lemma Collect_in_Vfrom: "\X \ Vfrom(A,j); Transset(A)\ \ Collect(X,P) \ Vfrom(A, succ(j))" apply (drule Transset_Vfrom) apply (rule subset_mem_Vfrom) apply (unfold Transset_def, blast) done lemma Collect_in_VLimit: "\X \ Vfrom(A,i); Limit(i); Transset(A)\ \ Collect(X,P) \ Vfrom(A,i)" apply (rule Limit_VfromE, assumption+) apply (blast intro: Limit_has_succ VfromI Collect_in_Vfrom) done lemma Collect_in_univ: "\X \ univ(A); Transset(A)\ \ Collect(X,P) \ univ(A)" by (simp add: univ_def Collect_in_VLimit) lemma "separation(\x. x \ univ(0), P)" apply (simp add: separation_def, clarify) apply (rule_tac x = "Collect(z,P)" in bexI) apply (blast intro: Collect_in_univ Transset_0)+ done text\Unordered pairing axiom\ lemma "upair_ax(\x. x \ univ(0))" apply (simp add: upair_ax_def upair_def) apply (blast intro: doubleton_in_univ) done text\Union axiom\ lemma "Union_ax(\x. x \ univ(0))" apply (simp add: Union_ax_def big_union_def, clarify) apply (rule_tac x="\x" in bexI) apply (blast intro: univ0_downwards_mem) apply (blast intro: Union_in_univ Transset_0) done text\Powerset axiom\ lemma Pow_in_univ: "\X \ univ(A); Transset(A)\ \ Pow(X) \ univ(A)" apply (simp add: univ_def Pow_in_VLimit) done lemma "power_ax(\x. x \ univ(0))" apply (simp add: power_ax_def powerset_def subset_def, clarify) apply (rule_tac x="Pow(x)" in bexI) apply (blast intro: univ0_downwards_mem) apply (blast intro: Pow_in_univ Transset_0) done text\Foundation axiom\ lemma "foundation_ax(\x. x \ univ(0))" apply (simp add: foundation_ax_def, clarify) apply (cut_tac A=x in foundation) apply (blast intro: univ0_downwards_mem) done lemma "replacement(\x. x \ univ(0), P)" apply (simp add: replacement_def, clarify) oops text\no idea: maybe prove by induction on the rank of A?\ text\Still missing: Replacement, Choice\ subsection\Lemmas Needed to Reduce Some Set Constructions to Instances of Separation\ lemma image_iff_Collect: "r `` A = {y \ \(\(r)). \p\r. \x\A. p=\x,y\}" apply (rule equalityI, auto) apply (simp add: Pair_def, blast) done lemma vimage_iff_Collect: "r -`` A = {x \ \(\(r)). \p\r. \y\A. p=\x,y\}" apply (rule equalityI, auto) apply (simp add: Pair_def, blast) done text\These two lemmas lets us prove \domain_closed\ and \range_closed\ without new instances of separation\ lemma domain_eq_vimage: "domain(r) = r -`` Union(Union(r))" apply (rule equalityI, auto) apply (rule vimageI, assumption) apply (simp add: Pair_def, blast) done lemma range_eq_image: "range(r) = r `` Union(Union(r))" apply (rule equalityI, auto) apply (rule imageI, assumption) apply (simp add: Pair_def, blast) done lemma replacementD: "\replacement(M,P); M(A); univalent(M,A,P)\ \ \Y[M]. (\b[M]. ((\x[M]. x\A \ P(x,b)) \ b \ Y))" by (simp add: replacement_def) lemma strong_replacementD: "\strong_replacement(M,P); M(A); univalent(M,A,P)\ \ \Y[M]. (\b[M]. (b \ Y \ (\x[M]. x\A \ P(x,b))))" by (simp add: strong_replacement_def) lemma separationD: "\separation(M,P); M(z)\ \ \y[M]. \x[M]. x \ y \ x \ z \ P(x)" by (simp add: separation_def) text\More constants, for order types\ definition order_isomorphism :: "[i\o,i,i,i,i,i] \ o" where "order_isomorphism(M,A,r,B,s,f) \ bijection(M,A,B,f) \ (\x[M]. x\A \ (\y[M]. y\A \ (\p[M]. \fx[M]. \fy[M]. \q[M]. pair(M,x,y,p) \ fun_apply(M,f,x,fx) \ fun_apply(M,f,y,fy) \ pair(M,fx,fy,q) \ (p\r \ q\s))))" definition pred_set :: "[i\o,i,i,i,i] \ o" where "pred_set(M,A,x,r,B) \ \y[M]. y \ B \ (\p[M]. p\r \ y \ A \ pair(M,y,x,p))" definition membership :: "[i\o,i,i] \ o" where \ \membership relation\ "membership(M,A,r) \ \p[M]. p \ r \ (\x[M]. x\A \ (\y[M]. y\A \ x\y \ pair(M,x,y,p)))" subsection\Introducing a Transitive Class Model\ text\The class M is assumed to be transitive and inhabited\ locale M_trans = fixes M assumes transM: "\y\x; M(x)\ \ M(y)" and M_inhabited: "\x . M(x)" lemma (in M_trans) nonempty [simp]: "M(0)" proof - have "M(x) \ M(0)" for x proof (rule_tac P="\w. M(w) \ M(0)" in eps_induct) { fix x assume "\y\x. M(y) \ M(0)" "M(x)" consider (a) "\y. y\x" | (b) "x=0" by auto then have "M(x) \ M(0)" proof cases case a then show ?thesis using \\y\x._\ \M(x)\ transM by auto next case b then show ?thesis by simp qed } then show "M(x) \ M(0)" if "\y\x. M(y) \ M(0)" for x using that by auto qed with M_inhabited show "M(0)" using M_inhabited by blast qed text\The class M is assumed to be transitive and to satisfy some relativized ZF axioms\ locale M_trivial = M_trans + assumes upair_ax: "upair_ax(M)" and Union_ax: "Union_ax(M)" lemma (in M_trans) rall_abs [simp]: "M(A) \ (\x[M]. x\A \ P(x)) \ (\x\A. P(x))" by (blast intro: transM) lemma (in M_trans) rex_abs [simp]: "M(A) \ (\x[M]. x\A \ P(x)) \ (\x\A. P(x))" by (blast intro: transM) lemma (in M_trans) ball_iff_equiv: "M(A) \ (\x[M]. (x\A \ P(x))) \ (\x\A. P(x)) \ (\x. P(x) \ M(x) \ x\A)" by (blast intro: transM) text\Simplifies proofs of equalities when there's an iff-equality available for rewriting, universally quantified over M. But it's not the only way to prove such equalities: its premises \<^term>\M(A)\ and \<^term>\M(B)\ can be too strong.\ lemma (in M_trans) M_equalityI: "\\x. M(x) \ x\A \ x\B; M(A); M(B)\ \ A=B" by (blast dest: transM) subsubsection\Trivial Absoluteness Proofs: Empty Set, Pairs, etc.\ lemma (in M_trans) empty_abs [simp]: "M(z) \ empty(M,z) \ z=0" apply (simp add: empty_def) apply (blast intro: transM) done lemma (in M_trans) subset_abs [simp]: "M(A) \ subset(M,A,B) \ A \ B" apply (simp add: subset_def) apply (blast intro: transM) done lemma (in M_trans) upair_abs [simp]: "M(z) \ upair(M,a,b,z) \ z={a,b}" apply (simp add: upair_def) apply (blast intro: transM) done lemma (in M_trivial) upair_in_MI [intro!]: " M(a) \ M(b) \ M({a,b})" by (insert upair_ax, simp add: upair_ax_def) lemma (in M_trans) upair_in_MD [dest!]: "M({a,b}) \ M(a) \ M(b)" by (blast intro: transM) lemma (in M_trivial) upair_in_M_iff [simp]: "M({a,b}) \ M(a) \ M(b)" by blast lemma (in M_trivial) singleton_in_MI [intro!]: "M(a) \ M({a})" by (insert upair_in_M_iff [of a a], simp) lemma (in M_trans) singleton_in_MD [dest!]: "M({a}) \ M(a)" by (insert upair_in_MD [of a a], simp) lemma (in M_trivial) singleton_in_M_iff [simp]: "M({a}) \ M(a)" by blast lemma (in M_trans) pair_abs [simp]: "M(z) \ pair(M,a,b,z) \ z=\a,b\" apply (simp add: pair_def Pair_def) apply (blast intro: transM) done lemma (in M_trans) pair_in_MD [dest!]: "M(\a,b\) \ M(a) \ M(b)" by (simp add: Pair_def, blast intro: transM) lemma (in M_trivial) pair_in_MI [intro!]: "M(a) \ M(b) \ M(\a,b\)" by (simp add: Pair_def) lemma (in M_trivial) pair_in_M_iff [simp]: "M(\a,b\) \ M(a) \ M(b)" by blast lemma (in M_trans) pair_components_in_M: "\\x,y\ \ A; M(A)\ \ M(x) \ M(y)" by (blast dest: transM) lemma (in M_trivial) cartprod_abs [simp]: "\M(A); M(B); M(z)\ \ cartprod(M,A,B,z) \ z = A*B" apply (simp add: cartprod_def) apply (rule iffI) apply (blast intro!: equalityI intro: transM dest!: rspec) apply (blast dest: transM) done subsubsection\Absoluteness for Unions and Intersections\ lemma (in M_trans) union_abs [simp]: "\M(a); M(b); M(z)\ \ union(M,a,b,z) \ z = a \ b" unfolding union_def by (blast intro: transM ) lemma (in M_trans) inter_abs [simp]: "\M(a); M(b); M(z)\ \ inter(M,a,b,z) \ z = a \ b" unfolding inter_def by (blast intro: transM) lemma (in M_trans) setdiff_abs [simp]: "\M(a); M(b); M(z)\ \ setdiff(M,a,b,z) \ z = a-b" unfolding setdiff_def by (blast intro: transM) lemma (in M_trans) Union_abs [simp]: "\M(A); M(z)\ \ big_union(M,A,z) \ z = \(A)" unfolding big_union_def by (blast dest: transM) lemma (in M_trivial) Union_closed [intro,simp]: "M(A) \ M(\(A))" by (insert Union_ax, simp add: Union_ax_def) lemma (in M_trivial) Un_closed [intro,simp]: "\M(A); M(B)\ \ M(A \ B)" by (simp only: Un_eq_Union, blast) lemma (in M_trivial) cons_closed [intro,simp]: "\M(a); M(A)\ \ M(cons(a,A))" by (subst cons_eq [symmetric], blast) lemma (in M_trivial) cons_abs [simp]: "\M(b); M(z)\ \ is_cons(M,a,b,z) \ z = cons(a,b)" by (simp add: is_cons_def, blast intro: transM) lemma (in M_trivial) successor_abs [simp]: "\M(a); M(z)\ \ successor(M,a,z) \ z = succ(a)" by (simp add: successor_def, blast) lemma (in M_trans) succ_in_MD [dest!]: "M(succ(a)) \ M(a)" unfolding succ_def by (blast intro: transM) lemma (in M_trivial) succ_in_MI [intro!]: "M(a) \ M(succ(a))" unfolding succ_def by (blast intro: transM) lemma (in M_trivial) succ_in_M_iff [simp]: "M(succ(a)) \ M(a)" by blast subsubsection\Absoluteness for Separation and Replacement\ lemma (in M_trans) separation_closed [intro,simp]: "\separation(M,P); M(A)\ \ M(Collect(A,P))" apply (insert separation, simp add: separation_def) apply (drule rspec, assumption, clarify) apply (subgoal_tac "y = Collect(A,P)", blast) apply (blast dest: transM) done lemma separation_iff: "separation(M,P) \ (\z[M]. \y[M]. is_Collect(M,z,P,y))" by (simp add: separation_def is_Collect_def) lemma (in M_trans) Collect_abs [simp]: "\M(A); M(z)\ \ is_Collect(M,A,P,z) \ z = Collect(A,P)" unfolding is_Collect_def by (blast dest: transM) subsubsection\The Operator \<^term>\is_Replace\\ lemma is_Replace_cong [cong]: "\A=A'; \x y. \M(x); M(y)\ \ P(x,y) \ P'(x,y); z=z'\ \ is_Replace(M, A, \x y. P(x,y), z) \ is_Replace(M, A', \x y. P'(x,y), z')" by (simp add: is_Replace_def) lemma (in M_trans) univalent_Replace_iff: "\M(A); univalent(M,A,P); \x y. \x\A; P(x,y)\ \ M(y)\ \ u \ Replace(A,P) \ (\x. x\A \ P(x,u))" unfolding Replace_iff univalent_def by (blast dest: transM) (*The last premise expresses that P takes M to M*) lemma (in M_trans) strong_replacement_closed [intro,simp]: "\strong_replacement(M,P); M(A); univalent(M,A,P); \x y. \x\A; P(x,y)\ \ M(y)\ \ M(Replace(A,P))" apply (simp add: strong_replacement_def) apply (drule_tac x=A in rspec, safe) apply (subgoal_tac "Replace(A,P) = Y") apply simp apply (rule equality_iffI) apply (simp add: univalent_Replace_iff) apply (blast dest: transM) done lemma (in M_trans) Replace_abs: "\M(A); M(z); univalent(M,A,P); \x y. \x\A; P(x,y)\ \ M(y)\ \ is_Replace(M,A,P,z) \ z = Replace(A,P)" apply (simp add: is_Replace_def) apply (rule iffI) apply (rule equality_iffI) apply (simp_all add: univalent_Replace_iff) apply (blast dest: transM)+ done (*The first premise can't simply be assumed as a schema. It is essential to take care when asserting instances of Replacement. Let K be a nonconstructible subset of nat and define f(x) = x if x \ K and f(x)=0 otherwise. Then RepFun(nat,f) = cons(0,K), a nonconstructible set. So we cannot assume that M(X) implies M(RepFun(X,f)) even for f \ M -> M. *) lemma (in M_trans) RepFun_closed: "\strong_replacement(M, \x y. y = f(x)); M(A); \x\A. M(f(x))\ \ M(RepFun(A,f))" apply (simp add: RepFun_def) done lemma Replace_conj_eq: "{y . x \ A, x\A \ y=f(x)} = {y . x\A, y=f(x)}" by simp text\Better than \RepFun_closed\ when having the formula \<^term>\x\A\ makes relativization easier.\ lemma (in M_trans) RepFun_closed2: "\strong_replacement(M, \x y. x\A \ y = f(x)); M(A); \x\A. M(f(x))\ \ M(RepFun(A, \x. f(x)))" apply (simp add: RepFun_def) apply (frule strong_replacement_closed, assumption) apply (auto dest: transM simp add: Replace_conj_eq univalent_def) done subsubsection \Absoluteness for \<^term>\Lambda\\ definition is_lambda :: "[i\o, i, [i,i]\o, i] \ o" where "is_lambda(M, A, is_b, z) \ \p[M]. p \ z \ (\u[M]. \v[M]. u\A \ pair(M,u,v,p) \ is_b(u,v))" lemma (in M_trivial) lam_closed: "\strong_replacement(M, \x y. y = ); M(A); \x\A. M(b(x))\ \ M(\x\A. b(x))" by (simp add: lam_def, blast intro: RepFun_closed dest: transM) text\Better than \lam_closed\: has the formula \<^term>\x\A\\ lemma (in M_trivial) lam_closed2: "\strong_replacement(M, \x y. x\A \ y = \x, b(x)\); M(A); \m[M]. m\A \ M(b(m))\ \ M(Lambda(A,b))" apply (simp add: lam_def) apply (blast intro: RepFun_closed2 dest: transM) done lemma (in M_trivial) lambda_abs2: "\Relation1(M,A,is_b,b); M(A); \m[M]. m\A \ M(b(m)); M(z)\ \ is_lambda(M,A,is_b,z) \ z = Lambda(A,b)" apply (simp add: Relation1_def is_lambda_def) apply (rule iffI) prefer 2 apply (simp add: lam_def) apply (rule equality_iffI) apply (simp add: lam_def) apply (rule iffI) apply (blast dest: transM) apply (auto simp add: transM [of _ A]) done lemma is_lambda_cong [cong]: "\A=A'; z=z'; \x y. \x\A; M(x); M(y)\ \ is_b(x,y) \ is_b'(x,y)\ \ is_lambda(M, A, \x y. is_b(x,y), z) \ is_lambda(M, A', \x y. is_b'(x,y), z')" by (simp add: is_lambda_def) lemma (in M_trans) image_abs [simp]: "\M(r); M(A); M(z)\ \ image(M,r,A,z) \ z = r``A" apply (simp add: image_def) apply (rule iffI) apply (blast intro!: equalityI dest: transM, blast) done subsubsection\Relativization of Powerset\ text\What about \Pow_abs\? Powerset is NOT absolute! This result is one direction of absoluteness.\ lemma (in M_trans) powerset_Pow: "powerset(M, x, Pow(x))" by (simp add: powerset_def) text\But we can't prove that the powerset in \M\ includes the real powerset.\ lemma (in M_trans) powerset_imp_subset_Pow: "\powerset(M,x,y); M(y)\ \ y \ Pow(x)" apply (simp add: powerset_def) apply (blast dest: transM) done lemma (in M_trans) powerset_abs: assumes "M(y)" shows "powerset(M,x,y) \ y = {a\Pow(x) . M(a)}" proof (intro iffI equalityI) (* First show the converse implication by double inclusion *) assume "powerset(M,x,y)" with \M(y)\ show "y \ {a\Pow(x) . M(a)}" using powerset_imp_subset_Pow transM by blast from \powerset(M,x,y)\ show "{a\Pow(x) . M(a)} \ y" using transM unfolding powerset_def by auto next (* we show the direct implication *) assume "y = {a \ Pow(x) . M(a)}" then show "powerset(M, x, y)" unfolding powerset_def subset_def using transM by blast qed subsubsection\Absoluteness for the Natural Numbers\ lemma (in M_trivial) nat_into_M [intro]: "n \ nat \ M(n)" by (induct n rule: nat_induct, simp_all) lemma (in M_trans) nat_case_closed [intro,simp]: "\M(k); M(a); \m[M]. M(b(m))\ \ M(nat_case(a,b,k))" apply (case_tac "k=0", simp) apply (case_tac "\m. k = succ(m)", force) apply (simp add: nat_case_def) done lemma (in M_trivial) quasinat_abs [simp]: "M(z) \ is_quasinat(M,z) \ quasinat(z)" by (auto simp add: is_quasinat_def quasinat_def) lemma (in M_trivial) nat_case_abs [simp]: "\relation1(M,is_b,b); M(k); M(z)\ \ is_nat_case(M,a,is_b,k,z) \ z = nat_case(a,b,k)" apply (case_tac "quasinat(k)") prefer 2 apply (simp add: is_nat_case_def non_nat_case) apply (force simp add: quasinat_def) apply (simp add: quasinat_def is_nat_case_def) apply (elim disjE exE) apply (simp_all add: relation1_def) done (*NOT for the simplifier. The assumption M(z') is apparently necessary, but causes the error "Failed congruence proof!" It may be better to replace is_nat_case by nat_case before attempting congruence reasoning.*) lemma is_nat_case_cong: "\a = a'; k = k'; z = z'; M(z'); \x y. \M(x); M(y)\ \ is_b(x,y) \ is_b'(x,y)\ \ is_nat_case(M, a, is_b, k, z) \ is_nat_case(M, a', is_b', k', z')" by (simp add: is_nat_case_def) subsection\Absoluteness for Ordinals\ text\These results constitute Theorem IV 5.1 of Kunen (page 126).\ lemma (in M_trans) lt_closed: "\j \ M(j)" by (blast dest: ltD intro: transM) lemma (in M_trans) transitive_set_abs [simp]: "M(a) \ transitive_set(M,a) \ Transset(a)" by (simp add: transitive_set_def Transset_def) lemma (in M_trans) ordinal_abs [simp]: "M(a) \ ordinal(M,a) \ Ord(a)" by (simp add: ordinal_def Ord_def) lemma (in M_trivial) limit_ordinal_abs [simp]: "M(a) \ limit_ordinal(M,a) \ Limit(a)" -apply (unfold Limit_def limit_ordinal_def) + unfolding Limit_def limit_ordinal_def apply (simp add: Ord_0_lt_iff) apply (simp add: lt_def, blast) done lemma (in M_trivial) successor_ordinal_abs [simp]: "M(a) \ successor_ordinal(M,a) \ Ord(a) \ (\b[M]. a = succ(b))" apply (simp add: successor_ordinal_def, safe) apply (drule Ord_cases_disj, auto) done lemma finite_Ord_is_nat: "\Ord(a); \ Limit(a); \x\a. \ Limit(x)\ \ a \ nat" by (induct a rule: trans_induct3, simp_all) lemma (in M_trivial) finite_ordinal_abs [simp]: "M(a) \ finite_ordinal(M,a) \ a \ nat" apply (simp add: finite_ordinal_def) apply (blast intro: finite_Ord_is_nat intro: nat_into_Ord dest: Ord_trans naturals_not_limit) done lemma Limit_non_Limit_implies_nat: "\Limit(a); \x\a. \ Limit(x)\ \ a = nat" apply (rule le_anti_sym) apply (rule all_lt_imp_le, blast, blast intro: Limit_is_Ord) apply (simp add: lt_def) apply (blast intro: Ord_in_Ord Ord_trans finite_Ord_is_nat) apply (erule nat_le_Limit) done lemma (in M_trivial) omega_abs [simp]: "M(a) \ omega(M,a) \ a = nat" apply (simp add: omega_def) apply (blast intro: Limit_non_Limit_implies_nat dest: naturals_not_limit) done lemma (in M_trivial) number1_abs [simp]: "M(a) \ number1(M,a) \ a = 1" by (simp add: number1_def) lemma (in M_trivial) number2_abs [simp]: "M(a) \ number2(M,a) \ a = succ(1)" by (simp add: number2_def) lemma (in M_trivial) number3_abs [simp]: "M(a) \ number3(M,a) \ a = succ(succ(1))" by (simp add: number3_def) text\Kunen continued to 20...\ (*Could not get this to work. The \x\nat is essential because everything but the recursion variable must stay unchanged. But then the recursion equations only hold for x\nat (or in some other set) and not for the whole of the class M. consts natnumber_aux :: "[i\o,i] \ i" primrec "natnumber_aux(M,0) = (\x\nat. if empty(M,x) then 1 else 0)" "natnumber_aux(M,succ(n)) = (\x\nat. if (\y[M]. natnumber_aux(M,n)`y=1 \ successor(M,y,x)) then 1 else 0)" definition natnumber :: "[i\o,i,i] \ o" "natnumber(M,n,x) \ natnumber_aux(M,n)`x = 1" lemma (in M_trivial) [simp]: "natnumber(M,0,x) \ x=0" *) subsection\Some instances of separation and strong replacement\ locale M_basic = M_trivial + assumes Inter_separation: "M(A) \ separation(M, \x. \y[M]. y\A \ x\y)" and Diff_separation: "M(B) \ separation(M, \x. x \ B)" and cartprod_separation: "\M(A); M(B)\ \ separation(M, \z. \x[M]. x\A \ (\y[M]. y\B \ pair(M,x,y,z)))" and image_separation: "\M(A); M(r)\ \ separation(M, \y. \p[M]. p\r \ (\x[M]. x\A \ pair(M,x,y,p)))" and converse_separation: "M(r) \ separation(M, \z. \p[M]. p\r \ (\x[M]. \y[M]. pair(M,x,y,p) \ pair(M,y,x,z)))" and restrict_separation: "M(A) \ separation(M, \z. \x[M]. x\A \ (\y[M]. pair(M,x,y,z)))" and comp_separation: "\M(r); M(s)\ \ separation(M, \xz. \x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. pair(M,x,z,xz) \ pair(M,x,y,xy) \ pair(M,y,z,yz) \ xy\s \ yz\r)" and pred_separation: "\M(r); M(x)\ \ separation(M, \y. \p[M]. p\r \ pair(M,y,x,p))" and Memrel_separation: "separation(M, \z. \x[M]. \y[M]. pair(M,x,y,z) \ x \ y)" and funspace_succ_replacement: "M(n) \ strong_replacement(M, \p z. \f[M]. \b[M]. \nb[M]. \cnbf[M]. pair(M,f,b,p) \ pair(M,n,b,nb) \ is_cons(M,nb,f,cnbf) \ upair(M,cnbf,cnbf,z))" and is_recfun_separation: \ \for well-founded recursion: used to prove \is_recfun_equal\\ "\M(r); M(f); M(g); M(a); M(b)\ \ separation(M, \x. \xa[M]. \xb[M]. pair(M,x,a,xa) \ xa \ r \ pair(M,x,b,xb) \ xb \ r \ (\fx[M]. \gx[M]. fun_apply(M,f,x,fx) \ fun_apply(M,g,x,gx) \ fx \ gx))" and power_ax: "power_ax(M)" lemma (in M_trivial) cartprod_iff_lemma: "\M(C); \u[M]. u \ C \ (\x\A. \y\B. u = {{x}, {x,y}}); powerset(M, A \ B, p1); powerset(M, p1, p2); M(p2)\ \ C = {u \ p2 . \x\A. \y\B. u = {{x}, {x,y}}}" apply (simp add: powerset_def) apply (rule equalityI, clarify, simp) apply (frule transM, assumption) apply (frule transM, assumption, simp (no_asm_simp)) apply blast apply clarify apply (frule transM, assumption, force) done lemma (in M_basic) cartprod_iff: "\M(A); M(B); M(C)\ \ cartprod(M,A,B,C) \ (\p1[M]. \p2[M]. powerset(M,A \ B,p1) \ powerset(M,p1,p2) \ C = {z \ p2. \x\A. \y\B. z = \x,y\})" apply (simp add: Pair_def cartprod_def, safe) defer 1 apply (simp add: powerset_def) apply blast txt\Final, difficult case: the left-to-right direction of the theorem.\ apply (insert power_ax, simp add: power_ax_def) apply (frule_tac x="A \ B" and P="\x. rex(M,Q(x))" for Q in rspec) apply (blast, clarify) apply (drule_tac x=z and P="\x. rex(M,Q(x))" for Q in rspec) apply assumption apply (blast intro: cartprod_iff_lemma) done lemma (in M_basic) cartprod_closed_lemma: "\M(A); M(B)\ \ \C[M]. cartprod(M,A,B,C)" apply (simp del: cartprod_abs add: cartprod_iff) apply (insert power_ax, simp add: power_ax_def) apply (frule_tac x="A \ B" and P="\x. rex(M,Q(x))" for Q in rspec) apply (blast, clarify) apply (drule_tac x=z and P="\x. rex(M,Q(x))" for Q in rspec, auto) apply (intro rexI conjI, simp+) apply (insert cartprod_separation [of A B], simp) done text\All the lemmas above are necessary because Powerset is not absolute. I should have used Replacement instead!\ lemma (in M_basic) cartprod_closed [intro,simp]: "\M(A); M(B)\ \ M(A*B)" by (frule cartprod_closed_lemma, assumption, force) lemma (in M_basic) sum_closed [intro,simp]: "\M(A); M(B)\ \ M(A+B)" by (simp add: sum_def) lemma (in M_basic) sum_abs [simp]: "\M(A); M(B); M(Z)\ \ is_sum(M,A,B,Z) \ (Z = A+B)" by (simp add: is_sum_def sum_def singleton_0 nat_into_M) lemma (in M_trivial) Inl_in_M_iff [iff]: "M(Inl(a)) \ M(a)" by (simp add: Inl_def) lemma (in M_trivial) Inl_abs [simp]: "M(Z) \ is_Inl(M,a,Z) \ (Z = Inl(a))" by (simp add: is_Inl_def Inl_def) lemma (in M_trivial) Inr_in_M_iff [iff]: "M(Inr(a)) \ M(a)" by (simp add: Inr_def) lemma (in M_trivial) Inr_abs [simp]: "M(Z) \ is_Inr(M,a,Z) \ (Z = Inr(a))" by (simp add: is_Inr_def Inr_def) subsubsection \converse of a relation\ lemma (in M_basic) M_converse_iff: "M(r) \ converse(r) = {z \ \(\(r)) * \(\(r)). \p\r. \x[M]. \y[M]. p = \x,y\ \ z = \y,x\}" apply (rule equalityI) prefer 2 apply (blast dest: transM, clarify, simp) apply (simp add: Pair_def) apply (blast dest: transM) done lemma (in M_basic) converse_closed [intro,simp]: "M(r) \ M(converse(r))" apply (simp add: M_converse_iff) apply (insert converse_separation [of r], simp) done lemma (in M_basic) converse_abs [simp]: "\M(r); M(z)\ \ is_converse(M,r,z) \ z = converse(r)" apply (simp add: is_converse_def) apply (rule iffI) prefer 2 apply blast apply (rule M_equalityI) apply simp apply (blast dest: transM)+ done subsubsection \image, preimage, domain, range\ lemma (in M_basic) image_closed [intro,simp]: "\M(A); M(r)\ \ M(r``A)" apply (simp add: image_iff_Collect) apply (insert image_separation [of A r], simp) done lemma (in M_basic) vimage_abs [simp]: "\M(r); M(A); M(z)\ \ pre_image(M,r,A,z) \ z = r-``A" apply (simp add: pre_image_def) apply (rule iffI) apply (blast intro!: equalityI dest: transM, blast) done lemma (in M_basic) vimage_closed [intro,simp]: "\M(A); M(r)\ \ M(r-``A)" by (simp add: vimage_def) subsubsection\Domain, range and field\ lemma (in M_trans) domain_abs [simp]: "\M(r); M(z)\ \ is_domain(M,r,z) \ z = domain(r)" apply (simp add: is_domain_def) apply (blast intro!: equalityI dest: transM) done lemma (in M_basic) domain_closed [intro,simp]: "M(r) \ M(domain(r))" apply (simp add: domain_eq_vimage) done lemma (in M_trans) range_abs [simp]: "\M(r); M(z)\ \ is_range(M,r,z) \ z = range(r)" apply (simp add: is_range_def) apply (blast intro!: equalityI dest: transM) done lemma (in M_basic) range_closed [intro,simp]: "M(r) \ M(range(r))" apply (simp add: range_eq_image) done lemma (in M_basic) field_abs [simp]: "\M(r); M(z)\ \ is_field(M,r,z) \ z = field(r)" by (simp add: is_field_def field_def) lemma (in M_basic) field_closed [intro,simp]: "M(r) \ M(field(r))" by (simp add: field_def) subsubsection\Relations, functions and application\ lemma (in M_trans) relation_abs [simp]: "M(r) \ is_relation(M,r) \ relation(r)" apply (simp add: is_relation_def relation_def) apply (blast dest!: bspec dest: pair_components_in_M)+ done lemma (in M_trivial) function_abs [simp]: "M(r) \ is_function(M,r) \ function(r)" apply (simp add: is_function_def function_def, safe) apply (frule transM, assumption) apply (blast dest: pair_components_in_M)+ done lemma (in M_basic) apply_closed [intro,simp]: "\M(f); M(a)\ \ M(f`a)" by (simp add: apply_def) lemma (in M_basic) apply_abs [simp]: "\M(f); M(x); M(y)\ \ fun_apply(M,f,x,y) \ f`x = y" apply (simp add: fun_apply_def apply_def, blast) done lemma (in M_trivial) typed_function_abs [simp]: "\M(A); M(f)\ \ typed_function(M,A,B,f) \ f \ A -> B" apply (auto simp add: typed_function_def relation_def Pi_iff) apply (blast dest: pair_components_in_M)+ done lemma (in M_basic) injection_abs [simp]: "\M(A); M(f)\ \ injection(M,A,B,f) \ f \ inj(A,B)" apply (simp add: injection_def apply_iff inj_def) apply (blast dest: transM [of _ A]) done lemma (in M_basic) surjection_abs [simp]: "\M(A); M(B); M(f)\ \ surjection(M,A,B,f) \ f \ surj(A,B)" by (simp add: surjection_def surj_def) lemma (in M_basic) bijection_abs [simp]: "\M(A); M(B); M(f)\ \ bijection(M,A,B,f) \ f \ bij(A,B)" by (simp add: bijection_def bij_def) subsubsection\Composition of relations\ lemma (in M_basic) M_comp_iff: "\M(r); M(s)\ \ r O s = {xz \ domain(s) * range(r). \x[M]. \y[M]. \z[M]. xz = \x,z\ \ \x,y\ \ s \ \y,z\ \ r}" apply (simp add: comp_def) apply (rule equalityI) apply clarify apply simp apply (blast dest: transM)+ done lemma (in M_basic) comp_closed [intro,simp]: "\M(r); M(s)\ \ M(r O s)" apply (simp add: M_comp_iff) apply (insert comp_separation [of r s], simp) done lemma (in M_basic) composition_abs [simp]: "\M(r); M(s); M(t)\ \ composition(M,r,s,t) \ t = r O s" apply safe txt\Proving \<^term>\composition(M, r, s, r O s)\\ prefer 2 apply (simp add: composition_def comp_def) apply (blast dest: transM) txt\Opposite implication\ apply (rule M_equalityI) apply (simp add: composition_def comp_def) apply (blast del: allE dest: transM)+ done text\no longer needed\ lemma (in M_basic) restriction_is_function: "\restriction(M,f,A,z); function(f); M(f); M(A); M(z)\ \ function(z)" apply (simp add: restriction_def ball_iff_equiv) apply (unfold function_def, blast) done lemma (in M_trans) restriction_abs [simp]: "\M(f); M(A); M(z)\ \ restriction(M,f,A,z) \ z = restrict(f,A)" apply (simp add: ball_iff_equiv restriction_def restrict_def) apply (blast intro!: equalityI dest: transM) done lemma (in M_trans) M_restrict_iff: "M(r) \ restrict(r,A) = {z \ r . \x\A. \y[M]. z = \x, y\}" by (simp add: restrict_def, blast dest: transM) lemma (in M_basic) restrict_closed [intro,simp]: "\M(A); M(r)\ \ M(restrict(r,A))" apply (simp add: M_restrict_iff) apply (insert restrict_separation [of A], simp) done lemma (in M_trans) Inter_abs [simp]: "\M(A); M(z)\ \ big_inter(M,A,z) \ z = \(A)" apply (simp add: big_inter_def Inter_def) apply (blast intro!: equalityI dest: transM) done lemma (in M_basic) Inter_closed [intro,simp]: "M(A) \ M(\(A))" by (insert Inter_separation, simp add: Inter_def) lemma (in M_basic) Int_closed [intro,simp]: "\M(A); M(B)\ \ M(A \ B)" apply (subgoal_tac "M({A,B})") apply (frule Inter_closed, force+) done lemma (in M_basic) Diff_closed [intro,simp]: "\M(A); M(B)\ \ M(A-B)" by (insert Diff_separation, simp add: Diff_def) subsubsection\Some Facts About Separation Axioms\ lemma (in M_basic) separation_conj: "\separation(M,P); separation(M,Q)\ \ separation(M, \z. P(z) \ Q(z))" by (simp del: separation_closed add: separation_iff Collect_Int_Collect_eq [symmetric]) (*???equalities*) lemma Collect_Un_Collect_eq: "Collect(A,P) \ Collect(A,Q) = Collect(A, \x. P(x) \ Q(x))" by blast lemma Diff_Collect_eq: "A - Collect(A,P) = Collect(A, \x. \ P(x))" by blast lemma (in M_trans) Collect_rall_eq: "M(Y) \ Collect(A, \x. \y[M]. y\Y \ P(x,y)) = (if Y=0 then A else (\y \ Y. {x \ A. P(x,y)}))" by (simp,blast dest: transM) lemma (in M_basic) separation_disj: "\separation(M,P); separation(M,Q)\ \ separation(M, \z. P(z) \ Q(z))" by (simp del: separation_closed add: separation_iff Collect_Un_Collect_eq [symmetric]) lemma (in M_basic) separation_neg: "separation(M,P) \ separation(M, \z. \P(z))" by (simp del: separation_closed add: separation_iff Diff_Collect_eq [symmetric]) lemma (in M_basic) separation_imp: "\separation(M,P); separation(M,Q)\ \ separation(M, \z. P(z) \ Q(z))" by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric]) text\This result is a hint of how little can be done without the Reflection Theorem. The quantifier has to be bounded by a set. We also need another instance of Separation!\ lemma (in M_basic) separation_rall: "\M(Y); \y[M]. separation(M, \x. P(x,y)); \z[M]. strong_replacement(M, \x y. y = {u \ z . P(u,x)})\ \ separation(M, \x. \y[M]. y\Y \ P(x,y))" apply (simp del: separation_closed rall_abs add: separation_iff Collect_rall_eq) apply (blast intro!: RepFun_closed dest: transM) done subsubsection\Functions and function space\ text\The assumption \<^term>\M(A->B)\ is unusual, but essential: in all but trivial cases, A->B cannot be expected to belong to \<^term>\M\.\ lemma (in M_trivial) is_funspace_abs [simp]: "\M(A); M(B); M(F); M(A->B)\ \ is_funspace(M,A,B,F) \ F = A->B" apply (simp add: is_funspace_def) apply (rule iffI) prefer 2 apply blast apply (rule M_equalityI) apply simp_all done lemma (in M_basic) succ_fun_eq2: "\M(B); M(n->B)\ \ succ(n) -> B = \{z. p \ (n->B)*B, \f[M]. \b[M]. p = \f,b\ \ z = {cons(\n,b\, f)}}" apply (simp add: succ_fun_eq) apply (blast dest: transM) done lemma (in M_basic) funspace_succ: "\M(n); M(B); M(n->B)\ \ M(succ(n) -> B)" apply (insert funspace_succ_replacement [of n], simp) apply (force simp add: succ_fun_eq2 univalent_def) done text\\<^term>\M\ contains all finite function spaces. Needed to prove the absoluteness of transitive closure. See the definition of \rtrancl_alt\ in in \WF_absolute.thy\.\ lemma (in M_basic) finite_funspace_closed [intro,simp]: "\n\nat; M(B)\ \ M(n->B)" apply (induct_tac n, simp) apply (simp add: funspace_succ nat_into_M) done subsection\Relativization and Absoluteness for Boolean Operators\ definition is_bool_of_o :: "[i\o, o, i] \ o" where "is_bool_of_o(M,P,z) \ (P \ number1(M,z)) \ (\P \ empty(M,z))" definition is_not :: "[i\o, i, i] \ o" where "is_not(M,a,z) \ (number1(M,a) \ empty(M,z)) | (\number1(M,a) \ number1(M,z))" definition is_and :: "[i\o, i, i, i] \ o" where "is_and(M,a,b,z) \ (number1(M,a) \ z=b) | (\number1(M,a) \ empty(M,z))" definition is_or :: "[i\o, i, i, i] \ o" where "is_or(M,a,b,z) \ (number1(M,a) \ number1(M,z)) | (\number1(M,a) \ z=b)" lemma (in M_trivial) bool_of_o_abs [simp]: "M(z) \ is_bool_of_o(M,P,z) \ z = bool_of_o(P)" by (simp add: is_bool_of_o_def bool_of_o_def) lemma (in M_trivial) not_abs [simp]: "\M(a); M(z)\ \ is_not(M,a,z) \ z = not(a)" by (simp add: Bool.not_def cond_def is_not_def) lemma (in M_trivial) and_abs [simp]: "\M(a); M(b); M(z)\ \ is_and(M,a,b,z) \ z = a and b" by (simp add: Bool.and_def cond_def is_and_def) lemma (in M_trivial) or_abs [simp]: "\M(a); M(b); M(z)\ \ is_or(M,a,b,z) \ z = a or b" by (simp add: Bool.or_def cond_def is_or_def) lemma (in M_trivial) bool_of_o_closed [intro,simp]: "M(bool_of_o(P))" by (simp add: bool_of_o_def) lemma (in M_trivial) and_closed [intro,simp]: "\M(p); M(q)\ \ M(p and q)" by (simp add: and_def cond_def) lemma (in M_trivial) or_closed [intro,simp]: "\M(p); M(q)\ \ M(p or q)" by (simp add: or_def cond_def) lemma (in M_trivial) not_closed [intro,simp]: "M(p) \ M(not(p))" by (simp add: Bool.not_def cond_def) subsection\Relativization and Absoluteness for List Operators\ definition is_Nil :: "[i\o, i] \ o" where \ \because \<^prop>\[] \ Inl(0)\\ "is_Nil(M,xs) \ \zero[M]. empty(M,zero) \ is_Inl(M,zero,xs)" definition is_Cons :: "[i\o,i,i,i] \ o" where \ \because \<^prop>\Cons(a, l) \ Inr(\a,l\)\\ "is_Cons(M,a,l,Z) \ \p[M]. pair(M,a,l,p) \ is_Inr(M,p,Z)" lemma (in M_trivial) Nil_in_M [intro,simp]: "M(Nil)" by (simp add: Nil_def) lemma (in M_trivial) Nil_abs [simp]: "M(Z) \ is_Nil(M,Z) \ (Z = Nil)" by (simp add: is_Nil_def Nil_def) lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) \ M(a) \ M(l)" by (simp add: Cons_def) lemma (in M_trivial) Cons_abs [simp]: "\M(a); M(l); M(Z)\ \ is_Cons(M,a,l,Z) \ (Z = Cons(a,l))" by (simp add: is_Cons_def Cons_def) definition quasilist :: "i \ o" where "quasilist(xs) \ xs=Nil \ (\x l. xs = Cons(x,l))" definition is_quasilist :: "[i\o,i] \ o" where "is_quasilist(M,z) \ is_Nil(M,z) \ (\x[M]. \l[M]. is_Cons(M,x,l,z))" definition list_case' :: "[i, [i,i]\i, i] \ i" where \ \A version of \<^term>\list_case\ that's always defined.\ "list_case'(a,b,xs) \ if quasilist(xs) then list_case(a,b,xs) else 0" definition is_list_case :: "[i\o, i, [i,i,i]\o, i, i] \ o" where \ \Returns 0 for non-lists\ "is_list_case(M, a, is_b, xs, z) \ (is_Nil(M,xs) \ z=a) \ (\x[M]. \l[M]. is_Cons(M,x,l,xs) \ is_b(x,l,z)) \ (is_quasilist(M,xs) \ empty(M,z))" definition hd' :: "i \ i" where \ \A version of \<^term>\hd\ that's always defined.\ "hd'(xs) \ if quasilist(xs) then hd(xs) else 0" definition tl' :: "i \ i" where \ \A version of \<^term>\tl\ that's always defined.\ "tl'(xs) \ if quasilist(xs) then tl(xs) else 0" definition is_hd :: "[i\o,i,i] \ o" where \ \\<^term>\hd([]) = 0\ no constraints if not a list. Avoiding implication prevents the simplifier's looping.\ "is_hd(M,xs,H) \ (is_Nil(M,xs) \ empty(M,H)) \ (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) \ H=x) \ (is_quasilist(M,xs) \ empty(M,H))" definition is_tl :: "[i\o,i,i] \ o" where \ \\<^term>\tl([]) = []\; see comments about \<^term>\is_hd\\ "is_tl(M,xs,T) \ (is_Nil(M,xs) \ T=xs) \ (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) \ T=l) \ (is_quasilist(M,xs) \ empty(M,T))" subsubsection\\<^term>\quasilist\: For Case-Splitting with \<^term>\list_case'\\ lemma [iff]: "quasilist(Nil)" by (simp add: quasilist_def) lemma [iff]: "quasilist(Cons(x,l))" by (simp add: quasilist_def) lemma list_imp_quasilist: "l \ list(A) \ quasilist(l)" by (erule list.cases, simp_all) subsubsection\\<^term>\list_case'\, the Modified Version of \<^term>\list_case\\ lemma list_case'_Nil [simp]: "list_case'(a,b,Nil) = a" by (simp add: list_case'_def quasilist_def) lemma list_case'_Cons [simp]: "list_case'(a,b,Cons(x,l)) = b(x,l)" by (simp add: list_case'_def quasilist_def) lemma non_list_case: "\ quasilist(x) \ list_case'(a,b,x) = 0" by (simp add: quasilist_def list_case'_def) lemma list_case'_eq_list_case [simp]: "xs \ list(A) \list_case'(a,b,xs) = list_case(a,b,xs)" by (erule list.cases, simp_all) lemma (in M_basic) list_case'_closed [intro,simp]: "\M(k); M(a); \x[M]. \y[M]. M(b(x,y))\ \ M(list_case'(a,b,k))" apply (case_tac "quasilist(k)") apply (simp add: quasilist_def, force) apply (simp add: non_list_case) done lemma (in M_trivial) quasilist_abs [simp]: "M(z) \ is_quasilist(M,z) \ quasilist(z)" by (auto simp add: is_quasilist_def quasilist_def) lemma (in M_trivial) list_case_abs [simp]: "\relation2(M,is_b,b); M(k); M(z)\ \ is_list_case(M,a,is_b,k,z) \ z = list_case'(a,b,k)" apply (case_tac "quasilist(k)") prefer 2 apply (simp add: is_list_case_def non_list_case) apply (force simp add: quasilist_def) apply (simp add: quasilist_def is_list_case_def) apply (elim disjE exE) apply (simp_all add: relation2_def) done subsubsection\The Modified Operators \<^term>\hd'\ and \<^term>\tl'\\ lemma (in M_trivial) is_hd_Nil: "is_hd(M,[],Z) \ empty(M,Z)" by (simp add: is_hd_def) lemma (in M_trivial) is_hd_Cons: "\M(a); M(l)\ \ is_hd(M,Cons(a,l),Z) \ Z = a" by (force simp add: is_hd_def) lemma (in M_trivial) hd_abs [simp]: "\M(x); M(y)\ \ is_hd(M,x,y) \ y = hd'(x)" apply (simp add: hd'_def) apply (intro impI conjI) prefer 2 apply (force simp add: is_hd_def) apply (simp add: quasilist_def is_hd_def) apply (elim disjE exE, auto) done lemma (in M_trivial) is_tl_Nil: "is_tl(M,[],Z) \ Z = []" by (simp add: is_tl_def) lemma (in M_trivial) is_tl_Cons: "\M(a); M(l)\ \ is_tl(M,Cons(a,l),Z) \ Z = l" by (force simp add: is_tl_def) lemma (in M_trivial) tl_abs [simp]: "\M(x); M(y)\ \ is_tl(M,x,y) \ y = tl'(x)" apply (simp add: tl'_def) apply (intro impI conjI) prefer 2 apply (force simp add: is_tl_def) apply (simp add: quasilist_def is_tl_def) apply (elim disjE exE, auto) done lemma (in M_trivial) relation1_tl: "relation1(M, is_tl(M), tl')" by (simp add: relation1_def) lemma hd'_Nil: "hd'([]) = 0" by (simp add: hd'_def) lemma hd'_Cons: "hd'(Cons(a,l)) = a" by (simp add: hd'_def) lemma tl'_Nil: "tl'([]) = []" by (simp add: tl'_def) lemma tl'_Cons: "tl'(Cons(a,l)) = l" by (simp add: tl'_def) lemma iterates_tl_Nil: "n \ nat \ tl'^n ([]) = []" apply (induct_tac n) apply (simp_all add: tl'_Nil) done lemma (in M_basic) tl'_closed: "M(x) \ M(tl'(x))" apply (simp add: tl'_def) apply (force simp add: quasilist_def) done end diff --git a/src/ZF/Constructible/Satisfies_absolute.thy b/src/ZF/Constructible/Satisfies_absolute.thy --- a/src/ZF/Constructible/Satisfies_absolute.thy +++ b/src/ZF/Constructible/Satisfies_absolute.thy @@ -1,1039 +1,1039 @@ (* Title: ZF/Constructible/Satisfies_absolute.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) section \Absoluteness for the Satisfies Relation on Formulas\ theory Satisfies_absolute imports Datatype_absolute Rec_Separation begin subsection \More Internalization\ subsubsection\The Formula \<^term>\is_depth\, Internalized\ (* "is_depth(M,p,n) \ \sn[M]. \formula_n[M]. \formula_sn[M]. 2 1 0 is_formula_N(M,n,formula_n) \ p \ formula_n \ successor(M,n,sn) \ is_formula_N(M,sn,formula_sn) \ p \ formula_sn" *) definition depth_fm :: "[i,i]\i" where "depth_fm(p,n) \ Exists(Exists(Exists( And(formula_N_fm(n#+3,1), And(Neg(Member(p#+3,1)), And(succ_fm(n#+3,2), And(formula_N_fm(2,0), Member(p#+3,0))))))))" lemma depth_fm_type [TC]: "\x \ nat; y \ nat\ \ depth_fm(x,y) \ formula" by (simp add: depth_fm_def) lemma sats_depth_fm [simp]: "\x \ nat; y < length(env); env \ list(A)\ \ sats(A, depth_fm(x,y), env) \ is_depth(##A, nth(x,env), nth(y,env))" apply (frule_tac x=y in lt_length_in_nat, assumption) apply (simp add: depth_fm_def is_depth_def) done lemma depth_iff_sats: "\nth(i,env) = x; nth(j,env) = y; i \ nat; j < length(env); env \ list(A)\ \ is_depth(##A, x, y) \ sats(A, depth_fm(i,j), env)" by (simp) theorem depth_reflection: "REFLECTS[\x. is_depth(L, f(x), g(x)), \i x. is_depth(##Lset(i), f(x), g(x))]" apply (simp only: is_depth_def) apply (intro FOL_reflections function_reflections formula_N_reflection) done subsubsection\The Operator \<^term>\is_formula_case\\ text\The arguments of \<^term>\is_a\ are always 2, 1, 0, and the formula will be enclosed by three quantifiers.\ (* is_formula_case :: "[i\o, [i,i,i]\o, [i,i,i]\o, [i,i,i]\o, [i,i]\o, i, i] \ o" "is_formula_case(M, is_a, is_b, is_c, is_d, v, z) \ (\x[M]. \y[M]. x\nat \ y\nat \ is_Member(M,x,y,v) \ is_a(x,y,z)) \ (\x[M]. \y[M]. x\nat \ y\nat \ is_Equal(M,x,y,v) \ is_b(x,y,z)) \ (\x[M]. \y[M]. x\formula \ y\formula \ is_Nand(M,x,y,v) \ is_c(x,y,z)) \ (\x[M]. x\formula \ is_Forall(M,x,v) \ is_d(x,z))" *) definition formula_case_fm :: "[i, i, i, i, i, i]\i" where "formula_case_fm(is_a, is_b, is_c, is_d, v, z) \ And(Forall(Forall(Implies(finite_ordinal_fm(1), Implies(finite_ordinal_fm(0), Implies(Member_fm(1,0,v#+2), Forall(Implies(Equal(0,z#+3), is_a))))))), And(Forall(Forall(Implies(finite_ordinal_fm(1), Implies(finite_ordinal_fm(0), Implies(Equal_fm(1,0,v#+2), Forall(Implies(Equal(0,z#+3), is_b))))))), And(Forall(Forall(Implies(mem_formula_fm(1), Implies(mem_formula_fm(0), Implies(Nand_fm(1,0,v#+2), Forall(Implies(Equal(0,z#+3), is_c))))))), Forall(Implies(mem_formula_fm(0), Implies(Forall_fm(0,succ(v)), Forall(Implies(Equal(0,z#+2), is_d))))))))" lemma is_formula_case_type [TC]: "\is_a \ formula; is_b \ formula; is_c \ formula; is_d \ formula; x \ nat; y \ nat\ \ formula_case_fm(is_a, is_b, is_c, is_d, x, y) \ formula" by (simp add: formula_case_fm_def) lemma sats_formula_case_fm: assumes is_a_iff_sats: "\a0 a1 a2. \a0\A; a1\A; a2\A\ \ ISA(a2, a1, a0) \ sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))" and is_b_iff_sats: "\a0 a1 a2. \a0\A; a1\A; a2\A\ \ ISB(a2, a1, a0) \ sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))" and is_c_iff_sats: "\a0 a1 a2. \a0\A; a1\A; a2\A\ \ ISC(a2, a1, a0) \ sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))" and is_d_iff_sats: "\a0 a1. \a0\A; a1\A\ \ ISD(a1, a0) \ sats(A, is_d, Cons(a0,Cons(a1,env)))" shows "\x \ nat; y < length(env); env \ list(A)\ \ sats(A, formula_case_fm(is_a,is_b,is_c,is_d,x,y), env) \ is_formula_case(##A, ISA, ISB, ISC, ISD, nth(x,env), nth(y,env))" apply (frule_tac x=y in lt_length_in_nat, assumption) apply (simp add: formula_case_fm_def is_formula_case_def is_a_iff_sats [THEN iff_sym] is_b_iff_sats [THEN iff_sym] is_c_iff_sats [THEN iff_sym] is_d_iff_sats [THEN iff_sym]) done lemma formula_case_iff_sats: assumes is_a_iff_sats: "\a0 a1 a2. \a0\A; a1\A; a2\A\ \ ISA(a2, a1, a0) \ sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))" and is_b_iff_sats: "\a0 a1 a2. \a0\A; a1\A; a2\A\ \ ISB(a2, a1, a0) \ sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))" and is_c_iff_sats: "\a0 a1 a2. \a0\A; a1\A; a2\A\ \ ISC(a2, a1, a0) \ sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))" and is_d_iff_sats: "\a0 a1. \a0\A; a1\A\ \ ISD(a1, a0) \ sats(A, is_d, Cons(a0,Cons(a1,env)))" shows "\nth(i,env) = x; nth(j,env) = y; i \ nat; j < length(env); env \ list(A)\ \ is_formula_case(##A, ISA, ISB, ISC, ISD, x, y) \ sats(A, formula_case_fm(is_a,is_b,is_c,is_d,i,j), env)" by (simp add: sats_formula_case_fm [OF is_a_iff_sats is_b_iff_sats is_c_iff_sats is_d_iff_sats]) text\The second argument of \<^term>\is_a\ gives it direct access to \<^term>\x\, which is essential for handling free variable references. Treatment is based on that of \is_nat_case_reflection\.\ theorem is_formula_case_reflection: assumes is_a_reflection: "\h f g g'. REFLECTS[\x. is_a(L, h(x), f(x), g(x), g'(x)), \i x. is_a(##Lset(i), h(x), f(x), g(x), g'(x))]" and is_b_reflection: "\h f g g'. REFLECTS[\x. is_b(L, h(x), f(x), g(x), g'(x)), \i x. is_b(##Lset(i), h(x), f(x), g(x), g'(x))]" and is_c_reflection: "\h f g g'. REFLECTS[\x. is_c(L, h(x), f(x), g(x), g'(x)), \i x. is_c(##Lset(i), h(x), f(x), g(x), g'(x))]" and is_d_reflection: "\h f g g'. REFLECTS[\x. is_d(L, h(x), f(x), g(x)), \i x. is_d(##Lset(i), h(x), f(x), g(x))]" shows "REFLECTS[\x. is_formula_case(L, is_a(L,x), is_b(L,x), is_c(L,x), is_d(L,x), g(x), h(x)), \i x. is_formula_case(##Lset(i), is_a(##Lset(i), x), is_b(##Lset(i), x), is_c(##Lset(i), x), is_d(##Lset(i), x), g(x), h(x))]" apply (simp (no_asm_use) only: is_formula_case_def) apply (intro FOL_reflections function_reflections finite_ordinal_reflection mem_formula_reflection Member_reflection Equal_reflection Nand_reflection Forall_reflection is_a_reflection is_b_reflection is_c_reflection is_d_reflection) done subsection \Absoluteness for the Function \<^term>\satisfies\\ definition is_depth_apply :: "[i\o,i,i,i] \ o" where \ \Merely a useful abbreviation for the sequel.\ "is_depth_apply(M,h,p,z) \ \dp[M]. \sdp[M]. \hsdp[M]. finite_ordinal(M,dp) \ is_depth(M,p,dp) \ successor(M,dp,sdp) \ fun_apply(M,h,sdp,hsdp) \ fun_apply(M,hsdp,p,z)" lemma (in M_datatypes) is_depth_apply_abs [simp]: "\M(h); p \ formula; M(z)\ \ is_depth_apply(M,h,p,z) \ z = h ` succ(depth(p)) ` p" by (simp add: is_depth_apply_def formula_into_M depth_type eq_commute) text\There is at present some redundancy between the relativizations in e.g. \satisfies_is_a\ and those in e.g. \Member_replacement\.\ text\These constants let us instantiate the parameters \<^term>\a\, \<^term>\b\, \<^term>\c\, \<^term>\d\, etc., of the locale \Formula_Rec\.\ definition satisfies_a :: "[i,i,i]\i" where "satisfies_a(A) \ \x y. \env \ list(A). bool_of_o (nth(x,env) \ nth(y,env))" definition satisfies_is_a :: "[i\o,i,i,i,i]\o" where "satisfies_is_a(M,A) \ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. is_bool_of_o(M, \nx[M]. \ny[M]. is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ nx \ ny, z), zz)" definition satisfies_b :: "[i,i,i]\i" where "satisfies_b(A) \ \x y. \env \ list(A). bool_of_o (nth(x,env) = nth(y,env))" definition satisfies_is_b :: "[i\o,i,i,i,i]\o" where \ \We simplify the formula to have just \<^term>\nx\ rather than introducing \<^term>\ny\ with \<^term>\nx=ny\\ "satisfies_is_b(M,A) \ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. is_bool_of_o(M, \nx[M]. is_nth(M,x,env,nx) \ is_nth(M,y,env,nx), z), zz)" definition satisfies_c :: "[i,i,i,i,i]\i" where "satisfies_c(A) \ \p q rp rq. \env \ list(A). not(rp ` env and rq ` env)" definition satisfies_is_c :: "[i\o,i,i,i,i,i]\o" where "satisfies_is_c(M,A,h) \ \p q zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \hp[M]. \hq[M]. (\rp[M]. is_depth_apply(M,h,p,rp) \ fun_apply(M,rp,env,hp)) \ (\rq[M]. is_depth_apply(M,h,q,rq) \ fun_apply(M,rq,env,hq)) \ (\pq[M]. is_and(M,hp,hq,pq) \ is_not(M,pq,z)), zz)" definition satisfies_d :: "[i,i,i]\i" where "satisfies_d(A) \ \p rp. \env \ list(A). bool_of_o (\x\A. rp ` (Cons(x,env)) = 1)" definition satisfies_is_d :: "[i\o,i,i,i,i]\o" where "satisfies_is_d(M,A,h) \ \p zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \rp[M]. is_depth_apply(M,h,p,rp) \ is_bool_of_o(M, \x[M]. \xenv[M]. \hp[M]. x\A \ is_Cons(M,x,env,xenv) \ fun_apply(M,rp,xenv,hp) \ number1(M,hp), z), zz)" definition satisfies_MH :: "[i\o,i,i,i,i]\o" where \ \The variable \<^term>\u\ is unused, but gives \<^term>\satisfies_MH\ the correct arity.\ "satisfies_MH \ \M A u f z. \fml[M]. is_formula(M,fml) \ is_lambda (M, fml, is_formula_case (M, satisfies_is_a(M,A), satisfies_is_b(M,A), satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)), z)" definition is_satisfies :: "[i\o,i,i,i]\o" where "is_satisfies(M,A) \ is_formula_rec (M, satisfies_MH(M,A))" text\This lemma relates the fragments defined above to the original primitive recursion in \<^term>\satisfies\. Induction is not required: the definitions are directly equal!\ lemma satisfies_eq: "satisfies(A,p) = formula_rec (satisfies_a(A), satisfies_b(A), satisfies_c(A), satisfies_d(A), p)" by (simp add: satisfies_formula_def satisfies_a_def satisfies_b_def satisfies_c_def satisfies_d_def) text\Further constraints on the class \<^term>\M\ in order to prove absoluteness for the constants defined above. The ultimate goal is the absoluteness of the function \<^term>\satisfies\.\ locale M_satisfies = M_eclose + assumes Member_replacement: "\M(A); x \ nat; y \ nat\ \ strong_replacement (M, \env z. \bo[M]. \nx[M]. \ny[M]. env \ list(A) \ is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ is_bool_of_o(M, nx \ ny, bo) \ pair(M, env, bo, z))" and Equal_replacement: "\M(A); x \ nat; y \ nat\ \ strong_replacement (M, \env z. \bo[M]. \nx[M]. \ny[M]. env \ list(A) \ is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ is_bool_of_o(M, nx = ny, bo) \ pair(M, env, bo, z))" and Nand_replacement: "\M(A); M(rp); M(rq)\ \ strong_replacement (M, \env z. \rpe[M]. \rqe[M]. \andpq[M]. \notpq[M]. fun_apply(M,rp,env,rpe) \ fun_apply(M,rq,env,rqe) \ is_and(M,rpe,rqe,andpq) \ is_not(M,andpq,notpq) \ env \ list(A) \ pair(M, env, notpq, z))" and Forall_replacement: "\M(A); M(rp)\ \ strong_replacement (M, \env z. \bo[M]. env \ list(A) \ is_bool_of_o (M, \a[M]. \co[M]. \rpco[M]. a\A \ is_Cons(M,a,env,co) \ fun_apply(M,rp,co,rpco) \ number1(M, rpco), bo) \ pair(M,env,bo,z))" and formula_rec_replacement: \ \For the \<^term>\transrec\\ "\n \ nat; M(A)\ \ transrec_replacement(M, satisfies_MH(M,A), n)" and formula_rec_lambda_replacement: \ \For the \\-abstraction\ in the \<^term>\transrec\ body\ "\M(g); M(A)\ \ strong_replacement (M, \x y. mem_formula(M,x) \ (\c[M]. is_formula_case(M, satisfies_is_a(M,A), satisfies_is_b(M,A), satisfies_is_c(M,A,g), satisfies_is_d(M,A,g), x, c) \ pair(M, x, c, y)))" lemma (in M_satisfies) Member_replacement': "\M(A); x \ nat; y \ nat\ \ strong_replacement (M, \env z. env \ list(A) \ z = \env, bool_of_o(nth(x, env) \ nth(y, env))\)" by (insert Member_replacement, simp) lemma (in M_satisfies) Equal_replacement': "\M(A); x \ nat; y \ nat\ \ strong_replacement (M, \env z. env \ list(A) \ z = \env, bool_of_o(nth(x, env) = nth(y, env))\)" by (insert Equal_replacement, simp) lemma (in M_satisfies) Nand_replacement': "\M(A); M(rp); M(rq)\ \ strong_replacement (M, \env z. env \ list(A) \ z = \env, not(rp`env and rq`env)\)" by (insert Nand_replacement, simp) lemma (in M_satisfies) Forall_replacement': "\M(A); M(rp)\ \ strong_replacement (M, \env z. env \ list(A) \ z = \env, bool_of_o (\a\A. rp ` Cons(a,env) = 1)\)" by (insert Forall_replacement, simp) lemma (in M_satisfies) a_closed: "\M(A); x\nat; y\nat\ \ M(satisfies_a(A,x,y))" apply (simp add: satisfies_a_def) apply (blast intro: lam_closed2 Member_replacement') done lemma (in M_satisfies) a_rel: "M(A) \ Relation2(M, nat, nat, satisfies_is_a(M,A), satisfies_a(A))" apply (simp add: Relation2_def satisfies_is_a_def satisfies_a_def) apply (auto del: iffI intro!: lambda_abs2 simp add: Relation1_def) done lemma (in M_satisfies) b_closed: "\M(A); x\nat; y\nat\ \ M(satisfies_b(A,x,y))" apply (simp add: satisfies_b_def) apply (blast intro: lam_closed2 Equal_replacement') done lemma (in M_satisfies) b_rel: "M(A) \ Relation2(M, nat, nat, satisfies_is_b(M,A), satisfies_b(A))" apply (simp add: Relation2_def satisfies_is_b_def satisfies_b_def) apply (auto del: iffI intro!: lambda_abs2 simp add: Relation1_def) done lemma (in M_satisfies) c_closed: "\M(A); x \ formula; y \ formula; M(rx); M(ry)\ \ M(satisfies_c(A,x,y,rx,ry))" apply (simp add: satisfies_c_def) apply (rule lam_closed2) apply (rule Nand_replacement') apply (simp_all add: formula_into_M list_into_M [of _ A]) done lemma (in M_satisfies) c_rel: "\M(A); M(f)\ \ Relation2 (M, formula, formula, satisfies_is_c(M,A,f), \u v. satisfies_c(A, u, v, f ` succ(depth(u)) ` u, f ` succ(depth(v)) ` v))" apply (simp add: Relation2_def satisfies_is_c_def satisfies_c_def) apply (auto del: iffI intro!: lambda_abs2 simp add: Relation1_def formula_into_M) done lemma (in M_satisfies) d_closed: "\M(A); x \ formula; M(rx)\ \ M(satisfies_d(A,x,rx))" apply (simp add: satisfies_d_def) apply (rule lam_closed2) apply (rule Forall_replacement') apply (simp_all add: formula_into_M list_into_M [of _ A]) done lemma (in M_satisfies) d_rel: "\M(A); M(f)\ \ Relation1(M, formula, satisfies_is_d(M,A,f), \u. satisfies_d(A, u, f ` succ(depth(u)) ` u))" apply (simp del: rall_abs add: Relation1_def satisfies_is_d_def satisfies_d_def) apply (auto del: iffI intro!: lambda_abs2 simp add: Relation1_def) done lemma (in M_satisfies) fr_replace: "\n \ nat; M(A)\ \ transrec_replacement(M,satisfies_MH(M,A),n)" by (blast intro: formula_rec_replacement) lemma (in M_satisfies) formula_case_satisfies_closed: "\M(g); M(A); x \ formula\ \ M(formula_case (satisfies_a(A), satisfies_b(A), \u v. satisfies_c(A, u, v, g ` succ(depth(u)) ` u, g ` succ(depth(v)) ` v), \u. satisfies_d (A, u, g ` succ(depth(u)) ` u), x))" by (blast intro: a_closed b_closed c_closed d_closed) lemma (in M_satisfies) fr_lam_replace: "\M(g); M(A)\ \ strong_replacement (M, \x y. x \ formula \ y = \x, formula_rec_case(satisfies_a(A), satisfies_b(A), satisfies_c(A), satisfies_d(A), g, x)\)" apply (insert formula_rec_lambda_replacement) apply (simp add: formula_rec_case_def formula_case_satisfies_closed formula_case_abs [OF a_rel b_rel c_rel d_rel]) done text\Instantiate locale \Formula_Rec\ for the Function \<^term>\satisfies\\ lemma (in M_satisfies) Formula_Rec_axioms_M: "M(A) \ Formula_Rec_axioms(M, satisfies_a(A), satisfies_is_a(M,A), satisfies_b(A), satisfies_is_b(M,A), satisfies_c(A), satisfies_is_c(M,A), satisfies_d(A), satisfies_is_d(M,A))" apply (rule Formula_Rec_axioms.intro) apply (assumption | rule a_closed a_rel b_closed b_rel c_closed c_rel d_closed d_rel fr_replace [unfolded satisfies_MH_def] fr_lam_replace) + done theorem (in M_satisfies) Formula_Rec_M: "M(A) \ Formula_Rec(M, satisfies_a(A), satisfies_is_a(M,A), satisfies_b(A), satisfies_is_b(M,A), satisfies_c(A), satisfies_is_c(M,A), satisfies_d(A), satisfies_is_d(M,A))" apply (rule Formula_Rec.intro) apply (rule M_satisfies.axioms, rule M_satisfies_axioms) apply (erule Formula_Rec_axioms_M) done lemmas (in M_satisfies) satisfies_closed' = Formula_Rec.formula_rec_closed [OF Formula_Rec_M] and satisfies_abs' = Formula_Rec.formula_rec_abs [OF Formula_Rec_M] lemma (in M_satisfies) satisfies_closed: "\M(A); p \ formula\ \ M(satisfies(A,p))" by (simp add: Formula_Rec.formula_rec_closed [OF Formula_Rec_M] satisfies_eq) lemma (in M_satisfies) satisfies_abs: "\M(A); M(z); p \ formula\ \ is_satisfies(M,A,p,z) \ z = satisfies(A,p)" by (simp only: Formula_Rec.formula_rec_abs [OF Formula_Rec_M] satisfies_eq is_satisfies_def satisfies_MH_def) subsection\Internalizations Needed to Instantiate \M_satisfies\\ subsubsection\The Operator \<^term>\is_depth_apply\, Internalized\ (* is_depth_apply(M,h,p,z) \ \dp[M]. \sdp[M]. \hsdp[M]. 2 1 0 finite_ordinal(M,dp) \ is_depth(M,p,dp) \ successor(M,dp,sdp) \ fun_apply(M,h,sdp,hsdp) \ fun_apply(M,hsdp,p,z) *) definition depth_apply_fm :: "[i,i,i]\i" where "depth_apply_fm(h,p,z) \ Exists(Exists(Exists( And(finite_ordinal_fm(2), And(depth_fm(p#+3,2), And(succ_fm(2,1), And(fun_apply_fm(h#+3,1,0), fun_apply_fm(0,p#+3,z#+3))))))))" lemma depth_apply_type [TC]: "\x \ nat; y \ nat; z \ nat\ \ depth_apply_fm(x,y,z) \ formula" by (simp add: depth_apply_fm_def) lemma sats_depth_apply_fm [simp]: "\x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, depth_apply_fm(x,y,z), env) \ is_depth_apply(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: depth_apply_fm_def is_depth_apply_def) lemma depth_apply_iff_sats: "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)\ \ is_depth_apply(##A, x, y, z) \ sats(A, depth_apply_fm(i,j,k), env)" by simp lemma depth_apply_reflection: "REFLECTS[\x. is_depth_apply(L,f(x),g(x),h(x)), \i x. is_depth_apply(##Lset(i),f(x),g(x),h(x))]" apply (simp only: is_depth_apply_def) apply (intro FOL_reflections function_reflections depth_reflection finite_ordinal_reflection) done subsubsection\The Operator \<^term>\satisfies_is_a\, Internalized\ (* satisfies_is_a(M,A) \ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. is_bool_of_o(M, \nx[M]. \ny[M]. is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ nx \ ny, z), zz) *) definition satisfies_is_a_fm :: "[i,i,i,i]\i" where "satisfies_is_a_fm(A,x,y,z) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( bool_of_o_fm(Exists( Exists(And(nth_fm(x#+6,3,1), And(nth_fm(y#+6,3,0), Member(1,0))))), 0), 0, succ(z))))" lemma satisfies_is_a_type [TC]: "\A \ nat; x \ nat; y \ nat; z \ nat\ \ satisfies_is_a_fm(A,x,y,z) \ formula" by (simp add: satisfies_is_a_fm_def) lemma sats_satisfies_is_a_fm [simp]: "\u \ nat; x < length(env); y < length(env); z \ nat; env \ list(A)\ \ sats(A, satisfies_is_a_fm(u,x,y,z), env) \ satisfies_is_a(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" apply (frule_tac x=x in lt_length_in_nat, assumption) apply (frule_tac x=y in lt_length_in_nat, assumption) apply (simp add: satisfies_is_a_fm_def satisfies_is_a_def sats_lambda_fm sats_bool_of_o_fm) done lemma satisfies_is_a_iff_sats: "\nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; u \ nat; x < length(env); y < length(env); z \ nat; env \ list(A)\ \ satisfies_is_a(##A,nu,nx,ny,nz) \ sats(A, satisfies_is_a_fm(u,x,y,z), env)" by simp theorem satisfies_is_a_reflection: "REFLECTS[\x. satisfies_is_a(L,f(x),g(x),h(x),g'(x)), \i x. satisfies_is_a(##Lset(i),f(x),g(x),h(x),g'(x))]" unfolding satisfies_is_a_def apply (intro FOL_reflections is_lambda_reflection bool_of_o_reflection nth_reflection is_list_reflection) done subsubsection\The Operator \<^term>\satisfies_is_b\, Internalized\ (* satisfies_is_b(M,A) \ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. is_bool_of_o(M, \nx[M]. is_nth(M,x,env,nx) \ is_nth(M,y,env,nx), z), zz) *) definition satisfies_is_b_fm :: "[i,i,i,i]\i" where "satisfies_is_b_fm(A,x,y,z) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( bool_of_o_fm(Exists(And(nth_fm(x#+5,2,0), nth_fm(y#+5,2,0))), 0), 0, succ(z))))" lemma satisfies_is_b_type [TC]: "\A \ nat; x \ nat; y \ nat; z \ nat\ \ satisfies_is_b_fm(A,x,y,z) \ formula" by (simp add: satisfies_is_b_fm_def) lemma sats_satisfies_is_b_fm [simp]: "\u \ nat; x < length(env); y < length(env); z \ nat; env \ list(A)\ \ sats(A, satisfies_is_b_fm(u,x,y,z), env) \ satisfies_is_b(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" apply (frule_tac x=x in lt_length_in_nat, assumption) apply (frule_tac x=y in lt_length_in_nat, assumption) apply (simp add: satisfies_is_b_fm_def satisfies_is_b_def sats_lambda_fm sats_bool_of_o_fm) done lemma satisfies_is_b_iff_sats: "\nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; u \ nat; x < length(env); y < length(env); z \ nat; env \ list(A)\ \ satisfies_is_b(##A,nu,nx,ny,nz) \ sats(A, satisfies_is_b_fm(u,x,y,z), env)" by simp theorem satisfies_is_b_reflection: "REFLECTS[\x. satisfies_is_b(L,f(x),g(x),h(x),g'(x)), \i x. satisfies_is_b(##Lset(i),f(x),g(x),h(x),g'(x))]" unfolding satisfies_is_b_def apply (intro FOL_reflections is_lambda_reflection bool_of_o_reflection nth_reflection is_list_reflection) done subsubsection\The Operator \<^term>\satisfies_is_c\, Internalized\ (* satisfies_is_c(M,A,h) \ \p q zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \hp[M]. \hq[M]. (\rp[M]. is_depth_apply(M,h,p,rp) \ fun_apply(M,rp,env,hp)) \ (\rq[M]. is_depth_apply(M,h,q,rq) \ fun_apply(M,rq,env,hq)) \ (\pq[M]. is_and(M,hp,hq,pq) \ is_not(M,pq,z)), zz) *) definition satisfies_is_c_fm :: "[i,i,i,i,i]\i" where "satisfies_is_c_fm(A,h,p,q,zz) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( Exists(Exists( And(Exists(And(depth_apply_fm(h#+7,p#+7,0), fun_apply_fm(0,4,2))), And(Exists(And(depth_apply_fm(h#+7,q#+7,0), fun_apply_fm(0,4,1))), Exists(And(and_fm(2,1,0), not_fm(0,3))))))), 0, succ(zz))))" lemma satisfies_is_c_type [TC]: "\A \ nat; h \ nat; x \ nat; y \ nat; z \ nat\ \ satisfies_is_c_fm(A,h,x,y,z) \ formula" by (simp add: satisfies_is_c_fm_def) lemma sats_satisfies_is_c_fm [simp]: "\u \ nat; v \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, satisfies_is_c_fm(u,v,x,y,z), env) \ satisfies_is_c(##A, nth(u,env), nth(v,env), nth(x,env), nth(y,env), nth(z,env))" by (simp add: satisfies_is_c_fm_def satisfies_is_c_def sats_lambda_fm) lemma satisfies_is_c_iff_sats: "\nth(u,env) = nu; nth(v,env) = nv; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; u \ nat; v \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ \ satisfies_is_c(##A,nu,nv,nx,ny,nz) \ sats(A, satisfies_is_c_fm(u,v,x,y,z), env)" by simp theorem satisfies_is_c_reflection: "REFLECTS[\x. satisfies_is_c(L,f(x),g(x),h(x),g'(x),h'(x)), \i x. satisfies_is_c(##Lset(i),f(x),g(x),h(x),g'(x),h'(x))]" unfolding satisfies_is_c_def apply (intro FOL_reflections function_reflections is_lambda_reflection extra_reflections nth_reflection depth_apply_reflection is_list_reflection) done subsubsection\The Operator \<^term>\satisfies_is_d\, Internalized\ (* satisfies_is_d(M,A,h) \ \p zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \rp[M]. is_depth_apply(M,h,p,rp) \ is_bool_of_o(M, \x[M]. \xenv[M]. \hp[M]. x\A \ is_Cons(M,x,env,xenv) \ fun_apply(M,rp,xenv,hp) \ number1(M,hp), z), zz) *) definition satisfies_is_d_fm :: "[i,i,i,i]\i" where "satisfies_is_d_fm(A,h,p,zz) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( Exists( And(depth_apply_fm(h#+5,p#+5,0), bool_of_o_fm( Forall(Forall(Forall( Implies(Member(2,A#+8), Implies(Cons_fm(2,5,1), Implies(fun_apply_fm(3,1,0), number1_fm(0))))))), 1))), 0, succ(zz))))" lemma satisfies_is_d_type [TC]: "\A \ nat; h \ nat; x \ nat; z \ nat\ \ satisfies_is_d_fm(A,h,x,z) \ formula" by (simp add: satisfies_is_d_fm_def) lemma sats_satisfies_is_d_fm [simp]: "\u \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, satisfies_is_d_fm(u,x,y,z), env) \ satisfies_is_d(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" by (simp add: satisfies_is_d_fm_def satisfies_is_d_def sats_lambda_fm sats_bool_of_o_fm) lemma satisfies_is_d_iff_sats: "\nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; u \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ \ satisfies_is_d(##A,nu,nx,ny,nz) \ sats(A, satisfies_is_d_fm(u,x,y,z), env)" by simp theorem satisfies_is_d_reflection: "REFLECTS[\x. satisfies_is_d(L,f(x),g(x),h(x),g'(x)), \i x. satisfies_is_d(##Lset(i),f(x),g(x),h(x),g'(x))]" unfolding satisfies_is_d_def apply (intro FOL_reflections function_reflections is_lambda_reflection extra_reflections nth_reflection depth_apply_reflection is_list_reflection) done subsubsection\The Operator \<^term>\satisfies_MH\, Internalized\ (* satisfies_MH \ \M A u f zz. \fml[M]. is_formula(M,fml) \ is_lambda (M, fml, is_formula_case (M, satisfies_is_a(M,A), satisfies_is_b(M,A), satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)), zz) *) definition satisfies_MH_fm :: "[i,i,i,i]\i" where "satisfies_MH_fm(A,u,f,zz) \ Forall( Implies(is_formula_fm(0), lambda_fm( formula_case_fm(satisfies_is_a_fm(A#+7,2,1,0), satisfies_is_b_fm(A#+7,2,1,0), satisfies_is_c_fm(A#+7,f#+7,2,1,0), satisfies_is_d_fm(A#+6,f#+6,1,0), 1, 0), 0, succ(zz))))" lemma satisfies_MH_type [TC]: "\A \ nat; u \ nat; x \ nat; z \ nat\ \ satisfies_MH_fm(A,u,x,z) \ formula" by (simp add: satisfies_MH_fm_def) lemma sats_satisfies_MH_fm [simp]: "\u \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ \ sats(A, satisfies_MH_fm(u,x,y,z), env) \ satisfies_MH(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" by (simp add: satisfies_MH_fm_def satisfies_MH_def sats_lambda_fm sats_formula_case_fm) lemma satisfies_MH_iff_sats: "\nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; u \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ \ satisfies_MH(##A,nu,nx,ny,nz) \ sats(A, satisfies_MH_fm(u,x,y,z), env)" by simp lemmas satisfies_reflections = is_lambda_reflection is_formula_reflection is_formula_case_reflection satisfies_is_a_reflection satisfies_is_b_reflection satisfies_is_c_reflection satisfies_is_d_reflection theorem satisfies_MH_reflection: "REFLECTS[\x. satisfies_MH(L,f(x),g(x),h(x),g'(x)), \i x. satisfies_MH(##Lset(i),f(x),g(x),h(x),g'(x))]" unfolding satisfies_MH_def apply (intro FOL_reflections satisfies_reflections) done subsection\Lemmas for Instantiating the Locale \M_satisfies\\ subsubsection\The \<^term>\Member\ Case\ lemma Member_Reflects: "REFLECTS[\u. \v[L]. v \ B \ (\bo[L]. \nx[L]. \ny[L]. v \ lstA \ is_nth(L,x,v,nx) \ is_nth(L,y,v,ny) \ is_bool_of_o(L, nx \ ny, bo) \ pair(L,v,bo,u)), \i u. \v \ Lset(i). v \ B \ (\bo \ Lset(i). \nx \ Lset(i). \ny \ Lset(i). v \ lstA \ is_nth(##Lset(i), x, v, nx) \ is_nth(##Lset(i), y, v, ny) \ is_bool_of_o(##Lset(i), nx \ ny, bo) \ pair(##Lset(i), v, bo, u))]" by (intro FOL_reflections function_reflections nth_reflection bool_of_o_reflection) lemma Member_replacement: "\L(A); x \ nat; y \ nat\ \ strong_replacement (L, \env z. \bo[L]. \nx[L]. \ny[L]. env \ list(A) \ is_nth(L,x,env,nx) \ is_nth(L,y,env,ny) \ is_bool_of_o(L, nx \ ny, bo) \ pair(L, env, bo, z))" apply (rule strong_replacementI) apply (rule_tac u="{list(A),B,x,y}" in gen_separation_multi [OF Member_Reflects], auto) apply (rule_tac env="[list(A),B,x,y]" in DPow_LsetI) apply (rule sep_rules nth_iff_sats is_bool_of_o_iff_sats | simp)+ done subsubsection\The \<^term>\Equal\ Case\ lemma Equal_Reflects: "REFLECTS[\u. \v[L]. v \ B \ (\bo[L]. \nx[L]. \ny[L]. v \ lstA \ is_nth(L, x, v, nx) \ is_nth(L, y, v, ny) \ is_bool_of_o(L, nx = ny, bo) \ pair(L, v, bo, u)), \i u. \v \ Lset(i). v \ B \ (\bo \ Lset(i). \nx \ Lset(i). \ny \ Lset(i). v \ lstA \ is_nth(##Lset(i), x, v, nx) \ is_nth(##Lset(i), y, v, ny) \ is_bool_of_o(##Lset(i), nx = ny, bo) \ pair(##Lset(i), v, bo, u))]" by (intro FOL_reflections function_reflections nth_reflection bool_of_o_reflection) lemma Equal_replacement: "\L(A); x \ nat; y \ nat\ \ strong_replacement (L, \env z. \bo[L]. \nx[L]. \ny[L]. env \ list(A) \ is_nth(L,x,env,nx) \ is_nth(L,y,env,ny) \ is_bool_of_o(L, nx = ny, bo) \ pair(L, env, bo, z))" apply (rule strong_replacementI) apply (rule_tac u="{list(A),B,x,y}" in gen_separation_multi [OF Equal_Reflects], auto) apply (rule_tac env="[list(A),B,x,y]" in DPow_LsetI) apply (rule sep_rules nth_iff_sats is_bool_of_o_iff_sats | simp)+ done subsubsection\The \<^term>\Nand\ Case\ lemma Nand_Reflects: "REFLECTS [\x. \u[L]. u \ B \ (\rpe[L]. \rqe[L]. \andpq[L]. \notpq[L]. fun_apply(L, rp, u, rpe) \ fun_apply(L, rq, u, rqe) \ is_and(L, rpe, rqe, andpq) \ is_not(L, andpq, notpq) \ u \ list(A) \ pair(L, u, notpq, x)), \i x. \u \ Lset(i). u \ B \ (\rpe \ Lset(i). \rqe \ Lset(i). \andpq \ Lset(i). \notpq \ Lset(i). fun_apply(##Lset(i), rp, u, rpe) \ fun_apply(##Lset(i), rq, u, rqe) \ is_and(##Lset(i), rpe, rqe, andpq) \ is_not(##Lset(i), andpq, notpq) \ u \ list(A) \ pair(##Lset(i), u, notpq, x))]" -apply (unfold is_and_def is_not_def) + unfolding is_and_def is_not_def apply (intro FOL_reflections function_reflections) done lemma Nand_replacement: "\L(A); L(rp); L(rq)\ \ strong_replacement (L, \env z. \rpe[L]. \rqe[L]. \andpq[L]. \notpq[L]. fun_apply(L,rp,env,rpe) \ fun_apply(L,rq,env,rqe) \ is_and(L,rpe,rqe,andpq) \ is_not(L,andpq,notpq) \ env \ list(A) \ pair(L, env, notpq, z))" apply (rule strong_replacementI) apply (rule_tac u="{list(A),B,rp,rq}" in gen_separation_multi [OF Nand_Reflects], auto) apply (rule_tac env="[list(A),B,rp,rq]" in DPow_LsetI) apply (rule sep_rules is_and_iff_sats is_not_iff_sats | simp)+ done subsubsection\The \<^term>\Forall\ Case\ lemma Forall_Reflects: "REFLECTS [\x. \u[L]. u \ B \ (\bo[L]. u \ list(A) \ is_bool_of_o (L, \a[L]. \co[L]. \rpco[L]. a \ A \ is_Cons(L,a,u,co) \ fun_apply(L,rp,co,rpco) \ number1(L,rpco), bo) \ pair(L,u,bo,x)), \i x. \u \ Lset(i). u \ B \ (\bo \ Lset(i). u \ list(A) \ is_bool_of_o (##Lset(i), \a \ Lset(i). \co \ Lset(i). \rpco \ Lset(i). a \ A \ is_Cons(##Lset(i),a,u,co) \ fun_apply(##Lset(i),rp,co,rpco) \ number1(##Lset(i),rpco), bo) \ pair(##Lset(i),u,bo,x))]" unfolding is_bool_of_o_def apply (intro FOL_reflections function_reflections Cons_reflection) done lemma Forall_replacement: "\L(A); L(rp)\ \ strong_replacement (L, \env z. \bo[L]. env \ list(A) \ is_bool_of_o (L, \a[L]. \co[L]. \rpco[L]. a\A \ is_Cons(L,a,env,co) \ fun_apply(L,rp,co,rpco) \ number1(L, rpco), bo) \ pair(L,env,bo,z))" apply (rule strong_replacementI) apply (rule_tac u="{A,list(A),B,rp}" in gen_separation_multi [OF Forall_Reflects], auto) apply (rule_tac env="[A,list(A),B,rp]" in DPow_LsetI) apply (rule sep_rules is_bool_of_o_iff_sats Cons_iff_sats | simp)+ done subsubsection\The \<^term>\transrec_replacement\ Case\ lemma formula_rec_replacement_Reflects: "REFLECTS [\x. \u[L]. u \ B \ (\y[L]. pair(L, u, y, x) \ is_wfrec (L, satisfies_MH(L,A), mesa, u, y)), \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(##Lset(i), u, y, x) \ is_wfrec (##Lset(i), satisfies_MH(##Lset(i),A), mesa, u, y))]" by (intro FOL_reflections function_reflections satisfies_MH_reflection is_wfrec_reflection) lemma formula_rec_replacement: \ \For the \<^term>\transrec\\ "\n \ nat; L(A)\ \ transrec_replacement(L, satisfies_MH(L,A), n)" apply (rule L.transrec_replacementI, simp add: L.nat_into_M) apply (rule strong_replacementI) apply (rule_tac u="{B,A,n,Memrel(eclose({n}))}" in gen_separation_multi [OF formula_rec_replacement_Reflects], auto simp add: L.nat_into_M) apply (rule_tac env="[B,A,n,Memrel(eclose({n}))]" in DPow_LsetI) apply (rule sep_rules satisfies_MH_iff_sats is_wfrec_iff_sats | simp)+ done subsubsection\The Lambda Replacement Case\ lemma formula_rec_lambda_replacement_Reflects: "REFLECTS [\x. \u[L]. u \ B \ mem_formula(L,u) \ (\c[L]. is_formula_case (L, satisfies_is_a(L,A), satisfies_is_b(L,A), satisfies_is_c(L,A,g), satisfies_is_d(L,A,g), u, c) \ pair(L,u,c,x)), \i x. \u \ Lset(i). u \ B \ mem_formula(##Lset(i),u) \ (\c \ Lset(i). is_formula_case (##Lset(i), satisfies_is_a(##Lset(i),A), satisfies_is_b(##Lset(i),A), satisfies_is_c(##Lset(i),A,g), satisfies_is_d(##Lset(i),A,g), u, c) \ pair(##Lset(i),u,c,x))]" by (intro FOL_reflections function_reflections mem_formula_reflection is_formula_case_reflection satisfies_is_a_reflection satisfies_is_b_reflection satisfies_is_c_reflection satisfies_is_d_reflection) lemma formula_rec_lambda_replacement: \ \For the \<^term>\transrec\\ "\L(g); L(A)\ \ strong_replacement (L, \x y. mem_formula(L,x) \ (\c[L]. is_formula_case(L, satisfies_is_a(L,A), satisfies_is_b(L,A), satisfies_is_c(L,A,g), satisfies_is_d(L,A,g), x, c) \ pair(L, x, c, y)))" apply (rule strong_replacementI) apply (rule_tac u="{B,A,g}" in gen_separation_multi [OF formula_rec_lambda_replacement_Reflects], auto) apply (rule_tac env="[A,g,B]" in DPow_LsetI) apply (rule sep_rules mem_formula_iff_sats formula_case_iff_sats satisfies_is_a_iff_sats satisfies_is_b_iff_sats satisfies_is_c_iff_sats satisfies_is_d_iff_sats | simp)+ done subsection\Instantiating \M_satisfies\\ lemma M_satisfies_axioms_L: "M_satisfies_axioms(L)" apply (rule M_satisfies_axioms.intro) apply (assumption | rule Member_replacement Equal_replacement Nand_replacement Forall_replacement formula_rec_replacement formula_rec_lambda_replacement)+ done theorem M_satisfies_L: "M_satisfies(L)" apply (rule M_satisfies.intro) apply (rule M_eclose_L) apply (rule M_satisfies_axioms_L) done text\Finally: the point of the whole theory!\ lemmas satisfies_closed = M_satisfies.satisfies_closed [OF M_satisfies_L] and satisfies_abs = M_satisfies.satisfies_abs [OF M_satisfies_L] end diff --git a/src/ZF/Epsilon.thy b/src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy +++ b/src/ZF/Epsilon.thy @@ -1,399 +1,399 @@ (* Title: ZF/Epsilon.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge *) section\Epsilon Induction and Recursion\ theory Epsilon imports Nat begin definition eclose :: "i\i" where "eclose(A) \ \n\nat. nat_rec(n, A, \m r. \(r))" definition transrec :: "[i, [i,i]\i] \i" where "transrec(a,H) \ wfrec(Memrel(eclose({a})), a, H)" definition rank :: "i\i" where "rank(a) \ transrec(a, \x f. \y\x. succ(f`y))" definition transrec2 :: "[i, i, [i,i]\i] \i" where "transrec2(k, a, b) \ transrec(k, \i r. if(i=0, a, if(\j. i=succ(j), b(THE j. i=succ(j), r`(THE j. i=succ(j))), \ji, i]\i" where "recursor(a,b,k) \ transrec(k, \n f. nat_case(a, \m. b(m, f`m), n))" definition rec :: "[i, i, [i,i]\i]\i" where "rec(k,a,b) \ recursor(a,b,k)" subsection\Basic Closure Properties\ lemma arg_subset_eclose: "A \ eclose(A)" unfolding eclose_def apply (rule nat_rec_0 [THEN equalityD2, THEN subset_trans]) apply (rule nat_0I [THEN UN_upper]) done lemmas arg_into_eclose = arg_subset_eclose [THEN subsetD] lemma Transset_eclose: "Transset(eclose(A))" -apply (unfold eclose_def Transset_def) + unfolding eclose_def Transset_def apply (rule subsetI [THEN ballI]) apply (erule UN_E) apply (rule nat_succI [THEN UN_I], assumption) apply (erule nat_rec_succ [THEN ssubst]) apply (erule UnionI, assumption) done (* @{term"x \ eclose(A) \ x \ eclose(A)"} *) lemmas eclose_subset = Transset_eclose [unfolded Transset_def, THEN bspec] (* @{term"\A \ eclose(B); c \ A\ \ c \ eclose(B)"} *) lemmas ecloseD = eclose_subset [THEN subsetD] lemmas arg_in_eclose_sing = arg_subset_eclose [THEN singleton_subsetD] lemmas arg_into_eclose_sing = arg_in_eclose_sing [THEN ecloseD] (* This is epsilon-induction for eclose(A); see also eclose_induct_down... \a \ eclose(A); \x. \x \ eclose(A); \y\x. P(y)\ \ P(x) \ \ P(a) *) lemmas eclose_induct = Transset_induct [OF _ Transset_eclose, induct set: eclose] (*Epsilon induction*) lemma eps_induct: "\\x. \y\x. P(y) \ P(x)\ \ P(a)" by (rule arg_in_eclose_sing [THEN eclose_induct], blast) subsection\Leastness of \<^term>\eclose\\ (** eclose(A) is the least transitive set including A as a subset. **) lemma eclose_least_lemma: "\Transset(X); A<=X; n \ nat\ \ nat_rec(n, A, \m r. \(r)) \ X" unfolding Transset_def apply (erule nat_induct) apply (simp add: nat_rec_0) apply (simp add: nat_rec_succ, blast) done lemma eclose_least: "\Transset(X); A<=X\ \ eclose(A) \ X" unfolding eclose_def apply (rule eclose_least_lemma [THEN UN_least], assumption+) done (*COMPLETELY DIFFERENT induction principle from eclose_induct\*) lemma eclose_induct_down [consumes 1]: "\a \ eclose(b); \y. \y \ b\ \ P(y); \y z. \y \ eclose(b); P(y); z \ y\ \ P(z) \ \ P(a)" apply (rule eclose_least [THEN subsetD, THEN CollectD2, of "eclose(b)"]) prefer 3 apply assumption unfolding Transset_def apply (blast intro: ecloseD) apply (blast intro: arg_subset_eclose [THEN subsetD]) done lemma Transset_eclose_eq_arg: "Transset(X) \ eclose(X) = X" apply (erule equalityI [OF eclose_least arg_subset_eclose]) apply (rule subset_refl) done text\A transitive set either is empty or contains the empty set.\ lemma Transset_0_lemma [rule_format]: "Transset(A) \ x\A \ 0\A" apply (simp add: Transset_def) apply (rule_tac a=x in eps_induct, clarify) apply (drule bspec, assumption) apply (case_tac "x=0", auto) done lemma Transset_0_disj: "Transset(A) \ A=0 | 0\A" by (blast dest: Transset_0_lemma) subsection\Epsilon Recursion\ (*Unused...*) lemma mem_eclose_trans: "\A \ eclose(B); B \ eclose(C)\ \ A \ eclose(C)" by (rule eclose_least [OF Transset_eclose eclose_subset, THEN subsetD], assumption+) (*Variant of the previous lemma in a useable form for the sequel*) lemma mem_eclose_sing_trans: "\A \ eclose({B}); B \ eclose({C})\ \ A \ eclose({C})" by (rule eclose_least [OF Transset_eclose singleton_subsetI, THEN subsetD], assumption+) lemma under_Memrel: "\Transset(i); j \ i\ \ Memrel(i)-``{j} = j" by (unfold Transset_def, blast) lemma lt_Memrel: "j < i \ Memrel(i) -`` {j} = j" by (simp add: lt_def Ord_def under_Memrel) (* @{term"j \ eclose(A) \ Memrel(eclose(A)) -`` j = j"} *) lemmas under_Memrel_eclose = Transset_eclose [THEN under_Memrel] lemmas wfrec_ssubst = wf_Memrel [THEN wfrec, THEN ssubst] lemma wfrec_eclose_eq: "\k \ eclose({j}); j \ eclose({i})\ \ wfrec(Memrel(eclose({i})), k, H) = wfrec(Memrel(eclose({j})), k, H)" apply (erule eclose_induct) apply (rule wfrec_ssubst) apply (rule wfrec_ssubst) apply (simp add: under_Memrel_eclose mem_eclose_sing_trans [of _ j i]) done lemma wfrec_eclose_eq2: "k \ i \ wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)" apply (rule arg_in_eclose_sing [THEN wfrec_eclose_eq]) apply (erule arg_into_eclose_sing) done lemma transrec: "transrec(a,H) = H(a, \x\a. transrec(x,H))" unfolding transrec_def apply (rule wfrec_ssubst) apply (simp add: wfrec_eclose_eq2 arg_in_eclose_sing under_Memrel_eclose) done (*Avoids explosions in proofs; resolve it with a meta-level definition.*) lemma def_transrec: "\\x. f(x)\transrec(x,H)\ \ f(a) = H(a, \x\a. f(x))" apply simp apply (rule transrec) done lemma transrec_type: "\\x u. \x \ eclose({a}); u \ Pi(x,B)\ \ H(x,u) \ B(x)\ \ transrec(a,H) \ B(a)" apply (rule_tac i = a in arg_in_eclose_sing [THEN eclose_induct]) apply (subst transrec) apply (simp add: lam_type) done lemma eclose_sing_Ord: "Ord(i) \ eclose({i}) \ succ(i)" apply (erule Ord_is_Transset [THEN Transset_succ, THEN eclose_least]) apply (rule succI1 [THEN singleton_subsetI]) done lemma succ_subset_eclose_sing: "succ(i) \ eclose({i})" apply (insert arg_subset_eclose [of "{i}"], simp) apply (frule eclose_subset, blast) done lemma eclose_sing_Ord_eq: "Ord(i) \ eclose({i}) = succ(i)" apply (rule equalityI) apply (erule eclose_sing_Ord) apply (rule succ_subset_eclose_sing) done lemma Ord_transrec_type: assumes jini: "j \ i" and ordi: "Ord(i)" and minor: " \x u. \x \ i; u \ Pi(x,B)\ \ H(x,u) \ B(x)" shows "transrec(j,H) \ B(j)" apply (rule transrec_type) apply (insert jini ordi) apply (blast intro!: minor intro: Ord_trans dest: Ord_in_Ord [THEN eclose_sing_Ord, THEN subsetD]) done subsection\Rank\ (*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) lemma rank: "rank(a) = (\y\a. succ(rank(y)))" by (subst rank_def [THEN def_transrec], simp) lemma Ord_rank [simp]: "Ord(rank(a))" apply (rule_tac a=a in eps_induct) apply (subst rank) apply (rule Ord_succ [THEN Ord_UN]) apply (erule bspec, assumption) done lemma rank_of_Ord: "Ord(i) \ rank(i) = i" apply (erule trans_induct) apply (subst rank) apply (simp add: Ord_equality) done lemma rank_lt: "a \ b \ rank(a) < rank(b)" apply (rule_tac a1 = b in rank [THEN ssubst]) apply (erule UN_I [THEN ltI]) apply (rule_tac [2] Ord_UN, auto) done lemma eclose_rank_lt: "a \ eclose(b) \ rank(a) < rank(b)" apply (erule eclose_induct_down) apply (erule rank_lt) apply (erule rank_lt [THEN lt_trans], assumption) done lemma rank_mono: "a<=b \ rank(a) \ rank(b)" apply (rule subset_imp_le) apply (auto simp add: rank [of a] rank [of b]) done lemma rank_Pow: "rank(Pow(a)) = succ(rank(a))" apply (rule rank [THEN trans]) apply (rule le_anti_sym) apply (rule_tac [2] UN_upper_le) apply (rule UN_least_le) apply (auto intro: rank_mono simp add: Ord_UN) done lemma rank_0 [simp]: "rank(0) = 0" by (rule rank [THEN trans], blast) lemma rank_succ [simp]: "rank(succ(x)) = succ(rank(x))" apply (rule rank [THEN trans]) apply (rule equalityI [OF UN_least succI1 [THEN UN_upper]]) apply (erule succE, blast) apply (erule rank_lt [THEN leI, THEN succ_leI, THEN le_imp_subset]) done lemma rank_Union: "rank(\(A)) = (\x\A. rank(x))" apply (rule equalityI) apply (rule_tac [2] rank_mono [THEN le_imp_subset, THEN UN_least]) apply (erule_tac [2] Union_upper) apply (subst rank) apply (rule UN_least) apply (erule UnionE) apply (rule subset_trans) apply (erule_tac [2] RepFunI [THEN Union_upper]) apply (erule rank_lt [THEN succ_leI, THEN le_imp_subset]) done lemma rank_eclose: "rank(eclose(a)) = rank(a)" apply (rule le_anti_sym) apply (rule_tac [2] arg_subset_eclose [THEN rank_mono]) apply (rule_tac a1 = "eclose (a) " in rank [THEN ssubst]) apply (rule Ord_rank [THEN UN_least_le]) apply (erule eclose_rank_lt [THEN succ_leI]) done lemma rank_pair1: "rank(a) < rank(\a,b\)" unfolding Pair_def apply (rule consI1 [THEN rank_lt, THEN lt_trans]) apply (rule consI1 [THEN consI2, THEN rank_lt]) done lemma rank_pair2: "rank(b) < rank(\a,b\)" unfolding Pair_def apply (rule consI1 [THEN consI2, THEN rank_lt, THEN lt_trans]) apply (rule consI1 [THEN consI2, THEN rank_lt]) done (*Not clear how to remove the P(a) condition, since the "then" part must refer to "a"*) lemma the_equality_if: "P(a) \ (THE x. P(x)) = (if (\!x. P(x)) then a else 0)" by (simp add: the_0 the_equality2) (*The first premise not only fixs i but ensures @{term"f\0"}. The second premise is now essential. Consider otherwise the relation r = {\0,0\,\0,1\,\0,2\,...}. Then f`0 = \(f``{0}) = \(nat) = nat, whose rank equals that of r.*) lemma rank_apply: "\i \ domain(f); function(f)\ \ rank(f`i) < rank(f)" apply clarify apply (simp add: function_apply_equality) apply (blast intro: lt_trans rank_lt rank_pair2) done subsection\Corollaries of Leastness\ lemma mem_eclose_subset: "A \ B \ eclose(A)<=eclose(B)" apply (rule Transset_eclose [THEN eclose_least]) apply (erule arg_into_eclose [THEN eclose_subset]) done lemma eclose_mono: "A<=B \ eclose(A) \ eclose(B)" apply (rule Transset_eclose [THEN eclose_least]) apply (erule subset_trans) apply (rule arg_subset_eclose) done (** Idempotence of eclose **) lemma eclose_idem: "eclose(eclose(A)) = eclose(A)" apply (rule equalityI) apply (rule eclose_least [OF Transset_eclose subset_refl]) apply (rule arg_subset_eclose) done (** Transfinite recursion for definitions based on the three cases of ordinals **) lemma transrec2_0 [simp]: "transrec2(0,a,b) = a" by (rule transrec2_def [THEN def_transrec, THEN trans], simp) lemma transrec2_succ [simp]: "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))" apply (rule transrec2_def [THEN def_transrec, THEN trans]) apply (simp add: the_equality if_P) done lemma transrec2_Limit: "Limit(i) \ transrec2(i,a,b) = (\jx. f(x)\transrec2(x,a,b)) \ f(0) = a \ f(succ(i)) = b(i, f(i)) \ (Limit(K) \ f(K) = (\jn \ nat; a \ C(0); \m z. \m \ nat; z \ C(m)\ \ b(m,z): C(succ(m))\ \ rec(n,a,b) \ C(n)" by (erule nat_induct, auto) end diff --git a/src/ZF/EquivClass.thy b/src/ZF/EquivClass.thy --- a/src/ZF/EquivClass.thy +++ b/src/ZF/EquivClass.thy @@ -1,233 +1,233 @@ (* Title: ZF/EquivClass.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Equivalence Relations\ theory EquivClass imports Trancl Perm begin definition quotient :: "[i,i]\i" (infixl \'/'/\ 90) (*set of equiv classes*) where "A//r \ {r``{x} . x \ A}" definition congruent :: "[i,i\i]\o" where "congruent(r,b) \ \y z. \y,z\:r \ b(y)=b(z)" definition congruent2 :: "[i,i,[i,i]\i]\o" where "congruent2(r1,r2,b) \ \y1 z1 y2 z2. \y1,z1\:r1 \ \y2,z2\:r2 \ b(y1,y2) = b(z1,z2)" abbreviation RESPECTS ::"[i\i, i] \ o" (infixr \respects\ 80) where "f respects r \ congruent(r,f)" abbreviation RESPECTS2 ::"[i\i\i, i] \ o" (infixr \respects2 \ 80) where "f respects2 r \ congruent2(r,r,f)" \ \Abbreviation for the common case where the relations are identical\ subsection\Suppes, Theorem 70: \<^term>\r\ is an equiv relation iff \<^term>\converse(r) O r = r\\ (** first half: equiv(A,r) \ converse(r) O r = r **) lemma sym_trans_comp_subset: "\sym(r); trans(r)\ \ converse(r) O r \ r" by (unfold trans_def sym_def, blast) lemma refl_comp_subset: "\refl(A,r); r \ A*A\ \ r \ converse(r) O r" by (unfold refl_def, blast) lemma equiv_comp_eq: "equiv(A,r) \ converse(r) O r = r" unfolding equiv_def apply (blast del: subsetI intro!: sym_trans_comp_subset refl_comp_subset) done (*second half*) lemma comp_equivI: "\converse(r) O r = r; domain(r) = A\ \ equiv(A,r)" -apply (unfold equiv_def refl_def sym_def trans_def) + unfolding equiv_def refl_def sym_def trans_def apply (erule equalityE) apply (subgoal_tac "\x y. \x,y\ \ r \ \y,x\ \ r", blast+) done (** Equivalence classes **) (*Lemma for the next result*) lemma equiv_class_subset: "\sym(r); trans(r); \a,b\: r\ \ r``{a} \ r``{b}" by (unfold trans_def sym_def, blast) lemma equiv_class_eq: "\equiv(A,r); \a,b\: r\ \ r``{a} = r``{b}" unfolding equiv_def apply (safe del: subsetI intro!: equalityI equiv_class_subset) apply (unfold sym_def, blast) done lemma equiv_class_self: "\equiv(A,r); a \ A\ \ a \ r``{a}" by (unfold equiv_def refl_def, blast) (*Lemma for the next result*) lemma subset_equiv_class: "\equiv(A,r); r``{b} \ r``{a}; b \ A\ \ \a,b\: r" by (unfold equiv_def refl_def, blast) lemma eq_equiv_class: "\r``{a} = r``{b}; equiv(A,r); b \ A\ \ \a,b\: r" by (assumption | rule equalityD2 subset_equiv_class)+ (*thus r``{a} = r``{b} as well*) lemma equiv_class_nondisjoint: "\equiv(A,r); x: (r``{a} \ r``{b})\ \ \a,b\: r" by (unfold equiv_def trans_def sym_def, blast) lemma equiv_type: "equiv(A,r) \ r \ A*A" by (unfold equiv_def, blast) lemma equiv_class_eq_iff: "equiv(A,r) \ \x,y\: r \ r``{x} = r``{y} \ x \ A \ y \ A" by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) lemma eq_equiv_class_iff: "\equiv(A,r); x \ A; y \ A\ \ r``{x} = r``{y} \ \x,y\: r" by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) (*** Quotients ***) (** Introduction/elimination rules -- needed? **) lemma quotientI [TC]: "x \ A \ r``{x}: A//r" unfolding quotient_def apply (erule RepFunI) done lemma quotientE: "\X \ A//r; \x. \X = r``{x}; x \ A\ \ P\ \ P" by (unfold quotient_def, blast) lemma Union_quotient: "equiv(A,r) \ \(A//r) = A" by (unfold equiv_def refl_def quotient_def, blast) lemma quotient_disj: "\equiv(A,r); X \ A//r; Y \ A//r\ \ X=Y | (X \ Y \ 0)" unfolding quotient_def apply (safe intro!: equiv_class_eq, assumption) apply (unfold equiv_def trans_def sym_def, blast) done subsection\Defining Unary Operations upon Equivalence Classes\ (** Could have a locale with the premises equiv(A,r) and congruent(r,b) **) (*Conversion rule*) lemma UN_equiv_class: "\equiv(A,r); b respects r; a \ A\ \ (\x\r``{a}. b(x)) = b(a)" apply (subgoal_tac "\x \ r``{a}. b(x) = b(a)") apply simp apply (blast intro: equiv_class_self) apply (unfold equiv_def sym_def congruent_def, blast) done (*type checking of @{term"\x\r``{a}. b(x)"} *) lemma UN_equiv_class_type: "\equiv(A,r); b respects r; X \ A//r; \x. x \ A \ b(x) \ B\ \ (\x\X. b(x)) \ B" apply (unfold quotient_def, safe) apply (simp (no_asm_simp) add: UN_equiv_class) done (*Sufficient conditions for injectiveness. Could weaken premises! major premise could be an inclusion; bcong could be \y. y \ A \ b(y):B *) lemma UN_equiv_class_inject: "\equiv(A,r); b respects r; (\x\X. b(x))=(\y\Y. b(y)); X \ A//r; Y \ A//r; \x y. \x \ A; y \ A; b(x)=b(y)\ \ \x,y\:r\ \ X=Y" apply (unfold quotient_def, safe) apply (rule equiv_class_eq, assumption) apply (simp add: UN_equiv_class [of A r b]) done subsection\Defining Binary Operations upon Equivalence Classes\ lemma congruent2_implies_congruent: "\equiv(A,r1); congruent2(r1,r2,b); a \ A\ \ congruent(r2,b(a))" by (unfold congruent_def congruent2_def equiv_def refl_def, blast) lemma congruent2_implies_congruent_UN: "\equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a \ A2\ \ congruent(r1, \x1. \x2 \ r2``{a}. b(x1,x2))" apply (unfold congruent_def, safe) apply (frule equiv_type [THEN subsetD], assumption) apply clarify apply (simp add: UN_equiv_class congruent2_implies_congruent) apply (unfold congruent2_def equiv_def refl_def, blast) done lemma UN_equiv_class2: "\equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a1: A1; a2: A2\ \ (\x1 \ r1``{a1}. \x2 \ r2``{a2}. b(x1,x2)) = b(a1,a2)" by (simp add: UN_equiv_class congruent2_implies_congruent congruent2_implies_congruent_UN) (*type checking*) lemma UN_equiv_class_type2: "\equiv(A,r); b respects2 r; X1: A//r; X2: A//r; \x1 x2. \x1: A; x2: A\ \ b(x1,x2) \ B \ \ (\x1\X1. \x2\X2. b(x1,x2)) \ B" apply (unfold quotient_def, safe) apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN congruent2_implies_congruent quotientI) done (*Suggested by John Harrison -- the two subproofs may be MUCH simpler than the direct proof*) lemma congruent2I: "\equiv(A1,r1); equiv(A2,r2); \y z w. \w \ A2; \y,z\ \ r1\ \ b(y,w) = b(z,w); \y z w. \w \ A1; \y,z\ \ r2\ \ b(w,y) = b(w,z) \ \ congruent2(r1,r2,b)" apply (unfold congruent2_def equiv_def refl_def, safe) apply (blast intro: trans) done lemma congruent2_commuteI: assumes equivA: "equiv(A,r)" and commute: "\y z. \y \ A; z \ A\ \ b(y,z) = b(z,y)" and congt: "\y z w. \w \ A; \y,z\: r\ \ b(w,y) = b(w,z)" shows "b respects2 r" apply (insert equivA [THEN equiv_type, THEN subsetD]) apply (rule congruent2I [OF equivA equivA]) apply (rule commute [THEN trans]) apply (rule_tac [3] commute [THEN trans, symmetric]) apply (rule_tac [5] sym) apply (blast intro: congt)+ done (*Obsolete?*) lemma congruent_commuteI: "\equiv(A,r); Z \ A//r; \w. \w \ A\ \ congruent(r, \z. b(w,z)); \x y. \x \ A; y \ A\ \ b(y,x) = b(x,y) \ \ congruent(r, \w. \z\Z. b(w,z))" apply (simp (no_asm) add: congruent_def) apply (safe elim!: quotientE) apply (frule equiv_type [THEN subsetD], assumption) apply (simp add: UN_equiv_class [of A r]) apply (simp add: congruent_def) done end diff --git a/src/ZF/Finite.thy b/src/ZF/Finite.thy --- a/src/ZF/Finite.thy +++ b/src/ZF/Finite.thy @@ -1,209 +1,209 @@ (* Title: ZF/Finite.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge prove: b \ Fin(A) \ inj(b,b) \ surj(b,b) *) section\Finite Powerset Operator and Finite Function Space\ theory Finite imports Inductive Epsilon Nat begin (*The natural numbers as a datatype*) rep_datatype elimination natE induction nat_induct case_eqns nat_case_0 nat_case_succ recursor_eqns recursor_0 recursor_succ consts Fin :: "i\i" FiniteFun :: "[i,i]\i" (\(_ -||>/ _)\ [61, 60] 60) inductive domains "Fin(A)" \ "Pow(A)" intros emptyI: "0 \ Fin(A)" consI: "\a \ A; b \ Fin(A)\ \ cons(a,b) \ Fin(A)" type_intros empty_subsetI cons_subsetI PowI type_elims PowD [elim_format] inductive domains "FiniteFun(A,B)" \ "Fin(A*B)" intros emptyI: "0 \ A -||> B" consI: "\a \ A; b \ B; h \ A -||> B; a \ domain(h)\ \ cons(\a,b\,h) \ A -||> B" type_intros Fin.intros subsection \Finite Powerset Operator\ lemma Fin_mono: "A<=B \ Fin(A) \ Fin(B)" -apply (unfold Fin.defs) + unfolding Fin.defs apply (rule lfp_mono) apply (rule Fin.bnd_mono)+ apply blast done (* @{term"A \ Fin(B) \ A \ B"} *) lemmas FinD = Fin.dom_subset [THEN subsetD, THEN PowD] (** Induction on finite sets **) (*Discharging @{term"x\y"} entails extra work*) lemma Fin_induct [case_names 0 cons, induct set: Fin]: "\b \ Fin(A); P(0); \x y. \x \ A; y \ Fin(A); x\y; P(y)\ \ P(cons(x,y)) \ \ P(b)" apply (erule Fin.induct, simp) apply (case_tac "a \ b") apply (erule cons_absorb [THEN ssubst], assumption) (*backtracking!*) apply simp done (** Simplification for Fin **) declare Fin.intros [simp] lemma Fin_0: "Fin(0) = {0}" by (blast intro: Fin.emptyI dest: FinD) (*The union of two finite sets is finite.*) lemma Fin_UnI [simp]: "\b \ Fin(A); c \ Fin(A)\ \ b \ c \ Fin(A)" apply (erule Fin_induct) apply (simp_all add: Un_cons) done (*The union of a set of finite sets is finite.*) lemma Fin_UnionI: "C \ Fin(Fin(A)) \ \(C) \ Fin(A)" by (erule Fin_induct, simp_all) (*Every subset of a finite set is finite.*) lemma Fin_subset_lemma [rule_format]: "b \ Fin(A) \ \z. z<=b \ z \ Fin(A)" apply (erule Fin_induct) apply (simp add: subset_empty_iff) apply (simp add: subset_cons_iff distrib_simps, safe) apply (erule_tac b = z in cons_Diff [THEN subst], simp) done lemma Fin_subset: "\c<=b; b \ Fin(A)\ \ c \ Fin(A)" by (blast intro: Fin_subset_lemma) lemma Fin_IntI1 [intro,simp]: "b \ Fin(A) \ b \ c \ Fin(A)" by (blast intro: Fin_subset) lemma Fin_IntI2 [intro,simp]: "c \ Fin(A) \ b \ c \ Fin(A)" by (blast intro: Fin_subset) lemma Fin_0_induct_lemma [rule_format]: "\c \ Fin(A); b \ Fin(A); P(b); \x y. \x \ A; y \ Fin(A); x \ y; P(y)\ \ P(y-{x}) \ \ c<=b \ P(b-c)" apply (erule Fin_induct, simp) apply (subst Diff_cons) apply (simp add: cons_subset_iff Diff_subset [THEN Fin_subset]) done lemma Fin_0_induct: "\b \ Fin(A); P(b); \x y. \x \ A; y \ Fin(A); x \ y; P(y)\ \ P(y-{x}) \ \ P(0)" apply (rule Diff_cancel [THEN subst]) apply (blast intro: Fin_0_induct_lemma) done (*Functions from a finite ordinal*) lemma nat_fun_subset_Fin: "n \ nat \ n->A \ Fin(nat*A)" apply (induct_tac "n") apply (simp add: subset_iff) apply (simp add: succ_def mem_not_refl [THEN cons_fun_eq]) apply (fast intro!: Fin.consI) done subsection\Finite Function Space\ lemma FiniteFun_mono: "\A<=C; B<=D\ \ A -||> B \ C -||> D" -apply (unfold FiniteFun.defs) + unfolding FiniteFun.defs apply (rule lfp_mono) apply (rule FiniteFun.bnd_mono)+ apply (intro Fin_mono Sigma_mono basic_monos, assumption+) done lemma FiniteFun_mono1: "A<=B \ A -||> A \ B -||> B" by (blast dest: FiniteFun_mono) lemma FiniteFun_is_fun: "h \ A -||>B \ h \ domain(h) -> B" apply (erule FiniteFun.induct, simp) apply (simp add: fun_extend3) done lemma FiniteFun_domain_Fin: "h \ A -||>B \ domain(h) \ Fin(A)" by (erule FiniteFun.induct, simp, simp) lemmas FiniteFun_apply_type = FiniteFun_is_fun [THEN apply_type] (*Every subset of a finite function is a finite function.*) lemma FiniteFun_subset_lemma [rule_format]: "b \ A-||>B \ \z. z<=b \ z \ A-||>B" apply (erule FiniteFun.induct) apply (simp add: subset_empty_iff FiniteFun.intros) apply (simp add: subset_cons_iff distrib_simps, safe) apply (erule_tac b = z in cons_Diff [THEN subst]) apply (drule spec [THEN mp], assumption) apply (fast intro!: FiniteFun.intros) done lemma FiniteFun_subset: "\c<=b; b \ A-||>B\ \ c \ A-||>B" by (blast intro: FiniteFun_subset_lemma) (** Some further results by Sidi O. Ehmety **) lemma fun_FiniteFunI [rule_format]: "A \ Fin(X) \ \f. f \ A->B \ f \ A-||>B" apply (erule Fin.induct) apply (simp add: FiniteFun.intros, clarify) apply (case_tac "a \ b") apply (simp add: cons_absorb) apply (subgoal_tac "restrict (f,b) \ b -||> B") prefer 2 apply (blast intro: restrict_type2) apply (subst fun_cons_restrict_eq, assumption) apply (simp add: restrict_def lam_def) apply (blast intro: apply_funtype FiniteFun.intros FiniteFun_mono [THEN [2] rev_subsetD]) done lemma lam_FiniteFun: "A \ Fin(X) \ (\x\A. b(x)) \ A -||> {b(x). x \ A}" by (blast intro: fun_FiniteFunI lam_funtype) lemma FiniteFun_Collect_iff: "f \ FiniteFun(A, {y \ B. P(y)}) \ f \ FiniteFun(A,B) \ (\x\domain(f). P(f`x))" apply auto apply (blast intro: FiniteFun_mono [THEN [2] rev_subsetD]) apply (blast dest: Pair_mem_PiD FiniteFun_is_fun) apply (rule_tac A1="domain(f)" in subset_refl [THEN [2] FiniteFun_mono, THEN subsetD]) apply (fast dest: FiniteFun_domain_Fin Fin.dom_subset [THEN subsetD]) apply (rule fun_FiniteFunI) apply (erule FiniteFun_domain_Fin) apply (rule_tac B = "range (f) " in fun_weaken_type) apply (blast dest: FiniteFun_is_fun range_of_fun range_type apply_equality)+ done subsection\The Contents of a Singleton Set\ definition contents :: "i\i" where "contents(X) \ THE x. X = {x}" lemma contents_eq [simp]: "contents ({x}) = x" by (simp add: contents_def) end diff --git a/src/ZF/Induct/Binary_Trees.thy b/src/ZF/Induct/Binary_Trees.thy --- a/src/ZF/Induct/Binary_Trees.thy +++ b/src/ZF/Induct/Binary_Trees.thy @@ -1,138 +1,138 @@ (* Title: ZF/Induct/Binary_Trees.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section \Binary trees\ theory Binary_Trees imports ZF begin subsection \Datatype definition\ consts bt :: "i \ i" datatype "bt(A)" = Lf | Br ("a \ A", "t1 \ bt(A)", "t2 \ bt(A)") declare bt.intros [simp] lemma Br_neq_left: "l \ bt(A) \ Br(x, l, r) \ l" by (induct arbitrary: x r set: bt) auto lemma Br_iff: "Br(a, l, r) = Br(a', l', r') \ a = a' \ l = l' \ r = r'" \ \Proving a freeness theorem.\ by (fast elim!: bt.free_elims) inductive_cases BrE: "Br(a, l, r) \ bt(A)" \ \An elimination rule, for type-checking.\ text \ \medskip Lemmas to justify using \<^term>\bt\ in other recursive type definitions. \ lemma bt_mono: "A \ B \ bt(A) \ bt(B)" - apply (unfold bt.defs) + unfolding bt.defs apply (rule lfp_mono) apply (rule bt.bnd_mono)+ apply (rule univ_mono basic_monos | assumption)+ done lemma bt_univ: "bt(univ(A)) \ univ(A)" - apply (unfold bt.defs bt.con_defs) + unfolding bt.defs bt.con_defs apply (rule lfp_lowerbound) apply (rule_tac [2] A_subset_univ [THEN univ_mono]) apply (fast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) done lemma bt_subset_univ: "A \ univ(B) \ bt(A) \ univ(B)" apply (rule subset_trans) apply (erule bt_mono) apply (rule bt_univ) done lemma bt_rec_type: "\t \ bt(A); c \ C(Lf); \x y z r s. \x \ A; y \ bt(A); z \ bt(A); r \ C(y); s \ C(z)\ \ h(x, y, z, r, s) \ C(Br(x, y, z)) \ \ bt_rec(c, h, t) \ C(t)" \ \Type checking for recursor -- example only; not really needed.\ apply (induct_tac t) apply simp_all done subsection \Number of nodes, with an example of tail-recursion\ consts n_nodes :: "i \ i" primrec "n_nodes(Lf) = 0" "n_nodes(Br(a, l, r)) = succ(n_nodes(l) #+ n_nodes(r))" lemma n_nodes_type [simp]: "t \ bt(A) \ n_nodes(t) \ nat" by (induct set: bt) auto consts n_nodes_aux :: "i \ i" primrec "n_nodes_aux(Lf) = (\k \ nat. k)" "n_nodes_aux(Br(a, l, r)) = (\k \ nat. n_nodes_aux(r) ` (n_nodes_aux(l) ` succ(k)))" lemma n_nodes_aux_eq: "t \ bt(A) \ k \ nat \ n_nodes_aux(t)`k = n_nodes(t) #+ k" apply (induct arbitrary: k set: bt) apply simp apply (atomize, simp) done definition n_nodes_tail :: "i \ i" where "n_nodes_tail(t) \ n_nodes_aux(t) ` 0" lemma "t \ bt(A) \ n_nodes_tail(t) = n_nodes(t)" by (simp add: n_nodes_tail_def n_nodes_aux_eq) subsection \Number of leaves\ consts n_leaves :: "i \ i" primrec "n_leaves(Lf) = 1" "n_leaves(Br(a, l, r)) = n_leaves(l) #+ n_leaves(r)" lemma n_leaves_type [simp]: "t \ bt(A) \ n_leaves(t) \ nat" by (induct set: bt) auto subsection \Reflecting trees\ consts bt_reflect :: "i \ i" primrec "bt_reflect(Lf) = Lf" "bt_reflect(Br(a, l, r)) = Br(a, bt_reflect(r), bt_reflect(l))" lemma bt_reflect_type [simp]: "t \ bt(A) \ bt_reflect(t) \ bt(A)" by (induct set: bt) auto text \ \medskip Theorems about \<^term>\n_leaves\. \ lemma n_leaves_reflect: "t \ bt(A) \ n_leaves(bt_reflect(t)) = n_leaves(t)" by (induct set: bt) (simp_all add: add_commute) lemma n_leaves_nodes: "t \ bt(A) \ n_leaves(t) = succ(n_nodes(t))" by (induct set: bt) simp_all text \ Theorems about \<^term>\bt_reflect\. \ lemma bt_reflect_bt_reflect_ident: "t \ bt(A) \ bt_reflect(bt_reflect(t)) = t" by (induct set: bt) simp_all end diff --git a/src/ZF/Induct/Datatypes.thy b/src/ZF/Induct/Datatypes.thy --- a/src/ZF/Induct/Datatypes.thy +++ b/src/ZF/Induct/Datatypes.thy @@ -1,71 +1,71 @@ (* Title: ZF/Induct/Datatypes.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Sample datatype definitions\ theory Datatypes imports ZF begin subsection \A type with four constructors\ text \ It has four contructors, of arities 0--3, and two parameters \A\ and \B\. \ consts data :: "[i, i] \ i" datatype "data(A, B)" = Con0 | Con1 ("a \ A") | Con2 ("a \ A", "b \ B") | Con3 ("a \ A", "b \ B", "d \ data(A, B)") lemma data_unfold: "data(A, B) = ({0} + A) + (A \ B + A \ B \ data(A, B))" by (fast intro!: data.intros [unfolded data.con_defs] elim: data.cases [unfolded data.con_defs]) text \ \medskip Lemmas to justify using \<^term>\data\ in other recursive type definitions. \ lemma data_mono: "\A \ C; B \ D\ \ data(A, B) \ data(C, D)" - apply (unfold data.defs) + unfolding data.defs apply (rule lfp_mono) apply (rule data.bnd_mono)+ apply (rule univ_mono Un_mono basic_monos | assumption)+ done lemma data_univ: "data(univ(A), univ(A)) \ univ(A)" - apply (unfold data.defs data.con_defs) + unfolding data.defs data.con_defs apply (rule lfp_lowerbound) apply (rule_tac [2] subset_trans [OF A_subset_univ Un_upper1, THEN univ_mono]) apply (fast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) done lemma data_subset_univ: "\A \ univ(C); B \ univ(C)\ \ data(A, B) \ univ(C)" by (rule subset_trans [OF data_mono data_univ]) subsection \Example of a big enumeration type\ text \ Can go up to at least 100 constructors, but it takes nearly 7 minutes \dots\ (back in 1994 that is). \ consts enum :: i datatype enum = C00 | C01 | C02 | C03 | C04 | C05 | C06 | C07 | C08 | C09 | C10 | C11 | C12 | C13 | C14 | C15 | C16 | C17 | C18 | C19 | C20 | C21 | C22 | C23 | C24 | C25 | C26 | C27 | C28 | C29 | C30 | C31 | C32 | C33 | C34 | C35 | C36 | C37 | C38 | C39 | C40 | C41 | C42 | C43 | C44 | C45 | C46 | C47 | C48 | C49 | C50 | C51 | C52 | C53 | C54 | C55 | C56 | C57 | C58 | C59 end diff --git a/src/ZF/Induct/ListN.thy b/src/ZF/Induct/ListN.thy --- a/src/ZF/Induct/ListN.thy +++ b/src/ZF/Induct/ListN.thy @@ -1,61 +1,61 @@ (* Title: ZF/Induct/ListN.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Lists of n elements\ theory ListN imports ZF begin text \ Inductive definition of lists of \n\ elements; see @{cite "paulin-tlca"}. \ consts listn :: "i\i" inductive domains "listn(A)" \ "nat \ list(A)" intros NilI: "\0,Nil\ \ listn(A)" ConsI: "\a \ A; \n,l\ \ listn(A)\ \ \ listn(A)" type_intros nat_typechecks list.intros lemma list_into_listn: "l \ list(A) \ \ listn(A)" by (induct set: list) (simp_all add: listn.intros) lemma listn_iff: "\n,l\ \ listn(A) \ l \ list(A) \ length(l)=n" apply (rule iffI) apply (erule listn.induct) apply auto apply (blast intro: list_into_listn) done lemma listn_image_eq: "listn(A)``{n} = {l \ list(A). length(l)=n}" apply (rule equality_iffI) apply (simp add: listn_iff separation image_singleton_iff) done lemma listn_mono: "A \ B \ listn(A) \ listn(B)" - apply (unfold listn.defs) + unfolding listn.defs apply (rule lfp_mono) apply (rule listn.bnd_mono)+ apply (assumption | rule univ_mono Sigma_mono list_mono basic_monos)+ done lemma listn_append: "\\n,l\ \ listn(A); \ listn(A)\ \ \ listn(A)" apply (erule listn.induct) apply (frule listn.dom_subset [THEN subsetD]) apply (simp_all add: listn.intros) done inductive_cases Nil_listn_case: "\i,Nil\ \ listn(A)" and Cons_listn_case: " \ listn(A)" inductive_cases zero_listn_case: "\0,l\ \ listn(A)" and succ_listn_case: " \ listn(A)" end diff --git a/src/ZF/Induct/Multiset.thy b/src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy +++ b/src/ZF/Induct/Multiset.thy @@ -1,1294 +1,1294 @@ (* Title: ZF/Induct/Multiset.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory A definitional theory of multisets, including a wellfoundedness proof for the multiset order. The theory features ordinal multisets and the usual ordering. *) theory Multiset imports FoldSet Acc begin abbreviation (input) \ \Short cut for multiset space\ Mult :: "i\i" where "Mult(A) \ A -||> nat-{0}" definition (* This is the original "restrict" from ZF.thy. Restricts the function f to the domain A FIXME: adapt Multiset to the new "restrict". *) funrestrict :: "[i,i] \ i" where "funrestrict(f,A) \ \x \ A. f`x" definition (* M is a multiset *) multiset :: "i \ o" where "multiset(M) \ \A. M \ A -> nat-{0} \ Finite(A)" definition mset_of :: "i\i" where "mset_of(M) \ domain(M)" definition munion :: "[i, i] \ i" (infixl \+#\ 65) where "M +# N \ \x \ mset_of(M) \ mset_of(N). if x \ mset_of(M) \ mset_of(N) then (M`x) #+ (N`x) else (if x \ mset_of(M) then M`x else N`x)" definition (*convert a function to a multiset by eliminating 0*) normalize :: "i \ i" where "normalize(f) \ if (\A. f \ A -> nat \ Finite(A)) then funrestrict(f, {x \ mset_of(f). 0 < f`x}) else 0" definition mdiff :: "[i, i] \ i" (infixl \-#\ 65) where "M -# N \ normalize(\x \ mset_of(M). if x \ mset_of(N) then M`x #- N`x else M`x)" definition (* set of elements of a multiset *) msingle :: "i \ i" (\{#_#}\) where "{#a#} \ {\a, 1\}" definition MCollect :: "[i, i\o] \ i" (*comprehension*) where "MCollect(M, P) \ funrestrict(M, {x \ mset_of(M). P(x)})" definition (* Counts the number of occurrences of an element in a multiset *) mcount :: "[i, i] \ i" where "mcount(M, a) \ if a \ mset_of(M) then M`a else 0" definition msize :: "i \ i" where "msize(M) \ setsum(\a. $# mcount(M,a), mset_of(M))" abbreviation melem :: "[i,i] \ o" (\(_/ :# _)\ [50, 51] 50) where "a :# M \ a \ mset_of(M)" syntax "_MColl" :: "[pttrn, i, o] \ i" (\(1{# _ \ _./ _#})\) translations "{#x \ M. P#}" == "CONST MCollect(M, \x. P)" (* multiset orderings *) definition (* multirel1 has to be a set (not a predicate) so that we can form its transitive closure and reason about wf(.) and acc(.) *) multirel1 :: "[i,i]\i" where "multirel1(A, r) \ {\M, N\ \ Mult(A)*Mult(A). \a \ A. \M0 \ Mult(A). \K \ Mult(A). N=M0 +# {#a#} \ M=M0 +# K \ (\b \ mset_of(K). \b,a\ \ r)}" definition multirel :: "[i, i] \ i" where "multirel(A, r) \ multirel1(A, r)^+" (* ordinal multiset orderings *) definition omultiset :: "i \ o" where "omultiset(M) \ \i. Ord(i) \ M \ Mult(field(Memrel(i)))" definition mless :: "[i, i] \ o" (infixl \<#\ 50) where "M <# N \ \i. Ord(i) \ \M, N\ \ multirel(field(Memrel(i)), Memrel(i))" definition mle :: "[i, i] \ o" (infixl \<#=\ 50) where "M <#= N \ (omultiset(M) \ M = N) | M <# N" subsection\Properties of the original "restrict" from ZF.thy\ lemma funrestrict_subset: "\f \ Pi(C,B); A\C\ \ funrestrict(f,A) \ f" by (auto simp add: funrestrict_def lam_def intro: apply_Pair) lemma funrestrict_type: "\\x. x \ A \ f`x \ B(x)\ \ funrestrict(f,A) \ Pi(A,B)" by (simp add: funrestrict_def lam_type) lemma funrestrict_type2: "\f \ Pi(C,B); A\C\ \ funrestrict(f,A) \ Pi(A,B)" by (blast intro: apply_type funrestrict_type) lemma funrestrict [simp]: "a \ A \ funrestrict(f,A) ` a = f`a" by (simp add: funrestrict_def) lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0" by (simp add: funrestrict_def) lemma domain_funrestrict [simp]: "domain(funrestrict(f,C)) = C" by (auto simp add: funrestrict_def lam_def) lemma fun_cons_funrestrict_eq: "f \ cons(a, b) -> B \ f = cons(, funrestrict(f, b))" apply (rule equalityI) prefer 2 apply (blast intro: apply_Pair funrestrict_subset [THEN subsetD]) apply (auto dest!: Pi_memberD simp add: funrestrict_def lam_def) done declare domain_of_fun [simp] declare domainE [rule del] text\A useful simplification rule\ lemma multiset_fun_iff: "(f \ A -> nat-{0}) \ f \ A->nat\(\a \ A. f`a \ nat \ 0 < f`a)" apply safe apply (rule_tac [4] B1 = "range (f) " in Pi_mono [THEN subsetD]) apply (auto intro!: Ord_0_lt dest: apply_type Diff_subset [THEN Pi_mono, THEN subsetD] simp add: range_of_fun apply_iff) done (** The multiset space **) lemma multiset_into_Mult: "\multiset(M); mset_of(M)\A\ \ M \ Mult(A)" apply (simp add: multiset_def) apply (auto simp add: multiset_fun_iff mset_of_def) apply (rule_tac B1 = "nat-{0}" in FiniteFun_mono [THEN subsetD], simp_all) apply (rule Finite_into_Fin [THEN [2] Fin_mono [THEN subsetD], THEN fun_FiniteFunI]) apply (simp_all (no_asm_simp) add: multiset_fun_iff) done lemma Mult_into_multiset: "M \ Mult(A) \ multiset(M) \ mset_of(M)\A" apply (simp add: multiset_def mset_of_def) apply (frule FiniteFun_is_fun) apply (drule FiniteFun_domain_Fin) apply (frule FinD, clarify) apply (rule_tac x = "domain (M) " in exI) apply (blast intro: Fin_into_Finite) done lemma Mult_iff_multiset: "M \ Mult(A) \ multiset(M) \ mset_of(M)\A" by (blast dest: Mult_into_multiset intro: multiset_into_Mult) lemma multiset_iff_Mult_mset_of: "multiset(M) \ M \ Mult(mset_of(M))" by (auto simp add: Mult_iff_multiset) text\The \<^term>\multiset\ operator\ (* the empty multiset is 0 *) lemma multiset_0 [simp]: "multiset(0)" by (auto intro: FiniteFun.intros simp add: multiset_iff_Mult_mset_of) text\The \<^term>\mset_of\ operator\ lemma multiset_set_of_Finite [simp]: "multiset(M) \ Finite(mset_of(M))" by (simp add: multiset_def mset_of_def, auto) lemma mset_of_0 [iff]: "mset_of(0) = 0" by (simp add: mset_of_def) lemma mset_is_0_iff: "multiset(M) \ mset_of(M)=0 \ M=0" by (auto simp add: multiset_def mset_of_def) lemma mset_of_single [iff]: "mset_of({#a#}) = {a}" by (simp add: msingle_def mset_of_def) lemma mset_of_union [iff]: "mset_of(M +# N) = mset_of(M) \ mset_of(N)" by (simp add: mset_of_def munion_def) lemma mset_of_diff [simp]: "mset_of(M)\A \ mset_of(M -# N) \ A" by (auto simp add: mdiff_def multiset_def normalize_def mset_of_def) (* msingle *) lemma msingle_not_0 [iff]: "{#a#} \ 0 \ 0 \ {#a#}" by (simp add: msingle_def) lemma msingle_eq_iff [iff]: "({#a#} = {#b#}) \ (a = b)" by (simp add: msingle_def) lemma msingle_multiset [iff,TC]: "multiset({#a#})" apply (simp add: multiset_def msingle_def) apply (rule_tac x = "{a}" in exI) apply (auto intro: Finite_cons Finite_0 fun_extend3) done (** normalize **) lemmas Collect_Finite = Collect_subset [THEN subset_Finite] lemma normalize_idem [simp]: "normalize(normalize(f)) = normalize(f)" apply (simp add: normalize_def funrestrict_def mset_of_def) apply (case_tac "\A. f \ A -> nat \ Finite (A) ") apply clarify apply (drule_tac x = "{x \ domain (f) . 0 < f ` x}" in spec) apply auto apply (auto intro!: lam_type simp add: Collect_Finite) done lemma normalize_multiset [simp]: "multiset(M) \ normalize(M) = M" by (auto simp add: multiset_def normalize_def mset_of_def funrestrict_def multiset_fun_iff) lemma multiset_normalize [simp]: "multiset(normalize(f))" apply (simp add: normalize_def) apply (simp add: normalize_def mset_of_def multiset_def, auto) apply (rule_tac x = "{x \ A . 0multiset(M); multiset(N)\ \ multiset(M +# N)" apply (unfold multiset_def munion_def mset_of_def, auto) apply (rule_tac x = "A \ Aa" in exI) apply (auto intro!: lam_type intro: Finite_Un simp add: multiset_fun_iff zero_less_add) done (* difference *) lemma mdiff_multiset [simp]: "multiset(M -# N)" by (simp add: mdiff_def) (** Algebraic properties of multisets **) (* Union *) lemma munion_0 [simp]: "multiset(M) \ M +# 0 = M \ 0 +# M = M" apply (simp add: multiset_def) apply (auto simp add: munion_def mset_of_def) done lemma munion_commute: "M +# N = N +# M" by (auto intro!: lam_cong simp add: munion_def) lemma munion_assoc: "(M +# N) +# K = M +# (N +# K)" -apply (unfold munion_def mset_of_def) + unfolding munion_def mset_of_def apply (rule lam_cong, auto) done lemma munion_lcommute: "M +# (N +# K) = N +# (M +# K)" -apply (unfold munion_def mset_of_def) + unfolding munion_def mset_of_def apply (rule lam_cong, auto) done lemmas munion_ac = munion_commute munion_assoc munion_lcommute (* Difference *) lemma mdiff_self_eq_0 [simp]: "M -# M = 0" by (simp add: mdiff_def normalize_def mset_of_def) lemma mdiff_0 [simp]: "0 -# M = 0" by (simp add: mdiff_def normalize_def) lemma mdiff_0_right [simp]: "multiset(M) \ M -# 0 = M" by (auto simp add: multiset_def mdiff_def normalize_def multiset_fun_iff mset_of_def funrestrict_def) lemma mdiff_union_inverse2 [simp]: "multiset(M) \ M +# {#a#} -# {#a#} = M" -apply (unfold multiset_def munion_def mdiff_def msingle_def normalize_def mset_of_def) + unfolding multiset_def munion_def mdiff_def msingle_def normalize_def mset_of_def apply (auto cong add: if_cong simp add: ltD multiset_fun_iff funrestrict_def subset_Un_iff2 [THEN iffD1]) prefer 2 apply (force intro!: lam_type) apply (subgoal_tac [2] "{x \ A \ {a} . x \ a \ x \ A} = A") apply (rule fun_extension, auto) apply (drule_tac x = "A \ {a}" in spec) apply (simp add: Finite_Un) apply (force intro!: lam_type) done (** Count of elements **) lemma mcount_type [simp,TC]: "multiset(M) \ mcount(M, a) \ nat" by (auto simp add: multiset_def mcount_def mset_of_def multiset_fun_iff) lemma mcount_0 [simp]: "mcount(0, a) = 0" by (simp add: mcount_def) lemma mcount_single [simp]: "mcount({#b#}, a) = (if a=b then 1 else 0)" by (simp add: mcount_def mset_of_def msingle_def) lemma mcount_union [simp]: "\multiset(M); multiset(N)\ \ mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)" apply (auto simp add: multiset_def multiset_fun_iff mcount_def munion_def mset_of_def) done lemma mcount_diff [simp]: "multiset(M) \ mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)" apply (simp add: multiset_def) apply (auto dest!: not_lt_imp_le simp add: mdiff_def multiset_fun_iff mcount_def normalize_def mset_of_def) apply (force intro!: lam_type) apply (force intro!: lam_type) done lemma mcount_elem: "\multiset(M); a \ mset_of(M)\ \ 0 < mcount(M, a)" apply (simp add: multiset_def, clarify) apply (simp add: mcount_def mset_of_def) apply (simp add: multiset_fun_iff) done (** msize **) lemma msize_0 [simp]: "msize(0) = #0" by (simp add: msize_def) lemma msize_single [simp]: "msize({#a#}) = #1" by (simp add: msize_def) lemma msize_type [simp,TC]: "msize(M) \ int" by (simp add: msize_def) lemma msize_zpositive: "multiset(M)\ #0 $\ msize(M)" by (auto simp add: msize_def intro: g_zpos_imp_setsum_zpos) lemma msize_int_of_nat: "multiset(M) \ \n \ nat. msize(M)= $# n" apply (rule not_zneg_int_of) apply (simp_all (no_asm_simp) add: msize_type [THEN znegative_iff_zless_0] not_zless_iff_zle msize_zpositive) done lemma not_empty_multiset_imp_exist: "\M\0; multiset(M)\ \ \a \ mset_of(M). 0 < mcount(M, a)" apply (simp add: multiset_def) apply (erule not_emptyE) apply (auto simp add: mset_of_def mcount_def multiset_fun_iff) apply (blast dest!: fun_is_rel) done lemma msize_eq_0_iff: "multiset(M) \ msize(M)=#0 \ M=0" apply (simp add: msize_def, auto) apply (rule_tac P = "setsum (u,v) \ #0" for u v in swap) apply blast apply (drule not_empty_multiset_imp_exist, assumption, clarify) apply (subgoal_tac "Finite (mset_of (M) - {a}) ") prefer 2 apply (simp add: Finite_Diff) apply (subgoal_tac "setsum (\x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0") prefer 2 apply (simp add: cons_Diff, simp) apply (subgoal_tac "#0 $\ setsum (\x. $# mcount (M, x), mset_of (M) - {a}) ") apply (rule_tac [2] g_zpos_imp_setsum_zpos) apply (auto simp add: Finite_Diff not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym]) apply (rule not_zneg_int_of [THEN bexE]) apply (auto simp del: int_of_0 simp add: int_of_add [symmetric] int_of_0 [symmetric]) done lemma setsum_mcount_Int: "Finite(A) \ setsum(\a. $# mcount(N, a), A \ mset_of(N)) = setsum(\a. $# mcount(N, a), A)" apply (induct rule: Finite_induct) apply auto apply (subgoal_tac "Finite (B \ mset_of (N))") prefer 2 apply (blast intro: subset_Finite) apply (auto simp add: mcount_def Int_cons_left) done lemma msize_union [simp]: "\multiset(M); multiset(N)\ \ msize(M +# N) = msize(M) $+ msize(N)" apply (simp add: msize_def setsum_Un setsum_addf int_of_add setsum_mcount_Int) apply (subst Int_commute) apply (simp add: setsum_mcount_Int) done lemma msize_eq_succ_imp_elem: "\msize(M)= $# succ(n); n \ nat\ \ \a. a \ mset_of(M)" unfolding msize_def apply (blast dest: setsum_succD) done (** Equality of multisets **) lemma equality_lemma: "\multiset(M); multiset(N); \a. mcount(M, a)=mcount(N, a)\ \ mset_of(M)=mset_of(N)" apply (simp add: multiset_def) apply (rule sym, rule equalityI) apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) apply (drule_tac [!] x=x in spec) apply (case_tac [2] "x \ Aa", case_tac "x \ A", auto) done lemma multiset_equality: "\multiset(M); multiset(N)\\ M=N\(\a. mcount(M, a)=mcount(N, a))" apply auto apply (subgoal_tac "mset_of (M) = mset_of (N) ") prefer 2 apply (blast intro: equality_lemma) apply (simp add: multiset_def mset_of_def) apply (auto simp add: multiset_fun_iff) apply (rule fun_extension) apply (blast, blast) apply (drule_tac x = x in spec) apply (auto simp add: mcount_def mset_of_def) done (** More algebraic properties of multisets **) lemma munion_eq_0_iff [simp]: "\multiset(M); multiset(N)\\(M +# N =0) \ (M=0 \ N=0)" by (auto simp add: multiset_equality) lemma empty_eq_munion_iff [simp]: "\multiset(M); multiset(N)\\(0=M +# N) \ (M=0 \ N=0)" apply (rule iffI, drule sym) apply (simp_all add: multiset_equality) done lemma munion_right_cancel [simp]: "\multiset(M); multiset(N); multiset(K)\\(M +# K = N +# K)\(M=N)" by (auto simp add: multiset_equality) lemma munion_left_cancel [simp]: "\multiset(K); multiset(M); multiset(N)\ \(K +# M = K +# N) \ (M = N)" by (auto simp add: multiset_equality) lemma nat_add_eq_1_cases: "\m \ nat; n \ nat\ \ (m #+ n = 1) \ (m=1 \ n=0) | (m=0 \ n=1)" by (induct_tac n) auto lemma munion_is_single: "\multiset(M); multiset(N)\ \ (M +# N = {#a#}) \ (M={#a#} \ N=0) | (M = 0 \ N = {#a#})" apply (simp (no_asm_simp) add: multiset_equality) apply safe apply simp_all apply (case_tac "aa=a") apply (drule_tac [2] x = aa in spec) apply (drule_tac x = a in spec) apply (simp add: nat_add_eq_1_cases, simp) apply (case_tac "aaa=aa", simp) apply (drule_tac x = aa in spec) apply (simp add: nat_add_eq_1_cases) apply (case_tac "aaa=a") apply (drule_tac [4] x = aa in spec) apply (drule_tac [3] x = a in spec) apply (drule_tac [2] x = aaa in spec) apply (drule_tac x = aa in spec) apply (simp_all add: nat_add_eq_1_cases) done lemma msingle_is_union: "\multiset(M); multiset(N)\ \ ({#a#} = M +# N) \ ({#a#} = M \ N=0 | M = 0 \ {#a#} = N)" apply (subgoal_tac " ({#a#} = M +# N) \ (M +# N = {#a#}) ") apply (simp (no_asm_simp) add: munion_is_single) apply blast apply (blast dest: sym) done (** Towards induction over multisets **) lemma setsum_decr: "Finite(A) \ (\M. multiset(M) \ (\a \ mset_of(M). setsum(\z. $# mcount(M(a:=M`a #- 1), z), A) = (if a \ A then setsum(\z. $# mcount(M, z), A) $- #1 else setsum(\z. $# mcount(M, z), A))))" unfolding multiset_def apply (erule Finite_induct) apply (auto simp add: multiset_fun_iff) -apply (unfold mset_of_def mcount_def) + unfolding mset_of_def mcount_def apply (case_tac "x \ A", auto) apply (subgoal_tac "$# M ` x $+ #-1 = $# M ` x $- $# 1") apply (erule ssubst) apply (rule int_of_diff, auto) done lemma setsum_decr2: "Finite(A) \ \M. multiset(M) \ (\a \ mset_of(M). setsum(\x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) = (if a \ A then setsum(\x. $# mcount(M, x), A) $- $# M`a else setsum(\x. $# mcount(M, x), A)))" apply (simp add: multiset_def) apply (erule Finite_induct) apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) done lemma setsum_decr3: "\Finite(A); multiset(M); a \ mset_of(M)\ \ setsum(\x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) = (if a \ A then setsum(\x. $# mcount(M, x), A) $- $# M`a else setsum(\x. $# mcount(M, x), A))" apply (subgoal_tac "setsum (\x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (\x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ") apply (rule_tac [2] setsum_Diff [symmetric]) apply (rule sym, rule ssubst, blast) apply (rule sym, drule setsum_decr2, auto) apply (simp add: mcount_def mset_of_def) done lemma nat_le_1_cases: "n \ nat \ n \ 1 \ (n=0 | n=1)" by (auto elim: natE) lemma succ_pred_eq_self: "\0 nat\ \ succ(n #- 1) = n" apply (subgoal_tac "1 \ n") apply (drule add_diff_inverse2, auto) done text\Specialized for use in the proof below.\ lemma multiset_funrestict: "\\a\A. M ` a \ nat \ 0 < M ` a; Finite(A)\ \ multiset(funrestrict(M, A - {a}))" apply (simp add: multiset_def multiset_fun_iff) apply (rule_tac x="A-{a}" in exI) apply (auto intro: Finite_Diff funrestrict_type) done lemma multiset_induct_aux: assumes prem1: "\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(\a, 1\, M))" and prem2: "\M b. \multiset(M); b \ mset_of(M); P(M)\ \ P(M(b:= M`b #+ 1))" shows "\n \ nat; P(0)\ \ (\M. multiset(M)\ (setsum(\x. $# mcount(M, x), {x \ mset_of(M). 0 < M`x}) = $# n) \ P(M))" apply (erule nat_induct, clarify) apply (frule msize_eq_0_iff) apply (auto simp add: mset_of_def multiset_def multiset_fun_iff msize_def) apply (subgoal_tac "setsum (\x. $# mcount (M, x), A) =$# succ (x) ") apply (drule setsum_succD, auto) apply (case_tac "1 cons (a, A) . x\a\0, funrestrict (M, A-{a}))") prefer 2 apply (rule fun_cons_funrestrict_eq) apply (subgoal_tac "cons (a, A-{a}) = A") apply force apply force apply (rule_tac a = "cons (\a, 1\, funrestrict (M, A - {a}))" in ssubst) apply simp apply (frule multiset_funrestict, assumption) apply (rule prem1, assumption) apply (simp add: mset_of_def) apply (drule_tac x = "funrestrict (M, A-{a}) " in spec) apply (drule mp) apply (rule_tac x = "A-{a}" in exI) apply (auto intro: Finite_Diff funrestrict_type simp add: funrestrict) apply (frule_tac A = A and M = M and a = a in setsum_decr3) apply (simp (no_asm_simp) add: multiset_def multiset_fun_iff) apply blast apply (simp (no_asm_simp) add: mset_of_def) apply (drule_tac b = "if u then v else w" for u v w in sym, simp_all) apply (subgoal_tac "{x \ A - {a} . 0 < funrestrict (M, A - {x}) ` x} = A - {a}") apply (auto intro!: setsum_cong simp add: zdiff_eq_iff zadd_commute multiset_def multiset_fun_iff mset_of_def) done lemma multiset_induct2: "\multiset(M); P(0); (\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(\a, 1\, M))); (\M b. \multiset(M); b \ mset_of(M); P(M)\ \ P(M(b:= M`b #+ 1)))\ \ P(M)" apply (subgoal_tac "\n \ nat. setsum (\x. $# mcount (M, x), {x \ mset_of (M) . 0 < M ` x}) = $# n") apply (rule_tac [2] not_zneg_int_of) apply (simp_all (no_asm_simp) add: znegative_iff_zless_0 not_zless_iff_zle) apply (rule_tac [2] g_zpos_imp_setsum_zpos) prefer 2 apply (blast intro: multiset_set_of_Finite Collect_subset [THEN subset_Finite]) prefer 2 apply (simp add: multiset_def multiset_fun_iff, clarify) apply (rule multiset_induct_aux [rule_format], auto) done lemma munion_single_case1: "\multiset(M); a \mset_of(M)\ \ M +# {#a#} = cons(\a, 1\, M)" apply (simp add: multiset_def msingle_def) apply (auto simp add: munion_def) apply (unfold mset_of_def, simp) apply (rule fun_extension, rule lam_type, simp_all) apply (auto simp add: multiset_fun_iff fun_extend_apply) apply (drule_tac c = a and b = 1 in fun_extend3) apply (auto simp add: cons_eq Un_commute [of _ "{a}"]) done lemma munion_single_case2: "\multiset(M); a \ mset_of(M)\ \ M +# {#a#} = M(a:=M`a #+ 1)" apply (simp add: multiset_def) apply (auto simp add: munion_def multiset_fun_iff msingle_def) apply (unfold mset_of_def, simp) apply (subgoal_tac "A \ {a} = A") apply (rule fun_extension) apply (auto dest: domain_type intro: lam_type update_type) done (* Induction principle for multisets *) lemma multiset_induct: assumes M: "multiset(M)" and P0: "P(0)" and step: "\M a. \multiset(M); P(M)\ \ P(M +# {#a#})" shows "P(M)" apply (rule multiset_induct2 [OF M]) apply (simp_all add: P0) apply (frule_tac [2] a = b in munion_single_case2 [symmetric]) apply (frule_tac a = a in munion_single_case1 [symmetric]) apply (auto intro: step) done (** MCollect **) lemma MCollect_multiset [simp]: "multiset(M) \ multiset({# x \ M. P(x)#})" apply (simp add: MCollect_def multiset_def mset_of_def, clarify) apply (rule_tac x = "{x \ A. P (x) }" in exI) apply (auto dest: CollectD1 [THEN [2] apply_type] intro: Collect_subset [THEN subset_Finite] funrestrict_type) done lemma mset_of_MCollect [simp]: "multiset(M) \ mset_of({# x \ M. P(x) #}) \ mset_of(M)" by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def) lemma MCollect_mem_iff [iff]: "x \ mset_of({#x \ M. P(x)#}) \ x \ mset_of(M) \ P(x)" by (simp add: MCollect_def mset_of_def) lemma mcount_MCollect [simp]: "mcount({# x \ M. P(x) #}, a) = (if P(a) then mcount(M,a) else 0)" by (simp add: mcount_def MCollect_def mset_of_def) lemma multiset_partition: "multiset(M) \ M = {# x \ M. P(x) #} +# {# x \ M. \ P(x) #}" by (simp add: multiset_equality) lemma natify_elem_is_self [simp]: "\multiset(M); a \ mset_of(M)\ \ natify(M`a) = M`a" by (auto simp add: multiset_def mset_of_def multiset_fun_iff) (* and more algebraic laws on multisets *) lemma munion_eq_conv_diff: "\multiset(M); multiset(N)\ \ (M +# {#a#} = N +# {#b#}) \ (M = N \ a = b | M = N -# {#a#} +# {#b#} \ N = M -# {#b#} +# {#a#})" apply (simp del: mcount_single add: multiset_equality) apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE) apply (case_tac "a=b", auto) apply (drule_tac x = a in spec) apply (drule_tac [2] x = b in spec) apply (drule_tac [3] x = aa in spec) apply (drule_tac [4] x = a in spec, auto) apply (subgoal_tac [!] "mcount (N,a) :nat") apply (erule_tac [3] natE, erule natE, auto) done lemma melem_diff_single: "multiset(M) \ k \ mset_of(M -# {#a#}) \ (k=a \ 1 < mcount(M,a)) | (k\ a \ k \ mset_of(M))" apply (simp add: multiset_def) apply (simp add: normalize_def mset_of_def msingle_def mdiff_def mcount_def) apply (auto dest: domain_type intro: zero_less_diff [THEN iffD1] simp add: multiset_fun_iff apply_iff) apply (force intro!: lam_type) apply (force intro!: lam_type) apply (force intro!: lam_type) done lemma munion_eq_conv_exist: "\M \ Mult(A); N \ Mult(A)\ \ (M +# {#a#} = N +# {#b#}) \ (M=N \ a=b | (\K \ Mult(A). M= K +# {#b#} \ N=K +# {#a#}))" by (auto simp add: Mult_iff_multiset melem_diff_single munion_eq_conv_diff) subsection\Multiset Orderings\ (* multiset on a domain A are finite functions from A to nat-{0} *) (* multirel1 type *) lemma multirel1_type: "multirel1(A, r) \ Mult(A)*Mult(A)" by (auto simp add: multirel1_def) lemma multirel1_0 [simp]: "multirel1(0, r) =0" by (auto simp add: multirel1_def) lemma multirel1_iff: " \N, M\ \ multirel1(A, r) \ (\a. a \ A \ (\M0. M0 \ Mult(A) \ (\K. K \ Mult(A) \ M=M0 +# {#a#} \ N=M0 +# K \ (\b \ mset_of(K). \b,a\ \ r))))" by (auto simp add: multirel1_def Mult_iff_multiset Bex_def) text\Monotonicity of \<^term>\multirel1\\ lemma multirel1_mono1: "A\B \ multirel1(A, r)\multirel1(B, r)" apply (auto simp add: multirel1_def) apply (auto simp add: Un_subset_iff Mult_iff_multiset) apply (rule_tac x = a in bexI) apply (rule_tac x = M0 in bexI, simp) apply (rule_tac x = K in bexI) apply (auto simp add: Mult_iff_multiset) done lemma multirel1_mono2: "r\s \ multirel1(A,r)\multirel1(A, s)" apply (simp add: multirel1_def, auto) apply (rule_tac x = a in bexI) apply (rule_tac x = M0 in bexI) apply (simp_all add: Mult_iff_multiset) apply (rule_tac x = K in bexI) apply (simp_all add: Mult_iff_multiset, auto) done lemma multirel1_mono: "\A\B; r\s\ \ multirel1(A, r) \ multirel1(B, s)" apply (rule subset_trans) apply (rule multirel1_mono1) apply (rule_tac [2] multirel1_mono2, auto) done subsection\Toward the proof of well-foundedness of multirel1\ lemma not_less_0 [iff]: "\M,0\ \ multirel1(A, r)" by (auto simp add: multirel1_def Mult_iff_multiset) lemma less_munion: "\ \ multirel1(A, r); M0 \ Mult(A)\ \ (\M. \M, M0\ \ multirel1(A, r) \ N = M +# {#a#}) | (\K. K \ Mult(A) \ (\b \ mset_of(K). \b, a\ \ r) \ N = M0 +# K)" apply (frule multirel1_type [THEN subsetD]) apply (simp add: multirel1_iff) apply (auto simp add: munion_eq_conv_exist) apply (rule_tac x="Ka +# K" in exI, auto, simp add: Mult_iff_multiset) apply (simp (no_asm_simp) add: munion_left_cancel munion_assoc) apply (auto simp add: munion_commute) done lemma multirel1_base: "\M \ Mult(A); a \ A\ \ \ multirel1(A, r)" apply (auto simp add: multirel1_iff) apply (simp add: Mult_iff_multiset) apply (rule_tac x = a in exI, clarify) apply (rule_tac x = M in exI, simp) apply (rule_tac x = 0 in exI, auto) done lemma acc_0: "acc(0)=0" by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD]) lemma lemma1: "\\b \ A. \b,a\ \ r \ (\M \ acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r))); M0 \ acc(multirel1(A, r)); a \ A; \M. \M,M0\ \ multirel1(A, r) \ M +# {#a#} \ acc(multirel1(A, r))\ \ M0 +# {#a#} \ acc(multirel1(A, r))" apply (subgoal_tac "M0 \ Mult(A) ") prefer 2 apply (erule acc.cases) apply (erule fieldE) apply (auto dest: multirel1_type [THEN subsetD]) apply (rule accI) apply (rename_tac "N") apply (drule less_munion, blast) apply (auto simp add: Mult_iff_multiset) apply (erule_tac P = "\x \ mset_of (K) . \x, a\ \ r" in rev_mp) apply (erule_tac P = "mset_of (K) \A" in rev_mp) apply (erule_tac M = K in multiset_induct) (* three subgoals *) (* subgoal 1 \ the induction base case *) apply (simp (no_asm_simp)) (* subgoal 2 \ the induction general case *) apply (simp add: Ball_def Un_subset_iff, clarify) apply (drule_tac x = aa in spec, simp) apply (subgoal_tac "aa \ A") prefer 2 apply blast apply (drule_tac x = "M0 +# M" and P = "\x. x \ acc(multirel1(A, r)) \ Q(x)" for Q in spec) apply (simp add: munion_assoc [symmetric]) (* subgoal 3 \ additional conditions *) apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset) done lemma lemma2: "\\b \ A. \b,a\ \ r \ (\M \ acc(multirel1(A, r)). M +# {#b#} :acc(multirel1(A, r))); M \ acc(multirel1(A, r)); a \ A\ \ M +# {#a#} \ acc(multirel1(A, r))" apply (erule acc_induct) apply (blast intro: lemma1) done lemma lemma3: "\wf[A](r); a \ A\ \ \M \ acc(multirel1(A, r)). M +# {#a#} \ acc(multirel1(A, r))" apply (erule_tac a = a in wf_on_induct, blast) apply (blast intro: lemma2) done lemma lemma4: "multiset(M) \ mset_of(M)\A \ wf[A](r) \ M \ field(multirel1(A, r)) \ M \ acc(multirel1(A, r))" apply (erule multiset_induct) (* proving the base case *) apply clarify apply (rule accI, force) apply (simp add: multirel1_def) (* Proving the general case *) apply clarify apply simp apply (subgoal_tac "mset_of (M) \A") prefer 2 apply blast apply clarify apply (drule_tac a = a in lemma3, blast) apply (subgoal_tac "M \ field (multirel1 (A,r))") apply blast apply (rule multirel1_base [THEN fieldI1]) apply (auto simp add: Mult_iff_multiset) done lemma all_accessible: "\wf[A](r); M \ Mult(A); A \ 0\ \ M \ acc(multirel1(A, r))" apply (erule not_emptyE) apply (rule lemma4 [THEN mp, THEN mp, THEN mp]) apply (rule_tac [4] multirel1_base [THEN fieldI1]) apply (auto simp add: Mult_iff_multiset) done lemma wf_on_multirel1: "wf[A](r) \ wf[A-||>nat-{0}](multirel1(A, r))" apply (case_tac "A=0") apply (simp (no_asm_simp)) apply (rule wf_imp_wf_on) apply (rule wf_on_field_imp_wf) apply (simp (no_asm_simp) add: wf_on_0) apply (rule_tac A = "acc (multirel1 (A,r))" in wf_on_subset_A) apply (rule wf_on_acc) apply (blast intro: all_accessible) done lemma wf_multirel1: "wf(r) \wf(multirel1(field(r), r))" apply (simp (no_asm_use) add: wf_iff_wf_on_field) apply (drule wf_on_multirel1) apply (rule_tac A = "field (r) -||> nat - {0}" in wf_on_subset_A) apply (simp (no_asm_simp)) apply (rule field_rel_subset) apply (rule multirel1_type) done (** multirel **) lemma multirel_type: "multirel(A, r) \ Mult(A)*Mult(A)" apply (simp add: multirel_def) apply (rule trancl_type [THEN subset_trans]) apply (auto dest: multirel1_type [THEN subsetD]) done (* Monotonicity of multirel *) lemma multirel_mono: "\A\B; r\s\ \ multirel(A, r)\multirel(B,s)" apply (simp add: multirel_def) apply (rule trancl_mono) apply (rule multirel1_mono, auto) done (* Equivalence of multirel with the usual (closure-free) definition *) lemma add_diff_eq: "k \ nat \ 0 < k \ n #+ k #- 1 = n #+ (k #- 1)" by (erule nat_induct, auto) lemma mdiff_union_single_conv: "\a \ mset_of(J); multiset(I); multiset(J)\ \ I +# J -# {#a#} = I +# (J-# {#a#})" apply (simp (no_asm_simp) add: multiset_equality) apply (case_tac "a \ mset_of (I) ") apply (auto simp add: mcount_def mset_of_def multiset_def multiset_fun_iff) apply (auto dest: domain_type simp add: add_diff_eq) done lemma diff_add_commute: "\n \ m; m \ nat; n \ nat; k \ nat\ \ m #- n #+ k = m #+ k #- n" by (auto simp add: le_iff less_iff_succ_add) (* One direction *) lemma multirel_implies_one_step: "\M,N\ \ multirel(A, r) \ trans[A](r) \ (\I J K. I \ Mult(A) \ J \ Mult(A) \ K \ Mult(A) \ N = I +# J \ M = I +# K \ J \ 0 \ (\k \ mset_of(K). \j \ mset_of(J). \k,j\ \ r))" apply (simp add: multirel_def Ball_def Bex_def) apply (erule converse_trancl_induct) apply (simp_all add: multirel1_iff Mult_iff_multiset) (* Two subgoals remain *) (* Subgoal 1 *) apply clarify apply (rule_tac x = M0 in exI, force) (* Subgoal 2 *) apply clarify apply hypsubst_thin apply (case_tac "a \ mset_of (Ka) ") apply (rule_tac x = I in exI, simp (no_asm_simp)) apply (rule_tac x = J in exI, simp (no_asm_simp)) apply (rule_tac x = " (Ka -# {#a#}) +# K" in exI, simp (no_asm_simp)) apply (simp_all add: Un_subset_iff) apply (simp (no_asm_simp) add: munion_assoc [symmetric]) apply (drule_tac t = "\M. M-#{#a#}" in subst_context) apply (simp add: mdiff_union_single_conv melem_diff_single, clarify) apply (erule disjE, simp) apply (erule disjE, simp) apply (drule_tac x = a and P = "\x. x :# Ka \ Q(x)" for Q in spec) apply clarify apply (rule_tac x = xa in exI) apply (simp (no_asm_simp)) apply (blast dest: trans_onD) (* new we know that a\mset_of(Ka) *) apply (subgoal_tac "a :# I") apply (rule_tac x = "I-#{#a#}" in exI, simp (no_asm_simp)) apply (rule_tac x = "J+#{#a#}" in exI) apply (simp (no_asm_simp) add: Un_subset_iff) apply (rule_tac x = "Ka +# K" in exI) apply (simp (no_asm_simp) add: Un_subset_iff) apply (rule conjI) apply (simp (no_asm_simp) add: multiset_equality mcount_elem [THEN succ_pred_eq_self]) apply (rule conjI) apply (drule_tac t = "\M. M-#{#a#}" in subst_context) apply (simp add: mdiff_union_inverse2) apply (simp_all (no_asm_simp) add: multiset_equality) apply (rule diff_add_commute [symmetric]) apply (auto intro: mcount_elem) apply (subgoal_tac "a \ mset_of (I +# Ka) ") apply (drule_tac [2] sym, auto) done lemma melem_imp_eq_diff_union [simp]: "\a \ mset_of(M); multiset(M)\ \ M -# {#a#} +# {#a#} = M" by (simp add: multiset_equality mcount_elem [THEN succ_pred_eq_self]) lemma msize_eq_succ_imp_eq_union: "\msize(M)=$# succ(n); M \ Mult(A); n \ nat\ \ \a N. M = N +# {#a#} \ N \ Mult(A) \ a \ A" apply (drule msize_eq_succ_imp_elem, auto) apply (rule_tac x = a in exI) apply (rule_tac x = "M -# {#a#}" in exI) apply (frule Mult_into_multiset) apply (simp (no_asm_simp)) apply (auto simp add: Mult_iff_multiset) done (* The second direction *) lemma one_step_implies_multirel_lemma [rule_format (no_asm)]: "n \ nat \ (\I J K. I \ Mult(A) \ J \ Mult(A) \ K \ Mult(A) \ (msize(J) = $# n \ J \0 \ (\k \ mset_of(K). \j \ mset_of(J). \k, j\ \ r)) \ \ multirel(A, r))" apply (simp add: Mult_iff_multiset) apply (erule nat_induct, clarify) apply (drule_tac M = J in msize_eq_0_iff, auto) (* one subgoal remains *) apply (subgoal_tac "msize (J) =$# succ (x) ") prefer 2 apply simp apply (frule_tac A = A in msize_eq_succ_imp_eq_union) apply (simp_all add: Mult_iff_multiset, clarify) apply (rename_tac "J'", simp) apply (case_tac "J' = 0") apply (simp add: multirel_def) apply (rule r_into_trancl, clarify) apply (simp add: multirel1_iff Mult_iff_multiset, force) (*Now we know J' \ 0*) apply (drule sym, rotate_tac -1, simp) apply (erule_tac V = "$# x = msize (J') " in thin_rl) apply (frule_tac M = K and P = "\x. \x,a\ \ r" in multiset_partition) apply (erule_tac P = "\k \ mset_of (K) . P(k)" for P in rev_mp) apply (erule ssubst) apply (simp add: Ball_def, auto) apply (subgoal_tac "< (I +# {# x \ K. \x, a\ \ r#}) +# {# x \ K. \x, a\ \ r#}, (I +# {# x \ K. \x, a\ \ r#}) +# J'> \ multirel(A, r) ") prefer 2 apply (drule_tac x = "I +# {# x \ K. \x, a\ \ r#}" in spec) apply (rotate_tac -1) apply (drule_tac x = "J'" in spec) apply (rotate_tac -1) apply (drule_tac x = "{# x \ K. \x, a\ \ r#}" in spec, simp) apply blast apply (simp add: munion_assoc [symmetric] multirel_def) apply (rule_tac b = "I +# {# x \ K. \x, a\ \ r#} +# J'" in trancl_trans, blast) apply (rule r_into_trancl) apply (simp add: multirel1_iff Mult_iff_multiset) apply (rule_tac x = a in exI) apply (simp (no_asm_simp)) apply (rule_tac x = "I +# J'" in exI) apply (auto simp add: munion_ac Un_subset_iff) done lemma one_step_implies_multirel: "\J \ 0; \k \ mset_of(K). \j \ mset_of(J). \k,j\ \ r; I \ Mult(A); J \ Mult(A); K \ Mult(A)\ \ \ multirel(A, r)" apply (subgoal_tac "multiset (J) ") prefer 2 apply (simp add: Mult_iff_multiset) apply (frule_tac M = J in msize_int_of_nat) apply (auto intro: one_step_implies_multirel_lemma) done (** Proving that multisets are partially ordered **) (*irreflexivity*) lemma multirel_irrefl_lemma: "Finite(A) \ part_ord(A, r) \ (\x \ A. \y \ A. \x,y\ \ r) \A=0" apply (erule Finite_induct) apply (auto dest: subset_consI [THEN [2] part_ord_subset]) apply (auto simp add: part_ord_def irrefl_def) apply (drule_tac x = xa in bspec) apply (drule_tac [2] a = xa and b = x in trans_onD, auto) done lemma irrefl_on_multirel: "part_ord(A, r) \ irrefl(Mult(A), multirel(A, r))" apply (simp add: irrefl_def) apply (subgoal_tac "trans[A](r) ") prefer 2 apply (simp add: part_ord_def, clarify) apply (drule multirel_implies_one_step, clarify) apply (simp add: Mult_iff_multiset, clarify) apply (subgoal_tac "Finite (mset_of (K))") apply (frule_tac r = r in multirel_irrefl_lemma) apply (frule_tac B = "mset_of (K) " in part_ord_subset) apply simp_all apply (auto simp add: multiset_def mset_of_def) done lemma trans_on_multirel: "trans[Mult(A)](multirel(A, r))" apply (simp add: multirel_def trans_on_def) apply (blast intro: trancl_trans) done lemma multirel_trans: "\\M, N\ \ multirel(A, r); \N, K\ \ multirel(A, r)\ \ \M, K\ \ multirel(A,r)" apply (simp add: multirel_def) apply (blast intro: trancl_trans) done lemma trans_multirel: "trans(multirel(A,r))" apply (simp add: multirel_def) apply (rule trans_trancl) done lemma part_ord_multirel: "part_ord(A,r) \ part_ord(Mult(A), multirel(A, r))" apply (simp (no_asm) add: part_ord_def) apply (blast intro: irrefl_on_multirel trans_on_multirel) done (** Monotonicity of multiset union **) lemma munion_multirel1_mono: "\\M,N\ \ multirel1(A, r); K \ Mult(A)\ \ \ multirel1(A, r)" apply (frule multirel1_type [THEN subsetD]) apply (auto simp add: multirel1_iff Mult_iff_multiset) apply (rule_tac x = a in exI) apply (simp (no_asm_simp)) apply (rule_tac x = "K+#M0" in exI) apply (simp (no_asm_simp) add: Un_subset_iff) apply (rule_tac x = Ka in exI) apply (simp (no_asm_simp) add: munion_assoc) done lemma munion_multirel_mono2: "\\M, N\ \ multirel(A, r); K \ Mult(A)\\ \ multirel(A, r)" apply (frule multirel_type [THEN subsetD]) apply (simp (no_asm_use) add: multirel_def) apply clarify apply (drule_tac psi = "\M,N\ \ multirel1 (A, r) ^+" in asm_rl) apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) apply (erule trancl_induct, clarify) apply (blast intro: munion_multirel1_mono r_into_trancl, clarify) apply (subgoal_tac "y \ Mult(A) ") prefer 2 apply (blast dest: multirel_type [unfolded multirel_def, THEN subsetD]) apply (subgoal_tac " \ multirel1 (A, r) ") prefer 2 apply (blast intro: munion_multirel1_mono) apply (blast intro: r_into_trancl trancl_trans) done lemma munion_multirel_mono1: "\\M, N\ \ multirel(A, r); K \ Mult(A)\ \ \ multirel(A, r)" apply (frule multirel_type [THEN subsetD]) apply (rule_tac P = "\x. \x,u\ \ multirel(A, r)" for u in munion_commute [THEN subst]) apply (subst munion_commute [of N]) apply (rule munion_multirel_mono2) apply (auto simp add: Mult_iff_multiset) done lemma munion_multirel_mono: "\\M,K\ \ multirel(A, r); \N,L\ \ multirel(A, r)\ \ \ multirel(A, r)" apply (subgoal_tac "M \ Mult(A) \ N \ Mult(A) \ K \ Mult(A) \ L \ Mult(A) ") prefer 2 apply (blast dest: multirel_type [THEN subsetD]) apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2) done subsection\Ordinal Multisets\ (* A \ B \ field(Memrel(A)) \ field(Memrel(B)) *) lemmas field_Memrel_mono = Memrel_mono [THEN field_mono] (* \Aa \ Ba; A \ B\ \ multirel(field(Memrel(Aa)), Memrel(A))\ multirel(field(Memrel(Ba)), Memrel(B)) *) lemmas multirel_Memrel_mono = multirel_mono [OF field_Memrel_mono Memrel_mono] lemma omultiset_is_multiset [simp]: "omultiset(M) \ multiset(M)" apply (simp add: omultiset_def) apply (auto simp add: Mult_iff_multiset) done lemma munion_omultiset [simp]: "\omultiset(M); omultiset(N)\ \ omultiset(M +# N)" apply (simp add: omultiset_def, clarify) apply (rule_tac x = "i \ ia" in exI) apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff) apply (blast intro: field_Memrel_mono) done lemma mdiff_omultiset [simp]: "omultiset(M) \ omultiset(M -# N)" apply (simp add: omultiset_def, clarify) apply (simp add: Mult_iff_multiset) apply (rule_tac x = i in exI) apply (simp (no_asm_simp)) done (** Proving that Memrel is a partial order **) lemma irrefl_Memrel: "Ord(i) \ irrefl(field(Memrel(i)), Memrel(i))" apply (rule irreflI, clarify) apply (subgoal_tac "Ord (x) ") prefer 2 apply (blast intro: Ord_in_Ord) apply (drule_tac i = x in ltI [THEN lt_irrefl], auto) done lemma trans_iff_trans_on: "trans(r) \ trans[field(r)](r)" by (simp add: trans_on_def trans_def, auto) lemma part_ord_Memrel: "Ord(i) \part_ord(field(Memrel(i)), Memrel(i))" apply (simp add: part_ord_def) apply (simp (no_asm) add: trans_iff_trans_on [THEN iff_sym]) apply (blast intro: trans_Memrel irrefl_Memrel) done (* Ord(i) \ part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i))) *) lemmas part_ord_mless = part_ord_Memrel [THEN part_ord_multirel] (*irreflexivity*) lemma mless_not_refl: "\(M <# M)" apply (simp add: mless_def, clarify) apply (frule multirel_type [THEN subsetD]) apply (drule part_ord_mless) apply (simp add: part_ord_def irrefl_def) done (* N R *) lemmas mless_irrefl = mless_not_refl [THEN notE, elim!] (*transitivity*) lemma mless_trans: "\K <# M; M <# N\ \ K <# N" apply (simp add: mless_def, clarify) apply (rule_tac x = "i \ ia" in exI) apply (blast dest: multirel_Memrel_mono [OF Un_upper1 Un_upper1, THEN subsetD] multirel_Memrel_mono [OF Un_upper2 Un_upper2, THEN subsetD] intro: multirel_trans Ord_Un) done (*asymmetry*) lemma mless_not_sym: "M <# N \ \ N <# M" apply clarify apply (rule mless_not_refl [THEN notE]) apply (erule mless_trans, assumption) done lemma mless_asym: "\M <# N; \P \ N <# M\ \ P" by (blast dest: mless_not_sym) lemma mle_refl [simp]: "omultiset(M) \ M <#= M" by (simp add: mle_def) (*anti-symmetry*) lemma mle_antisym: "\M <#= N; N <#= M\ \ M = N" apply (simp add: mle_def) apply (blast dest: mless_not_sym) done (*transitivity*) lemma mle_trans: "\K <#= M; M <#= N\ \ K <#= N" apply (simp add: mle_def) apply (blast intro: mless_trans) done lemma mless_le_iff: "M <# N \ (M <#= N \ M \ N)" by (simp add: mle_def, auto) (** Monotonicity of mless **) lemma munion_less_mono2: "\M <# N; omultiset(K)\ \ K +# M <# K +# N" apply (simp add: mless_def omultiset_def, clarify) apply (rule_tac x = "i \ ia" in exI) apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff) apply (rule munion_multirel_mono2) apply (blast intro: multirel_Memrel_mono [THEN subsetD]) apply (simp add: Mult_iff_multiset) apply (blast intro: field_Memrel_mono [THEN subsetD]) done lemma munion_less_mono1: "\M <# N; omultiset(K)\ \ M +# K <# N +# K" by (force dest: munion_less_mono2 simp add: munion_commute) lemma mless_imp_omultiset: "M <# N \ omultiset(M) \ omultiset(N)" by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD]) lemma munion_less_mono: "\M <# K; N <# L\ \ M +# N <# K +# L" apply (frule_tac M = M in mless_imp_omultiset) apply (frule_tac M = N in mless_imp_omultiset) apply (blast intro: munion_less_mono1 munion_less_mono2 mless_trans) done (* <#= *) lemma mle_imp_omultiset: "M <#= N \ omultiset(M) \ omultiset(N)" by (auto simp add: mle_def mless_imp_omultiset) lemma mle_mono: "\M <#= K; N <#= L\ \ M +# N <#= K +# L" apply (frule_tac M = M in mle_imp_omultiset) apply (frule_tac M = N in mle_imp_omultiset) apply (auto simp add: mle_def intro: munion_less_mono1 munion_less_mono2 munion_less_mono) done lemma omultiset_0 [iff]: "omultiset(0)" by (auto simp add: omultiset_def Mult_iff_multiset) lemma empty_leI [simp]: "omultiset(M) \ 0 <#= M" apply (simp add: mle_def mless_def) apply (subgoal_tac "\i. Ord (i) \ M \ Mult(field(Memrel(i))) ") prefer 2 apply (simp add: omultiset_def) apply (case_tac "M=0", simp_all, clarify) apply (subgoal_tac "<0 +# 0, 0 +# M> \ multirel(field (Memrel(i)), Memrel(i))") apply (rule_tac [2] one_step_implies_multirel) apply (auto simp add: Mult_iff_multiset) done lemma munion_upper1: "\omultiset(M); omultiset(N)\ \ M <#= M +# N" apply (subgoal_tac "M +# 0 <#= M +# N") apply (rule_tac [2] mle_mono, auto) done end diff --git a/src/ZF/Induct/Ntree.thy b/src/ZF/Induct/Ntree.thy --- a/src/ZF/Induct/Ntree.thy +++ b/src/ZF/Induct/Ntree.thy @@ -1,174 +1,174 @@ (* Title: ZF/Induct/Ntree.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Datatype definition n-ary branching trees\ theory Ntree imports ZF begin text \ Demonstrates a simple use of function space in a datatype definition. Based upon theory \Term\. \ consts ntree :: "i \ i" maptree :: "i \ i" maptree2 :: "[i, i] \ i" datatype "ntree(A)" = Branch ("a \ A", "h \ (\n \ nat. n -> ntree(A))") monos UN_mono [OF subset_refl Pi_mono] \ \MUST have this form\ type_intros nat_fun_univ [THEN subsetD] type_elims UN_E datatype "maptree(A)" = Sons ("a \ A", "h \ maptree(A) -||> maptree(A)") monos FiniteFun_mono1 \ \Use monotonicity in BOTH args\ type_intros FiniteFun_univ1 [THEN subsetD] datatype "maptree2(A, B)" = Sons2 ("a \ A", "h \ B -||> maptree2(A, B)") monos FiniteFun_mono [OF subset_refl] type_intros FiniteFun_in_univ' definition ntree_rec :: "[[i, i, i] \ i, i] \ i" where "ntree_rec(b) \ Vrecursor(\pr. ntree_case(\x h. b(x, h, \i \ domain(h). pr`(h`i))))" definition ntree_copy :: "i \ i" where "ntree_copy(z) \ ntree_rec(\x h r. Branch(x,r), z)" text \ \medskip \ntree\ \ lemma ntree_unfold: "ntree(A) = A \ (\n \ nat. n -> ntree(A))" by (blast intro: ntree.intros [unfolded ntree.con_defs] elim: ntree.cases [unfolded ntree.con_defs]) lemma ntree_induct [consumes 1, case_names Branch, induct set: ntree]: assumes t: "t \ ntree(A)" and step: "\x n h. \x \ A; n \ nat; h \ n -> ntree(A); \i \ n. P(h`i) \ \ P(Branch(x,h))" shows "P(t)" \ \A nicer induction rule than the standard one.\ using t apply induct apply (erule UN_E) apply (assumption | rule step)+ apply (fast elim: fun_weaken_type) apply (fast dest: apply_type) done lemma ntree_induct_eqn [consumes 1]: assumes t: "t \ ntree(A)" and f: "f \ ntree(A)->B" and g: "g \ ntree(A)->B" and step: "\x n h. \x \ A; n \ nat; h \ n -> ntree(A); f O h = g O h\ \ f ` Branch(x,h) = g ` Branch(x,h)" shows "f`t=g`t" \ \Induction on \<^term>\ntree(A)\ to prove an equation\ using t apply induct apply (assumption | rule step)+ apply (insert f g) apply (rule fun_extension) apply (assumption | rule comp_fun)+ apply (simp add: comp_fun_apply) done text \ \medskip Lemmas to justify using \Ntree\ in other recursive type definitions. \ lemma ntree_mono: "A \ B \ ntree(A) \ ntree(B)" - apply (unfold ntree.defs) + unfolding ntree.defs apply (rule lfp_mono) apply (rule ntree.bnd_mono)+ apply (assumption | rule univ_mono basic_monos)+ done lemma ntree_univ: "ntree(univ(A)) \ univ(A)" \ \Easily provable by induction also\ - apply (unfold ntree.defs ntree.con_defs) + unfolding ntree.defs ntree.con_defs apply (rule lfp_lowerbound) apply (rule_tac [2] A_subset_univ [THEN univ_mono]) apply (blast intro: Pair_in_univ nat_fun_univ [THEN subsetD]) done lemma ntree_subset_univ: "A \ univ(B) \ ntree(A) \ univ(B)" by (rule subset_trans [OF ntree_mono ntree_univ]) text \ \medskip \ntree\ recursion. \ lemma ntree_rec_Branch: "function(h) \ ntree_rec(b, Branch(x,h)) = b(x, h, \i \ domain(h). ntree_rec(b, h`i))" apply (rule ntree_rec_def [THEN def_Vrecursor, THEN trans]) apply (simp add: ntree.con_defs rank_pair2 [THEN [2] lt_trans] rank_apply) done lemma ntree_copy_Branch [simp]: "function(h) \ ntree_copy (Branch(x, h)) = Branch(x, \i \ domain(h). ntree_copy (h`i))" by (simp add: ntree_copy_def ntree_rec_Branch) lemma ntree_copy_is_ident: "z \ ntree(A) \ ntree_copy(z) = z" by (induct z set: ntree) (auto simp add: domain_of_fun Pi_Collect_iff fun_is_function) text \ \medskip \maptree\ \ lemma maptree_unfold: "maptree(A) = A \ (maptree(A) -||> maptree(A))" by (fast intro!: maptree.intros [unfolded maptree.con_defs] elim: maptree.cases [unfolded maptree.con_defs]) lemma maptree_induct [consumes 1, induct set: maptree]: assumes t: "t \ maptree(A)" and step: "\x n h. \x \ A; h \ maptree(A) -||> maptree(A); \y \ field(h). P(y) \ \ P(Sons(x,h))" shows "P(t)" \ \A nicer induction rule than the standard one.\ using t apply induct apply (assumption | rule step)+ apply (erule Collect_subset [THEN FiniteFun_mono1, THEN subsetD]) apply (drule FiniteFun.dom_subset [THEN subsetD]) apply (drule Fin.dom_subset [THEN subsetD]) apply fast done text \ \medskip \maptree2\ \ lemma maptree2_unfold: "maptree2(A, B) = A \ (B -||> maptree2(A, B))" by (fast intro!: maptree2.intros [unfolded maptree2.con_defs] elim: maptree2.cases [unfolded maptree2.con_defs]) lemma maptree2_induct [consumes 1, induct set: maptree2]: assumes t: "t \ maptree2(A, B)" and step: "\x n h. \x \ A; h \ B -||> maptree2(A,B); \y \ range(h). P(y) \ \ P(Sons2(x,h))" shows "P(t)" using t apply induct apply (assumption | rule step)+ apply (erule FiniteFun_mono [OF subset_refl Collect_subset, THEN subsetD]) apply (drule FiniteFun.dom_subset [THEN subsetD]) apply (drule Fin.dom_subset [THEN subsetD]) apply fast done end diff --git a/src/ZF/Induct/PropLog.thy b/src/ZF/Induct/PropLog.thy --- a/src/ZF/Induct/PropLog.thy +++ b/src/ZF/Induct/PropLog.thy @@ -1,340 +1,340 @@ (* Title: ZF/Induct/PropLog.thy Author: Tobias Nipkow \ Lawrence C Paulson Copyright 1993 University of Cambridge *) section \Meta-theory of propositional logic\ theory PropLog imports ZF begin text \ Datatype definition of propositional logic formulae and inductive definition of the propositional tautologies. Inductive definition of propositional logic. Soundness and completeness w.r.t.\ truth-tables. Prove: If \H |= p\ then \G |= p\ where \G \ Fin(H)\ \ subsection \The datatype of propositions\ consts propn :: i datatype propn = Fls | Var ("n \ nat") (\#_\ [100] 100) | Imp ("p \ propn", "q \ propn") (infixr \\\ 90) subsection \The proof system\ consts thms :: "i \ i" abbreviation thms_syntax :: "[i,i] \ o" (infixl \|-\ 50) where "H |- p \ p \ thms(H)" inductive domains "thms(H)" \ "propn" intros H: "\p \ H; p \ propn\ \ H |- p" K: "\p \ propn; q \ propn\ \ H |- p\q\p" S: "\p \ propn; q \ propn; r \ propn\ \ H |- (p\q\r) \ (p\q) \ p\r" DN: "p \ propn \ H |- ((p\Fls) \ Fls) \ p" MP: "\H |- p\q; H |- p; p \ propn; q \ propn\ \ H |- q" type_intros "propn.intros" declare propn.intros [simp] subsection \The semantics\ subsubsection \Semantics of propositional logic.\ consts is_true_fun :: "[i,i] \ i" primrec "is_true_fun(Fls, t) = 0" "is_true_fun(Var(v), t) = (if v \ t then 1 else 0)" "is_true_fun(p\q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)" definition is_true :: "[i,i] \ o" where "is_true(p,t) \ is_true_fun(p,t) = 1" \ \this definition is required since predicates can't be recursive\ lemma is_true_Fls [simp]: "is_true(Fls,t) \ False" by (simp add: is_true_def) lemma is_true_Var [simp]: "is_true(#v,t) \ v \ t" by (simp add: is_true_def) lemma is_true_Imp [simp]: "is_true(p\q,t) \ (is_true(p,t)\is_true(q,t))" by (simp add: is_true_def) subsubsection \Logical consequence\ text \ For every valuation, if all elements of \H\ are true then so is \p\. \ definition logcon :: "[i,i] \ o" (infixl \|=\ 50) where "H |= p \ \t. (\q \ H. is_true(q,t)) \ is_true(p,t)" text \ A finite set of hypotheses from \t\ and the \Var\s in \p\. \ consts hyps :: "[i,i] \ i" primrec "hyps(Fls, t) = 0" "hyps(Var(v), t) = (if v \ t then {#v} else {#v\Fls})" "hyps(p\q, t) = hyps(p,t) \ hyps(q,t)" subsection \Proof theory of propositional logic\ lemma thms_mono: "G \ H \ thms(G) \ thms(H)" - apply (unfold thms.defs) + unfolding thms.defs apply (rule lfp_mono) apply (rule thms.bnd_mono)+ apply (assumption | rule univ_mono basic_monos)+ done lemmas thms_in_pl = thms.dom_subset [THEN subsetD] inductive_cases ImpE: "p\q \ propn" lemma thms_MP: "\H |- p\q; H |- p\ \ H |- q" \ \Stronger Modus Ponens rule: no typechecking!\ apply (rule thms.MP) apply (erule asm_rl thms_in_pl thms_in_pl [THEN ImpE])+ done lemma thms_I: "p \ propn \ H |- p\p" \ \Rule is called \I\ for Identity Combinator, not for Introduction.\ apply (rule thms.S [THEN thms_MP, THEN thms_MP]) apply (rule_tac [5] thms.K) apply (rule_tac [4] thms.K) apply simp_all done subsubsection \Weakening, left and right\ lemma weaken_left: "\G \ H; G|-p\ \ H|-p" \ \Order of premises is convenient with \THEN\\ by (erule thms_mono [THEN subsetD]) lemma weaken_left_cons: "H |- p \ cons(a,H) |- p" by (erule subset_consI [THEN weaken_left]) lemmas weaken_left_Un1 = Un_upper1 [THEN weaken_left] lemmas weaken_left_Un2 = Un_upper2 [THEN weaken_left] lemma weaken_right: "\H |- q; p \ propn\ \ H |- p\q" by (simp_all add: thms.K [THEN thms_MP] thms_in_pl) subsubsection \The deduction theorem\ theorem deduction: "\cons(p,H) |- q; p \ propn\ \ H |- p\q" apply (erule thms.induct) apply (blast intro: thms_I thms.H [THEN weaken_right]) apply (blast intro: thms.K [THEN weaken_right]) apply (blast intro: thms.S [THEN weaken_right]) apply (blast intro: thms.DN [THEN weaken_right]) apply (blast intro: thms.S [THEN thms_MP [THEN thms_MP]]) done subsubsection \The cut rule\ lemma cut: "\H|-p; cons(p,H) |- q\ \ H |- q" apply (rule deduction [THEN thms_MP]) apply (simp_all add: thms_in_pl) done lemma thms_FlsE: "\H |- Fls; p \ propn\ \ H |- p" apply (rule thms.DN [THEN thms_MP]) apply (rule_tac [2] weaken_right) apply (simp_all add: propn.intros) done lemma thms_notE: "\H |- p\Fls; H |- p; q \ propn\ \ H |- q" by (erule thms_MP [THEN thms_FlsE]) subsubsection \Soundness of the rules wrt truth-table semantics\ theorem soundness: "H |- p \ H |= p" unfolding logcon_def apply (induct set: thms) apply auto done subsection \Completeness\ subsubsection \Towards the completeness proof\ lemma Fls_Imp: "\H |- p\Fls; q \ propn\ \ H |- p\q" apply (frule thms_in_pl) apply (rule deduction) apply (rule weaken_left_cons [THEN thms_notE]) apply (blast intro: thms.H elim: ImpE)+ done lemma Imp_Fls: "\H |- p; H |- q\Fls\ \ H |- (p\q)\Fls" apply (frule thms_in_pl) apply (frule thms_in_pl [of concl: "q\Fls"]) apply (rule deduction) apply (erule weaken_left_cons [THEN thms_MP]) apply (rule consI1 [THEN thms.H, THEN thms_MP]) apply (blast intro: weaken_left_cons elim: ImpE)+ done lemma hyps_thms_if: "p \ propn \ hyps(p,t) |- (if is_true(p,t) then p else p\Fls)" \ \Typical example of strengthening the induction statement.\ apply simp apply (induct_tac p) apply (simp_all add: thms_I thms.H) apply (safe elim!: Fls_Imp [THEN weaken_left_Un1] Fls_Imp [THEN weaken_left_Un2]) apply (blast intro: weaken_left_Un1 weaken_left_Un2 weaken_right Imp_Fls)+ done lemma logcon_thms_p: "\p \ propn; 0 |= p\ \ hyps(p,t) |- p" \ \Key lemma for completeness; yields a set of assumptions satisfying \p\\ apply (drule hyps_thms_if) apply (simp add: logcon_def) done text \ For proving certain theorems in our new propositional logic. \ lemmas propn_SIs = propn.intros deduction and propn_Is = thms_in_pl thms.H thms.H [THEN thms_MP] text \ The excluded middle in the form of an elimination rule. \ lemma thms_excluded_middle: "\p \ propn; q \ propn\ \ H |- (p\q) \ ((p\Fls)\q) \ q" apply (rule deduction [THEN deduction]) apply (rule thms.DN [THEN thms_MP]) apply (best intro!: propn_SIs intro: propn_Is)+ done lemma thms_excluded_middle_rule: "\cons(p,H) |- q; cons(p\Fls,H) |- q; p \ propn\ \ H |- q" \ \Hard to prove directly because it requires cuts\ apply (rule thms_excluded_middle [THEN thms_MP, THEN thms_MP]) apply (blast intro!: propn_SIs intro: propn_Is)+ done subsubsection \Completeness -- lemmas for reducing the set of assumptions\ text \ For the case \<^prop>\hyps(p,t)-cons(#v,Y) |- p\ we also have \<^prop>\hyps(p,t)-{#v} \ hyps(p, t-{v})\. \ lemma hyps_Diff: "p \ propn \ hyps(p, t-{v}) \ cons(#v\Fls, hyps(p,t)-{#v})" by (induct set: propn) auto text \ For the case \<^prop>\hyps(p,t)-cons(#v \ Fls,Y) |- p\ we also have \<^prop>\hyps(p,t)-{#v\Fls} \ hyps(p, cons(v,t))\. \ lemma hyps_cons: "p \ propn \ hyps(p, cons(v,t)) \ cons(#v, hyps(p,t)-{#v\Fls})" by (induct set: propn) auto text \Two lemmas for use with \weaken_left\\ lemma cons_Diff_same: "B-C \ cons(a, B-cons(a,C))" by blast lemma cons_Diff_subset2: "cons(a, B-{c}) - D \ cons(a, B-cons(c,D))" by blast text \ The set \<^term>\hyps(p,t)\ is finite, and elements have the form \<^term>\#v\ or \<^term>\#v\Fls\; could probably prove the stronger \<^prop>\hyps(p,t) \ Fin(hyps(p,0) \ hyps(p,nat))\. \ lemma hyps_finite: "p \ propn \ hyps(p,t) \ Fin(\v \ nat. {#v, #v\Fls})" by (induct set: propn) auto lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left] text \ Induction on the finite set of assumptions \<^term>\hyps(p,t0)\. We may repeatedly subtract assumptions until none are left! \ lemma completeness_0_lemma [rule_format]: "\p \ propn; 0 |= p\ \ \t. hyps(p,t) - hyps(p,t0) |- p" apply (frule hyps_finite) apply (erule Fin_induct) apply (simp add: logcon_thms_p Diff_0) txt \inductive step\ apply safe txt \Case \<^prop>\hyps(p,t)-cons(#v,Y) |- p\\ apply (rule thms_excluded_middle_rule) apply (erule_tac [3] propn.intros) apply (blast intro: cons_Diff_same [THEN weaken_left]) apply (blast intro: cons_Diff_subset2 [THEN weaken_left] hyps_Diff [THEN Diff_weaken_left]) txt \Case \<^prop>\hyps(p,t)-cons(#v \ Fls,Y) |- p\\ apply (rule thms_excluded_middle_rule) apply (erule_tac [3] propn.intros) apply (blast intro: cons_Diff_subset2 [THEN weaken_left] hyps_cons [THEN Diff_weaken_left]) apply (blast intro: cons_Diff_same [THEN weaken_left]) done subsubsection \Completeness theorem\ lemma completeness_0: "\p \ propn; 0 |= p\ \ 0 |- p" \ \The base case for completeness\ apply (rule Diff_cancel [THEN subst]) apply (blast intro: completeness_0_lemma) done lemma logcon_Imp: "\cons(p,H) |= q\ \ H |= p\q" \ \A semantic analogue of the Deduction Theorem\ by (simp add: logcon_def) lemma completeness: "H \ Fin(propn) \ p \ propn \ H |= p \ H |- p" apply (induct arbitrary: p set: Fin) apply (safe intro!: completeness_0) apply (rule weaken_left_cons [THEN thms_MP]) apply (blast intro!: logcon_Imp propn.intros) apply (blast intro: propn_Is) done theorem thms_iff: "H \ Fin(propn) \ H |- p \ H |= p \ p \ propn" by (blast intro: soundness completeness thms_in_pl) end diff --git a/src/ZF/Induct/Rmap.thy b/src/ZF/Induct/Rmap.thy --- a/src/ZF/Induct/Rmap.thy +++ b/src/ZF/Induct/Rmap.thy @@ -1,70 +1,70 @@ (* Title: ZF/Induct/Rmap.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \An operator to ``map'' a relation over a list\ theory Rmap imports ZF begin consts rmap :: "i\i" inductive domains "rmap(r)" \ "list(domain(r)) \ list(range(r))" intros NilI: "\Nil,Nil\ \ rmap(r)" ConsI: "\\x,y\: r; \xs,ys\ \ rmap(r)\ \ \ rmap(r)" type_intros domainI rangeI list.intros lemma rmap_mono: "r \ s \ rmap(r) \ rmap(s)" - apply (unfold rmap.defs) + unfolding rmap.defs apply (rule lfp_mono) apply (rule rmap.bnd_mono)+ apply (assumption | rule Sigma_mono list_mono domain_mono range_mono basic_monos)+ done inductive_cases Nil_rmap_case [elim!]: "\Nil,zs\ \ rmap(r)" and Cons_rmap_case [elim!]: " \ rmap(r)" declare rmap.intros [intro] lemma rmap_rel_type: "r \ A \ B \ rmap(r) \ list(A) \ list(B)" apply (rule rmap.dom_subset [THEN subset_trans]) apply (assumption | rule domain_rel_subset range_rel_subset Sigma_mono list_mono)+ done lemma rmap_total: "A \ domain(r) \ list(A) \ domain(rmap(r))" apply (rule subsetI) apply (erule list.induct) apply blast+ done lemma rmap_functional: "function(r) \ function(rmap(r))" unfolding function_def apply (rule impI [THEN allI, THEN allI]) apply (erule rmap.induct) apply blast+ done text \ \medskip If \f\ is a function then \rmap(f)\ behaves as expected. \ lemma rmap_fun_type: "f \ A->B \ rmap(f): list(A)->list(B)" by (simp add: Pi_iff rmap_rel_type rmap_functional rmap_total) lemma rmap_Nil: "rmap(f)`Nil = Nil" by (unfold apply_def) blast lemma rmap_Cons: "\f \ A->B; x \ A; xs: list(A)\ \ rmap(f) ` Cons(x,xs) = Cons(f`x, rmap(f)`xs)" by (blast intro: apply_equality apply_Pair rmap_fun_type rmap.intros) end diff --git a/src/ZF/Induct/Term.thy b/src/ZF/Induct/Term.thy --- a/src/ZF/Induct/Term.thy +++ b/src/ZF/Induct/Term.thy @@ -1,284 +1,284 @@ (* Title: ZF/Induct/Term.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Terms over an alphabet\ theory Term imports ZF begin text \ Illustrates the list functor (essentially the same type as in \Trees_Forest\). \ consts "term" :: "i \ i" datatype "term(A)" = Apply ("a \ A", "l \ list(term(A))") monos list_mono type_elims list_univ [THEN subsetD, elim_format] declare Apply [TC] definition term_rec :: "[i, [i, i, i] \ i] \ i" where "term_rec(t,d) \ Vrec(t, \t g. term_case(\x zs. d(x, zs, map(\z. g`z, zs)), t))" definition term_map :: "[i \ i, i] \ i" where "term_map(f,t) \ term_rec(t, \x zs rs. Apply(f(x), rs))" definition term_size :: "i \ i" where "term_size(t) \ term_rec(t, \x zs rs. succ(list_add(rs)))" definition reflect :: "i \ i" where "reflect(t) \ term_rec(t, \x zs rs. Apply(x, rev(rs)))" definition preorder :: "i \ i" where "preorder(t) \ term_rec(t, \x zs rs. Cons(x, flat(rs)))" definition postorder :: "i \ i" where "postorder(t) \ term_rec(t, \x zs rs. flat(rs) @ [x])" lemma term_unfold: "term(A) = A * list(term(A))" by (fast intro!: term.intros [unfolded term.con_defs] elim: term.cases [unfolded term.con_defs]) lemma term_induct2: "\t \ term(A); \x. \x \ A\ \ P(Apply(x,Nil)); \x z zs. \x \ A; z \ term(A); zs: list(term(A)); P(Apply(x,zs)) \ \ P(Apply(x, Cons(z,zs))) \ \ P(t)" \ \Induction on \<^term>\term(A)\ followed by induction on \<^term>\list\.\ apply (induct_tac t) apply (erule list.induct) apply (auto dest: list_CollectD) done lemma term_induct_eqn [consumes 1, case_names Apply]: "\t \ term(A); \x zs. \x \ A; zs: list(term(A)); map(f,zs) = map(g,zs)\ \ f(Apply(x,zs)) = g(Apply(x,zs)) \ \ f(t) = g(t)" \ \Induction on \<^term>\term(A)\ to prove an equation.\ apply (induct_tac t) apply (auto dest: map_list_Collect list_CollectD) done text \ \medskip Lemmas to justify using \<^term>\term\ in other recursive type definitions. \ lemma term_mono: "A \ B \ term(A) \ term(B)" - apply (unfold term.defs) + unfolding term.defs apply (rule lfp_mono) apply (rule term.bnd_mono)+ apply (rule univ_mono basic_monos| assumption)+ done lemma term_univ: "term(univ(A)) \ univ(A)" \ \Easily provable by induction also\ - apply (unfold term.defs term.con_defs) + unfolding term.defs term.con_defs apply (rule lfp_lowerbound) apply (rule_tac [2] A_subset_univ [THEN univ_mono]) apply safe apply (assumption | rule Pair_in_univ list_univ [THEN subsetD])+ done lemma term_subset_univ: "A \ univ(B) \ term(A) \ univ(B)" apply (rule subset_trans) apply (erule term_mono) apply (rule term_univ) done lemma term_into_univ: "\t \ term(A); A \ univ(B)\ \ t \ univ(B)" by (rule term_subset_univ [THEN subsetD]) text \ \medskip \term_rec\ -- by \Vset\ recursion. \ lemma map_lemma: "\l \ list(A); Ord(i); rank(l) \ map(\z. (\x \ Vset(i).h(x)) ` z, l) = map(h,l)" \ \\<^term>\map\ works correctly on the underlying list of terms.\ apply (induct set: list) apply simp apply (subgoal_tac "rank (a) rank (l) < i") apply (simp add: rank_of_Ord) apply (simp add: list.con_defs) apply (blast dest: rank_rls [THEN lt_trans]) done lemma term_rec [simp]: "ts \ list(A) \ term_rec(Apply(a,ts), d) = d(a, ts, map (\z. term_rec(z,d), ts))" \ \Typing premise is necessary to invoke \map_lemma\.\ apply (rule term_rec_def [THEN def_Vrec, THEN trans]) - apply (unfold term.con_defs) + unfolding term.con_defs apply (simp add: rank_pair2 map_lemma) done lemma term_rec_type: assumes t: "t \ term(A)" and a: "\x zs r. \x \ A; zs: list(term(A)); r \ list(\t \ term(A). C(t))\ \ d(x, zs, r): C(Apply(x,zs))" shows "term_rec(t,d) \ C(t)" \ \Slightly odd typing condition on \r\ in the second premise!\ using t apply induct apply (frule list_CollectD) apply (subst term_rec) apply (assumption | rule a)+ apply (erule list.induct) apply auto done lemma def_term_rec: "\\t. j(t)\term_rec(t,d); ts: list(A)\ \ j(Apply(a,ts)) = d(a, ts, map(\Z. j(Z), ts))" apply (simp only:) apply (erule term_rec) done lemma term_rec_simple_type [TC]: "\t \ term(A); \x zs r. \x \ A; zs: list(term(A)); r \ list(C)\ \ d(x, zs, r): C \ \ term_rec(t,d) \ C" apply (erule term_rec_type) apply (drule subset_refl [THEN UN_least, THEN list_mono, THEN subsetD]) apply simp done text \ \medskip \<^term>\term_map\. \ lemma term_map [simp]: "ts \ list(A) \ term_map(f, Apply(a, ts)) = Apply(f(a), map(term_map(f), ts))" by (rule term_map_def [THEN def_term_rec]) lemma term_map_type [TC]: "\t \ term(A); \x. x \ A \ f(x): B\ \ term_map(f,t) \ term(B)" unfolding term_map_def apply (erule term_rec_simple_type) apply fast done lemma term_map_type2 [TC]: "t \ term(A) \ term_map(f,t) \ term({f(u). u \ A})" apply (erule term_map_type) apply (erule RepFunI) done text \ \medskip \<^term>\term_size\. \ lemma term_size [simp]: "ts \ list(A) \ term_size(Apply(a, ts)) = succ(list_add(map(term_size, ts)))" by (rule term_size_def [THEN def_term_rec]) lemma term_size_type [TC]: "t \ term(A) \ term_size(t) \ nat" by (auto simp add: term_size_def) text \ \medskip \reflect\. \ lemma reflect [simp]: "ts \ list(A) \ reflect(Apply(a, ts)) = Apply(a, rev(map(reflect, ts)))" by (rule reflect_def [THEN def_term_rec]) lemma reflect_type [TC]: "t \ term(A) \ reflect(t) \ term(A)" by (auto simp add: reflect_def) text \ \medskip \preorder\. \ lemma preorder [simp]: "ts \ list(A) \ preorder(Apply(a, ts)) = Cons(a, flat(map(preorder, ts)))" by (rule preorder_def [THEN def_term_rec]) lemma preorder_type [TC]: "t \ term(A) \ preorder(t) \ list(A)" by (simp add: preorder_def) text \ \medskip \postorder\. \ lemma postorder [simp]: "ts \ list(A) \ postorder(Apply(a, ts)) = flat(map(postorder, ts)) @ [a]" by (rule postorder_def [THEN def_term_rec]) lemma postorder_type [TC]: "t \ term(A) \ postorder(t) \ list(A)" by (simp add: postorder_def) text \ \medskip Theorems about \term_map\. \ declare map_compose [simp] lemma term_map_ident: "t \ term(A) \ term_map(\u. u, t) = t" by (induct rule: term_induct_eqn) simp lemma term_map_compose: "t \ term(A) \ term_map(f, term_map(g,t)) = term_map(\u. f(g(u)), t)" by (induct rule: term_induct_eqn) simp lemma term_map_reflect: "t \ term(A) \ term_map(f, reflect(t)) = reflect(term_map(f,t))" by (induct rule: term_induct_eqn) (simp add: rev_map_distrib [symmetric]) text \ \medskip Theorems about \term_size\. \ lemma term_size_term_map: "t \ term(A) \ term_size(term_map(f,t)) = term_size(t)" by (induct rule: term_induct_eqn) simp lemma term_size_reflect: "t \ term(A) \ term_size(reflect(t)) = term_size(t)" by (induct rule: term_induct_eqn) (simp add: rev_map_distrib [symmetric] list_add_rev) lemma term_size_length: "t \ term(A) \ term_size(t) = length(preorder(t))" by (induct rule: term_induct_eqn) (simp add: length_flat) text \ \medskip Theorems about \reflect\. \ lemma reflect_reflect_ident: "t \ term(A) \ reflect(reflect(t)) = t" by (induct rule: term_induct_eqn) (simp add: rev_map_distrib) text \ \medskip Theorems about preorder. \ lemma preorder_term_map: "t \ term(A) \ preorder(term_map(f,t)) = map(f, preorder(t))" by (induct rule: term_induct_eqn) (simp add: map_flat) lemma preorder_reflect_eq_rev_postorder: "t \ term(A) \ preorder(reflect(t)) = rev(postorder(t))" by (induct rule: term_induct_eqn) (simp add: rev_app_distrib rev_flat rev_map_distrib [symmetric]) end diff --git a/src/ZF/Induct/Tree_Forest.thy b/src/ZF/Induct/Tree_Forest.thy --- a/src/ZF/Induct/Tree_Forest.thy +++ b/src/ZF/Induct/Tree_Forest.thy @@ -1,251 +1,251 @@ (* Title: ZF/Induct/Tree_Forest.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Trees and forests, a mutually recursive type definition\ theory Tree_Forest imports ZF begin subsection \Datatype definition\ consts tree :: "i \ i" forest :: "i \ i" tree_forest :: "i \ i" datatype "tree(A)" = Tcons ("a \ A", "f \ forest(A)") and "forest(A)" = Fnil | Fcons ("t \ tree(A)", "f \ forest(A)") (* FIXME *) lemmas tree'induct = tree_forest.mutual_induct [THEN conjunct1, THEN spec, THEN [2] rev_mp, of concl: _ t, consumes 1] and forest'induct = tree_forest.mutual_induct [THEN conjunct2, THEN spec, THEN [2] rev_mp, of concl: _ f, consumes 1] for t f declare tree_forest.intros [simp, TC] lemma tree_def: "tree(A) \ Part(tree_forest(A), Inl)" by (simp only: tree_forest.defs) lemma forest_def: "forest(A) \ Part(tree_forest(A), Inr)" by (simp only: tree_forest.defs) text \ \medskip \<^term>\tree_forest(A)\ as the union of \<^term>\tree(A)\ and \<^term>\forest(A)\. \ lemma tree_subset_TF: "tree(A) \ tree_forest(A)" - apply (unfold tree_forest.defs) + unfolding tree_forest.defs apply (rule Part_subset) done lemma treeI [TC]: "x \ tree(A) \ x \ tree_forest(A)" by (rule tree_subset_TF [THEN subsetD]) lemma forest_subset_TF: "forest(A) \ tree_forest(A)" - apply (unfold tree_forest.defs) + unfolding tree_forest.defs apply (rule Part_subset) done lemma treeI' [TC]: "x \ forest(A) \ x \ tree_forest(A)" by (rule forest_subset_TF [THEN subsetD]) lemma TF_equals_Un: "tree(A) \ forest(A) = tree_forest(A)" apply (insert tree_subset_TF forest_subset_TF) apply (auto intro!: equalityI tree_forest.intros elim: tree_forest.cases) done lemma tree_forest_unfold: "tree_forest(A) = (A \ forest(A)) + ({0} + tree(A) \ forest(A))" \ \NOT useful, but interesting \dots\ supply rews = tree_forest.con_defs tree_def forest_def - apply (unfold tree_def forest_def) + unfolding tree_def forest_def apply (fast intro!: tree_forest.intros [unfolded rews, THEN PartD1] elim: tree_forest.cases [unfolded rews]) done lemma tree_forest_unfold': "tree_forest(A) = A \ Part(tree_forest(A), \w. Inr(w)) + {0} + Part(tree_forest(A), \w. Inl(w)) * Part(tree_forest(A), \w. Inr(w))" by (rule tree_forest_unfold [unfolded tree_def forest_def]) lemma tree_unfold: "tree(A) = {Inl(x). x \ A \ forest(A)}" - apply (unfold tree_def forest_def) + unfolding tree_def forest_def apply (rule Part_Inl [THEN subst]) apply (rule tree_forest_unfold' [THEN subst_context]) done lemma forest_unfold: "forest(A) = {Inr(x). x \ {0} + tree(A)*forest(A)}" - apply (unfold tree_def forest_def) + unfolding tree_def forest_def apply (rule Part_Inr [THEN subst]) apply (rule tree_forest_unfold' [THEN subst_context]) done text \ \medskip Type checking for recursor: Not needed; possibly interesting? \ lemma TF_rec_type: "\z \ tree_forest(A); \x f r. \x \ A; f \ forest(A); r \ C(f) \ \ b(x,f,r) \ C(Tcons(x,f)); c \ C(Fnil); \t f r1 r2. \t \ tree(A); f \ forest(A); r1 \ C(t); r2 \ C(f) \ \ d(t,f,r1,r2) \ C(Fcons(t,f)) \ \ tree_forest_rec(b,c,d,z) \ C(z)" by (induct_tac z) simp_all lemma tree_forest_rec_type: "\\x f r. \x \ A; f \ forest(A); r \ D(f) \ \ b(x,f,r) \ C(Tcons(x,f)); c \ D(Fnil); \t f r1 r2. \t \ tree(A); f \ forest(A); r1 \ C(t); r2 \ D(f) \ \ d(t,f,r1,r2) \ D(Fcons(t,f)) \ \ (\t \ tree(A). tree_forest_rec(b,c,d,t) \ C(t)) \ (\f \ forest(A). tree_forest_rec(b,c,d,f) \ D(f))" \ \Mutually recursive version.\ unfolding Ball_def apply (rule tree_forest.mutual_induct) apply simp_all done subsection \Operations\ consts map :: "[i \ i, i] \ i" size :: "i \ i" preorder :: "i \ i" list_of_TF :: "i \ i" of_list :: "i \ i" reflect :: "i \ i" primrec "list_of_TF (Tcons(x,f)) = [Tcons(x,f)]" "list_of_TF (Fnil) = []" "list_of_TF (Fcons(t,tf)) = Cons (t, list_of_TF(tf))" primrec "of_list([]) = Fnil" "of_list(Cons(t,l)) = Fcons(t, of_list(l))" primrec "map (h, Tcons(x,f)) = Tcons(h(x), map(h,f))" "map (h, Fnil) = Fnil" "map (h, Fcons(t,tf)) = Fcons (map(h, t), map(h, tf))" primrec "size (Tcons(x,f)) = succ(size(f))" "size (Fnil) = 0" "size (Fcons(t,tf)) = size(t) #+ size(tf)" primrec "preorder (Tcons(x,f)) = Cons(x, preorder(f))" "preorder (Fnil) = Nil" "preorder (Fcons(t,tf)) = preorder(t) @ preorder(tf)" primrec "reflect (Tcons(x,f)) = Tcons(x, reflect(f))" "reflect (Fnil) = Fnil" "reflect (Fcons(t,tf)) = of_list (list_of_TF (reflect(tf)) @ Cons(reflect(t), Nil))" text \ \medskip \list_of_TF\ and \of_list\. \ lemma list_of_TF_type [TC]: "z \ tree_forest(A) \ list_of_TF(z) \ list(tree(A))" by (induct set: tree_forest) simp_all lemma of_list_type [TC]: "l \ list(tree(A)) \ of_list(l) \ forest(A)" by (induct set: list) simp_all text \ \medskip \map\. \ lemma assumes "\x. x \ A \ h(x): B" shows map_tree_type: "t \ tree(A) \ map(h,t) \ tree(B)" and map_forest_type: "f \ forest(A) \ map(h,f) \ forest(B)" using assms by (induct rule: tree'induct forest'induct) simp_all text \ \medskip \size\. \ lemma size_type [TC]: "z \ tree_forest(A) \ size(z) \ nat" by (induct set: tree_forest) simp_all text \ \medskip \preorder\. \ lemma preorder_type [TC]: "z \ tree_forest(A) \ preorder(z) \ list(A)" by (induct set: tree_forest) simp_all text \ \medskip Theorems about \list_of_TF\ and \of_list\. \ lemma forest_induct [consumes 1, case_names Fnil Fcons]: "\f \ forest(A); R(Fnil); \t f. \t \ tree(A); f \ forest(A); R(f)\ \ R(Fcons(t,f)) \ \ R(f)" \ \Essentially the same as list induction.\ apply (erule tree_forest.mutual_induct [THEN conjunct2, THEN spec, THEN [2] rev_mp]) apply (rule TrueI) apply simp apply simp done lemma forest_iso: "f \ forest(A) \ of_list(list_of_TF(f)) = f" by (induct rule: forest_induct) simp_all lemma tree_list_iso: "ts: list(tree(A)) \ list_of_TF(of_list(ts)) = ts" by (induct set: list) simp_all text \ \medskip Theorems about \map\. \ lemma map_ident: "z \ tree_forest(A) \ map(\u. u, z) = z" by (induct set: tree_forest) simp_all lemma map_compose: "z \ tree_forest(A) \ map(h, map(j,z)) = map(\u. h(j(u)), z)" by (induct set: tree_forest) simp_all text \ \medskip Theorems about \size\. \ lemma size_map: "z \ tree_forest(A) \ size(map(h,z)) = size(z)" by (induct set: tree_forest) simp_all lemma size_length: "z \ tree_forest(A) \ size(z) = length(preorder(z))" by (induct set: tree_forest) (simp_all add: length_app) text \ \medskip Theorems about \preorder\. \ lemma preorder_map: "z \ tree_forest(A) \ preorder(map(h,z)) = List.map(h, preorder(z))" by (induct set: tree_forest) (simp_all add: map_app_distrib) end diff --git a/src/ZF/List.thy b/src/ZF/List.thy --- a/src/ZF/List.thy +++ b/src/ZF/List.thy @@ -1,1270 +1,1270 @@ (* Title: ZF/List.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Lists in Zermelo-Fraenkel Set Theory\ theory List imports Datatype ArithSimp begin consts list :: "i\i" datatype "list(A)" = Nil | Cons ("a \ A", "l \ list(A)") syntax "_Nil" :: i (\[]\) "_List" :: "is \ i" (\[(_)]\) translations "[x, xs]" == "CONST Cons(x, [xs])" "[x]" == "CONST Cons(x, [])" "[]" == "CONST Nil" consts length :: "i\i" hd :: "i\i" tl :: "i\i" primrec "length([]) = 0" "length(Cons(a,l)) = succ(length(l))" primrec "hd([]) = 0" "hd(Cons(a,l)) = a" primrec "tl([]) = []" "tl(Cons(a,l)) = l" consts map :: "[i\i, i] \ i" set_of_list :: "i\i" app :: "[i,i]\i" (infixr \@\ 60) (*map is a binding operator -- it applies to meta-level functions, not object-level functions. This simplifies the final form of term_rec_conv, although complicating its derivation.*) primrec "map(f,[]) = []" "map(f,Cons(a,l)) = Cons(f(a), map(f,l))" primrec "set_of_list([]) = 0" "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))" primrec app_Nil: "[] @ ys = ys" app_Cons: "(Cons(a,l)) @ ys = Cons(a, l @ ys)" consts rev :: "i\i" flat :: "i\i" list_add :: "i\i" primrec "rev([]) = []" "rev(Cons(a,l)) = rev(l) @ [a]" primrec "flat([]) = []" "flat(Cons(l,ls)) = l @ flat(ls)" primrec "list_add([]) = 0" "list_add(Cons(a,l)) = a #+ list_add(l)" consts drop :: "[i,i]\i" primrec drop_0: "drop(0,l) = l" drop_succ: "drop(succ(i), l) = tl (drop(i,l))" (*** Thanks to Sidi Ehmety for the following ***) definition (* Function `take' returns the first n elements of a list *) take :: "[i,i]\i" where "take(n, as) \ list_rec(\n\nat. [], \a l r. \n\nat. nat_case([], \m. Cons(a, r`m), n), as)`n" definition nth :: "[i, i]\i" where \ \returns the (n+1)th element of a list, or 0 if the list is too short.\ "nth(n, as) \ list_rec(\n\nat. 0, \a l r. \n\nat. nat_case(a, \m. r`m, n), as) ` n" definition list_update :: "[i, i, i]\i" where "list_update(xs, i, v) \ list_rec(\n\nat. Nil, \u us vs. \n\nat. nat_case(Cons(v, us), \m. Cons(u, vs`m), n), xs)`i" consts filter :: "[i\o, i] \ i" upt :: "[i, i] \i" primrec "filter(P, Nil) = Nil" "filter(P, Cons(x, xs)) = (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))" primrec "upt(i, 0) = Nil" "upt(i, succ(j)) = (if i \ j then upt(i, j)@[j] else Nil)" definition min :: "[i,i] \i" where "min(x, y) \ (if x \ y then x else y)" definition max :: "[i, i] \i" where "max(x, y) \ (if x \ y then y else x)" (*** Aspects of the datatype definition ***) declare list.intros [simp,TC] (*An elimination rule, for type-checking*) inductive_cases ConsE: "Cons(a,l) \ list(A)" lemma Cons_type_iff [simp]: "Cons(a,l) \ list(A) \ a \ A \ l \ list(A)" by (blast elim: ConsE) (*Proving freeness results*) lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' \ l=l'" by auto lemma Nil_Cons_iff: "\ Nil=Cons(a,l)" by auto lemma list_unfold: "list(A) = {0} + (A * list(A))" by (blast intro!: list.intros [unfolded list.con_defs] elim: list.cases [unfolded list.con_defs]) (** Lemmas to justify using "list" in other recursive type definitions **) lemma list_mono: "A<=B \ list(A) \ list(B)" -apply (unfold list.defs ) + unfolding list.defs apply (rule lfp_mono) apply (simp_all add: list.bnd_mono) apply (assumption | rule univ_mono basic_monos)+ done (*There is a similar proof by list induction.*) lemma list_univ: "list(univ(A)) \ univ(A)" -apply (unfold list.defs list.con_defs) + unfolding list.defs list.con_defs apply (rule lfp_lowerbound) apply (rule_tac [2] A_subset_univ [THEN univ_mono]) apply (blast intro!: zero_in_univ Inl_in_univ Inr_in_univ Pair_in_univ) done (*These two theorems justify datatypes involving list(nat), list(A), ...*) lemmas list_subset_univ = subset_trans [OF list_mono list_univ] lemma list_into_univ: "\l \ list(A); A \ univ(B)\ \ l \ univ(B)" by (blast intro: list_subset_univ [THEN subsetD]) lemma list_case_type: "\l \ list(A); c \ C(Nil); \x y. \x \ A; y \ list(A)\ \ h(x,y): C(Cons(x,y)) \ \ list_case(c,h,l) \ C(l)" by (erule list.induct, auto) lemma list_0_triv: "list(0) = {Nil}" apply (rule equalityI, auto) apply (induct_tac x, auto) done (*** List functions ***) lemma tl_type: "l \ list(A) \ tl(l) \ list(A)" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: list.intros) done (** drop **) lemma drop_Nil [simp]: "i \ nat \ drop(i, Nil) = Nil" apply (induct_tac "i") apply (simp_all (no_asm_simp)) done lemma drop_succ_Cons [simp]: "i \ nat \ drop(succ(i), Cons(a,l)) = drop(i,l)" apply (rule sym) apply (induct_tac "i") apply (simp (no_asm)) apply (simp (no_asm_simp)) done lemma drop_type [simp,TC]: "\i \ nat; l \ list(A)\ \ drop(i,l) \ list(A)" apply (induct_tac "i") apply (simp_all (no_asm_simp) add: tl_type) done declare drop_succ [simp del] (** Type checking -- proved by induction, as usual **) lemma list_rec_type [TC]: "\l \ list(A); c \ C(Nil); \x y r. \x \ A; y \ list(A); r \ C(y)\ \ h(x,y,r): C(Cons(x,y)) \ \ list_rec(c,h,l) \ C(l)" by (induct_tac "l", auto) (** map **) lemma map_type [TC]: "\l \ list(A); \x. x \ A \ h(x): B\ \ map(h,l) \ list(B)" apply (simp add: map_list_def) apply (typecheck add: list.intros list_rec_type, blast) done lemma map_type2 [TC]: "l \ list(A) \ map(h,l) \ list({h(u). u \ A})" apply (erule map_type) apply (erule RepFunI) done (** length **) lemma length_type [TC]: "l \ list(A) \ length(l) \ nat" by (simp add: length_list_def) lemma lt_length_in_nat: "\x < length(xs); xs \ list(A)\ \ x \ nat" by (frule lt_nat_in_nat, typecheck) (** app **) lemma app_type [TC]: "\xs: list(A); ys: list(A)\ \ xs@ys \ list(A)" by (simp add: app_list_def) (** rev **) lemma rev_type [TC]: "xs: list(A) \ rev(xs) \ list(A)" by (simp add: rev_list_def) (** flat **) lemma flat_type [TC]: "ls: list(list(A)) \ flat(ls) \ list(A)" by (simp add: flat_list_def) (** set_of_list **) lemma set_of_list_type [TC]: "l \ list(A) \ set_of_list(l) \ Pow(A)" unfolding set_of_list_list_def apply (erule list_rec_type, auto) done lemma set_of_list_append: "xs: list(A) \ set_of_list (xs@ys) = set_of_list(xs) \ set_of_list(ys)" apply (erule list.induct) apply (simp_all (no_asm_simp) add: Un_cons) done (** list_add **) lemma list_add_type [TC]: "xs: list(nat) \ list_add(xs) \ nat" by (simp add: list_add_list_def) (*** theorems about map ***) lemma map_ident [simp]: "l \ list(A) \ map(\u. u, l) = l" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done lemma map_compose: "l \ list(A) \ map(h, map(j,l)) = map(\u. h(j(u)), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done lemma map_app_distrib: "xs: list(A) \ map(h, xs@ys) = map(h,xs) @ map(h,ys)" apply (induct_tac "xs") apply (simp_all (no_asm_simp)) done lemma map_flat: "ls: list(list(A)) \ map(h, flat(ls)) = flat(map(map(h),ls))" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: map_app_distrib) done lemma list_rec_map: "l \ list(A) \ list_rec(c, d, map(h,l)) = list_rec(c, \x xs r. d(h(x), map(h,xs), r), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done (** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **) (* @{term"c \ list(Collect(B,P)) \ c \ list"} *) lemmas list_CollectD = Collect_subset [THEN list_mono, THEN subsetD] lemma map_list_Collect: "l \ list({x \ A. h(x)=j(x)}) \ map(h,l) = map(j,l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done (*** theorems about length ***) lemma length_map [simp]: "xs: list(A) \ length(map(h,xs)) = length(xs)" by (induct_tac "xs", simp_all) lemma length_app [simp]: "\xs: list(A); ys: list(A)\ \ length(xs@ys) = length(xs) #+ length(ys)" by (induct_tac "xs", simp_all) lemma length_rev [simp]: "xs: list(A) \ length(rev(xs)) = length(xs)" apply (induct_tac "xs") apply (simp_all (no_asm_simp) add: length_app) done lemma length_flat: "ls: list(list(A)) \ length(flat(ls)) = list_add(map(length,ls))" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: length_app) done (** Length and drop **) (*Lemma for the inductive step of drop_length*) lemma drop_length_Cons [rule_format]: "xs: list(A) \ \x. \z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)" by (erule list.induct, simp_all) lemma drop_length [rule_format]: "l \ list(A) \ \i \ length(l). (\z zs. drop(i,l) = Cons(z,zs))" apply (erule list.induct, simp_all, safe) apply (erule drop_length_Cons) apply (rule natE) apply (erule Ord_trans [OF asm_rl length_type Ord_nat], assumption, simp_all) apply (blast intro: succ_in_naturalD length_type) done (*** theorems about app ***) lemma app_right_Nil [simp]: "xs: list(A) \ xs@Nil=xs" by (erule list.induct, simp_all) lemma app_assoc: "xs: list(A) \ (xs@ys)@zs = xs@(ys@zs)" by (induct_tac "xs", simp_all) lemma flat_app_distrib: "ls: list(list(A)) \ flat(ls@ms) = flat(ls)@flat(ms)" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: app_assoc) done (*** theorems about rev ***) lemma rev_map_distrib: "l \ list(A) \ rev(map(h,l)) = map(h,rev(l))" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: map_app_distrib) done (*Simplifier needs the premises as assumptions because rewriting will not instantiate the variable ?A in the rules' typing conditions; note that rev_type does not instantiate ?A. Only the premises do. *) lemma rev_app_distrib: "\xs: list(A); ys: list(A)\ \ rev(xs@ys) = rev(ys)@rev(xs)" apply (erule list.induct) apply (simp_all add: app_assoc) done lemma rev_rev_ident [simp]: "l \ list(A) \ rev(rev(l))=l" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: rev_app_distrib) done lemma rev_flat: "ls: list(list(A)) \ rev(flat(ls)) = flat(map(rev,rev(ls)))" apply (induct_tac "ls") apply (simp_all add: map_app_distrib flat_app_distrib rev_app_distrib) done (*** theorems about list_add ***) lemma list_add_app: "\xs: list(nat); ys: list(nat)\ \ list_add(xs@ys) = list_add(ys) #+ list_add(xs)" apply (induct_tac "xs", simp_all) done lemma list_add_rev: "l \ list(nat) \ list_add(rev(l)) = list_add(l)" apply (induct_tac "l") apply (simp_all (no_asm_simp) add: list_add_app) done lemma list_add_flat: "ls: list(list(nat)) \ list_add(flat(ls)) = list_add(map(list_add,ls))" apply (induct_tac "ls") apply (simp_all (no_asm_simp) add: list_add_app) done (** New induction rules **) lemma list_append_induct [case_names Nil snoc, consumes 1]: "\l \ list(A); P(Nil); \x y. \x \ A; y \ list(A); P(y)\ \ P(y @ [x]) \ \ P(l)" apply (subgoal_tac "P(rev(rev(l)))", simp) apply (erule rev_type [THEN list.induct], simp_all) done lemma list_complete_induct_lemma [rule_format]: assumes ih: "\l. \l \ list(A); \l' \ list(A). length(l') < length(l) \ P(l')\ \ P(l)" shows "n \ nat \ \l \ list(A). length(l) < n \ P(l)" apply (induct_tac n, simp) apply (blast intro: ih elim!: leE) done theorem list_complete_induct: "\l \ list(A); \l. \l \ list(A); \l' \ list(A). length(l') < length(l) \ P(l')\ \ P(l) \ \ P(l)" apply (rule list_complete_induct_lemma [of A]) prefer 4 apply (rule le_refl, simp) apply blast apply simp apply assumption done (*** Thanks to Sidi Ehmety for these results about min, take, etc. ***) (** min FIXME: replace by Int! **) (* Min theorems are also true for i, j ordinals *) lemma min_sym: "\i \ nat; j \ nat\ \ min(i,j)=min(j,i)" unfolding min_def apply (auto dest!: not_lt_imp_le dest: lt_not_sym intro: le_anti_sym) done lemma min_type [simp,TC]: "\i \ nat; j \ nat\ \ min(i,j):nat" by (unfold min_def, auto) lemma min_0 [simp]: "i \ nat \ min(0,i) = 0" unfolding min_def apply (auto dest: not_lt_imp_le) done lemma min_02 [simp]: "i \ nat \ min(i, 0) = 0" unfolding min_def apply (auto dest: not_lt_imp_le) done lemma lt_min_iff: "\i \ nat; j \ nat; k \ nat\ \ i i ii \ nat; j \ nat\ \ min(succ(i), succ(j))= succ(min(i, j))" apply (unfold min_def, auto) done (*** more theorems about lists ***) (** filter **) lemma filter_append [simp]: "xs:list(A) \ filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)" by (induct_tac "xs", auto) lemma filter_type [simp,TC]: "xs:list(A) \ filter(P, xs):list(A)" by (induct_tac "xs", auto) lemma length_filter: "xs:list(A) \ length(filter(P, xs)) \ length(xs)" apply (induct_tac "xs", auto) apply (rule_tac j = "length (l) " in le_trans) apply (auto simp add: le_iff) done lemma filter_is_subset: "xs:list(A) \ set_of_list(filter(P,xs)) \ set_of_list(xs)" by (induct_tac "xs", auto) lemma filter_False [simp]: "xs:list(A) \ filter(\p. False, xs) = Nil" by (induct_tac "xs", auto) lemma filter_True [simp]: "xs:list(A) \ filter(\p. True, xs) = xs" by (induct_tac "xs", auto) (** length **) lemma length_is_0_iff [simp]: "xs:list(A) \ length(xs)=0 \ xs=Nil" by (erule list.induct, auto) lemma length_is_0_iff2 [simp]: "xs:list(A) \ 0 = length(xs) \ xs=Nil" by (erule list.induct, auto) lemma length_tl [simp]: "xs:list(A) \ length(tl(xs)) = length(xs) #- 1" by (erule list.induct, auto) lemma length_greater_0_iff: "xs:list(A) \ 0 xs \ Nil" by (erule list.induct, auto) lemma length_succ_iff: "xs:list(A) \ length(xs)=succ(n) \ (\y ys. xs=Cons(y, ys) \ length(ys)=n)" by (erule list.induct, auto) (** more theorems about append **) lemma append_is_Nil_iff [simp]: "xs:list(A) \ (xs@ys = Nil) \ (xs=Nil \ ys = Nil)" by (erule list.induct, auto) lemma append_is_Nil_iff2 [simp]: "xs:list(A) \ (Nil = xs@ys) \ (xs=Nil \ ys = Nil)" by (erule list.induct, auto) lemma append_left_is_self_iff [simp]: "xs:list(A) \ (xs@ys = xs) \ (ys = Nil)" by (erule list.induct, auto) lemma append_left_is_self_iff2 [simp]: "xs:list(A) \ (xs = xs@ys) \ (ys = Nil)" by (erule list.induct, auto) (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff [rule_format]: "\xs:list(A); ys:list(A); zs:list(A)\ \ length(ys)=length(zs) \ (xs@ys=zs \ (xs=Nil \ ys=zs))" apply (erule list.induct) apply (auto simp add: length_app) done (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff2 [rule_format]: "\xs:list(A); ys:list(A); zs:list(A)\ \ length(ys)=length(zs) \ (zs=ys@xs \ (xs=Nil \ ys=zs))" apply (erule list.induct) apply (auto simp add: length_app) done lemma append_eq_append_iff [rule_format]: "xs:list(A) \ \ys \ list(A). length(xs)=length(ys) \ (xs@us = ys@vs) \ (xs=ys \ us=vs)" apply (erule list.induct) apply (simp (no_asm_simp)) apply clarify apply (erule_tac a = ys in list.cases, auto) done declare append_eq_append_iff [simp] lemma append_eq_append [rule_format]: "xs:list(A) \ \ys \ list(A). \us \ list(A). \vs \ list(A). length(us) = length(vs) \ (xs@us = ys@vs) \ (xs=ys \ us=vs)" apply (induct_tac "xs") apply (force simp add: length_app, clarify) apply (erule_tac a = ys in list.cases, simp) apply (subgoal_tac "Cons (a, l) @ us =vs") apply (drule rev_iffD1 [OF _ append_left_is_Nil_iff], simp_all, blast) done lemma append_eq_append_iff2 [simp]: "\xs:list(A); ys:list(A); us:list(A); vs:list(A); length(us)=length(vs)\ \ xs@us = ys@vs \ (xs=ys \ us=vs)" apply (rule iffI) apply (rule append_eq_append, auto) done lemma append_self_iff [simp]: "\xs:list(A); ys:list(A); zs:list(A)\ \ xs@ys=xs@zs \ ys=zs" by simp lemma append_self_iff2 [simp]: "\xs:list(A); ys:list(A); zs:list(A)\ \ ys@xs=zs@xs \ ys=zs" by simp (* Can also be proved from append_eq_append_iff2, but the proof requires two more hypotheses: x \ A and y \ A *) lemma append1_eq_iff [rule_format]: "xs:list(A) \ \ys \ list(A). xs@[x] = ys@[y] \ (xs = ys \ x=y)" apply (erule list.induct) apply clarify apply (erule list.cases) apply simp_all txt\Inductive step\ apply clarify apply (erule_tac a=ys in list.cases, simp_all) done declare append1_eq_iff [simp] lemma append_right_is_self_iff [simp]: "\xs:list(A); ys:list(A)\ \ (xs@ys = ys) \ (xs=Nil)" by (simp (no_asm_simp) add: append_left_is_Nil_iff) lemma append_right_is_self_iff2 [simp]: "\xs:list(A); ys:list(A)\ \ (ys = xs@ys) \ (xs=Nil)" apply (rule iffI) apply (drule sym, auto) done lemma hd_append [rule_format]: "xs:list(A) \ xs \ Nil \ hd(xs @ ys) = hd(xs)" by (induct_tac "xs", auto) declare hd_append [simp] lemma tl_append [rule_format]: "xs:list(A) \ xs\Nil \ tl(xs @ ys) = tl(xs)@ys" by (induct_tac "xs", auto) declare tl_append [simp] (** rev **) lemma rev_is_Nil_iff [simp]: "xs:list(A) \ (rev(xs) = Nil \ xs = Nil)" by (erule list.induct, auto) lemma Nil_is_rev_iff [simp]: "xs:list(A) \ (Nil = rev(xs) \ xs = Nil)" by (erule list.induct, auto) lemma rev_is_rev_iff [rule_format]: "xs:list(A) \ \ys \ list(A). rev(xs)=rev(ys) \ xs=ys" apply (erule list.induct, force, clarify) apply (erule_tac a = ys in list.cases, auto) done declare rev_is_rev_iff [simp] lemma rev_list_elim [rule_format]: "xs:list(A) \ (xs=Nil \ P) \ (\ys \ list(A). \y \ A. xs =ys@[y] \P)\P" by (erule list_append_induct, auto) (** more theorems about drop **) lemma length_drop [rule_format]: "n \ nat \ \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" apply (erule nat_induct) apply (auto elim: list.cases) done declare length_drop [simp] lemma drop_all [rule_format]: "n \ nat \ \xs \ list(A). length(xs) \ n \ drop(n, xs)=Nil" apply (erule nat_induct) apply (auto elim: list.cases) done declare drop_all [simp] lemma drop_append [rule_format]: "n \ nat \ \xs \ list(A). drop(n, xs@ys) = drop(n,xs) @ drop(n #- length(xs), ys)" apply (induct_tac "n") apply (auto elim: list.cases) done lemma drop_drop: "m \ nat \ \xs \ list(A). \n \ nat. drop(n, drop(m, xs))=drop(n #+ m, xs)" apply (induct_tac "m") apply (auto elim: list.cases) done (** take **) lemma take_0 [simp]: "xs:list(A) \ take(0, xs) = Nil" unfolding take_def apply (erule list.induct, auto) done lemma take_succ_Cons [simp]: "n \ nat \ take(succ(n), Cons(a, xs)) = Cons(a, take(n, xs))" by (simp add: take_def) (* Needed for proving take_all *) lemma take_Nil [simp]: "n \ nat \ take(n, Nil) = Nil" by (unfold take_def, auto) lemma take_all [rule_format]: "n \ nat \ \xs \ list(A). length(xs) \ n \ take(n, xs) = xs" apply (erule nat_induct) apply (auto elim: list.cases) done declare take_all [simp] lemma take_type [rule_format]: "xs:list(A) \ \n \ nat. take(n, xs):list(A)" apply (erule list.induct, simp, clarify) apply (erule natE, auto) done declare take_type [simp,TC] lemma take_append [rule_format]: "xs:list(A) \ \ys \ list(A). \n \ nat. take(n, xs @ ys) = take(n, xs) @ take(n #- length(xs), ys)" apply (erule list.induct, simp, clarify) apply (erule natE, auto) done declare take_append [simp] lemma take_take [rule_format]: "m \ nat \ \xs \ list(A). \n \ nat. take(n, take(m,xs))= take(min(n, m), xs)" apply (induct_tac "m", auto) apply (erule_tac a = xs in list.cases) apply (auto simp add: take_Nil) apply (erule_tac n=n in natE) apply (auto intro: take_0 take_type) done (** nth **) lemma nth_0 [simp]: "nth(0, Cons(a, l)) = a" by (simp add: nth_def) lemma nth_Cons [simp]: "n \ nat \ nth(succ(n), Cons(a,l)) = nth(n,l)" by (simp add: nth_def) lemma nth_empty [simp]: "nth(n, Nil) = 0" by (simp add: nth_def) lemma nth_type [rule_format]: "xs:list(A) \ \n. n < length(xs) \ nth(n,xs) \ A" apply (erule list.induct, simp, clarify) apply (subgoal_tac "n \ nat") apply (erule natE, auto dest!: le_in_nat) done declare nth_type [simp,TC] lemma nth_eq_0 [rule_format]: "xs:list(A) \ \n \ nat. length(xs) \ n \ nth(n,xs) = 0" apply (erule list.induct, simp, clarify) apply (erule natE, auto) done lemma nth_append [rule_format]: "xs:list(A) \ \n \ nat. nth(n, xs @ ys) = (if n < length(xs) then nth(n,xs) else nth(n #- length(xs), ys))" apply (induct_tac "xs", simp, clarify) apply (erule natE, auto) done lemma set_of_list_conv_nth: "xs:list(A) \ set_of_list(xs) = {x \ A. \i\nat. i x = nth(i,xs)}" apply (induct_tac "xs", simp_all) apply (rule equalityI, auto) apply (rule_tac x = 0 in bexI, auto) apply (erule natE, auto) done (* Other theorems about lists *) lemma nth_take_lemma [rule_format]: "k \ nat \ \xs \ list(A). (\ys \ list(A). k \ length(xs) \ k \ length(ys) \ (\i \ nat. i nth(i,xs) = nth(i,ys))\ take(k,xs) = take(k,ys))" apply (induct_tac "k") apply (simp_all (no_asm_simp) add: lt_succ_eq_0_disj all_conj_distrib) apply clarify (*Both lists are non-empty*) apply (erule_tac a=xs in list.cases, simp) apply (erule_tac a=ys in list.cases, clarify) apply (simp (no_asm_use) ) apply clarify apply (simp (no_asm_simp)) apply (rule conjI, force) apply (rename_tac y ys z zs) apply (drule_tac x = zs and x1 = ys in bspec [THEN bspec], auto) done lemma nth_equalityI [rule_format]: "\xs:list(A); ys:list(A); length(xs) = length(ys); \i \ nat. i < length(xs) \ nth(i,xs) = nth(i,ys)\ \ xs = ys" apply (subgoal_tac "length (xs) \ length (ys) ") apply (cut_tac k="length(xs)" and xs=xs and ys=ys in nth_take_lemma) apply (simp_all add: take_all) done (*The famous take-lemma*) lemma take_equalityI [rule_format]: "\xs:list(A); ys:list(A); (\i \ nat. take(i, xs) = take(i,ys))\ \ xs = ys" apply (case_tac "length (xs) \ length (ys) ") apply (drule_tac x = "length (ys) " in bspec) apply (drule_tac [3] not_lt_imp_le) apply (subgoal_tac [5] "length (ys) \ length (xs) ") apply (rule_tac [6] j = "succ (length (ys))" in le_trans) apply (rule_tac [6] leI) apply (drule_tac [5] x = "length (xs) " in bspec) apply (simp_all add: take_all) done lemma nth_drop [rule_format]: "n \ nat \ \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" apply (induct_tac "n", simp_all, clarify) apply (erule list.cases, auto) done lemma take_succ [rule_format]: "xs\list(A) \ \i. i < length(xs) \ take(succ(i), xs) = take(i,xs) @ [nth(i, xs)]" apply (induct_tac "xs", auto) apply (subgoal_tac "i\nat") apply (erule natE) apply (auto simp add: le_in_nat) done lemma take_add [rule_format]: "\xs\list(A); j\nat\ \ \i\nat. take(i #+ j, xs) = take(i,xs) @ take(j, drop(i,xs))" apply (induct_tac "xs", simp_all, clarify) apply (erule_tac n = i in natE, simp_all) done lemma length_take: "l\list(A) \ \n\nat. length(take(n,l)) = min(n, length(l))" apply (induct_tac "l", safe, simp_all) apply (erule natE, simp_all) done subsection\The function zip\ text\Crafty definition to eliminate a type argument\ consts zip_aux :: "[i,i]\i" primrec (*explicit lambda is required because both arguments of "un" vary*) "zip_aux(B,[]) = (\ys \ list(B). list_case([], \y l. [], ys))" "zip_aux(B,Cons(x,l)) = (\ys \ list(B). list_case(Nil, \y zs. Cons(\x,y\, zip_aux(B,l)`zs), ys))" definition zip :: "[i, i]\i" where "zip(xs, ys) \ zip_aux(set_of_list(ys),xs)`ys" (* zip equations *) lemma list_on_set_of_list: "xs \ list(A) \ xs \ list(set_of_list(xs))" apply (induct_tac xs, simp_all) apply (blast intro: list_mono [THEN subsetD]) done lemma zip_Nil [simp]: "ys:list(A) \ zip(Nil, ys)=Nil" apply (simp add: zip_def list_on_set_of_list [of _ A]) apply (erule list.cases, simp_all) done lemma zip_Nil2 [simp]: "xs:list(A) \ zip(xs, Nil)=Nil" apply (simp add: zip_def list_on_set_of_list [of _ A]) apply (erule list.cases, simp_all) done lemma zip_aux_unique [rule_format]: "\B<=C; xs \ list(A)\ \ \ys \ list(B). zip_aux(C,xs) ` ys = zip_aux(B,xs) ` ys" apply (induct_tac xs) apply simp_all apply (blast intro: list_mono [THEN subsetD], clarify) apply (erule_tac a=ys in list.cases, auto) apply (blast intro: list_mono [THEN subsetD]) done lemma zip_Cons_Cons [simp]: "\xs:list(A); ys:list(B); x \ A; y \ B\ \ zip(Cons(x,xs), Cons(y, ys)) = Cons(\x,y\, zip(xs, ys))" apply (simp add: zip_def, auto) apply (rule zip_aux_unique, auto) apply (simp add: list_on_set_of_list [of _ B]) apply (blast intro: list_on_set_of_list list_mono [THEN subsetD]) done lemma zip_type [rule_format]: "xs:list(A) \ \ys \ list(B). zip(xs, ys):list(A*B)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule_tac a = ys in list.cases, auto) done declare zip_type [simp,TC] (* zip length *) lemma length_zip [rule_format]: "xs:list(A) \ \ys \ list(B). length(zip(xs,ys)) = min(length(xs), length(ys))" unfolding min_def apply (induct_tac "xs", simp_all, clarify) apply (erule_tac a = ys in list.cases, auto) done declare length_zip [simp] lemma zip_append1 [rule_format]: "\ys:list(A); zs:list(B)\ \ \xs \ list(A). zip(xs @ ys, zs) = zip(xs, take(length(xs), zs)) @ zip(ys, drop(length(xs),zs))" apply (induct_tac "zs", force, clarify) apply (erule_tac a = xs in list.cases, simp_all) done lemma zip_append2 [rule_format]: "\xs:list(A); zs:list(B)\ \ \ys \ list(B). zip(xs, ys@zs) = zip(take(length(ys), xs), ys) @ zip(drop(length(ys), xs), zs)" apply (induct_tac "xs", force, clarify) apply (erule_tac a = ys in list.cases, auto) done lemma zip_append [simp]: "\length(xs) = length(us); length(ys) = length(vs); xs:list(A); us:list(B); ys:list(A); vs:list(B)\ \ zip(xs@ys,us@vs) = zip(xs, us) @ zip(ys, vs)" by (simp (no_asm_simp) add: zip_append1 drop_append diff_self_eq_0) lemma zip_rev [rule_format]: "ys:list(B) \ \xs \ list(A). length(xs) = length(ys) \ zip(rev(xs), rev(ys)) = rev(zip(xs, ys))" apply (induct_tac "ys", force, clarify) apply (erule_tac a = xs in list.cases) apply (auto simp add: length_rev) done declare zip_rev [simp] lemma nth_zip [rule_format]: "ys:list(B) \ \i \ nat. \xs \ list(A). i < length(xs) \ i < length(ys) \ nth(i,zip(xs, ys)) = " apply (induct_tac "ys", force, clarify) apply (erule_tac a = xs in list.cases, simp) apply (auto elim: natE) done declare nth_zip [simp] lemma set_of_list_zip [rule_format]: "\xs:list(A); ys:list(B); i \ nat\ \ set_of_list(zip(xs, ys)) = {\x, y\:A*B. \i\nat. i < min(length(xs), length(ys)) \ x = nth(i, xs) \ y = nth(i, ys)}" by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth) (** list_update **) lemma list_update_Nil [simp]: "i \ nat \list_update(Nil, i, v) = Nil" by (unfold list_update_def, auto) lemma list_update_Cons_0 [simp]: "list_update(Cons(x, xs), 0, v)= Cons(v, xs)" by (unfold list_update_def, auto) lemma list_update_Cons_succ [simp]: "n \ nat \ list_update(Cons(x, xs), succ(n), v)= Cons(x, list_update(xs, n, v))" apply (unfold list_update_def, auto) done lemma list_update_type [rule_format]: "\xs:list(A); v \ A\ \ \n \ nat. list_update(xs, n, v):list(A)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done declare list_update_type [simp,TC] lemma length_list_update [rule_format]: "xs:list(A) \ \i \ nat. length(list_update(xs, i, v))=length(xs)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done declare length_list_update [simp] lemma nth_list_update [rule_format]: "\xs:list(A)\ \ \i \ nat. \j \ nat. i < length(xs) \ nth(j, list_update(xs, i, x)) = (if i=j then x else nth(j, xs))" apply (induct_tac "xs") apply simp_all apply clarify apply (rename_tac i j) apply (erule_tac n=i in natE) apply (erule_tac [2] n=j in natE) apply (erule_tac n=j in natE, simp_all, force) done lemma nth_list_update_eq [simp]: "\i < length(xs); xs:list(A)\ \ nth(i, list_update(xs, i,x)) = x" by (simp (no_asm_simp) add: lt_length_in_nat nth_list_update) lemma nth_list_update_neq [rule_format]: "xs:list(A) \ \i \ nat. \j \ nat. i \ j \ nth(j, list_update(xs,i,x)) = nth(j,xs)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE) apply (erule_tac [2] natE, simp_all) apply (erule natE, simp_all) done declare nth_list_update_neq [simp] lemma list_update_overwrite [rule_format]: "xs:list(A) \ \i \ nat. i < length(xs) \ list_update(list_update(xs, i, x), i, y) = list_update(xs, i,y)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done declare list_update_overwrite [simp] lemma list_update_same_conv [rule_format]: "xs:list(A) \ \i \ nat. i < length(xs) \ (list_update(xs, i, x) = xs) \ (nth(i, xs) = x)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify apply (erule natE, auto) done lemma update_zip [rule_format]: "ys:list(B) \ \i \ nat. \xy \ A*B. \xs \ list(A). length(xs) = length(ys) \ list_update(zip(xs, ys), i, xy) = zip(list_update(xs, i, fst(xy)), list_update(ys, i, snd(xy)))" apply (induct_tac "ys") apply auto apply (erule_tac a = xs in list.cases) apply (auto elim: natE) done lemma set_update_subset_cons [rule_format]: "xs:list(A) \ \i \ nat. set_of_list(list_update(xs, i, x)) \ cons(x, set_of_list(xs))" apply (induct_tac "xs") apply simp apply (rule ballI) apply (erule natE, simp_all, auto) done lemma set_of_list_update_subsetI: "\set_of_list(xs) \ A; xs:list(A); x \ A; i \ nat\ \ set_of_list(list_update(xs, i,x)) \ A" apply (rule subset_trans) apply (rule set_update_subset_cons, auto) done (** upt **) lemma upt_rec: "j \ nat \ upt(i,j) = (if ij \ i; j \ nat\ \ upt(i,j) = Nil" apply (subst upt_rec, auto) apply (auto simp add: le_iff) apply (drule lt_asym [THEN notE], auto) done (*Only needed if upt_Suc is deleted from the simpset*) lemma upt_succ_append: "\i \ j; j \ nat\ \ upt(i,succ(j)) = upt(i, j)@[j]" by simp lemma upt_conv_Cons: "\i nat\ \ upt(i,j) = Cons(i,upt(succ(i),j))" apply (rule trans) apply (rule upt_rec, auto) done lemma upt_type [simp,TC]: "j \ nat \ upt(i,j):list(nat)" by (induct_tac "j", auto) (*LOOPS as a simprule, since j<=j*) lemma upt_add_eq_append: "\i \ j; j \ nat; k \ nat\ \ upt(i, j #+k) = upt(i,j)@upt(j,j#+k)" apply (induct_tac "k") apply (auto simp add: app_assoc app_type) apply (rule_tac j = j in le_trans, auto) done lemma length_upt [simp]: "\i \ nat; j \ nat\ \length(upt(i,j)) = j #- i" apply (induct_tac "j") apply (rule_tac [2] sym) apply (auto dest!: not_lt_imp_le simp add: diff_succ diff_is_0_iff) done lemma nth_upt [simp]: "\i \ nat; j \ nat; k \ nat; i #+ k < j\ \ nth(k, upt(i,j)) = i #+ k" apply (rotate_tac -1, erule rev_mp) apply (induct_tac "j", simp) apply (auto dest!: not_lt_imp_le simp add: nth_append le_iff less_diff_conv add_commute) done lemma take_upt [rule_format]: "\m \ nat; n \ nat\ \ \i \ nat. i #+ m \ n \ take(m, upt(i,n)) = upt(i,i#+m)" apply (induct_tac "m") apply (simp (no_asm_simp) add: take_0) apply clarify apply (subst upt_rec, simp) apply (rule sym) apply (subst upt_rec, simp) apply (simp_all del: upt.simps) apply (rule_tac j = "succ (i #+ x) " in lt_trans2) apply auto done declare take_upt [simp] lemma map_succ_upt: "\m \ nat; n \ nat\ \ map(succ, upt(m,n))= upt(succ(m), succ(n))" apply (induct_tac "n") apply (auto simp add: map_app_distrib) done lemma nth_map [rule_format]: "xs:list(A) \ \n \ nat. n < length(xs) \ nth(n, map(f, xs)) = f(nth(n, xs))" apply (induct_tac "xs", simp) apply (rule ballI) apply (induct_tac "n", auto) done declare nth_map [simp] lemma nth_map_upt [rule_format]: "\m \ nat; n \ nat\ \ \i \ nat. i < n #- m \ nth(i, map(f, upt(m,n))) = f(m #+ i)" apply (rule_tac n = m and m = n in diff_induct, typecheck, simp, simp) apply (subst map_succ_upt [symmetric], simp_all, clarify) apply (subgoal_tac "i < length (upt (0, x))") prefer 2 apply (simp add: less_diff_conv) apply (rule_tac j = "succ (i #+ y) " in lt_trans2) apply simp apply simp apply (subgoal_tac "i < length (upt (y, x))") apply (simp_all add: add_commute less_diff_conv) done (** sublist (a generalization of nth to sets) **) definition sublist :: "[i, i] \ i" where "sublist(xs, A)\ map(fst, (filter(\p. snd(p): A, zip(xs, upt(0,length(xs))))))" lemma sublist_0 [simp]: "xs:list(A) \sublist(xs, 0) =Nil" by (unfold sublist_def, auto) lemma sublist_Nil [simp]: "sublist(Nil, A) = Nil" by (unfold sublist_def, auto) lemma sublist_shift_lemma: "\xs:list(B); i \ nat\ \ map(fst, filter(\p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) = map(fst, filter(\p. snd(p):nat \ snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" apply (erule list_append_induct) apply (simp (no_asm_simp)) apply (auto simp add: add_commute length_app filter_append map_app_distrib) done lemma sublist_type [simp,TC]: "xs:list(B) \ sublist(xs, A):list(B)" unfolding sublist_def apply (induct_tac "xs") apply (auto simp add: filter_append map_app_distrib) done lemma upt_add_eq_append2: "\i \ nat; j \ nat\ \ upt(0, i #+ j) = upt(0, i) @ upt(i, i #+ j)" by (simp add: upt_add_eq_append [of 0] nat_0_le) lemma sublist_append: "\xs:list(B); ys:list(B)\ \ sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \ nat. j #+ length(xs): A})" unfolding sublist_def apply (erule_tac l = ys in list_append_induct, simp) apply (simp (no_asm_simp) add: upt_add_eq_append2 app_assoc [symmetric]) apply (auto simp add: sublist_shift_lemma length_type map_app_distrib app_assoc) apply (simp_all add: add_commute) done lemma sublist_Cons: "\xs:list(B); x \ B\ \ sublist(Cons(x, xs), A) = (if 0 \ A then [x] else []) @ sublist(xs, {j \ nat. succ(j) \ A})" apply (erule_tac l = xs in list_append_induct) apply (simp (no_asm_simp) add: sublist_def) apply (simp del: app_Cons add: app_Cons [symmetric] sublist_append, simp) done lemma sublist_singleton [simp]: "sublist([x], A) = (if 0 \ A then [x] else [])" by (simp add: sublist_Cons) lemma sublist_upt_eq_take [rule_format]: "xs:list(A) \ \n\nat. sublist(xs,n) = take(n,xs)" apply (erule list.induct, simp) apply (clarify ) apply (erule natE) apply (simp_all add: nat_eq_Collect_lt Ord_mem_iff_lt sublist_Cons) done declare sublist_upt_eq_take [simp] lemma sublist_Int_eq: "xs \ list(B) \ sublist(xs, A \ nat) = sublist(xs, A)" apply (erule list.induct) apply (simp_all add: sublist_Cons) done text\Repetition of a List Element\ consts repeat :: "[i,i]\i" primrec "repeat(a,0) = []" "repeat(a,succ(n)) = Cons(a,repeat(a,n))" lemma length_repeat: "n \ nat \ length(repeat(a,n)) = n" by (induct_tac n, auto) lemma repeat_succ_app: "n \ nat \ repeat(a,succ(n)) = repeat(a,n) @ [a]" apply (induct_tac n) apply (simp_all del: app_Cons add: app_Cons [symmetric]) done lemma repeat_type [TC]: "\a \ A; n \ nat\ \ repeat(a,n) \ list(A)" by (induct_tac n, auto) end diff --git a/src/ZF/Order.thy b/src/ZF/Order.thy --- a/src/ZF/Order.thy +++ b/src/ZF/Order.thy @@ -1,714 +1,714 @@ (* Title: ZF/Order.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Results from the book "Set Theory: an Introduction to Independence Proofs" by Kenneth Kunen. Chapter 1, section 6. Additional definitions and lemmas for reflexive orders. *) section\Partial and Total Orderings: Basic Definitions and Properties\ theory Order imports WF Perm begin text \We adopt the following convention: \ord\ is used for strict orders and \order\ is used for their reflexive counterparts.\ definition part_ord :: "[i,i]\o" (*Strict partial ordering*) where "part_ord(A,r) \ irrefl(A,r) \ trans[A](r)" definition linear :: "[i,i]\o" (*Strict total ordering*) where "linear(A,r) \ (\x\A. \y\A. \x,y\:r | x=y | \y,x\:r)" definition tot_ord :: "[i,i]\o" (*Strict total ordering*) where "tot_ord(A,r) \ part_ord(A,r) \ linear(A,r)" definition "preorder_on(A, r) \ refl(A, r) \ trans[A](r)" definition (*Partial ordering*) "partial_order_on(A, r) \ preorder_on(A, r) \ antisym(r)" abbreviation "Preorder(r) \ preorder_on(field(r), r)" abbreviation "Partial_order(r) \ partial_order_on(field(r), r)" definition well_ord :: "[i,i]\o" (*Well-ordering*) where "well_ord(A,r) \ tot_ord(A,r) \ wf[A](r)" definition mono_map :: "[i,i,i,i]\i" (*Order-preserving maps*) where "mono_map(A,r,B,s) \ {f \ A->B. \x\A. \y\A. \x,y\:r \ :s}" definition ord_iso :: "[i,i,i,i]\i" (\(\_, _\ \/ \_, _\)\ 51) (*Order isomorphisms*) where "\A,r\ \ \B,s\ \ {f \ bij(A,B). \x\A. \y\A. \x,y\:r \ :s}" definition pred :: "[i,i,i]\i" (*Set of predecessors*) where "pred(A,x,r) \ {y \ A. \y,x\:r}" definition ord_iso_map :: "[i,i,i,i]\i" (*Construction for linearity theorem*) where "ord_iso_map(A,r,B,s) \ \x\A. \y\B. \f \ ord_iso(pred(A,x,r), r, pred(B,y,s), s). {\x,y\}" definition first :: "[i, i, i] \ o" where "first(u, X, R) \ u \ X \ (\v\X. v\u \ \u,v\ \ R)" subsection\Immediate Consequences of the Definitions\ lemma part_ord_Imp_asym: "part_ord(A,r) \ asym(r \ A*A)" by (unfold part_ord_def irrefl_def trans_on_def asym_def, blast) lemma linearE: "\linear(A,r); x \ A; y \ A; \x,y\:r \ P; x=y \ P; \y,x\:r \ P\ \ P" by (simp add: linear_def, blast) (** General properties of well_ord **) lemma well_ordI: "\wf[A](r); linear(A,r)\ \ well_ord(A,r)" apply (simp add: irrefl_def part_ord_def tot_ord_def trans_on_def well_ord_def wf_on_not_refl) apply (fast elim: linearE wf_on_asym wf_on_chain3) done lemma well_ord_is_wf: "well_ord(A,r) \ wf[A](r)" by (unfold well_ord_def, safe) lemma well_ord_is_trans_on: "well_ord(A,r) \ trans[A](r)" by (unfold well_ord_def tot_ord_def part_ord_def, safe) lemma well_ord_is_linear: "well_ord(A,r) \ linear(A,r)" by (unfold well_ord_def tot_ord_def, blast) (** Derived rules for pred(A,x,r) **) lemma pred_iff: "y \ pred(A,x,r) \ \y,x\:r \ y \ A" by (unfold pred_def, blast) lemmas predI = conjI [THEN pred_iff [THEN iffD2]] lemma predE: "\y \ pred(A,x,r); \y \ A; \y,x\:r\ \ P\ \ P" by (simp add: pred_def) lemma pred_subset_under: "pred(A,x,r) \ r -`` {x}" by (simp add: pred_def, blast) lemma pred_subset: "pred(A,x,r) \ A" by (simp add: pred_def, blast) lemma pred_pred_eq: "pred(pred(A,x,r), y, r) = pred(A,x,r) \ pred(A,y,r)" by (simp add: pred_def, blast) lemma trans_pred_pred_eq: "\trans[A](r); \y,x\:r; x \ A; y \ A\ \ pred(pred(A,x,r), y, r) = pred(A,y,r)" by (unfold trans_on_def pred_def, blast) subsection\Restricting an Ordering's Domain\ (** The ordering's properties hold over all subsets of its domain [including initial segments of the form pred(A,x,r) **) (*Note: a relation s such that s<=r need not be a partial ordering*) lemma part_ord_subset: "\part_ord(A,r); B<=A\ \ part_ord(B,r)" by (unfold part_ord_def irrefl_def trans_on_def, blast) lemma linear_subset: "\linear(A,r); B<=A\ \ linear(B,r)" by (unfold linear_def, blast) lemma tot_ord_subset: "\tot_ord(A,r); B<=A\ \ tot_ord(B,r)" unfolding tot_ord_def apply (fast elim!: part_ord_subset linear_subset) done lemma well_ord_subset: "\well_ord(A,r); B<=A\ \ well_ord(B,r)" unfolding well_ord_def apply (fast elim!: tot_ord_subset wf_on_subset_A) done (** Relations restricted to a smaller domain, by Krzysztof Grabczewski **) lemma irrefl_Int_iff: "irrefl(A,r \ A*A) \ irrefl(A,r)" by (unfold irrefl_def, blast) lemma trans_on_Int_iff: "trans[A](r \ A*A) \ trans[A](r)" by (unfold trans_on_def, blast) lemma part_ord_Int_iff: "part_ord(A,r \ A*A) \ part_ord(A,r)" unfolding part_ord_def apply (simp add: irrefl_Int_iff trans_on_Int_iff) done lemma linear_Int_iff: "linear(A,r \ A*A) \ linear(A,r)" by (unfold linear_def, blast) lemma tot_ord_Int_iff: "tot_ord(A,r \ A*A) \ tot_ord(A,r)" unfolding tot_ord_def apply (simp add: part_ord_Int_iff linear_Int_iff) done lemma wf_on_Int_iff: "wf[A](r \ A*A) \ wf[A](r)" apply (unfold wf_on_def wf_def, fast) (*10 times faster than blast!*) done lemma well_ord_Int_iff: "well_ord(A,r \ A*A) \ well_ord(A,r)" unfolding well_ord_def apply (simp add: tot_ord_Int_iff wf_on_Int_iff) done subsection\Empty and Unit Domains\ (*The empty relation is well-founded*) lemma wf_on_any_0: "wf[A](0)" by (simp add: wf_on_def wf_def, fast) subsubsection\Relations over the Empty Set\ lemma irrefl_0: "irrefl(0,r)" by (unfold irrefl_def, blast) lemma trans_on_0: "trans[0](r)" by (unfold trans_on_def, blast) lemma part_ord_0: "part_ord(0,r)" unfolding part_ord_def apply (simp add: irrefl_0 trans_on_0) done lemma linear_0: "linear(0,r)" by (unfold linear_def, blast) lemma tot_ord_0: "tot_ord(0,r)" unfolding tot_ord_def apply (simp add: part_ord_0 linear_0) done lemma wf_on_0: "wf[0](r)" by (unfold wf_on_def wf_def, blast) lemma well_ord_0: "well_ord(0,r)" unfolding well_ord_def apply (simp add: tot_ord_0 wf_on_0) done subsubsection\The Empty Relation Well-Orders the Unit Set\ text\by Grabczewski\ lemma tot_ord_unit: "tot_ord({a},0)" by (simp add: irrefl_def trans_on_def part_ord_def linear_def tot_ord_def) lemma well_ord_unit: "well_ord({a},0)" unfolding well_ord_def apply (simp add: tot_ord_unit wf_on_any_0) done subsection\Order-Isomorphisms\ text\Suppes calls them "similarities"\ (** Order-preserving (monotone) maps **) lemma mono_map_is_fun: "f \ mono_map(A,r,B,s) \ f \ A->B" by (simp add: mono_map_def) lemma mono_map_is_inj: "\linear(A,r); wf[B](s); f \ mono_map(A,r,B,s)\ \ f \ inj(A,B)" apply (unfold mono_map_def inj_def, clarify) apply (erule_tac x=w and y=x in linearE, assumption+) apply (force intro: apply_type dest: wf_on_not_refl)+ done lemma ord_isoI: "\f \ bij(A, B); \x y. \x \ A; y \ A\ \ \x, y\ \ r \ \ s\ \ f \ ord_iso(A,r,B,s)" by (simp add: ord_iso_def) lemma ord_iso_is_mono_map: "f \ ord_iso(A,r,B,s) \ f \ mono_map(A,r,B,s)" apply (simp add: ord_iso_def mono_map_def) apply (blast dest!: bij_is_fun) done lemma ord_iso_is_bij: "f \ ord_iso(A,r,B,s) \ f \ bij(A,B)" by (simp add: ord_iso_def) (*Needed? But ord_iso_converse is!*) lemma ord_iso_apply: "\f \ ord_iso(A,r,B,s); \x,y\: r; x \ A; y \ A\ \ \ s" by (simp add: ord_iso_def) lemma ord_iso_converse: "\f \ ord_iso(A,r,B,s); \x,y\: s; x \ B; y \ B\ \ \ r" apply (simp add: ord_iso_def, clarify) apply (erule bspec [THEN bspec, THEN iffD2]) apply (erule asm_rl bij_converse_bij [THEN bij_is_fun, THEN apply_type])+ apply (auto simp add: right_inverse_bij) done (** Symmetry and Transitivity Rules **) (*Reflexivity of similarity*) lemma ord_iso_refl: "id(A): ord_iso(A,r,A,r)" by (rule id_bij [THEN ord_isoI], simp) (*Symmetry of similarity*) lemma ord_iso_sym: "f \ ord_iso(A,r,B,s) \ converse(f): ord_iso(B,s,A,r)" apply (simp add: ord_iso_def) apply (auto simp add: right_inverse_bij bij_converse_bij bij_is_fun [THEN apply_funtype]) done (*Transitivity of similarity*) lemma mono_map_trans: "\g \ mono_map(A,r,B,s); f \ mono_map(B,s,C,t)\ \ (f O g): mono_map(A,r,C,t)" unfolding mono_map_def apply (auto simp add: comp_fun) done (*Transitivity of similarity: the order-isomorphism relation*) lemma ord_iso_trans: "\g \ ord_iso(A,r,B,s); f \ ord_iso(B,s,C,t)\ \ (f O g): ord_iso(A,r,C,t)" apply (unfold ord_iso_def, clarify) apply (frule bij_is_fun [of f]) apply (frule bij_is_fun [of g]) apply (auto simp add: comp_bij) done (** Two monotone maps can make an order-isomorphism **) lemma mono_ord_isoI: "\f \ mono_map(A,r,B,s); g \ mono_map(B,s,A,r); f O g = id(B); g O f = id(A)\ \ f \ ord_iso(A,r,B,s)" apply (simp add: ord_iso_def mono_map_def, safe) apply (intro fg_imp_bijective, auto) apply (subgoal_tac " \ r") apply (simp add: comp_eq_id_iff [THEN iffD1]) apply (blast intro: apply_funtype) done lemma well_ord_mono_ord_isoI: "\well_ord(A,r); well_ord(B,s); f \ mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r)\ \ f \ ord_iso(A,r,B,s)" apply (intro mono_ord_isoI, auto) apply (frule mono_map_is_fun [THEN fun_is_rel]) apply (erule converse_converse [THEN subst], rule left_comp_inverse) apply (blast intro: left_comp_inverse mono_map_is_inj well_ord_is_linear well_ord_is_wf)+ done (** Order-isomorphisms preserve the ordering's properties **) lemma part_ord_ord_iso: "\part_ord(B,s); f \ ord_iso(A,r,B,s)\ \ part_ord(A,r)" apply (simp add: part_ord_def irrefl_def trans_on_def ord_iso_def) apply (fast intro: bij_is_fun [THEN apply_type]) done lemma linear_ord_iso: "\linear(B,s); f \ ord_iso(A,r,B,s)\ \ linear(A,r)" apply (simp add: linear_def ord_iso_def, safe) apply (drule_tac x1 = "f`x" and x = "f`y" in bspec [THEN bspec]) apply (safe elim!: bij_is_fun [THEN apply_type]) apply (drule_tac t = "(`) (converse (f))" in subst_context) apply (simp add: left_inverse_bij) done lemma wf_on_ord_iso: "\wf[B](s); f \ ord_iso(A,r,B,s)\ \ wf[A](r)" apply (simp add: wf_on_def wf_def ord_iso_def, safe) apply (drule_tac x = "{f`z. z \ Z \ A}" in spec) apply (safe intro!: equalityI) apply (blast dest!: equalityD1 intro: bij_is_fun [THEN apply_type])+ done lemma well_ord_ord_iso: "\well_ord(B,s); f \ ord_iso(A,r,B,s)\ \ well_ord(A,r)" -apply (unfold well_ord_def tot_ord_def) + unfolding well_ord_def tot_ord_def apply (fast elim!: part_ord_ord_iso linear_ord_iso wf_on_ord_iso) done subsection\Main results of Kunen, Chapter 1 section 6\ (*Inductive argument for Kunen's Lemma 6.1, etc. Simple proof from Halmos, page 72*) lemma well_ord_iso_subset_lemma: "\well_ord(A,r); f \ ord_iso(A,r, A',r); A'<= A; y \ A\ \ \ : r" apply (simp add: well_ord_def ord_iso_def) apply (elim conjE CollectE) apply (rule_tac a=y in wf_on_induct, assumption+) apply (blast dest: bij_is_fun [THEN apply_type]) done (*Kunen's Lemma 6.1 \ there's no order-isomorphism to an initial segment of a well-ordering*) lemma well_ord_iso_predE: "\well_ord(A,r); f \ ord_iso(A, r, pred(A,x,r), r); x \ A\ \ P" apply (insert well_ord_iso_subset_lemma [of A r f "pred(A,x,r)" x]) apply (simp add: pred_subset) (*Now we know f`x < x *) apply (drule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption) (*Now we also know @{term"f`x \ pred(A,x,r)"}: contradiction! *) apply (simp add: well_ord_def pred_def) done (*Simple consequence of Lemma 6.1*) lemma well_ord_iso_pred_eq: "\well_ord(A,r); f \ ord_iso(pred(A,a,r), r, pred(A,c,r), r); a \ A; c \ A\ \ a=c" apply (frule well_ord_is_trans_on) apply (frule well_ord_is_linear) apply (erule_tac x=a and y=c in linearE, assumption+) apply (drule ord_iso_sym) (*two symmetric cases*) apply (auto elim!: well_ord_subset [OF _ pred_subset, THEN well_ord_iso_predE] intro!: predI simp add: trans_pred_pred_eq) done (*Does not assume r is a wellordering!*) lemma ord_iso_image_pred: "\f \ ord_iso(A,r,B,s); a \ A\ \ f `` pred(A,a,r) = pred(B, f`a, s)" -apply (unfold ord_iso_def pred_def) + unfolding ord_iso_def pred_def apply (erule CollectE) apply (simp (no_asm_simp) add: image_fun [OF bij_is_fun Collect_subset]) apply (rule equalityI) apply (safe elim!: bij_is_fun [THEN apply_type]) apply (rule RepFun_eqI) apply (blast intro!: right_inverse_bij [symmetric]) apply (auto simp add: right_inverse_bij bij_is_fun [THEN apply_funtype]) done lemma ord_iso_restrict_image: "\f \ ord_iso(A,r,B,s); C<=A\ \ restrict(f,C) \ ord_iso(C, r, f``C, s)" apply (simp add: ord_iso_def) apply (blast intro: bij_is_inj restrict_bij) done (*But in use, A and B may themselves be initial segments. Then use trans_pred_pred_eq to simplify the pred(pred...) terms. See just below.*) lemma ord_iso_restrict_pred: "\f \ ord_iso(A,r,B,s); a \ A\ \ restrict(f, pred(A,a,r)) \ ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)" apply (simp add: ord_iso_image_pred [symmetric]) apply (blast intro: ord_iso_restrict_image elim: predE) done (*Tricky; a lot of forward proof!*) lemma well_ord_iso_preserving: "\well_ord(A,r); well_ord(B,s); \a,c\: r; f \ ord_iso(pred(A,a,r), r, pred(B,b,s), s); g \ ord_iso(pred(A,c,r), r, pred(B,d,s), s); a \ A; c \ A; b \ B; d \ B\ \ \b,d\: s" apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], (erule asm_rl predI predE)+) apply (subgoal_tac "b = g`a") apply (simp (no_asm_simp)) apply (rule well_ord_iso_pred_eq, auto) apply (frule ord_iso_restrict_pred, (erule asm_rl predI)+) apply (simp add: well_ord_is_trans_on trans_pred_pred_eq) apply (erule ord_iso_sym [THEN ord_iso_trans], assumption) done (*See Halmos, page 72*) lemma well_ord_iso_unique_lemma: "\well_ord(A,r); f \ ord_iso(A,r, B,s); g \ ord_iso(A,r, B,s); y \ A\ \ \ \ s" apply (frule well_ord_iso_subset_lemma) apply (rule_tac f = "converse (f) " and g = g in ord_iso_trans) apply auto apply (blast intro: ord_iso_sym) apply (frule ord_iso_is_bij [of f]) apply (frule ord_iso_is_bij [of g]) apply (frule ord_iso_converse) apply (blast intro!: bij_converse_bij intro: bij_is_fun apply_funtype)+ apply (erule notE) apply (simp add: left_inverse_bij bij_is_fun comp_fun_apply [of _ A B]) done (*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*) lemma well_ord_iso_unique: "\well_ord(A,r); f \ ord_iso(A,r, B,s); g \ ord_iso(A,r, B,s)\ \ f = g" apply (rule fun_extension) apply (erule ord_iso_is_bij [THEN bij_is_fun])+ apply (subgoal_tac "f`x \ B \ g`x \ B \ linear(B,s)") apply (simp add: linear_def) apply (blast dest: well_ord_iso_unique_lemma) apply (blast intro: ord_iso_is_bij bij_is_fun apply_funtype well_ord_is_linear well_ord_ord_iso ord_iso_sym) done subsection\Towards Kunen's Theorem 6.3: Linearity of the Similarity Relation\ lemma ord_iso_map_subset: "ord_iso_map(A,r,B,s) \ A*B" by (unfold ord_iso_map_def, blast) lemma domain_ord_iso_map: "domain(ord_iso_map(A,r,B,s)) \ A" by (unfold ord_iso_map_def, blast) lemma range_ord_iso_map: "range(ord_iso_map(A,r,B,s)) \ B" by (unfold ord_iso_map_def, blast) lemma converse_ord_iso_map: "converse(ord_iso_map(A,r,B,s)) = ord_iso_map(B,s,A,r)" unfolding ord_iso_map_def apply (blast intro: ord_iso_sym) done lemma function_ord_iso_map: "well_ord(B,s) \ function(ord_iso_map(A,r,B,s))" -apply (unfold ord_iso_map_def function_def) + unfolding ord_iso_map_def function_def apply (blast intro: well_ord_iso_pred_eq ord_iso_sym ord_iso_trans) done lemma ord_iso_map_fun: "well_ord(B,s) \ ord_iso_map(A,r,B,s) \ domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))" by (simp add: Pi_iff function_ord_iso_map ord_iso_map_subset [THEN domain_times_range]) lemma ord_iso_map_mono_map: "\well_ord(A,r); well_ord(B,s)\ \ ord_iso_map(A,r,B,s) \ mono_map(domain(ord_iso_map(A,r,B,s)), r, range(ord_iso_map(A,r,B,s)), s)" unfolding mono_map_def apply (simp (no_asm_simp) add: ord_iso_map_fun) apply safe apply (subgoal_tac "x \ A \ ya:A \ y \ B \ yb:B") apply (simp add: apply_equality [OF _ ord_iso_map_fun]) unfolding ord_iso_map_def apply (blast intro: well_ord_iso_preserving, blast) done lemma ord_iso_map_ord_iso: "\well_ord(A,r); well_ord(B,s)\ \ ord_iso_map(A,r,B,s) \ ord_iso(domain(ord_iso_map(A,r,B,s)), r, range(ord_iso_map(A,r,B,s)), s)" apply (rule well_ord_mono_ord_isoI) prefer 4 apply (rule converse_ord_iso_map [THEN subst]) apply (simp add: ord_iso_map_mono_map ord_iso_map_subset [THEN converse_converse]) apply (blast intro!: domain_ord_iso_map range_ord_iso_map intro: well_ord_subset ord_iso_map_mono_map)+ done (*One way of saying that domain(ord_iso_map(A,r,B,s)) is downwards-closed*) lemma domain_ord_iso_map_subset: "\well_ord(A,r); well_ord(B,s); a \ A; a \ domain(ord_iso_map(A,r,B,s))\ \ domain(ord_iso_map(A,r,B,s)) \ pred(A, a, r)" unfolding ord_iso_map_def apply (safe intro!: predI) (*Case analysis on xa vs a in r *) apply (simp (no_asm_simp)) apply (frule_tac A = A in well_ord_is_linear) apply (rename_tac b y f) apply (erule_tac x=b and y=a in linearE, assumption+) (*Trivial case: b=a*) apply clarify apply blast (*Harder case: \a, xa\: r*) apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], (erule asm_rl predI predE)+) apply (frule ord_iso_restrict_pred) apply (simp add: pred_iff) apply (simp split: split_if_asm add: well_ord_is_trans_on trans_pred_pred_eq domain_UN domain_Union, blast) done (*For the 4-way case analysis in the main result*) lemma domain_ord_iso_map_cases: "\well_ord(A,r); well_ord(B,s)\ \ domain(ord_iso_map(A,r,B,s)) = A | (\x\A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))" apply (frule well_ord_is_wf) -apply (unfold wf_on_def wf_def) + unfolding wf_on_def wf_def apply (drule_tac x = "A-domain (ord_iso_map (A,r,B,s))" in spec) apply safe (*The first case: the domain equals A*) apply (rule domain_ord_iso_map [THEN equalityI]) apply (erule Diff_eq_0_iff [THEN iffD1]) (*The other case: the domain equals an initial segment*) apply (blast del: domainI subsetI elim!: predE intro!: domain_ord_iso_map_subset intro: subsetI)+ done (*As above, by duality*) lemma range_ord_iso_map_cases: "\well_ord(A,r); well_ord(B,s)\ \ range(ord_iso_map(A,r,B,s)) = B | (\y\B. range(ord_iso_map(A,r,B,s)) = pred(B,y,s))" apply (rule converse_ord_iso_map [THEN subst]) apply (simp add: domain_ord_iso_map_cases) done text\Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets\ theorem well_ord_trichotomy: "\well_ord(A,r); well_ord(B,s)\ \ ord_iso_map(A,r,B,s) \ ord_iso(A, r, B, s) | (\x\A. ord_iso_map(A,r,B,s) \ ord_iso(pred(A,x,r), r, B, s)) | (\y\B. ord_iso_map(A,r,B,s) \ ord_iso(A, r, pred(B,y,s), s))" apply (frule_tac B = B in domain_ord_iso_map_cases, assumption) apply (frule_tac B = B in range_ord_iso_map_cases, assumption) apply (drule ord_iso_map_ord_iso, assumption) apply (elim disjE bexE) apply (simp_all add: bexI) apply (rule wf_on_not_refl [THEN notE]) apply (erule well_ord_is_wf) apply assumption apply (subgoal_tac "\x,y\: ord_iso_map (A,r,B,s) ") apply (drule rangeI) apply (simp add: pred_def) apply (unfold ord_iso_map_def, blast) done subsection\Miscellaneous Results by Krzysztof Grabczewski\ (** Properties of converse(r) **) lemma irrefl_converse: "irrefl(A,r) \ irrefl(A,converse(r))" by (unfold irrefl_def, blast) lemma trans_on_converse: "trans[A](r) \ trans[A](converse(r))" by (unfold trans_on_def, blast) lemma part_ord_converse: "part_ord(A,r) \ part_ord(A,converse(r))" unfolding part_ord_def apply (blast intro!: irrefl_converse trans_on_converse) done lemma linear_converse: "linear(A,r) \ linear(A,converse(r))" by (unfold linear_def, blast) lemma tot_ord_converse: "tot_ord(A,r) \ tot_ord(A,converse(r))" unfolding tot_ord_def apply (blast intro!: part_ord_converse linear_converse) done (** By Krzysztof Grabczewski. Lemmas involving the first element of a well ordered set **) lemma first_is_elem: "first(b,B,r) \ b \ B" by (unfold first_def, blast) lemma well_ord_imp_ex1_first: "\well_ord(A,r); B<=A; B\0\ \ (\!b. first(b,B,r))" -apply (unfold well_ord_def wf_on_def wf_def first_def) + unfolding well_ord_def wf_on_def wf_def first_def apply (elim conjE allE disjE, blast) apply (erule bexE) apply (rule_tac a = x in ex1I, auto) apply (unfold tot_ord_def linear_def, blast) done lemma the_first_in: "\well_ord(A,r); B<=A; B\0\ \ (THE b. first(b,B,r)) \ B" apply (drule well_ord_imp_ex1_first, assumption+) apply (rule first_is_elem) apply (erule theI) done subsection \Lemmas for the Reflexive Orders\ lemma subset_vimage_vimage_iff: "\Preorder(r); A \ field(r); B \ field(r)\ \ r -`` A \ r -`` B \ (\a\A. \b\B. \a, b\ \ r)" apply (auto simp: subset_def preorder_on_def refl_def vimage_def image_def) apply blast unfolding trans_on_def apply (erule_tac P = "(\x. \y\field(r). \z\field(r). \x, y\ \ r \ \y, z\ \ r \ \x, z\ \ r)" for r in rev_ballE) (* instance obtained from proof term generated by best *) apply best apply blast done lemma subset_vimage1_vimage1_iff: "\Preorder(r); a \ field(r); b \ field(r)\ \ r -`` {a} \ r -`` {b} \ \a, b\ \ r" by (simp add: subset_vimage_vimage_iff) lemma Refl_antisym_eq_Image1_Image1_iff: "\refl(field(r), r); antisym(r); a \ field(r); b \ field(r)\ \ r `` {a} = r `` {b} \ a = b" apply rule apply (frule equality_iffD) apply (drule equality_iffD) apply (simp add: antisym_def refl_def) apply best apply (simp add: antisym_def refl_def) done lemma Partial_order_eq_Image1_Image1_iff: "\Partial_order(r); a \ field(r); b \ field(r)\ \ r `` {a} = r `` {b} \ a = b" by (simp add: partial_order_on_def preorder_on_def Refl_antisym_eq_Image1_Image1_iff) lemma Refl_antisym_eq_vimage1_vimage1_iff: "\refl(field(r), r); antisym(r); a \ field(r); b \ field(r)\ \ r -`` {a} = r -`` {b} \ a = b" apply rule apply (frule equality_iffD) apply (drule equality_iffD) apply (simp add: antisym_def refl_def) apply best apply (simp add: antisym_def refl_def) done lemma Partial_order_eq_vimage1_vimage1_iff: "\Partial_order(r); a \ field(r); b \ field(r)\ \ r -`` {a} = r -`` {b} \ a = b" by (simp add: partial_order_on_def preorder_on_def Refl_antisym_eq_vimage1_vimage1_iff) end diff --git a/src/ZF/OrderArith.thy b/src/ZF/OrderArith.thy --- a/src/ZF/OrderArith.thy +++ b/src/ZF/OrderArith.thy @@ -1,568 +1,568 @@ (* Title: ZF/OrderArith.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Combining Orderings: Foundations of Ordinal Arithmetic\ theory OrderArith imports Order Sum Ordinal begin definition (*disjoint sum of two relations; underlies ordinal addition*) radd :: "[i,i,i,i]\i" where "radd(A,r,B,s) \ {z: (A+B) * (A+B). (\x y. z = \Inl(x), Inr(y)\) | (\x' x. z = \Inl(x'), Inl(x)\ \ \x',x\:r) | (\y' y. z = \Inr(y'), Inr(y)\ \ \y',y\:s)}" definition (*lexicographic product of two relations; underlies ordinal multiplication*) rmult :: "[i,i,i,i]\i" where "rmult(A,r,B,s) \ {z: (A*B) * (A*B). \x' y' x y. z = \\x',y'\, \x,y\\ \ (\x',x\: r | (x'=x \ \y',y\: s))}" definition (*inverse image of a relation*) rvimage :: "[i,i,i]\i" where "rvimage(A,f,r) \ {z \ A*A. \x y. z = \x,y\ \ \f`x,f`y\: r}" definition measure :: "[i, i\i] \ i" where "measure(A,f) \ {\x,y\: A*A. f(x) < f(y)}" subsection\Addition of Relations -- Disjoint Sum\ subsubsection\Rewrite rules. Can be used to obtain introduction rules\ lemma radd_Inl_Inr_iff [iff]: "\Inl(a), Inr(b)\ \ radd(A,r,B,s) \ a \ A \ b \ B" by (unfold radd_def, blast) lemma radd_Inl_iff [iff]: "\Inl(a'), Inl(a)\ \ radd(A,r,B,s) \ a':A \ a \ A \ \a',a\:r" by (unfold radd_def, blast) lemma radd_Inr_iff [iff]: "\Inr(b'), Inr(b)\ \ radd(A,r,B,s) \ b':B \ b \ B \ \b',b\:s" by (unfold radd_def, blast) lemma radd_Inr_Inl_iff [simp]: "\Inr(b), Inl(a)\ \ radd(A,r,B,s) \ False" by (unfold radd_def, blast) declare radd_Inr_Inl_iff [THEN iffD1, dest!] subsubsection\Elimination Rule\ lemma raddE: "\\p',p\ \ radd(A,r,B,s); \x y. \p'=Inl(x); x \ A; p=Inr(y); y \ B\ \ Q; \x' x. \p'=Inl(x'); p=Inl(x); \x',x\: r; x':A; x \ A\ \ Q; \y' y. \p'=Inr(y'); p=Inr(y); \y',y\: s; y':B; y \ B\ \ Q \ \ Q" by (unfold radd_def, blast) subsubsection\Type checking\ lemma radd_type: "radd(A,r,B,s) \ (A+B) * (A+B)" unfolding radd_def apply (rule Collect_subset) done lemmas field_radd = radd_type [THEN field_rel_subset] subsubsection\Linearity\ lemma linear_radd: "\linear(A,r); linear(B,s)\ \ linear(A+B,radd(A,r,B,s))" by (unfold linear_def, blast) subsubsection\Well-foundedness\ lemma wf_on_radd: "\wf[A](r); wf[B](s)\ \ wf[A+B](radd(A,r,B,s))" apply (rule wf_onI2) apply (subgoal_tac "\x\A. Inl (x) \ Ba") \ \Proving the lemma, which is needed twice!\ prefer 2 apply (erule_tac V = "y \ A + B" in thin_rl) apply (rule_tac ballI) apply (erule_tac r = r and a = x in wf_on_induct, assumption) apply blast txt\Returning to main part of proof\ apply safe apply blast apply (erule_tac r = s and a = ya in wf_on_induct, assumption, blast) done lemma wf_radd: "\wf(r); wf(s)\ \ wf(radd(field(r),r,field(s),s))" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A [OF _ field_radd]) apply (blast intro: wf_on_radd) done lemma well_ord_radd: "\well_ord(A,r); well_ord(B,s)\ \ well_ord(A+B, radd(A,r,B,s))" apply (rule well_ordI) apply (simp add: well_ord_def wf_on_radd) apply (simp add: well_ord_def tot_ord_def linear_radd) done subsubsection\An \<^term>\ord_iso\ congruence law\ lemma sum_bij: "\f \ bij(A,C); g \ bij(B,D)\ \ (\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z)) \ bij(A+B, C+D)" apply (rule_tac d = "case (\x. Inl (converse(f)`x), \y. Inr(converse(g)`y))" in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) done lemma sum_ord_iso_cong: "\f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s')\ \ (\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z)) \ ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))" unfolding ord_iso_def apply (safe intro!: sum_bij) (*Do the beta-reductions now*) apply (auto cong add: conj_cong simp add: bij_is_fun [THEN apply_type]) done (*Could we prove an ord_iso result? Perhaps ord_iso(A+B, radd(A,r,B,s), A \ B, r \ s) *) lemma sum_disjoint_bij: "A \ B = 0 \ (\z\A+B. case(\x. x, \y. y, z)) \ bij(A+B, A \ B)" apply (rule_tac d = "\z. if z \ A then Inl (z) else Inr (z) " in lam_bijective) apply auto done subsubsection\Associativity\ lemma sum_assoc_bij: "(\z\(A+B)+C. case(case(Inl, \y. Inr(Inl(y))), \y. Inr(Inr(y)), z)) \ bij((A+B)+C, A+(B+C))" apply (rule_tac d = "case (\x. Inl (Inl (x)), case (\x. Inl (Inr (x)), Inr))" in lam_bijective) apply auto done lemma sum_assoc_ord_iso: "(\z\(A+B)+C. case(case(Inl, \y. Inr(Inl(y))), \y. Inr(Inr(y)), z)) \ ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))" by (rule sum_assoc_bij [THEN ord_isoI], auto) subsection\Multiplication of Relations -- Lexicographic Product\ subsubsection\Rewrite rule. Can be used to obtain introduction rules\ lemma rmult_iff [iff]: "\\a',b'\, \a,b\\ \ rmult(A,r,B,s) \ (\a',a\: r \ a':A \ a \ A \ b': B \ b \ B) | (\b',b\: s \ a'=a \ a \ A \ b': B \ b \ B)" by (unfold rmult_def, blast) lemma rmultE: "\\\a',b'\, \a,b\\ \ rmult(A,r,B,s); \\a',a\: r; a':A; a \ A; b':B; b \ B\ \ Q; \\b',b\: s; a \ A; a'=a; b':B; b \ B\ \ Q \ \ Q" by blast subsubsection\Type checking\ lemma rmult_type: "rmult(A,r,B,s) \ (A*B) * (A*B)" by (unfold rmult_def, rule Collect_subset) lemmas field_rmult = rmult_type [THEN field_rel_subset] subsubsection\Linearity\ lemma linear_rmult: "\linear(A,r); linear(B,s)\ \ linear(A*B,rmult(A,r,B,s))" by (simp add: linear_def, blast) subsubsection\Well-foundedness\ lemma wf_on_rmult: "\wf[A](r); wf[B](s)\ \ wf[A*B](rmult(A,r,B,s))" apply (rule wf_onI2) apply (erule SigmaE) apply (erule ssubst) apply (subgoal_tac "\b\B. \x,b\: Ba", blast) apply (erule_tac a = x in wf_on_induct, assumption) apply (rule ballI) apply (erule_tac a = b in wf_on_induct, assumption) apply (best elim!: rmultE bspec [THEN mp]) done lemma wf_rmult: "\wf(r); wf(s)\ \ wf(rmult(field(r),r,field(s),s))" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A [OF _ field_rmult]) apply (blast intro: wf_on_rmult) done lemma well_ord_rmult: "\well_ord(A,r); well_ord(B,s)\ \ well_ord(A*B, rmult(A,r,B,s))" apply (rule well_ordI) apply (simp add: well_ord_def wf_on_rmult) apply (simp add: well_ord_def tot_ord_def linear_rmult) done subsubsection\An \<^term>\ord_iso\ congruence law\ lemma prod_bij: "\f \ bij(A,C); g \ bij(B,D)\ \ (lam \x,y\:A*B. \f`x, g`y\) \ bij(A*B, C*D)" apply (rule_tac d = "\\x,y\. \converse (f) `x, converse (g) `y\" in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) done lemma prod_ord_iso_cong: "\f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s')\ \ (lam \x,y\:A*B. \f`x, g`y\) \ ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))" unfolding ord_iso_def apply (safe intro!: prod_bij) apply (simp_all add: bij_is_fun [THEN apply_type]) apply (blast intro: bij_is_inj [THEN inj_apply_equality]) done lemma singleton_prod_bij: "(\z\A. \x,z\) \ bij(A, {x}*A)" by (rule_tac d = snd in lam_bijective, auto) (*Used??*) lemma singleton_prod_ord_iso: "well_ord({x},xr) \ (\z\A. \x,z\) \ ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))" apply (rule singleton_prod_bij [THEN ord_isoI]) apply (simp (no_asm_simp)) apply (blast dest: well_ord_is_wf [THEN wf_on_not_refl]) done (*Here we build a complicated function term, then simplify it using case_cong, id_conv, comp_lam, case_case.*) lemma prod_sum_singleton_bij: "a\C \ (\x\C*B + D. case(\x. x, \y.\a,y\, x)) \ bij(C*B + D, C*B \ {a}*D)" apply (rule subst_elem) apply (rule id_bij [THEN sum_bij, THEN comp_bij]) apply (rule singleton_prod_bij) apply (rule sum_disjoint_bij, blast) apply (simp (no_asm_simp) cong add: case_cong) apply (rule comp_lam [THEN trans, symmetric]) apply (fast elim!: case_type) apply (simp (no_asm_simp) add: case_case) done lemma prod_sum_singleton_ord_iso: "\a \ A; well_ord(A,r)\ \ (\x\pred(A,a,r)*B + pred(B,b,s). case(\x. x, \y.\a,y\, x)) \ ord_iso(pred(A,a,r)*B + pred(B,b,s), radd(A*B, rmult(A,r,B,s), B, s), pred(A,a,r)*B \ {a}*pred(B,b,s), rmult(A,r,B,s))" apply (rule prod_sum_singleton_bij [THEN ord_isoI]) apply (simp (no_asm_simp) add: pred_iff well_ord_is_wf [THEN wf_on_not_refl]) apply (auto elim!: well_ord_is_wf [THEN wf_on_asym] predE) done subsubsection\Distributive law\ lemma sum_prod_distrib_bij: "(lam \x,z\:(A+B)*C. case(\y. Inl(\y,z\), \y. Inr(\y,z\), x)) \ bij((A+B)*C, (A*C)+(B*C))" by (rule_tac d = "case (\\x,y\.\Inl (x),y\, \\x,y\.\Inr (x),y\) " in lam_bijective, auto) lemma sum_prod_distrib_ord_iso: "(lam \x,z\:(A+B)*C. case(\y. Inl(\y,z\), \y. Inr(\y,z\), x)) \ ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t), (A*C)+(B*C), radd(A*C, rmult(A,r,C,t), B*C, rmult(B,s,C,t)))" by (rule sum_prod_distrib_bij [THEN ord_isoI], auto) subsubsection\Associativity\ lemma prod_assoc_bij: "(lam \\x,y\, z\:(A*B)*C. \x,\y,z\\) \ bij((A*B)*C, A*(B*C))" by (rule_tac d = "\\x, \y,z\\. \\x,y\, z\" in lam_bijective, auto) lemma prod_assoc_ord_iso: "(lam \\x,y\, z\:(A*B)*C. \x,\y,z\\) \ ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))" by (rule prod_assoc_bij [THEN ord_isoI], auto) subsection\Inverse Image of a Relation\ subsubsection\Rewrite rule\ lemma rvimage_iff: "\a,b\ \ rvimage(A,f,r) \ \f`a,f`b\: r \ a \ A \ b \ A" by (unfold rvimage_def, blast) subsubsection\Type checking\ lemma rvimage_type: "rvimage(A,f,r) \ A*A" by (unfold rvimage_def, rule Collect_subset) lemmas field_rvimage = rvimage_type [THEN field_rel_subset] lemma rvimage_converse: "rvimage(A,f, converse(r)) = converse(rvimage(A,f,r))" by (unfold rvimage_def, blast) subsubsection\Partial Ordering Properties\ lemma irrefl_rvimage: "\f \ inj(A,B); irrefl(B,r)\ \ irrefl(A, rvimage(A,f,r))" -apply (unfold irrefl_def rvimage_def) + unfolding irrefl_def rvimage_def apply (blast intro: inj_is_fun [THEN apply_type]) done lemma trans_on_rvimage: "\f \ inj(A,B); trans[B](r)\ \ trans[A](rvimage(A,f,r))" -apply (unfold trans_on_def rvimage_def) + unfolding trans_on_def rvimage_def apply (blast intro: inj_is_fun [THEN apply_type]) done lemma part_ord_rvimage: "\f \ inj(A,B); part_ord(B,r)\ \ part_ord(A, rvimage(A,f,r))" unfolding part_ord_def apply (blast intro!: irrefl_rvimage trans_on_rvimage) done subsubsection\Linearity\ lemma linear_rvimage: "\f \ inj(A,B); linear(B,r)\ \ linear(A,rvimage(A,f,r))" apply (simp add: inj_def linear_def rvimage_iff) apply (blast intro: apply_funtype) done lemma tot_ord_rvimage: "\f \ inj(A,B); tot_ord(B,r)\ \ tot_ord(A, rvimage(A,f,r))" unfolding tot_ord_def apply (blast intro!: part_ord_rvimage linear_rvimage) done subsubsection\Well-foundedness\ lemma wf_rvimage [intro!]: "wf(r) \ wf(rvimage(A,f,r))" apply (simp (no_asm_use) add: rvimage_def wf_eq_minimal) apply clarify apply (subgoal_tac "\w. w \ {w: {f`x. x \ Q}. \x. x \ Q \ (f`x = w) }") apply (erule allE) apply (erule impE) apply assumption apply blast apply blast done text\But note that the combination of \wf_imp_wf_on\ and \wf_rvimage\ gives \<^prop>\wf(r) \ wf[C](rvimage(A,f,r))\\ lemma wf_on_rvimage: "\f \ A\B; wf[B](r)\ \ wf[A](rvimage(A,f,r))" apply (rule wf_onI2) apply (subgoal_tac "\z\A. f`z=f`y \ z \ Ba") apply blast apply (erule_tac a = "f`y" in wf_on_induct) apply (blast intro!: apply_funtype) apply (blast intro!: apply_funtype dest!: rvimage_iff [THEN iffD1]) done (*Note that we need only wf[A](...) and linear(A,...) to get the result!*) lemma well_ord_rvimage: "\f \ inj(A,B); well_ord(B,r)\ \ well_ord(A, rvimage(A,f,r))" apply (rule well_ordI) -apply (unfold well_ord_def tot_ord_def) + unfolding well_ord_def tot_ord_def apply (blast intro!: wf_on_rvimage inj_is_fun) apply (blast intro!: linear_rvimage) done lemma ord_iso_rvimage: "f \ bij(A,B) \ f \ ord_iso(A, rvimage(A,f,s), B, s)" unfolding ord_iso_def apply (simp add: rvimage_iff) done lemma ord_iso_rvimage_eq: "f \ ord_iso(A,r, B,s) \ rvimage(A,f,s) = r \ A*A" by (unfold ord_iso_def rvimage_def, blast) subsection\Every well-founded relation is a subset of some inverse image of an ordinal\ lemma wf_rvimage_Ord: "Ord(i) \ wf(rvimage(A, f, Memrel(i)))" by (blast intro: wf_rvimage wf_Memrel) definition wfrank :: "[i,i]\i" where "wfrank(r,a) \ wfrec(r, a, \x f. \y \ r-``{x}. succ(f`y))" definition wftype :: "i\i" where "wftype(r) \ \y \ range(r). succ(wfrank(r,y))" lemma wfrank: "wf(r) \ wfrank(r,a) = (\y \ r-``{a}. succ(wfrank(r,y)))" by (subst wfrank_def [THEN def_wfrec], simp_all) lemma Ord_wfrank: "wf(r) \ Ord(wfrank(r,a))" apply (rule_tac a=a in wf_induct, assumption) apply (subst wfrank, assumption) apply (rule Ord_succ [THEN Ord_UN], blast) done lemma wfrank_lt: "\wf(r); \a,b\ \ r\ \ wfrank(r,a) < wfrank(r,b)" apply (rule_tac a1 = b in wfrank [THEN ssubst], assumption) apply (rule UN_I [THEN ltI]) apply (simp add: Ord_wfrank vimage_iff)+ done lemma Ord_wftype: "wf(r) \ Ord(wftype(r))" by (simp add: wftype_def Ord_wfrank) lemma wftypeI: "\wf(r); x \ field(r)\ \ wfrank(r,x) \ wftype(r)" apply (simp add: wftype_def) apply (blast intro: wfrank_lt [THEN ltD]) done lemma wf_imp_subset_rvimage: "\wf(r); r \ A*A\ \ \i f. Ord(i) \ r \ rvimage(A, f, Memrel(i))" apply (rule_tac x="wftype(r)" in exI) apply (rule_tac x="\x\A. wfrank(r,x)" in exI) apply (simp add: Ord_wftype, clarify) apply (frule subsetD, assumption, clarify) apply (simp add: rvimage_iff wfrank_lt [THEN ltD]) apply (blast intro: wftypeI) done theorem wf_iff_subset_rvimage: "relation(r) \ wf(r) \ (\i f A. Ord(i) \ r \ rvimage(A, f, Memrel(i)))" by (blast dest!: relation_field_times_field wf_imp_subset_rvimage intro: wf_rvimage_Ord [THEN wf_subset]) subsection\Other Results\ lemma wf_times: "A \ B = 0 \ wf(A*B)" by (simp add: wf_def, blast) text\Could also be used to prove \wf_radd\\ lemma wf_Un: "\range(r) \ domain(s) = 0; wf(r); wf(s)\ \ wf(r \ s)" apply (simp add: wf_def, clarify) apply (rule equalityI) prefer 2 apply blast apply clarify apply (drule_tac x=Z in spec) apply (drule_tac x="Z \ domain(s)" in spec) apply simp apply (blast intro: elim: equalityE) done subsubsection\The Empty Relation\ lemma wf0: "wf(0)" by (simp add: wf_def, blast) lemma linear0: "linear(0,0)" by (simp add: linear_def) lemma well_ord0: "well_ord(0,0)" by (blast intro: wf_imp_wf_on well_ordI wf0 linear0) subsubsection\The "measure" relation is useful with wfrec\ lemma measure_eq_rvimage_Memrel: "measure(A,f) = rvimage(A,Lambda(A,f),Memrel(Collect(RepFun(A,f),Ord)))" apply (simp (no_asm) add: measure_def rvimage_def Memrel_iff) apply (rule equalityI, auto) apply (auto intro: Ord_in_Ord simp add: lt_def) done lemma wf_measure [iff]: "wf(measure(A,f))" by (simp (no_asm) add: measure_eq_rvimage_Memrel wf_Memrel wf_rvimage) lemma measure_iff [iff]: "\x,y\ \ measure(A,f) \ x \ A \ y \ A \ f(x)x. x \ A \ Ord(f(x))" and inj: "\x y. \x \ A; y \ A; f(x) = f(y)\ \ x=y" shows "linear(A, measure(A,f))" apply (auto simp add: linear_def) apply (rule_tac i="f(x)" and j="f(y)" in Ord_linear_lt) apply (simp_all add: Ordf) apply (blast intro: inj) done lemma wf_on_measure: "wf[B](measure(A,f))" by (rule wf_imp_wf_on [OF wf_measure]) lemma well_ord_measure: assumes Ordf: "\x. x \ A \ Ord(f(x))" and inj: "\x y. \x \ A; y \ A; f(x) = f(y)\ \ x=y" shows "well_ord(A, measure(A,f))" apply (rule well_ordI) apply (rule wf_on_measure) apply (blast intro: linear_measure Ordf inj) done lemma measure_type: "measure(A,f) \ A*A" by (auto simp add: measure_def) subsubsection\Well-foundedness of Unions\ lemma wf_on_Union: assumes wfA: "wf[A](r)" and wfB: "\a. a\A \ wf[B(a)](s)" and ok: "\a u v. \\u,v\ \ s; v \ B(a); a \ A\ \ (\a'\A. \a',a\ \ r \ u \ B(a')) | u \ B(a)" shows "wf[\a\A. B(a)](s)" apply (rule wf_onI2) apply (erule UN_E) apply (subgoal_tac "\z \ B(a). z \ Ba", blast) apply (rule_tac a = a in wf_on_induct [OF wfA], assumption) apply (rule ballI) apply (rule_tac a = z in wf_on_induct [OF wfB], assumption, assumption) apply (rename_tac u) apply (drule_tac x=u in bspec, blast) apply (erule mp, clarify) apply (frule ok, assumption+, blast) done subsubsection\Bijections involving Powersets\ lemma Pow_sum_bij: "(\Z \ Pow(A+B). \{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}\) \ bij(Pow(A+B), Pow(A)*Pow(B))" apply (rule_tac d = "\\X,Y\. {Inl (x). x \ X} \ {Inr (y). y \ Y}" in lam_bijective) apply force+ done text\As a special case, we have \<^term>\bij(Pow(A*B), A \ Pow(B))\\ lemma Pow_Sigma_bij: "(\r \ Pow(Sigma(A,B)). \x \ A. r``{x}) \ bij(Pow(Sigma(A,B)), \x \ A. Pow(B(x)))" apply (rule_tac d = "\f. \x \ A. \y \ f`x. {\x,y\}" in lam_bijective) apply (blast intro: lam_type) apply (blast dest: apply_type, simp_all) apply fast (*strange, but blast can't do it*) apply (rule fun_extension, auto) by blast end diff --git a/src/ZF/OrderType.thy b/src/ZF/OrderType.thy --- a/src/ZF/OrderType.thy +++ b/src/ZF/OrderType.thy @@ -1,1028 +1,1028 @@ (* Title: ZF/OrderType.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Order Types and Ordinal Arithmetic\ theory OrderType imports OrderArith OrdQuant Nat begin text\The order type of a well-ordering is the least ordinal isomorphic to it. Ordinal arithmetic is traditionally defined in terms of order types, as it is here. But a definition by transfinite recursion would be much simpler!\ definition ordermap :: "[i,i]\i" where "ordermap(A,r) \ \x\A. wfrec[A](r, x, \x f. f `` pred(A,x,r))" definition ordertype :: "[i,i]\i" where "ordertype(A,r) \ ordermap(A,r)``A" definition (*alternative definition of ordinal numbers*) Ord_alt :: "i \ o" where "Ord_alt(X) \ well_ord(X, Memrel(X)) \ (\u\X. u=pred(X, u, Memrel(X)))" definition (*coercion to ordinal: if not, just 0*) ordify :: "i\i" where "ordify(x) \ if Ord(x) then x else 0" definition (*ordinal multiplication*) omult :: "[i,i]\i" (infixl \**\ 70) where "i ** j \ ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))" definition (*ordinal addition*) raw_oadd :: "[i,i]\i" where "raw_oadd(i,j) \ ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))" definition oadd :: "[i,i]\i" (infixl \++\ 65) where "i ++ j \ raw_oadd(ordify(i),ordify(j))" definition (*ordinal subtraction*) odiff :: "[i,i]\i" (infixl \--\ 65) where "i -- j \ ordertype(i-j, Memrel(i))" subsection\Proofs needing the combination of Ordinal.thy and Order.thy\ lemma le_well_ord_Memrel: "j \ i \ well_ord(j, Memrel(i))" apply (rule well_ordI) apply (rule wf_Memrel [THEN wf_imp_wf_on]) apply (simp add: ltD lt_Ord linear_def ltI [THEN lt_trans2 [of _ j i]]) apply (intro ballI Ord_linear) apply (blast intro: Ord_in_Ord lt_Ord)+ done (*"Ord(i) \ well_ord(i, Memrel(i))"*) lemmas well_ord_Memrel = le_refl [THEN le_well_ord_Memrel] (*Kunen's Theorem 7.3 (i), page 16; see also Ordinal/Ord_in_Ord The smaller ordinal is an initial segment of the larger *) lemma lt_pred_Memrel: "j pred(i, j, Memrel(i)) = j" apply (simp add: pred_def lt_def) apply (blast intro: Ord_trans) done lemma pred_Memrel: "x \ A \ pred(A, x, Memrel(A)) = A \ x" by (unfold pred_def Memrel_def, blast) lemma Ord_iso_implies_eq_lemma: "\j ord_iso(i,Memrel(i),j,Memrel(j))\ \ R" apply (frule lt_pred_Memrel) apply (erule ltE) apply (rule well_ord_Memrel [THEN well_ord_iso_predE, of i f j], auto) unfolding ord_iso_def (*Combining the two simplifications causes looping*) apply (simp (no_asm_simp)) apply (blast intro: bij_is_fun [THEN apply_type] Ord_trans) done (*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*) lemma Ord_iso_implies_eq: "\Ord(i); Ord(j); f \ ord_iso(i,Memrel(i),j,Memrel(j))\ \ i=j" apply (rule_tac i = i and j = j in Ord_linear_lt) apply (blast intro: ord_iso_sym Ord_iso_implies_eq_lemma)+ done subsection\Ordermap and ordertype\ lemma ordermap_type: "ordermap(A,r) \ A -> ordertype(A,r)" -apply (unfold ordermap_def ordertype_def) + unfolding ordermap_def ordertype_def apply (rule lam_type) apply (rule lamI [THEN imageI], assumption+) done subsubsection\Unfolding of ordermap\ (*Useful for cardinality reasoning; see CardinalArith.ML*) lemma ordermap_eq_image: "\wf[A](r); x \ A\ \ ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)" -apply (unfold ordermap_def pred_def) + unfolding ordermap_def pred_def apply (simp (no_asm_simp)) apply (erule wfrec_on [THEN trans], assumption) apply (simp (no_asm_simp) add: subset_iff image_lam vimage_singleton_iff) done (*Useful for rewriting PROVIDED pred is not unfolded until later!*) lemma ordermap_pred_unfold: "\wf[A](r); x \ A\ \ ordermap(A,r) ` x = {ordermap(A,r)`y . y \ pred(A,x,r)}" by (simp add: ordermap_eq_image pred_subset ordermap_type [THEN image_fun]) (*pred-unfolded version. NOT suitable for rewriting -- loops!*) lemmas ordermap_unfold = ordermap_pred_unfold [simplified pred_def] (*The theorem above is \wf[A](r); x \ A\ \ ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \ A . \y,x\ \ r}} NOTE: the definition of ordermap used here delivers ordinals only if r is transitive. If r is the predecessor relation on the naturals then ordermap(nat,predr) ` n equals {n-1} and not n. A more complicated definition, like ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y \ A . \y,x\ \ r}}, might eliminate the need for r to be transitive. *) subsubsection\Showing that ordermap, ordertype yield ordinals\ lemma Ord_ordermap: "\well_ord(A,r); x \ A\ \ Ord(ordermap(A,r) ` x)" apply (unfold well_ord_def tot_ord_def part_ord_def, safe) apply (rule_tac a=x in wf_on_induct, assumption+) apply (simp (no_asm_simp) add: ordermap_pred_unfold) apply (rule OrdI [OF _ Ord_is_Transset]) -apply (unfold pred_def Transset_def) + unfolding pred_def Transset_def apply (blast intro: trans_onD dest!: ordermap_unfold [THEN equalityD1])+ done lemma Ord_ordertype: "well_ord(A,r) \ Ord(ordertype(A,r))" unfolding ordertype_def apply (subst image_fun [OF ordermap_type subset_refl]) apply (rule OrdI [OF _ Ord_is_Transset]) prefer 2 apply (blast intro: Ord_ordermap) -apply (unfold Transset_def well_ord_def) + unfolding Transset_def well_ord_def apply (blast intro: trans_onD dest!: ordermap_unfold [THEN equalityD1]) done subsubsection\ordermap preserves the orderings in both directions\ lemma ordermap_mono: "\\w,x\: r; wf[A](r); w \ A; x \ A\ \ ordermap(A,r)`w \ ordermap(A,r)`x" apply (erule_tac x1 = x in ordermap_unfold [THEN ssubst], assumption, blast) done (*linearity of r is crucial here*) lemma converse_ordermap_mono: "\ordermap(A,r)`w \ ordermap(A,r)`x; well_ord(A,r); w \ A; x \ A\ \ \w,x\: r" apply (unfold well_ord_def tot_ord_def, safe) apply (erule_tac x=w and y=x in linearE, assumption+) apply (blast elim!: mem_not_refl [THEN notE]) apply (blast dest: ordermap_mono intro: mem_asym) done lemma ordermap_surj: "ordermap(A, r) \ surj(A, ordertype(A, r))" unfolding ordertype_def by (rule surj_image) (rule ordermap_type) lemma ordermap_bij: "well_ord(A,r) \ ordermap(A,r) \ bij(A, ordertype(A,r))" -apply (unfold well_ord_def tot_ord_def bij_def inj_def) + unfolding well_ord_def tot_ord_def bij_def inj_def apply (force intro!: ordermap_type ordermap_surj elim: linearE dest: ordermap_mono simp add: mem_not_refl) done subsubsection\Isomorphisms involving ordertype\ lemma ordertype_ord_iso: "well_ord(A,r) \ ordermap(A,r) \ ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))" unfolding ord_iso_def apply (safe elim!: well_ord_is_wf intro!: ordermap_type [THEN apply_type] ordermap_mono ordermap_bij) apply (blast dest!: converse_ordermap_mono) done lemma ordertype_eq: "\f \ ord_iso(A,r,B,s); well_ord(B,s)\ \ ordertype(A,r) = ordertype(B,s)" apply (frule well_ord_ord_iso, assumption) apply (rule Ord_iso_implies_eq, (erule Ord_ordertype)+) apply (blast intro: ord_iso_trans ord_iso_sym ordertype_ord_iso) done lemma ordertype_eq_imp_ord_iso: "\ordertype(A,r) = ordertype(B,s); well_ord(A,r); well_ord(B,s)\ \ \f. f \ ord_iso(A,r,B,s)" apply (rule exI) apply (rule ordertype_ord_iso [THEN ord_iso_trans], assumption) apply (erule ssubst) apply (erule ordertype_ord_iso [THEN ord_iso_sym]) done subsubsection\Basic equalities for ordertype\ (*Ordertype of Memrel*) lemma le_ordertype_Memrel: "j \ i \ ordertype(j,Memrel(i)) = j" apply (rule Ord_iso_implies_eq [symmetric]) apply (erule ltE, assumption) apply (blast intro: le_well_ord_Memrel Ord_ordertype) apply (rule ord_iso_trans) apply (erule_tac [2] le_well_ord_Memrel [THEN ordertype_ord_iso]) apply (rule id_bij [THEN ord_isoI]) apply (simp (no_asm_simp)) apply (fast elim: ltE Ord_in_Ord Ord_trans) done (*"Ord(i) \ ordertype(i, Memrel(i)) = i"*) lemmas ordertype_Memrel = le_refl [THEN le_ordertype_Memrel] lemma ordertype_0 [simp]: "ordertype(0,r) = 0" apply (rule id_bij [THEN ord_isoI, THEN ordertype_eq, THEN trans]) apply (erule emptyE) apply (rule well_ord_0) apply (rule Ord_0 [THEN ordertype_Memrel]) done (*Ordertype of rvimage: \f \ bij(A,B); well_ord(B,s)\ \ ordertype(A, rvimage(A,f,s)) = ordertype(B,s) *) lemmas bij_ordertype_vimage = ord_iso_rvimage [THEN ordertype_eq] subsubsection\A fundamental unfolding law for ordertype.\ (*Ordermap returns the same result if applied to an initial segment*) lemma ordermap_pred_eq_ordermap: "\well_ord(A,r); y \ A; z \ pred(A,y,r)\ \ ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z" apply (frule wf_on_subset_A [OF well_ord_is_wf pred_subset]) apply (rule_tac a=z in wf_on_induct, assumption+) apply (safe elim!: predE) apply (simp (no_asm_simp) add: ordermap_pred_unfold well_ord_is_wf pred_iff) (*combining these two simplifications LOOPS! *) apply (simp (no_asm_simp) add: pred_pred_eq) apply (simp add: pred_def) apply (rule RepFun_cong [OF _ refl]) apply (drule well_ord_is_trans_on) apply (fast elim!: trans_onD) done lemma ordertype_unfold: "ordertype(A,r) = {ordermap(A,r)`y . y \ A}" unfolding ordertype_def apply (rule image_fun [OF ordermap_type subset_refl]) done text\Theorems by Krzysztof Grabczewski; proofs simplified by lcp\ lemma ordertype_pred_subset: "\well_ord(A,r); x \ A\ \ ordertype(pred(A,x,r),r) \ ordertype(A,r)" apply (simp add: ordertype_unfold well_ord_subset [OF _ pred_subset]) apply (fast intro: ordermap_pred_eq_ordermap elim: predE) done lemma ordertype_pred_lt: "\well_ord(A,r); x \ A\ \ ordertype(pred(A,x,r),r) < ordertype(A,r)" apply (rule ordertype_pred_subset [THEN subset_imp_le, THEN leE]) apply (simp_all add: Ord_ordertype well_ord_subset [OF _ pred_subset]) apply (erule sym [THEN ordertype_eq_imp_ord_iso, THEN exE]) apply (erule_tac [3] well_ord_iso_predE) apply (simp_all add: well_ord_subset [OF _ pred_subset]) done (*May rewrite with this -- provided no rules are supplied for proving that well_ord(pred(A,x,r), r) *) lemma ordertype_pred_unfold: "well_ord(A,r) \ ordertype(A,r) = {ordertype(pred(A,x,r),r). x \ A}" apply (rule equalityI) apply (safe intro!: ordertype_pred_lt [THEN ltD]) apply (auto simp add: ordertype_def well_ord_is_wf [THEN ordermap_eq_image] ordermap_type [THEN image_fun] ordermap_pred_eq_ordermap pred_subset) done subsection\Alternative definition of ordinal\ (*proof by Krzysztof Grabczewski*) lemma Ord_is_Ord_alt: "Ord(i) \ Ord_alt(i)" unfolding Ord_alt_def apply (rule conjI) apply (erule well_ord_Memrel) apply (unfold Ord_def Transset_def pred_def Memrel_def, blast) done (*proof by lcp*) lemma Ord_alt_is_Ord: "Ord_alt(i) \ Ord(i)" apply (unfold Ord_alt_def Ord_def Transset_def well_ord_def tot_ord_def part_ord_def trans_on_def) apply (simp add: pred_Memrel) apply (blast elim!: equalityE) done subsection\Ordinal Addition\ subsubsection\Order Type calculations for radd\ text\Addition with 0\ lemma bij_sum_0: "(\z\A+0. case(\x. x, \y. y, z)) \ bij(A+0, A)" apply (rule_tac d = Inl in lam_bijective, safe) apply (simp_all (no_asm_simp)) done lemma ordertype_sum_0_eq: "well_ord(A,r) \ ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)" apply (rule bij_sum_0 [THEN ord_isoI, THEN ordertype_eq]) prefer 2 apply assumption apply force done lemma bij_0_sum: "(\z\0+A. case(\x. x, \y. y, z)) \ bij(0+A, A)" apply (rule_tac d = Inr in lam_bijective, safe) apply (simp_all (no_asm_simp)) done lemma ordertype_0_sum_eq: "well_ord(A,r) \ ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)" apply (rule bij_0_sum [THEN ord_isoI, THEN ordertype_eq]) prefer 2 apply assumption apply force done text\Initial segments of radd. Statements by Grabczewski\ (*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *) lemma pred_Inl_bij: "a \ A \ (\x\pred(A,a,r). Inl(x)) \ bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))" unfolding pred_def apply (rule_tac d = "case (\x. x, \y. y) " in lam_bijective) apply auto done lemma ordertype_pred_Inl_eq: "\a \ A; well_ord(A,r)\ \ ordertype(pred(A+B, Inl(a), radd(A,r,B,s)), radd(A,r,B,s)) = ordertype(pred(A,a,r), r)" apply (rule pred_Inl_bij [THEN ord_isoI, THEN ord_iso_sym, THEN ordertype_eq]) apply (simp_all add: well_ord_subset [OF _ pred_subset]) apply (simp add: pred_def) done lemma pred_Inr_bij: "b \ B \ id(A+pred(B,b,s)) \ bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))" -apply (unfold pred_def id_def) + unfolding pred_def id_def apply (rule_tac d = "\z. z" in lam_bijective, auto) done lemma ordertype_pred_Inr_eq: "\b \ B; well_ord(A,r); well_ord(B,s)\ \ ordertype(pred(A+B, Inr(b), radd(A,r,B,s)), radd(A,r,B,s)) = ordertype(A+pred(B,b,s), radd(A,r,pred(B,b,s),s))" apply (rule pred_Inr_bij [THEN ord_isoI, THEN ord_iso_sym, THEN ordertype_eq]) prefer 2 apply (force simp add: pred_def id_def, assumption) apply (blast intro: well_ord_radd well_ord_subset [OF _ pred_subset]) done subsubsection\ordify: trivial coercion to an ordinal\ lemma Ord_ordify [iff, TC]: "Ord(ordify(x))" by (simp add: ordify_def) (*Collapsing*) lemma ordify_idem [simp]: "ordify(ordify(x)) = ordify(x)" by (simp add: ordify_def) subsubsection\Basic laws for ordinal addition\ lemma Ord_raw_oadd: "\Ord(i); Ord(j)\ \ Ord(raw_oadd(i,j))" by (simp add: raw_oadd_def ordify_def Ord_ordertype well_ord_radd well_ord_Memrel) lemma Ord_oadd [iff,TC]: "Ord(i++j)" by (simp add: oadd_def Ord_raw_oadd) text\Ordinal addition with zero\ lemma raw_oadd_0: "Ord(i) \ raw_oadd(i,0) = i" by (simp add: raw_oadd_def ordify_def ordertype_sum_0_eq ordertype_Memrel well_ord_Memrel) lemma oadd_0 [simp]: "Ord(i) \ i++0 = i" apply (simp (no_asm_simp) add: oadd_def raw_oadd_0 ordify_def) done lemma raw_oadd_0_left: "Ord(i) \ raw_oadd(0,i) = i" by (simp add: raw_oadd_def ordify_def ordertype_0_sum_eq ordertype_Memrel well_ord_Memrel) lemma oadd_0_left [simp]: "Ord(i) \ 0++i = i" by (simp add: oadd_def raw_oadd_0_left ordify_def) lemma oadd_eq_if_raw_oadd: "i++j = (if Ord(i) then (if Ord(j) then raw_oadd(i,j) else i) else (if Ord(j) then j else 0))" by (simp add: oadd_def ordify_def raw_oadd_0_left raw_oadd_0) lemma raw_oadd_eq_oadd: "\Ord(i); Ord(j)\ \ raw_oadd(i,j) = i++j" by (simp add: oadd_def ordify_def) (*** Further properties of ordinal addition. Statements by Grabczewski, proofs by lcp. ***) (*Surely also provable by transfinite induction on j?*) lemma lt_oadd1: "k k < i++j" apply (simp add: oadd_def ordify_def lt_Ord2 raw_oadd_0, clarify) apply (simp add: raw_oadd_def) apply (rule ltE, assumption) apply (rule ltI) apply (force simp add: ordertype_pred_unfold well_ord_radd well_ord_Memrel ordertype_pred_Inl_eq lt_pred_Memrel leI [THEN le_ordertype_Memrel]) apply (blast intro: Ord_ordertype well_ord_radd well_ord_Memrel) done (*Thus also we obtain the rule @{term"i++j = k \ i \ k"} *) lemma oadd_le_self: "Ord(i) \ i \ i++j" apply (rule all_lt_imp_le) apply (auto simp add: Ord_oadd lt_oadd1) done text\Various other results\ lemma id_ord_iso_Memrel: "A<=B \ id(A) \ ord_iso(A, Memrel(A), A, Memrel(B))" apply (rule id_bij [THEN ord_isoI]) apply (simp (no_asm_simp)) apply blast done lemma subset_ord_iso_Memrel: "\f \ ord_iso(A,Memrel(B),C,r); A<=B\ \ f \ ord_iso(A,Memrel(A),C,r)" apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN fun_is_rel]) apply (frule ord_iso_trans [OF id_ord_iso_Memrel], assumption) apply (simp add: right_comp_id) done lemma restrict_ord_iso: "\f \ ord_iso(i, Memrel(i), Order.pred(A,a,r), r); a \ A; j < i; trans[A](r)\ \ restrict(f,j) \ ord_iso(j, Memrel(j), Order.pred(A,f`j,r), r)" apply (frule ltD) apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], assumption) apply (frule ord_iso_restrict_pred, assumption) apply (simp add: pred_iff trans_pred_pred_eq lt_pred_Memrel) apply (blast intro!: subset_ord_iso_Memrel le_imp_subset [OF leI]) done lemma restrict_ord_iso2: "\f \ ord_iso(Order.pred(A,a,r), r, i, Memrel(i)); a \ A; j < i; trans[A](r)\ \ converse(restrict(converse(f), j)) \ ord_iso(Order.pred(A, converse(f)`j, r), r, j, Memrel(j))" by (blast intro: restrict_ord_iso ord_iso_sym ltI) lemma ordertype_sum_Memrel: "\well_ord(A,r); k \ ordertype(A+k, radd(A, r, k, Memrel(j))) = ordertype(A+k, radd(A, r, k, Memrel(k)))" apply (erule ltE) apply (rule ord_iso_refl [THEN sum_ord_iso_cong, THEN ordertype_eq]) apply (erule OrdmemD [THEN id_ord_iso_Memrel, THEN ord_iso_sym]) apply (simp_all add: well_ord_radd well_ord_Memrel) done lemma oadd_lt_mono2: "k i++k < i++j" apply (simp add: oadd_def ordify_def raw_oadd_0_left lt_Ord lt_Ord2, clarify) apply (simp add: raw_oadd_def) apply (rule ltE, assumption) apply (rule ordertype_pred_unfold [THEN equalityD2, THEN subsetD, THEN ltI]) apply (simp_all add: Ord_ordertype well_ord_radd well_ord_Memrel) apply (rule bexI) apply (erule_tac [2] InrI) apply (simp add: ordertype_pred_Inr_eq well_ord_Memrel lt_pred_Memrel leI [THEN le_ordertype_Memrel] ordertype_sum_Memrel) done lemma oadd_lt_cancel2: "\i++j < i++k; Ord(j)\ \ j i++j < i++k \ ji++j = i++k; Ord(j); Ord(k)\ \ j=k" apply (simp add: oadd_eq_if_raw_oadd split: split_if_asm) apply (simp add: raw_oadd_eq_oadd) apply (rule Ord_linear_lt, auto) apply (force dest: oadd_lt_mono2 [of concl: i] simp add: lt_not_refl)+ done lemma lt_oadd_disj: "k < i++j \ kl\j. k = i++l )" apply (simp add: Ord_in_Ord' [of _ j] oadd_eq_if_raw_oadd split: split_if_asm) prefer 2 apply (simp add: Ord_in_Ord' [of _ j] lt_def) apply (simp add: ordertype_pred_unfold well_ord_radd well_ord_Memrel raw_oadd_def) apply (erule ltD [THEN RepFunE]) apply (force simp add: ordertype_pred_Inl_eq well_ord_Memrel ltI lt_pred_Memrel le_ordertype_Memrel leI ordertype_pred_Inr_eq ordertype_sum_Memrel) done subsubsection\Ordinal addition with successor -- via associativity!\ lemma oadd_assoc: "(i++j)++k = i++(j++k)" apply (simp add: oadd_eq_if_raw_oadd Ord_raw_oadd raw_oadd_0 raw_oadd_0_left, clarify) apply (simp add: raw_oadd_def) apply (rule ordertype_eq [THEN trans]) apply (rule sum_ord_iso_cong [OF ordertype_ord_iso [THEN ord_iso_sym] ord_iso_refl]) apply (simp_all add: Ord_ordertype well_ord_radd well_ord_Memrel) apply (rule sum_assoc_ord_iso [THEN ordertype_eq, THEN trans]) apply (rule_tac [2] ordertype_eq) apply (rule_tac [2] sum_ord_iso_cong [OF ord_iso_refl ordertype_ord_iso]) apply (blast intro: Ord_ordertype well_ord_radd well_ord_Memrel)+ done lemma oadd_unfold: "\Ord(i); Ord(j)\ \ i++j = i \ (\k\j. {i++k})" apply (rule subsetI [THEN equalityI]) apply (erule ltI [THEN lt_oadd_disj, THEN disjE]) apply (blast intro: Ord_oadd) apply (blast elim!: ltE, blast) apply (force intro: lt_oadd1 oadd_lt_mono2 simp add: Ord_mem_iff_lt) done lemma oadd_1: "Ord(i) \ i++1 = succ(i)" apply (simp (no_asm_simp) add: oadd_unfold Ord_1 oadd_0) apply blast done lemma oadd_succ [simp]: "Ord(j) \ i++succ(j) = succ(i++j)" apply (simp add: oadd_eq_if_raw_oadd, clarify) apply (simp add: raw_oadd_eq_oadd) apply (simp add: oadd_1 [of j, symmetric] oadd_1 [of "i++j", symmetric] oadd_assoc) done text\Ordinal addition with limit ordinals\ lemma oadd_UN: "\\x. x \ A \ Ord(j(x)); a \ A\ \ i ++ (\x\A. j(x)) = (\x\A. i++j(x))" by (blast intro: ltI Ord_UN Ord_oadd lt_oadd1 [THEN ltD] oadd_lt_mono2 [THEN ltD] elim!: ltE dest!: ltI [THEN lt_oadd_disj]) lemma oadd_Limit: "Limit(j) \ i++j = (\k\j. i++k)" apply (frule Limit_has_0 [THEN ltD]) apply (simp add: Limit_is_Ord [THEN Ord_in_Ord] oadd_UN [symmetric] Union_eq_UN [symmetric] Limit_Union_eq) done lemma oadd_eq_0_iff: "\Ord(i); Ord(j)\ \ (i ++ j) = 0 \ i=0 \ j=0" apply (erule trans_induct3 [of j]) apply (simp_all add: oadd_Limit) apply (simp add: Union_empty_iff Limit_def lt_def, blast) done lemma oadd_eq_lt_iff: "\Ord(i); Ord(j)\ \ 0 < (i ++ j) \ 0Ord(i); Limit(j)\ \ Limit(i ++ j)" apply (simp add: oadd_Limit) apply (frule Limit_has_1 [THEN ltD]) apply (rule increasing_LimitI) apply (rule Ord_0_lt) apply (blast intro: Ord_in_Ord [OF Limit_is_Ord]) apply (force simp add: Union_empty_iff oadd_eq_0_iff Limit_is_Ord [of j, THEN Ord_in_Ord], auto) apply (rule_tac x="succ(y)" in bexI) apply (simp add: ltI Limit_is_Ord [of j, THEN Ord_in_Ord]) apply (simp add: Limit_def lt_def) done text\Order/monotonicity properties of ordinal addition\ lemma oadd_le_self2: "Ord(i) \ i \ j++i" proof (induct i rule: trans_induct3) case 0 thus ?case by (simp add: Ord_0_le) next case (succ i) thus ?case by (simp add: oadd_succ succ_leI) next case (limit l) hence "l = (\x\l. x)" by (simp add: Union_eq_UN [symmetric] Limit_Union_eq) also have "... \ (\x\l. j++x)" by (rule le_implies_UN_le_UN) (rule limit.hyps) finally have "l \ (\x\l. j++x)" . thus ?case using limit.hyps by (simp add: oadd_Limit) qed lemma oadd_le_mono1: "k \ j \ k++i \ j++i" apply (frule lt_Ord) apply (frule le_Ord2) apply (simp add: oadd_eq_if_raw_oadd, clarify) apply (simp add: raw_oadd_eq_oadd) apply (erule_tac i = i in trans_induct3) apply (simp (no_asm_simp)) apply (simp (no_asm_simp) add: oadd_succ succ_le_iff) apply (simp (no_asm_simp) add: oadd_Limit) apply (rule le_implies_UN_le_UN, blast) done lemma oadd_lt_mono: "\i' \ i; j' \ i'++j' < i++j" by (blast intro: lt_trans1 oadd_le_mono1 oadd_lt_mono2 Ord_succD elim: ltE) lemma oadd_le_mono: "\i' \ i; j' \ j\ \ i'++j' \ i++j" by (simp del: oadd_succ add: oadd_succ [symmetric] le_Ord2 oadd_lt_mono) lemma oadd_le_iff2: "\Ord(j); Ord(k)\ \ i++j \ i++k \ j \ k" by (simp del: oadd_succ add: oadd_lt_iff2 oadd_succ [symmetric] Ord_succ) lemma oadd_lt_self: "\Ord(i); 0 \ i < i++j" apply (rule lt_trans2) apply (erule le_refl) apply (simp only: lt_Ord2 oadd_1 [of i, symmetric]) apply (blast intro: succ_leI oadd_le_mono) done text\Every ordinal is exceeded by some limit ordinal.\ lemma Ord_imp_greater_Limit: "Ord(i) \ \k. i Limit(k)" apply (rule_tac x="i ++ nat" in exI) apply (blast intro: oadd_LimitI oadd_lt_self Limit_nat [THEN Limit_has_0]) done lemma Ord2_imp_greater_Limit: "\Ord(i); Ord(j)\ \ \k. i j Limit(k)" apply (insert Ord_Un [of i j, THEN Ord_imp_greater_Limit]) apply (simp add: Un_least_lt_iff) done subsection\Ordinal Subtraction\ text\The difference is \<^term>\ordertype(j-i, Memrel(j))\. It's probably simpler to define the difference recursively!\ lemma bij_sum_Diff: "A<=B \ (\y\B. if(y \ A, Inl(y), Inr(y))) \ bij(B, A+(B-A))" apply (rule_tac d = "case (\x. x, \y. y) " in lam_bijective) apply (blast intro!: if_type) apply (fast intro!: case_type) apply (erule_tac [2] sumE) apply (simp_all (no_asm_simp)) done lemma ordertype_sum_Diff: "i \ j \ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = ordertype(j, Memrel(j))" apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule bij_sum_Diff [THEN ord_isoI, THEN ord_iso_sym, THEN ordertype_eq]) apply (erule_tac [3] well_ord_Memrel, assumption) apply (simp (no_asm_simp)) apply (frule_tac j = y in Ord_in_Ord, assumption) apply (frule_tac j = x in Ord_in_Ord, assumption) apply (simp (no_asm_simp) add: Ord_mem_iff_lt lt_Ord not_lt_iff_le) apply (blast intro: lt_trans2 lt_trans) done lemma Ord_odiff [simp,TC]: "\Ord(i); Ord(j)\ \ Ord(i--j)" unfolding odiff_def apply (blast intro: Ord_ordertype Diff_subset well_ord_subset well_ord_Memrel) done lemma raw_oadd_ordertype_Diff: "i \ j \ raw_oadd(i,j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))" apply (simp add: raw_oadd_def odiff_def) apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule sum_ord_iso_cong [THEN ordertype_eq]) apply (erule id_ord_iso_Memrel) apply (rule ordertype_ord_iso [THEN ord_iso_sym]) apply (blast intro: well_ord_radd Diff_subset well_ord_subset well_ord_Memrel)+ done lemma oadd_odiff_inverse: "i \ j \ i ++ (j--i) = j" by (simp add: lt_Ord le_Ord2 oadd_def ordify_def raw_oadd_ordertype_Diff ordertype_sum_Diff ordertype_Memrel lt_Ord2 [THEN Ord_succD]) (*By oadd_inject, the difference between i and j is unique. Note that we get i++j = k \ j = k--i. *) lemma odiff_oadd_inverse: "\Ord(i); Ord(j)\ \ (i++j) -- i = j" apply (rule oadd_inject) apply (blast intro: oadd_odiff_inverse oadd_le_self) apply (blast intro: Ord_ordertype Ord_oadd Ord_odiff)+ done lemma odiff_lt_mono2: "\i i\ \ i--k < j--k" apply (rule_tac i = k in oadd_lt_cancel2) apply (simp add: oadd_odiff_inverse) apply (subst oadd_odiff_inverse) apply (blast intro: le_trans leI, assumption) apply (simp (no_asm_simp) add: lt_Ord le_Ord2) done subsection\Ordinal Multiplication\ lemma Ord_omult [simp,TC]: "\Ord(i); Ord(j)\ \ Ord(i**j)" unfolding omult_def apply (blast intro: Ord_ordertype well_ord_rmult well_ord_Memrel) done subsubsection\A useful unfolding law\ lemma pred_Pair_eq: "\a \ A; b \ B\ \ pred(A*B, \a,b\, rmult(A,r,B,s)) = pred(A,a,r)*B \ ({a} * pred(B,b,s))" apply (unfold pred_def, blast) done lemma ordertype_pred_Pair_eq: "\a \ A; b \ B; well_ord(A,r); well_ord(B,s)\ \ ordertype(pred(A*B, \a,b\, rmult(A,r,B,s)), rmult(A,r,B,s)) = ordertype(pred(A,a,r)*B + pred(B,b,s), radd(A*B, rmult(A,r,B,s), B, s))" apply (simp (no_asm_simp) add: pred_Pair_eq) apply (rule ordertype_eq [symmetric]) apply (rule prod_sum_singleton_ord_iso) apply (simp_all add: pred_subset well_ord_rmult [THEN well_ord_subset]) apply (blast intro: pred_subset well_ord_rmult [THEN well_ord_subset] elim!: predE) done lemma ordertype_pred_Pair_lemma: "\i' \ ordertype(pred(i*j, , rmult(i,Memrel(i),j,Memrel(j))), rmult(i,Memrel(i),j,Memrel(j))) = raw_oadd (j**i', j')" -apply (unfold raw_oadd_def omult_def) + unfolding raw_oadd_def omult_def apply (simp add: ordertype_pred_Pair_eq lt_pred_Memrel ltD lt_Ord2 well_ord_Memrel) apply (rule trans) apply (rule_tac [2] ordertype_ord_iso [THEN sum_ord_iso_cong, THEN ordertype_eq]) apply (rule_tac [3] ord_iso_refl) apply (rule id_bij [THEN ord_isoI, THEN ordertype_eq]) apply (elim SigmaE sumE ltE ssubst) apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel Ord_ordertype lt_Ord lt_Ord2) apply (blast intro: Ord_trans)+ done lemma lt_omult: "\Ord(i); Ord(j); k \ \j' i'. k = j**i' ++ j' \ j' i'j' \ j**i' ++ j' < j**i" unfolding omult_def apply (rule ltI) prefer 2 apply (simp add: Ord_ordertype well_ord_rmult well_ord_Memrel lt_Ord2) apply (simp add: ordertype_pred_unfold well_ord_rmult well_ord_Memrel lt_Ord2) apply (rule bexI [of _ i']) apply (rule bexI [of _ j']) apply (simp add: ordertype_pred_Pair_lemma ltI omult_def [symmetric]) apply (simp add: lt_Ord lt_Ord2 raw_oadd_eq_oadd) apply (simp_all add: lt_def) done lemma omult_unfold: "\Ord(i); Ord(j)\ \ j**i = (\j'\j. \i'\i. {j**i' ++ j'})" apply (rule subsetI [THEN equalityI]) apply (rule lt_omult [THEN exE]) apply (erule_tac [3] ltI) apply (simp_all add: Ord_omult) apply (blast elim!: ltE) apply (blast intro: omult_oadd_lt [THEN ltD] ltI) done subsubsection\Basic laws for ordinal multiplication\ text\Ordinal multiplication by zero\ lemma omult_0 [simp]: "i**0 = 0" unfolding omult_def apply (simp (no_asm_simp)) done lemma omult_0_left [simp]: "0**i = 0" unfolding omult_def apply (simp (no_asm_simp)) done text\Ordinal multiplication by 1\ lemma omult_1 [simp]: "Ord(i) \ i**1 = i" unfolding omult_def apply (rule_tac s1="Memrel(i)" in ord_isoI [THEN ordertype_eq, THEN trans]) apply (rule_tac c = snd and d = "\z.\0,z\" in lam_bijective) apply (auto elim!: snd_type well_ord_Memrel ordertype_Memrel) done lemma omult_1_left [simp]: "Ord(i) \ 1**i = i" unfolding omult_def apply (rule_tac s1="Memrel(i)" in ord_isoI [THEN ordertype_eq, THEN trans]) apply (rule_tac c = fst and d = "\z.\z,0\" in lam_bijective) apply (auto elim!: fst_type well_ord_Memrel ordertype_Memrel) done text\Distributive law for ordinal multiplication and addition\ lemma oadd_omult_distrib: "\Ord(i); Ord(j); Ord(k)\ \ i**(j++k) = (i**j)++(i**k)" apply (simp add: oadd_eq_if_raw_oadd) apply (simp add: omult_def raw_oadd_def) apply (rule ordertype_eq [THEN trans]) apply (rule prod_ord_iso_cong [OF ordertype_ord_iso [THEN ord_iso_sym] ord_iso_refl]) apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel Ord_ordertype) apply (rule sum_prod_distrib_ord_iso [THEN ordertype_eq, THEN trans]) apply (rule_tac [2] ordertype_eq) apply (rule_tac [2] sum_ord_iso_cong [OF ordertype_ord_iso ordertype_ord_iso]) apply (simp_all add: well_ord_rmult well_ord_radd well_ord_Memrel Ord_ordertype) done lemma omult_succ: "\Ord(i); Ord(j)\ \ i**succ(j) = (i**j)++i" by (simp del: oadd_succ add: oadd_1 [of j, symmetric] oadd_omult_distrib) text\Associative law\ lemma omult_assoc: "\Ord(i); Ord(j); Ord(k)\ \ (i**j)**k = i**(j**k)" unfolding omult_def apply (rule ordertype_eq [THEN trans]) apply (rule prod_ord_iso_cong [OF ord_iso_refl ordertype_ord_iso [THEN ord_iso_sym]]) apply (blast intro: well_ord_rmult well_ord_Memrel)+ apply (rule prod_assoc_ord_iso [THEN ord_iso_sym, THEN ordertype_eq, THEN trans]) apply (rule_tac [2] ordertype_eq) apply (rule_tac [2] prod_ord_iso_cong [OF ordertype_ord_iso ord_iso_refl]) apply (blast intro: well_ord_rmult well_ord_Memrel Ord_ordertype)+ done text\Ordinal multiplication with limit ordinals\ lemma omult_UN: "\Ord(i); \x. x \ A \ Ord(j(x))\ \ i ** (\x\A. j(x)) = (\x\A. i**j(x))" by (simp (no_asm_simp) add: Ord_UN omult_unfold, blast) lemma omult_Limit: "\Ord(i); Limit(j)\ \ i**j = (\k\j. i**k)" by (simp add: Limit_is_Ord [THEN Ord_in_Ord] omult_UN [symmetric] Union_eq_UN [symmetric] Limit_Union_eq) subsubsection\Ordering/monotonicity properties of ordinal multiplication\ (*As a special case we have "\0 \ 0 < i**j" *) lemma lt_omult1: "\k \ k < i**j" apply (safe elim!: ltE intro!: ltI Ord_omult) apply (force simp add: omult_unfold) done lemma omult_le_self: "\Ord(i); 0 \ i \ i**j" by (blast intro: all_lt_imp_le Ord_omult lt_omult1 lt_Ord2) lemma omult_le_mono1: assumes kj: "k \ j" and i: "Ord(i)" shows "k**i \ j**i" proof - have o: "Ord(k)" "Ord(j)" by (rule lt_Ord [OF kj] le_Ord2 [OF kj])+ show ?thesis using i proof (induct i rule: trans_induct3) case 0 thus ?case by simp next case (succ i) thus ?case by (simp add: o kj omult_succ oadd_le_mono) next case (limit l) thus ?case by (auto simp add: o kj omult_Limit le_implies_UN_le_UN) qed qed lemma omult_lt_mono2: "\k \ i**k < i**j" apply (rule ltI) apply (simp (no_asm_simp) add: omult_unfold lt_Ord2) apply (safe elim!: ltE intro!: Ord_omult) apply (force simp add: Ord_omult) done lemma omult_le_mono2: "\k \ j; Ord(i)\ \ i**k \ i**j" apply (rule subset_imp_le) apply (safe elim!: ltE dest!: Ord_succD intro!: Ord_omult) apply (simp add: omult_unfold) apply (blast intro: Ord_trans) done lemma omult_le_mono: "\i' \ i; j' \ j\ \ i'**j' \ i**j" by (blast intro: le_trans omult_le_mono1 omult_le_mono2 Ord_succD elim: ltE) lemma omult_lt_mono: "\i' \ i; j' \ i'**j' < i**j" by (blast intro: lt_trans1 omult_le_mono1 omult_lt_mono2 Ord_succD elim: ltE) lemma omult_le_self2: assumes i: "Ord(i)" and j: "0 j**i" proof - have oj: "Ord(j)" by (rule lt_Ord2 [OF j]) show ?thesis using i proof (induct i rule: trans_induct3) case 0 thus ?case by simp next case (succ i) have "j ** i ++ 0 < j ** i ++ j" by (rule oadd_lt_mono2 [OF j]) with succ.hyps show ?case by (simp add: oj j omult_succ ) (rule lt_trans1) next case (limit l) hence "l = (\x\l. x)" by (simp add: Union_eq_UN [symmetric] Limit_Union_eq) also have "... \ (\x\l. j**x)" by (rule le_implies_UN_le_UN) (rule limit.hyps) finally have "l \ (\x\l. j**x)" . thus ?case using limit.hyps by (simp add: oj omult_Limit) qed qed text\Further properties of ordinal multiplication\ lemma omult_inject: "\i**j = i**k; 0 \ j=k" apply (rule Ord_linear_lt) prefer 4 apply assumption apply auto apply (force dest: omult_lt_mono2 simp add: lt_not_refl)+ done subsection\The Relation \<^term>\Lt\\ lemma wf_Lt: "wf(Lt)" apply (rule wf_subset) apply (rule wf_Memrel) apply (auto simp add: Lt_def Memrel_def lt_def) done lemma irrefl_Lt: "irrefl(A,Lt)" by (auto simp add: Lt_def irrefl_def) lemma trans_Lt: "trans[A](Lt)" apply (simp add: Lt_def trans_on_def) apply (blast intro: lt_trans) done lemma part_ord_Lt: "part_ord(A,Lt)" by (simp add: part_ord_def irrefl_Lt trans_Lt) lemma linear_Lt: "linear(nat,Lt)" apply (auto dest!: not_lt_imp_le simp add: Lt_def linear_def le_iff) apply (drule lt_asym, auto) done lemma tot_ord_Lt: "tot_ord(nat,Lt)" by (simp add: tot_ord_def linear_Lt part_ord_Lt) lemma well_ord_Lt: "well_ord(nat,Lt)" by (simp add: well_ord_def wf_Lt wf_imp_wf_on tot_ord_Lt) end diff --git a/src/ZF/Perm.thy b/src/ZF/Perm.thy --- a/src/ZF/Perm.thy +++ b/src/ZF/Perm.thy @@ -1,554 +1,554 @@ (* Title: ZF/Perm.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge The theory underlying permutation groups -- Composition of relations, the identity relation -- Injections, surjections, bijections -- Lemmas for the Schroeder-Bernstein Theorem *) section\Injections, Surjections, Bijections, Composition\ theory Perm imports func begin definition (*composition of relations and functions; NOT Suppes's relative product*) comp :: "[i,i]\i" (infixr \O\ 60) where "r O s \ {xz \ domain(s)*range(r) . \x y z. xz=\x,z\ \ \x,y\:s \ \y,z\:r}" definition (*the identity function for A*) id :: "i\i" where "id(A) \ (\x\A. x)" definition (*one-to-one functions from A to B*) inj :: "[i,i]\i" where "inj(A,B) \ { f \ A->B. \w\A. \x\A. f`w=f`x \ w=x}" definition (*onto functions from A to B*) surj :: "[i,i]\i" where "surj(A,B) \ { f \ A->B . \y\B. \x\A. f`x=y}" definition (*one-to-one and onto functions*) bij :: "[i,i]\i" where "bij(A,B) \ inj(A,B) \ surj(A,B)" subsection\Surjective Function Space\ lemma surj_is_fun: "f \ surj(A,B) \ f \ A->B" unfolding surj_def apply (erule CollectD1) done lemma fun_is_surj: "f \ Pi(A,B) \ f \ surj(A,range(f))" unfolding surj_def apply (blast intro: apply_equality range_of_fun domain_type) done lemma surj_range: "f \ surj(A,B) \ range(f)=B" unfolding surj_def apply (best intro: apply_Pair elim: range_type) done text\A function with a right inverse is a surjection\ lemma f_imp_surjective: "\f \ A->B; \y. y \ B \ d(y): A; \y. y \ B \ f`d(y) = y\ \ f \ surj(A,B)" by (simp add: surj_def, blast) lemma lam_surjective: "\\x. x \ A \ c(x): B; \y. y \ B \ d(y): A; \y. y \ B \ c(d(y)) = y \ \ (\x\A. c(x)) \ surj(A,B)" apply (rule_tac d = d in f_imp_surjective) apply (simp_all add: lam_type) done text\Cantor's theorem revisited\ lemma cantor_surj: "f \ surj(A,Pow(A))" apply (unfold surj_def, safe) apply (cut_tac cantor) apply (best del: subsetI) done subsection\Injective Function Space\ lemma inj_is_fun: "f \ inj(A,B) \ f \ A->B" unfolding inj_def apply (erule CollectD1) done text\Good for dealing with sets of pairs, but a bit ugly in use [used in AC]\ lemma inj_equality: "\\a,b\:f; \c,b\:f; f \ inj(A,B)\ \ a=c" unfolding inj_def apply (blast dest: Pair_mem_PiD) done lemma inj_apply_equality: "\f \ inj(A,B); f`a=f`b; a \ A; b \ A\ \ a=b" by (unfold inj_def, blast) text\A function with a left inverse is an injection\ lemma f_imp_injective: "\f \ A->B; \x\A. d(f`x)=x\ \ f \ inj(A,B)" apply (simp (no_asm_simp) add: inj_def) apply (blast intro: subst_context [THEN box_equals]) done lemma lam_injective: "\\x. x \ A \ c(x): B; \x. x \ A \ d(c(x)) = x\ \ (\x\A. c(x)) \ inj(A,B)" apply (rule_tac d = d in f_imp_injective) apply (simp_all add: lam_type) done subsection\Bijections\ lemma bij_is_inj: "f \ bij(A,B) \ f \ inj(A,B)" unfolding bij_def apply (erule IntD1) done lemma bij_is_surj: "f \ bij(A,B) \ f \ surj(A,B)" unfolding bij_def apply (erule IntD2) done lemma bij_is_fun: "f \ bij(A,B) \ f \ A->B" by (rule bij_is_inj [THEN inj_is_fun]) lemma lam_bijective: "\\x. x \ A \ c(x): B; \y. y \ B \ d(y): A; \x. x \ A \ d(c(x)) = x; \y. y \ B \ c(d(y)) = y \ \ (\x\A. c(x)) \ bij(A,B)" unfolding bij_def apply (blast intro!: lam_injective lam_surjective) done lemma RepFun_bijective: "(\y\x. \!y'. f(y') = f(y)) \ (\z\{f(y). y \ x}. THE y. f(y) = z) \ bij({f(y). y \ x}, x)" apply (rule_tac d = f in lam_bijective) apply (auto simp add: the_equality2) done subsection\Identity Function\ lemma idI [intro!]: "a \ A \ \a,a\ \ id(A)" unfolding id_def apply (erule lamI) done lemma idE [elim!]: "\p \ id(A); \x.\x \ A; p=\x,x\\ \ P\ \ P" by (simp add: id_def lam_def, blast) lemma id_type: "id(A) \ A->A" unfolding id_def apply (rule lam_type, assumption) done lemma id_conv [simp]: "x \ A \ id(A)`x = x" unfolding id_def apply (simp (no_asm_simp)) done lemma id_mono: "A<=B \ id(A) \ id(B)" unfolding id_def apply (erule lam_mono) done lemma id_subset_inj: "A<=B \ id(A): inj(A,B)" apply (simp add: inj_def id_def) apply (blast intro: lam_type) done lemmas id_inj = subset_refl [THEN id_subset_inj] lemma id_surj: "id(A): surj(A,A)" -apply (unfold id_def surj_def) + unfolding id_def surj_def apply (simp (no_asm_simp)) done lemma id_bij: "id(A): bij(A,A)" unfolding bij_def apply (blast intro: id_inj id_surj) done lemma subset_iff_id: "A \ B \ id(A) \ A->B" unfolding id_def apply (force intro!: lam_type dest: apply_type) done text\\<^term>\id\ as the identity relation\ lemma id_iff [simp]: "\x,y\ \ id(A) \ x=y \ y \ A" by auto subsection\Converse of a Function\ lemma inj_converse_fun: "f \ inj(A,B) \ converse(f) \ range(f)->A" unfolding inj_def apply (simp (no_asm_simp) add: Pi_iff function_def) apply (erule CollectE) apply (simp (no_asm_simp) add: apply_iff) apply (blast dest: fun_is_rel) done text\Equations for converse(f)\ text\The premises are equivalent to saying that f is injective...\ lemma left_inverse_lemma: "\f \ A->B; converse(f): C->A; a \ A\ \ converse(f)`(f`a) = a" by (blast intro: apply_Pair apply_equality converseI) lemma left_inverse [simp]: "\f \ inj(A,B); a \ A\ \ converse(f)`(f`a) = a" by (blast intro: left_inverse_lemma inj_converse_fun inj_is_fun) lemma left_inverse_eq: "\f \ inj(A,B); f ` x = y; x \ A\ \ converse(f) ` y = x" by auto lemmas left_inverse_bij = bij_is_inj [THEN left_inverse] lemma right_inverse_lemma: "\f \ A->B; converse(f): C->A; b \ C\ \ f`(converse(f)`b) = b" by (rule apply_Pair [THEN converseD [THEN apply_equality]], auto) (*Should the premises be f \ surj(A,B), b \ B for symmetry with left_inverse? No: they would not imply that converse(f) was a function! *) lemma right_inverse [simp]: "\f \ inj(A,B); b \ range(f)\ \ f`(converse(f)`b) = b" by (blast intro: right_inverse_lemma inj_converse_fun inj_is_fun) lemma right_inverse_bij: "\f \ bij(A,B); b \ B\ \ f`(converse(f)`b) = b" by (force simp add: bij_def surj_range) subsection\Converses of Injections, Surjections, Bijections\ lemma inj_converse_inj: "f \ inj(A,B) \ converse(f): inj(range(f), A)" apply (rule f_imp_injective) apply (erule inj_converse_fun, clarify) apply (rule right_inverse) apply assumption apply blast done lemma inj_converse_surj: "f \ inj(A,B) \ converse(f): surj(range(f), A)" by (blast intro: f_imp_surjective inj_converse_fun left_inverse inj_is_fun range_of_fun [THEN apply_type]) text\Adding this as an intro! rule seems to cause looping\ lemma bij_converse_bij [TC]: "f \ bij(A,B) \ converse(f): bij(B,A)" unfolding bij_def apply (fast elim: surj_range [THEN subst] inj_converse_inj inj_converse_surj) done subsection\Composition of Two Relations\ text\The inductive definition package could derive these theorems for \<^term>\r O s\\ lemma compI [intro]: "\\a,b\:s; \b,c\:r\ \ \a,c\ \ r O s" by (unfold comp_def, blast) lemma compE [elim!]: "\xz \ r O s; \x y z. \xz=\x,z\; \x,y\:s; \y,z\:r\ \ P\ \ P" by (unfold comp_def, blast) lemma compEpair: "\\a,c\ \ r O s; \y. \\a,y\:s; \y,c\:r\ \ P\ \ P" by (erule compE, simp) lemma converse_comp: "converse(R O S) = converse(S) O converse(R)" by blast subsection\Domain and Range -- see Suppes, Section 3.1\ text\Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327\ lemma range_comp: "range(r O s) \ range(r)" by blast lemma range_comp_eq: "domain(r) \ range(s) \ range(r O s) = range(r)" by (rule range_comp [THEN equalityI], blast) lemma domain_comp: "domain(r O s) \ domain(s)" by blast lemma domain_comp_eq: "range(s) \ domain(r) \ domain(r O s) = domain(s)" by (rule domain_comp [THEN equalityI], blast) lemma image_comp: "(r O s)``A = r``(s``A)" by blast lemma inj_inj_range: "f \ inj(A,B) \ f \ inj(A,range(f))" by (auto simp add: inj_def Pi_iff function_def) lemma inj_bij_range: "f \ inj(A,B) \ f \ bij(A,range(f))" by (auto simp add: bij_def intro: inj_inj_range inj_is_fun fun_is_surj) subsection\Other Results\ lemma comp_mono: "\r'<=r; s'<=s\ \ (r' O s') \ (r O s)" by blast text\composition preserves relations\ lemma comp_rel: "\s<=A*B; r<=B*C\ \ (r O s) \ A*C" by blast text\associative law for composition\ lemma comp_assoc: "(r O s) O t = r O (s O t)" by blast (*left identity of composition; provable inclusions are id(A) O r \ r and \r<=A*B; B<=C\ \ r \ id(C) O r *) lemma left_comp_id: "r<=A*B \ id(B) O r = r" by blast (*right identity of composition; provable inclusions are r O id(A) \ r and \r<=A*B; A<=C\ \ r \ r O id(C) *) lemma right_comp_id: "r<=A*B \ r O id(A) = r" by blast subsection\Composition Preserves Functions, Injections, and Surjections\ lemma comp_function: "\function(g); function(f)\ \ function(f O g)" by (unfold function_def, blast) text\Don't think the premises can be weakened much\ lemma comp_fun: "\g \ A->B; f \ B->C\ \ (f O g) \ A->C" apply (auto simp add: Pi_def comp_function Pow_iff comp_rel) apply (subst range_rel_subset [THEN domain_comp_eq], auto) done (*Thanks to the new definition of "apply", the premise f \ B->C is gone!*) lemma comp_fun_apply [simp]: "\g \ A->B; a \ A\ \ (f O g)`a = f`(g`a)" apply (frule apply_Pair, assumption) apply (simp add: apply_def image_comp) apply (blast dest: apply_equality) done text\Simplifies compositions of lambda-abstractions\ lemma comp_lam: "\\x. x \ A \ b(x): B\ \ (\y\B. c(y)) O (\x\A. b(x)) = (\x\A. c(b(x)))" apply (subgoal_tac "(\x\A. b(x)) \ A -> B") apply (rule fun_extension) apply (blast intro: comp_fun lam_funtype) apply (rule lam_funtype) apply simp apply (simp add: lam_type) done lemma comp_inj: "\g \ inj(A,B); f \ inj(B,C)\ \ (f O g) \ inj(A,C)" apply (frule inj_is_fun [of g]) apply (frule inj_is_fun [of f]) apply (rule_tac d = "\y. converse (g) ` (converse (f) ` y)" in f_imp_injective) apply (blast intro: comp_fun, simp) done lemma comp_surj: "\g \ surj(A,B); f \ surj(B,C)\ \ (f O g) \ surj(A,C)" unfolding surj_def apply (blast intro!: comp_fun comp_fun_apply) done lemma comp_bij: "\g \ bij(A,B); f \ bij(B,C)\ \ (f O g) \ bij(A,C)" unfolding bij_def apply (blast intro: comp_inj comp_surj) done subsection\Dual Properties of \<^term>\inj\ and \<^term>\surj\\ text\Useful for proofs from D Pastre. Automatic theorem proving in set theory. Artificial Intelligence, 10:1--27, 1978.\ lemma comp_mem_injD1: "\(f O g): inj(A,C); g \ A->B; f \ B->C\ \ g \ inj(A,B)" by (unfold inj_def, force) lemma comp_mem_injD2: "\(f O g): inj(A,C); g \ surj(A,B); f \ B->C\ \ f \ inj(B,C)" apply (unfold inj_def surj_def, safe) apply (rule_tac x1 = x in bspec [THEN bexE]) apply (erule_tac [3] x1 = w in bspec [THEN bexE], assumption+, safe) apply (rule_tac t = "(`) (g) " in subst_context) apply (erule asm_rl bspec [THEN bspec, THEN mp])+ apply (simp (no_asm_simp)) done lemma comp_mem_surjD1: "\(f O g): surj(A,C); g \ A->B; f \ B->C\ \ f \ surj(B,C)" unfolding surj_def apply (blast intro!: comp_fun_apply [symmetric] apply_funtype) done lemma comp_mem_surjD2: "\(f O g): surj(A,C); g \ A->B; f \ inj(B,C)\ \ g \ surj(A,B)" apply (unfold inj_def surj_def, safe) apply (drule_tac x = "f`y" in bspec, auto) apply (blast intro: apply_funtype) done subsubsection\Inverses of Composition\ text\left inverse of composition; one inclusion is \<^term>\f \ A->B \ id(A) \ converse(f) O f\\ lemma left_comp_inverse: "f \ inj(A,B) \ converse(f) O f = id(A)" apply (unfold inj_def, clarify) apply (rule equalityI) apply (auto simp add: apply_iff, blast) done text\right inverse of composition; one inclusion is \<^term>\f \ A->B \ f O converse(f) \ id(B)\\ lemma right_comp_inverse: "f \ surj(A,B) \ f O converse(f) = id(B)" apply (simp add: surj_def, clarify) apply (rule equalityI) apply (best elim: domain_type range_type dest: apply_equality2) apply (blast intro: apply_Pair) done subsubsection\Proving that a Function is a Bijection\ lemma comp_eq_id_iff: "\f \ A->B; g \ B->A\ \ f O g = id(B) \ (\y\B. f`(g`y)=y)" apply (unfold id_def, safe) apply (drule_tac t = "\h. h`y " in subst_context) apply simp apply (rule fun_extension) apply (blast intro: comp_fun lam_type) apply auto done lemma fg_imp_bijective: "\f \ A->B; g \ B->A; f O g = id(B); g O f = id(A)\ \ f \ bij(A,B)" unfolding bij_def apply (simp add: comp_eq_id_iff) apply (blast intro: f_imp_injective f_imp_surjective apply_funtype) done lemma nilpotent_imp_bijective: "\f \ A->A; f O f = id(A)\ \ f \ bij(A,A)" by (blast intro: fg_imp_bijective) lemma invertible_imp_bijective: "\converse(f): B->A; f \ A->B\ \ f \ bij(A,B)" by (simp add: fg_imp_bijective comp_eq_id_iff left_inverse_lemma right_inverse_lemma) subsubsection\Unions of Functions\ text\See similar theorems in func.thy\ text\Theorem by KG, proof by LCP\ lemma inj_disjoint_Un: "\f \ inj(A,B); g \ inj(C,D); B \ D = 0\ \ (\a\A \ C. if a \ A then f`a else g`a) \ inj(A \ C, B \ D)" apply (rule_tac d = "\z. if z \ B then converse (f) `z else converse (g) `z" in lam_injective) apply (auto simp add: inj_is_fun [THEN apply_type]) done lemma surj_disjoint_Un: "\f \ surj(A,B); g \ surj(C,D); A \ C = 0\ \ (f \ g) \ surj(A \ C, B \ D)" apply (simp add: surj_def fun_disjoint_Un) apply (blast dest!: domain_of_fun intro!: fun_disjoint_apply1 fun_disjoint_apply2) done text\A simple, high-level proof; the version for injections follows from it, using \<^term>\f \ inj(A,B) \ f \ bij(A,range(f))\\ lemma bij_disjoint_Un: "\f \ bij(A,B); g \ bij(C,D); A \ C = 0; B \ D = 0\ \ (f \ g) \ bij(A \ C, B \ D)" apply (rule invertible_imp_bijective) apply (subst converse_Un) apply (auto intro: fun_disjoint_Un bij_is_fun bij_converse_bij) done subsubsection\Restrictions as Surjections and Bijections\ lemma surj_image: "f \ Pi(A,B) \ f \ surj(A, f``A)" apply (simp add: surj_def) apply (blast intro: apply_equality apply_Pair Pi_type) done lemma surj_image_eq: "f \ surj(A, B) \ f``A = B" by (auto simp add: surj_def image_fun) (blast dest: apply_type) lemma restrict_image [simp]: "restrict(f,A) `` B = f `` (A \ B)" by (auto simp add: restrict_def) lemma restrict_inj: "\f \ inj(A,B); C<=A\ \ restrict(f,C): inj(C,B)" unfolding inj_def apply (safe elim!: restrict_type2, auto) done lemma restrict_surj: "\f \ Pi(A,B); C<=A\ \ restrict(f,C): surj(C, f``C)" apply (insert restrict_type2 [THEN surj_image]) apply (simp add: restrict_image) done lemma restrict_bij: "\f \ inj(A,B); C<=A\ \ restrict(f,C): bij(C, f``C)" apply (simp add: inj_def bij_def) apply (blast intro: restrict_surj surj_is_fun) done subsubsection\Lemmas for Ramsey's Theorem\ lemma inj_weaken_type: "\f \ inj(A,B); B<=D\ \ f \ inj(A,D)" unfolding inj_def apply (blast intro: fun_weaken_type) done lemma inj_succ_restrict: "\f \ inj(succ(m), A)\ \ restrict(f,m) \ inj(m, A-{f`m})" apply (rule restrict_bij [THEN bij_is_inj, THEN inj_weaken_type], assumption, blast) unfolding inj_def apply (fast elim: range_type mem_irrefl dest: apply_equality) done lemma inj_extend: "\f \ inj(A,B); a\A; b\B\ \ cons(\a,b\,f) \ inj(cons(a,A), cons(b,B))" unfolding inj_def apply (force intro: apply_type simp add: fun_extend) done end diff --git a/src/ZF/QUniv.thy b/src/ZF/QUniv.thy --- a/src/ZF/QUniv.thy +++ b/src/ZF/QUniv.thy @@ -1,203 +1,203 @@ (* Title: ZF/QUniv.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge *) section\A Small Universe for Lazy Recursive Types\ theory QUniv imports Univ QPair begin (*Disjoint sums as a datatype*) rep_datatype elimination sumE induction TrueI case_eqns case_Inl case_Inr (*Variant disjoint sums as a datatype*) rep_datatype elimination qsumE induction TrueI case_eqns qcase_QInl qcase_QInr definition quniv :: "i \ i" where "quniv(A) \ Pow(univ(eclose(A)))" subsection\Properties involving Transset and Sum\ lemma Transset_includes_summands: "\Transset(C); A+B \ C\ \ A \ C \ B \ C" apply (simp add: sum_def Un_subset_iff) apply (blast dest: Transset_includes_range) done lemma Transset_sum_Int_subset: "Transset(C) \ (A+B) \ C \ (A \ C) + (B \ C)" apply (simp add: sum_def Int_Un_distrib2) apply (blast dest: Transset_Pair_D) done subsection\Introduction and Elimination Rules\ lemma qunivI: "X \ univ(eclose(A)) \ X \ quniv(A)" by (simp add: quniv_def) lemma qunivD: "X \ quniv(A) \ X \ univ(eclose(A))" by (simp add: quniv_def) lemma quniv_mono: "A<=B \ quniv(A) \ quniv(B)" unfolding quniv_def apply (erule eclose_mono [THEN univ_mono, THEN Pow_mono]) done subsection\Closure Properties\ lemma univ_eclose_subset_quniv: "univ(eclose(A)) \ quniv(A)" apply (simp add: quniv_def Transset_iff_Pow [symmetric]) apply (rule Transset_eclose [THEN Transset_univ]) done (*Key property for proving A_subset_quniv; requires eclose in definition of quniv*) lemma univ_subset_quniv: "univ(A) \ quniv(A)" apply (rule arg_subset_eclose [THEN univ_mono, THEN subset_trans]) apply (rule univ_eclose_subset_quniv) done lemmas univ_into_quniv = univ_subset_quniv [THEN subsetD] lemma Pow_univ_subset_quniv: "Pow(univ(A)) \ quniv(A)" unfolding quniv_def apply (rule arg_subset_eclose [THEN univ_mono, THEN Pow_mono]) done lemmas univ_subset_into_quniv = PowI [THEN Pow_univ_subset_quniv [THEN subsetD]] lemmas zero_in_quniv = zero_in_univ [THEN univ_into_quniv] lemmas one_in_quniv = one_in_univ [THEN univ_into_quniv] lemmas two_in_quniv = two_in_univ [THEN univ_into_quniv] lemmas A_subset_quniv = subset_trans [OF A_subset_univ univ_subset_quniv] lemmas A_into_quniv = A_subset_quniv [THEN subsetD] (*** univ(A) closure for Quine-inspired pairs and injections ***) (*Quine ordered pairs*) lemma QPair_subset_univ: "\a \ univ(A); b \ univ(A)\ \ \ univ(A)" by (simp add: QPair_def sum_subset_univ) subsection\Quine Disjoint Sum\ lemma QInl_subset_univ: "a \ univ(A) \ QInl(a) \ univ(A)" unfolding QInl_def apply (erule empty_subsetI [THEN QPair_subset_univ]) done lemmas naturals_subset_nat = Ord_nat [THEN Ord_is_Transset, unfolded Transset_def, THEN bspec] lemmas naturals_subset_univ = subset_trans [OF naturals_subset_nat nat_subset_univ] lemma QInr_subset_univ: "a \ univ(A) \ QInr(a) \ univ(A)" unfolding QInr_def apply (erule nat_1I [THEN naturals_subset_univ, THEN QPair_subset_univ]) done subsection\Closure for Quine-Inspired Products and Sums\ (*Quine ordered pairs*) lemma QPair_in_quniv: "\a: quniv(A); b: quniv(A)\ \ \ quniv(A)" by (simp add: quniv_def QPair_def sum_subset_univ) lemma QSigma_quniv: "quniv(A) <*> quniv(A) \ quniv(A)" by (blast intro: QPair_in_quniv) lemmas QSigma_subset_quniv = subset_trans [OF QSigma_mono QSigma_quniv] (*The opposite inclusion*) lemma quniv_QPair_D: " \ quniv(A) \ a: quniv(A) \ b: quniv(A)" -apply (unfold quniv_def QPair_def) + unfolding quniv_def QPair_def apply (rule Transset_includes_summands [THEN conjE]) apply (rule Transset_eclose [THEN Transset_univ]) apply (erule PowD, blast) done lemmas quniv_QPair_E = quniv_QPair_D [THEN conjE] lemma quniv_QPair_iff: " \ quniv(A) \ a: quniv(A) \ b: quniv(A)" by (blast intro: QPair_in_quniv dest: quniv_QPair_D) subsection\Quine Disjoint Sum\ lemma QInl_in_quniv: "a: quniv(A) \ QInl(a) \ quniv(A)" by (simp add: QInl_def zero_in_quniv QPair_in_quniv) lemma QInr_in_quniv: "b: quniv(A) \ QInr(b) \ quniv(A)" by (simp add: QInr_def one_in_quniv QPair_in_quniv) lemma qsum_quniv: "quniv(C) <+> quniv(C) \ quniv(C)" by (blast intro: QInl_in_quniv QInr_in_quniv) lemmas qsum_subset_quniv = subset_trans [OF qsum_mono qsum_quniv] subsection\The Natural Numbers\ lemmas nat_subset_quniv = subset_trans [OF nat_subset_univ univ_subset_quniv] (* n:nat \ n:quniv(A) *) lemmas nat_into_quniv = nat_subset_quniv [THEN subsetD] lemmas bool_subset_quniv = subset_trans [OF bool_subset_univ univ_subset_quniv] lemmas bool_into_quniv = bool_subset_quniv [THEN subsetD] (*Intersecting with Vfrom...*) lemma QPair_Int_Vfrom_succ_subset: "Transset(X) \ \ Vfrom(X, succ(i)) \ Vfrom(X,i); b \ Vfrom(X,i)>" by (simp add: QPair_def sum_def Int_Un_distrib2 Un_mono product_Int_Vfrom_subset [THEN subset_trans] Sigma_mono [OF Int_lower1 subset_refl]) subsection\"Take-Lemma" Rules\ (*for proving a=b by coinduction and c: quniv(A)*) (*Rule for level i -- preserving the level, not decreasing it*) lemma QPair_Int_Vfrom_subset: "Transset(X) \ \ Vfrom(X,i) \ Vfrom(X,i); b \ Vfrom(X,i)>" unfolding QPair_def apply (erule Transset_Vfrom [THEN Transset_sum_Int_subset]) done (*@{term"\a \ Vset(i) \ c; b \ Vset(i) \ d\ \ \ Vset(i) \ "}*) lemmas QPair_Int_Vset_subset_trans = subset_trans [OF Transset_0 [THEN QPair_Int_Vfrom_subset] QPair_mono] lemma QPair_Int_Vset_subset_UN: "Ord(i) \ \ Vset(i) \ (\j\i. Vset(j); b \ Vset(j)>)" apply (erule Ord_cases) (*0 case*) apply (simp add: Vfrom_0) (*succ(j) case*) apply (erule ssubst) apply (rule Transset_0 [THEN QPair_Int_Vfrom_succ_subset, THEN subset_trans]) apply (rule succI1 [THEN UN_upper]) (*Limit(i) case*) apply (simp del: UN_simps add: Limit_Vfrom_eq Int_UN_distrib UN_mono QPair_Int_Vset_subset_trans) done end diff --git a/src/ZF/Resid/Confluence.thy b/src/ZF/Resid/Confluence.thy --- a/src/ZF/Resid/Confluence.thy +++ b/src/ZF/Resid/Confluence.thy @@ -1,118 +1,118 @@ (* Title: ZF/Resid/Confluence.thy Author: Ole Rasmussen Copyright 1995 University of Cambridge *) theory Confluence imports Reduction begin definition confluence :: "i\o" where "confluence(R) \ \x y. \x,y\ \ R \ (\z.\x,z\ \ R \ (\u.\y,u\ \ R \ \z,u\ \ R))" definition strip :: "o" where "strip \ \x y. (x =\ y) \ (\z.(x =1\ z) \ (\u.(y =1\ u) \ (z=\u)))" (* ------------------------------------------------------------------------- *) (* strip lemmas *) (* ------------------------------------------------------------------------- *) lemma strip_lemma_r: "\confluence(Spar_red1)\\ strip" -apply (unfold confluence_def strip_def) + unfolding confluence_def strip_def apply (rule impI [THEN allI, THEN allI]) apply (erule Spar_red.induct, fast) apply (fast intro: Spar_red.trans) done lemma strip_lemma_l: "strip\ confluence(Spar_red)" -apply (unfold confluence_def strip_def) + unfolding confluence_def strip_def apply (rule impI [THEN allI, THEN allI]) apply (erule Spar_red.induct, blast) apply clarify apply (blast intro: Spar_red.trans) done (* ------------------------------------------------------------------------- *) (* Confluence *) (* ------------------------------------------------------------------------- *) lemma parallel_moves: "confluence(Spar_red1)" apply (unfold confluence_def, clarify) apply (frule simulation) apply (frule_tac n = z in simulation, clarify) apply (frule_tac v = va in paving) apply (force intro: completeness)+ done lemmas confluence_parallel_reduction = parallel_moves [THEN strip_lemma_r, THEN strip_lemma_l] lemma lemma1: "\confluence(Spar_red)\\ confluence(Sred)" by (unfold confluence_def, blast intro: par_red_red red_par_red) lemmas confluence_beta_reduction = confluence_parallel_reduction [THEN lemma1] (**** Conversion ****) consts Sconv1 :: "i" Sconv :: "i" abbreviation Sconv1_rel (infixl \<-1->\ 50) where "a<-1->b \ \a,b\ \ Sconv1" abbreviation Sconv_rel (infixl \<-\\ 50) where "a<-\b \ \a,b\ \ Sconv" inductive domains "Sconv1" \ "lambda*lambda" intros red1: "m -1-> n \ m<-1->n" expl: "n -1-> m \ m<-1->n" type_intros red1D1 red1D2 lambda.intros bool_typechecks declare Sconv1.intros [intro] inductive domains "Sconv" \ "lambda*lambda" intros one_step: "m<-1->n \ m<-\n" refl: "m \ lambda \ m<-\m" trans: "\m<-\n; n<-\p\ \ m<-\p" type_intros Sconv1.dom_subset [THEN subsetD] lambda.intros bool_typechecks declare Sconv.intros [intro] lemma conv_sym: "m<-\n \ n<-\m" apply (erule Sconv.induct) apply (erule Sconv1.induct, blast+) done (* ------------------------------------------------------------------------- *) (* Church_Rosser Theorem *) (* ------------------------------------------------------------------------- *) lemma Church_Rosser: "m<-\n \ \p.(m -\p) \ (n -\ p)" apply (erule Sconv.induct) apply (erule Sconv1.induct) apply (blast intro: red1D1 redD2) apply (blast intro: red1D1 redD2) apply (blast intro: red1D1 redD2) apply (cut_tac confluence_beta_reduction) unfolding confluence_def apply (blast intro: Sred.trans) done end diff --git a/src/ZF/Trancl.thy b/src/ZF/Trancl.thy --- a/src/ZF/Trancl.thy +++ b/src/ZF/Trancl.thy @@ -1,374 +1,374 @@ (* Title: ZF/Trancl.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section\Relations: Their General Properties and Transitive Closure\ theory Trancl imports Fixedpt Perm begin definition refl :: "[i,i]\o" where "refl(A,r) \ (\x\A. \x,x\ \ r)" definition irrefl :: "[i,i]\o" where "irrefl(A,r) \ \x\A. \x,x\ \ r" definition sym :: "i\o" where "sym(r) \ \x y. \x,y\: r \ \y,x\: r" definition asym :: "i\o" where "asym(r) \ \x y. \x,y\:r \ \ \y,x\:r" definition antisym :: "i\o" where "antisym(r) \ \x y.\x,y\:r \ \y,x\:r \ x=y" definition trans :: "i\o" where "trans(r) \ \x y z. \x,y\: r \ \y,z\: r \ \x,z\: r" definition trans_on :: "[i,i]\o" (\trans[_]'(_')\) where "trans[A](r) \ \x\A. \y\A. \z\A. \x,y\: r \ \y,z\: r \ \x,z\: r" definition rtrancl :: "i\i" (\(_^*)\ [100] 100) (*refl/transitive closure*) where "r^* \ lfp(field(r)*field(r), \s. id(field(r)) \ (r O s))" definition trancl :: "i\i" (\(_^+)\ [100] 100) (*transitive closure*) where "r^+ \ r O r^*" definition equiv :: "[i,i]\o" where "equiv(A,r) \ r \ A*A \ refl(A,r) \ sym(r) \ trans(r)" subsection\General properties of relations\ subsubsection\irreflexivity\ lemma irreflI: "\\x. x \ A \ \x,x\ \ r\ \ irrefl(A,r)" by (simp add: irrefl_def) lemma irreflE: "\irrefl(A,r); x \ A\ \ \x,x\ \ r" by (simp add: irrefl_def) subsubsection\symmetry\ lemma symI: "\\x y.\x,y\: r \ \y,x\: r\ \ sym(r)" by (unfold sym_def, blast) lemma symE: "\sym(r); \x,y\: r\ \ \y,x\: r" by (unfold sym_def, blast) subsubsection\antisymmetry\ lemma antisymI: "\\x y.\\x,y\: r; \y,x\: r\ \ x=y\ \ antisym(r)" by (simp add: antisym_def, blast) lemma antisymE: "\antisym(r); \x,y\: r; \y,x\: r\ \ x=y" by (simp add: antisym_def, blast) subsubsection\transitivity\ lemma transD: "\trans(r); \a,b\:r; \b,c\:r\ \ \a,c\:r" by (unfold trans_def, blast) lemma trans_onD: "\trans[A](r); \a,b\:r; \b,c\:r; a \ A; b \ A; c \ A\ \ \a,c\:r" by (unfold trans_on_def, blast) lemma trans_imp_trans_on: "trans(r) \ trans[A](r)" by (unfold trans_def trans_on_def, blast) lemma trans_on_imp_trans: "\trans[A](r); r \ A*A\ \ trans(r)" by (simp add: trans_on_def trans_def, blast) subsection\Transitive closure of a relation\ lemma rtrancl_bnd_mono: "bnd_mono(field(r)*field(r), \s. id(field(r)) \ (r O s))" by (rule bnd_monoI, blast+) lemma rtrancl_mono: "r<=s \ r^* \ s^*" unfolding rtrancl_def apply (rule lfp_mono) apply (rule rtrancl_bnd_mono)+ apply blast done (* @{term"r^* = id(field(r)) \ ( r O r^* )"} *) lemmas rtrancl_unfold = rtrancl_bnd_mono [THEN rtrancl_def [THEN def_lfp_unfold]] (** The relation rtrancl **) (* @{term"r^* \ field(r) * field(r)"} *) lemmas rtrancl_type = rtrancl_def [THEN def_lfp_subset] lemma relation_rtrancl: "relation(r^*)" apply (simp add: relation_def) apply (blast dest: rtrancl_type [THEN subsetD]) done (*Reflexivity of rtrancl*) lemma rtrancl_refl: "\a \ field(r)\ \ \a,a\ \ r^*" apply (rule rtrancl_unfold [THEN ssubst]) apply (erule idI [THEN UnI1]) done (*Closure under composition with r *) lemma rtrancl_into_rtrancl: "\\a,b\ \ r^*; \b,c\ \ r\ \ \a,c\ \ r^*" apply (rule rtrancl_unfold [THEN ssubst]) apply (rule compI [THEN UnI2], assumption, assumption) done (*rtrancl of r contains all pairs in r *) lemma r_into_rtrancl: "\a,b\ \ r \ \a,b\ \ r^*" by (rule rtrancl_refl [THEN rtrancl_into_rtrancl], blast+) (*The premise ensures that r consists entirely of pairs*) lemma r_subset_rtrancl: "relation(r) \ r \ r^*" by (simp add: relation_def, blast intro: r_into_rtrancl) lemma rtrancl_field: "field(r^*) = field(r)" by (blast intro: r_into_rtrancl dest!: rtrancl_type [THEN subsetD]) (** standard induction rule **) lemma rtrancl_full_induct [case_names initial step, consumes 1]: "\\a,b\ \ r^*; \x. x \ field(r) \ P(\x,x\); \x y z.\P(\x,y\); \x,y\: r^*; \y,z\: r\ \ P(\x,z\)\ \ P(\a,b\)" by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast) (*nice induction rule. Tried adding the typing hypotheses y,z \ field(r), but these caused expensive case splits!*) lemma rtrancl_induct [case_names initial step, induct set: rtrancl]: "\\a,b\ \ r^*; P(a); \y z.\\a,y\ \ r^*; \y,z\ \ r; P(y)\ \ P(z) \ \ P(b)" (*by induction on this formula*) apply (subgoal_tac "\y. \a,b\ = \a,y\ \ P (y) ") (*now solve first subgoal: this formula is sufficient*) apply (erule spec [THEN mp], rule refl) (*now do the induction*) apply (erule rtrancl_full_induct, blast+) done (*transitivity of transitive closure\-- by induction.*) lemma trans_rtrancl: "trans(r^*)" unfolding trans_def apply (intro allI impI) apply (erule_tac b = z in rtrancl_induct, assumption) apply (blast intro: rtrancl_into_rtrancl) done lemmas rtrancl_trans = trans_rtrancl [THEN transD] (*elimination of rtrancl -- by induction on a special formula*) lemma rtranclE: "\\a,b\ \ r^*; (a=b) \ P; \y.\\a,y\ \ r^*; \y,b\ \ r\ \ P\ \ P" apply (subgoal_tac "a = b | (\y. \a,y\ \ r^* \ \y,b\ \ r) ") (*see HOL/trancl*) apply blast apply (erule rtrancl_induct, blast+) done (**** The relation trancl ****) (*Transitivity of r^+ is proved by transitivity of r^* *) lemma trans_trancl: "trans(r^+)" -apply (unfold trans_def trancl_def) + unfolding trans_def trancl_def apply (blast intro: rtrancl_into_rtrancl trans_rtrancl [THEN transD, THEN compI]) done lemmas trans_on_trancl = trans_trancl [THEN trans_imp_trans_on] lemmas trancl_trans = trans_trancl [THEN transD] (** Conversions between trancl and rtrancl **) lemma trancl_into_rtrancl: "\a,b\ \ r^+ \ \a,b\ \ r^*" unfolding trancl_def apply (blast intro: rtrancl_into_rtrancl) done (*r^+ contains all pairs in r *) lemma r_into_trancl: "\a,b\ \ r \ \a,b\ \ r^+" unfolding trancl_def apply (blast intro!: rtrancl_refl) done (*The premise ensures that r consists entirely of pairs*) lemma r_subset_trancl: "relation(r) \ r \ r^+" by (simp add: relation_def, blast intro: r_into_trancl) (*intro rule by definition: from r^* and r *) lemma rtrancl_into_trancl1: "\\a,b\ \ r^*; \b,c\ \ r\ \ \a,c\ \ r^+" by (unfold trancl_def, blast) (*intro rule from r and r^* *) lemma rtrancl_into_trancl2: "\\a,b\ \ r; \b,c\ \ r^*\ \ \a,c\ \ r^+" apply (erule rtrancl_induct) apply (erule r_into_trancl) apply (blast intro: r_into_trancl trancl_trans) done (*Nice induction rule for trancl*) lemma trancl_induct [case_names initial step, induct set: trancl]: "\\a,b\ \ r^+; \y. \\a,y\ \ r\ \ P(y); \y z.\\a,y\ \ r^+; \y,z\ \ r; P(y)\ \ P(z) \ \ P(b)" apply (rule compEpair) apply (unfold trancl_def, assumption) (*by induction on this formula*) apply (subgoal_tac "\z. \y,z\ \ r \ P (z) ") (*now solve first subgoal: this formula is sufficient*) apply blast apply (erule rtrancl_induct) apply (blast intro: rtrancl_into_trancl1)+ done (*elimination of r^+ -- NOT an induction rule*) lemma tranclE: "\\a,b\ \ r^+; \a,b\ \ r \ P; \y.\\a,y\ \ r^+; \y,b\ \ r\ \ P \ \ P" apply (subgoal_tac "\a,b\ \ r | (\y. \a,y\ \ r^+ \ \y,b\ \ r) ") apply blast apply (rule compEpair) apply (unfold trancl_def, assumption) apply (erule rtranclE) apply (blast intro: rtrancl_into_trancl1)+ done lemma trancl_type: "r^+ \ field(r)*field(r)" unfolding trancl_def apply (blast elim: rtrancl_type [THEN subsetD, THEN SigmaE2]) done lemma relation_trancl: "relation(r^+)" apply (simp add: relation_def) apply (blast dest: trancl_type [THEN subsetD]) done lemma trancl_subset_times: "r \ A * A \ r^+ \ A * A" by (insert trancl_type [of r], blast) lemma trancl_mono: "r<=s \ r^+ \ s^+" by (unfold trancl_def, intro comp_mono rtrancl_mono) lemma trancl_eq_r: "\relation(r); trans(r)\ \ r^+ = r" apply (rule equalityI) prefer 2 apply (erule r_subset_trancl, clarify) apply (frule trancl_type [THEN subsetD], clarify) apply (erule trancl_induct, assumption) apply (blast dest: transD) done (** Suggested by Sidi Ould Ehmety **) lemma rtrancl_idemp [simp]: "(r^*)^* = r^*" apply (rule equalityI, auto) prefer 2 apply (frule rtrancl_type [THEN subsetD]) apply (blast intro: r_into_rtrancl ) txt\converse direction\ apply (frule rtrancl_type [THEN subsetD], clarify) apply (erule rtrancl_induct) apply (simp add: rtrancl_refl rtrancl_field) apply (blast intro: rtrancl_trans) done lemma rtrancl_subset: "\R \ S; S \ R^*\ \ S^* = R^*" apply (drule rtrancl_mono) apply (drule rtrancl_mono, simp_all, blast) done lemma rtrancl_Un_rtrancl: "\relation(r); relation(s)\ \ (r^* \ s^*)^* = (r \ s)^*" apply (rule rtrancl_subset) apply (blast dest: r_subset_rtrancl) apply (blast intro: rtrancl_mono [THEN subsetD]) done (*** "converse" laws by Sidi Ould Ehmety ***) (** rtrancl **) lemma rtrancl_converseD: "\x,y\:converse(r)^* \ \x,y\:converse(r^*)" apply (rule converseI) apply (frule rtrancl_type [THEN subsetD]) apply (erule rtrancl_induct) apply (blast intro: rtrancl_refl) apply (blast intro: r_into_rtrancl rtrancl_trans) done lemma rtrancl_converseI: "\x,y\:converse(r^*) \ \x,y\:converse(r)^*" apply (drule converseD) apply (frule rtrancl_type [THEN subsetD]) apply (erule rtrancl_induct) apply (blast intro: rtrancl_refl) apply (blast intro: r_into_rtrancl rtrancl_trans) done lemma rtrancl_converse: "converse(r)^* = converse(r^*)" apply (safe intro!: equalityI) apply (frule rtrancl_type [THEN subsetD]) apply (safe dest!: rtrancl_converseD intro!: rtrancl_converseI) done (** trancl **) lemma trancl_converseD: "\a, b\:converse(r)^+ \ \a, b\:converse(r^+)" apply (erule trancl_induct) apply (auto intro: r_into_trancl trancl_trans) done lemma trancl_converseI: "\x,y\:converse(r^+) \ \x,y\:converse(r)^+" apply (drule converseD) apply (erule trancl_induct) apply (auto intro: r_into_trancl trancl_trans) done lemma trancl_converse: "converse(r)^+ = converse(r^+)" apply (safe intro!: equalityI) apply (frule trancl_type [THEN subsetD]) apply (safe dest!: trancl_converseD intro!: trancl_converseI) done lemma converse_trancl_induct [case_names initial step, consumes 1]: "\\a, b\:r^+; \y. \y, b\ :r \ P(y); \y z. \\y, z\ \ r; \z, b\ \ r^+; P(z)\ \ P(y)\ \ P(a)" apply (drule converseI) apply (simp (no_asm_use) add: trancl_converse [symmetric]) apply (erule trancl_induct) apply (auto simp add: trancl_converse) done end diff --git a/src/ZF/UNITY/AllocBase.thy b/src/ZF/UNITY/AllocBase.thy --- a/src/ZF/UNITY/AllocBase.thy +++ b/src/ZF/UNITY/AllocBase.thy @@ -1,398 +1,398 @@ (* Title: ZF/UNITY/AllocBase.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2001 University of Cambridge *) section\Common declarations for Chandy and Charpentier's Allocator\ theory AllocBase imports Follows MultisetSum Guar begin abbreviation (input) tokbag :: i (* tokbags could be multisets...or any ordered type?*) where "tokbag \ nat" axiomatization NbT :: i and (* Number of tokens in system *) Nclients :: i (* Number of clients *) where NbT_pos: "NbT \ nat-{0}" and Nclients_pos: "Nclients \ nat-{0}" text\This function merely sums the elements of a list\ consts tokens :: "i \i" item :: i (* Items to be merged/distributed *) primrec "tokens(Nil) = 0" "tokens (Cons(x,xs)) = x #+ tokens(xs)" consts bag_of :: "i \ i" primrec "bag_of(Nil) = 0" "bag_of(Cons(x,xs)) = {#x#} +# bag_of(xs)" text\Definitions needed in Client.thy. We define a recursive predicate using 0 and 1 to code the truth values.\ consts all_distinct0 :: "i\i" primrec "all_distinct0(Nil) = 1" "all_distinct0(Cons(a, l)) = (if a \ set_of_list(l) then 0 else all_distinct0(l))" definition all_distinct :: "i\o" where "all_distinct(l) \ all_distinct0(l)=1" definition state_of :: "i \i" \ \coersion from anyting to state\ where "state_of(s) \ if s \ state then s else st0" definition lift :: "i \(i\i)" \ \simplifies the expression of programs\ where "lift(x) \ \s. s`x" text\function to show that the set of variables is infinite\ consts nat_list_inj :: "i\i" var_inj :: "i\i" primrec "nat_list_inj(0) = Nil" "nat_list_inj(succ(n)) = Cons(n, nat_list_inj(n))" primrec "var_inj(Var(l)) = length(l)" definition nat_var_inj :: "i\i" where "nat_var_inj(n) \ Var(nat_list_inj(n))" subsection\Various simple lemmas\ lemma Nclients_NbT_gt_0 [simp]: "0 < Nclients \ 0 < NbT" apply (cut_tac Nclients_pos NbT_pos) apply (auto intro: Ord_0_lt) done lemma Nclients_NbT_not_0 [simp]: "Nclients \ 0 \ NbT \ 0" by (cut_tac Nclients_pos NbT_pos, auto) lemma Nclients_type [simp,TC]: "Nclients\nat" by (cut_tac Nclients_pos NbT_pos, auto) lemma NbT_type [simp,TC]: "NbT\nat" by (cut_tac Nclients_pos NbT_pos, auto) lemma INT_Nclient_iff [iff]: "b\\(RepFun(Nclients, B)) \ (\x\Nclients. b\B(x))" by (force simp add: INT_iff) lemma setsum_fun_mono [rule_format]: "n\nat \ (\i\nat. i f(i) $\ g(i)) \ setsum(f, n) $\ setsum(g,n)" apply (induct_tac "n", simp_all) apply (subgoal_tac "Finite(x) \ x\x") prefer 2 apply (simp add: nat_into_Finite mem_not_refl, clarify) apply (simp (no_asm_simp) add: succ_def) apply (subgoal_tac "\i\nat. i f(i) $\ g(i) ") prefer 2 apply (force dest: leI) apply (rule zadd_zle_mono, simp_all) done lemma tokens_type [simp,TC]: "l\list(A) \ tokens(l)\nat" by (erule list.induct, auto) lemma tokens_mono_aux [rule_format]: "xs\list(A) \ \ys\list(A). \xs, ys\\prefix(A) \ tokens(xs) \ tokens(ys)" apply (induct_tac "xs") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: prefix_def) done lemma tokens_mono: "\xs, ys\\prefix(A) \ tokens(xs) \ tokens(ys)" apply (cut_tac prefix_type) apply (blast intro: tokens_mono_aux) done lemma mono_tokens [iff]: "mono1(list(A), prefix(A), nat, Le,tokens)" unfolding mono1_def apply (auto intro: tokens_mono simp add: Le_def) done lemma tokens_append [simp]: "\xs\list(A); ys\list(A)\ \ tokens(xs@ys) = tokens(xs) #+ tokens(ys)" apply (induct_tac "xs", auto) done subsection\The function \<^term>\bag_of\\ lemma bag_of_type [simp,TC]: "l\list(A) \bag_of(l)\Mult(A)" apply (induct_tac "l") apply (auto simp add: Mult_iff_multiset) done lemma bag_of_multiset: "l\list(A) \ multiset(bag_of(l)) \ mset_of(bag_of(l))<=A" apply (drule bag_of_type) apply (auto simp add: Mult_iff_multiset) done lemma bag_of_append [simp]: "\xs\list(A); ys\list(A)\ \ bag_of(xs@ys) = bag_of(xs) +# bag_of(ys)" apply (induct_tac "xs") apply (auto simp add: bag_of_multiset munion_assoc) done lemma bag_of_mono_aux [rule_format]: "xs\list(A) \ \ys\list(A). \xs, ys\\prefix(A) \ \MultLe(A, r)" apply (induct_tac "xs", simp_all, clarify) apply (frule_tac l = ys in bag_of_multiset) apply (auto intro: empty_le_MultLe simp add: prefix_def) apply (rule munion_mono) apply (force simp add: MultLe_def Mult_iff_multiset) apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) done lemma bag_of_mono [intro]: "\\xs, ys\\prefix(A); xs\list(A); ys\list(A)\ \ \MultLe(A, r)" apply (blast intro: bag_of_mono_aux) done lemma mono_bag_of [simp]: "mono1(list(A), prefix(A), Mult(A), MultLe(A,r), bag_of)" by (auto simp add: mono1_def bag_of_type) subsection\The function \<^term>\msetsum\\ lemmas nat_into_Fin = eqpoll_refl [THEN [2] Fin_lemma] lemma list_Int_length_Fin: "l \ list(A) \ C \ length(l) \ Fin(length(l))" apply (drule length_type) apply (rule Fin_subset) apply (rule Int_lower2) apply (erule nat_into_Fin) done lemma mem_Int_imp_lt_length: "\xs \ list(A); k \ C \ length(xs)\ \ k < length(xs)" by (simp add: ltI) lemma Int_succ_right: "A \ succ(k) = (if k \ A then cons(k, A \ k) else A \ k)" by auto lemma bag_of_sublist_lemma: "\C \ nat; x \ A; xs \ list(A)\ \ msetsum(\i. {#nth(i, xs @ [x])#}, C \ succ(length(xs)), A) = (if length(xs) \ C then {#x#} +# msetsum(\x. {#nth(x, xs)#}, C \ length(xs), A) else msetsum(\x. {#nth(x, xs)#}, C \ length(xs), A))" apply (simp add: subsetD nth_append lt_not_refl mem_Int_imp_lt_length cong add: msetsum_cong) apply (simp add: Int_succ_right) apply (simp add: lt_not_refl mem_Int_imp_lt_length cong add: msetsum_cong, clarify) apply (subst msetsum_cons) apply (rule_tac [3] succI1) apply (blast intro: list_Int_length_Fin subset_succI [THEN Fin_mono, THEN subsetD]) apply (simp add: mem_not_refl) apply (simp add: nth_type lt_not_refl) apply (blast intro: nat_into_Ord ltI length_type) apply (simp add: lt_not_refl mem_Int_imp_lt_length cong add: msetsum_cong) done lemma bag_of_sublist_lemma2: "l\list(A) \ C \ nat \ bag_of(sublist(l, C)) = msetsum(\i. {#nth(i, l)#}, C \ length(l), A)" apply (erule list_append_induct) apply (simp (no_asm)) apply (simp (no_asm_simp) add: sublist_append nth_append bag_of_sublist_lemma munion_commute bag_of_sublist_lemma msetsum_multiset munion_0) done lemma nat_Int_length_eq: "l \ list(A) \ nat \ length(l) = length(l)" apply (rule Int_absorb1) apply (rule OrdmemD, auto) done (*eliminating the assumption C<=nat*) lemma bag_of_sublist: "l\list(A) \ bag_of(sublist(l, C)) = msetsum(\i. {#nth(i, l)#}, C \ length(l), A)" apply (subgoal_tac " bag_of (sublist (l, C \ nat)) = msetsum (\i. {#nth (i, l) #}, C \ length (l), A) ") apply (simp add: sublist_Int_eq) apply (simp add: bag_of_sublist_lemma2 Int_lower2 Int_assoc nat_Int_length_eq) done lemma bag_of_sublist_Un_Int: "l\list(A) \ bag_of(sublist(l, B \ C)) +# bag_of(sublist(l, B \ C)) = bag_of(sublist(l, B)) +# bag_of(sublist(l, C))" apply (subgoal_tac "B \ C \ length (l) = (B \ length (l)) \ (C \ length (l))") prefer 2 apply blast apply (simp (no_asm_simp) add: bag_of_sublist Int_Un_distrib2 msetsum_Un_Int) apply (rule msetsum_Un_Int) apply (erule list_Int_length_Fin)+ apply (simp add: ltI nth_type) done lemma bag_of_sublist_Un_disjoint: "\l\list(A); B \ C = 0\ \ bag_of(sublist(l, B \ C)) = bag_of(sublist(l, B)) +# bag_of(sublist(l, C))" by (simp add: bag_of_sublist_Un_Int [symmetric] bag_of_multiset) lemma bag_of_sublist_UN_disjoint [rule_format]: "\Finite(I); \i\I. \j\I. i\j \ A(i) \ A(j) = 0; l\list(B)\ \ bag_of(sublist(l, \i\I. A(i))) = (msetsum(\i. bag_of(sublist(l, A(i))), I, B)) " apply (simp (no_asm_simp) del: UN_simps add: UN_simps [symmetric] bag_of_sublist) apply (subst msetsum_UN_disjoint [of _ _ _ "length (l)"]) apply (drule Finite_into_Fin, assumption) prefer 3 apply force apply (auto intro!: Fin_IntI2 Finite_into_Fin simp add: ltI nth_type length_type nat_into_Finite) done lemma part_ord_Lt [simp]: "part_ord(nat, Lt)" -apply (unfold part_ord_def Lt_def irrefl_def trans_on_def) + unfolding part_ord_def Lt_def irrefl_def trans_on_def apply (auto intro: lt_trans) done subsubsection\The function \<^term>\all_distinct\\ lemma all_distinct_Nil [simp]: "all_distinct(Nil)" by (unfold all_distinct_def, auto) lemma all_distinct_Cons [simp]: "all_distinct(Cons(a,l)) \ (a\set_of_list(l) \ False) \ (a \ set_of_list(l) \ all_distinct(l))" unfolding all_distinct_def apply (auto elim: list.cases) done subsubsection\The function \<^term>\state_of\\ lemma state_of_state: "s\state \ state_of(s)=s" by (unfold state_of_def, auto) declare state_of_state [simp] lemma state_of_idem: "state_of(state_of(s))=state_of(s)" apply (unfold state_of_def, auto) done declare state_of_idem [simp] lemma state_of_type [simp,TC]: "state_of(s)\state" by (unfold state_of_def, auto) lemma lift_apply [simp]: "lift(x, s)=s`x" by (simp add: lift_def) (** Used in ClientImp **) lemma gen_Increains_state_of_eq: "Increasing(A, r, \s. f(state_of(s))) = Increasing(A, r, f)" apply (unfold Increasing_def, auto) done lemmas Increasing_state_ofD1 = gen_Increains_state_of_eq [THEN equalityD1, THEN subsetD] lemmas Increasing_state_ofD2 = gen_Increains_state_of_eq [THEN equalityD2, THEN subsetD] lemma Follows_state_of_eq: "Follows(A, r, \s. f(state_of(s)), \s. g(state_of(s))) = Follows(A, r, f, g)" apply (unfold Follows_def Increasing_def, auto) done lemmas Follows_state_ofD1 = Follows_state_of_eq [THEN equalityD1, THEN subsetD] lemmas Follows_state_ofD2 = Follows_state_of_eq [THEN equalityD2, THEN subsetD] lemma nat_list_inj_type: "n\nat \ nat_list_inj(n)\list(nat)" by (induct_tac "n", auto) lemma length_nat_list_inj: "n\nat \ length(nat_list_inj(n)) = n" by (induct_tac "n", auto) lemma var_infinite_lemma: "(\x\nat. nat_var_inj(x))\inj(nat, var)" unfolding nat_var_inj_def apply (rule_tac d = var_inj in lam_injective) apply (auto simp add: var.intros nat_list_inj_type) apply (simp add: length_nat_list_inj) done lemma nat_lepoll_var: "nat \ var" unfolding lepoll_def apply (rule_tac x = " (\x\nat. nat_var_inj (x))" in exI) apply (rule var_infinite_lemma) done lemma var_not_Finite: "\Finite(var)" apply (insert nat_not_Finite) apply (blast intro: lepoll_Finite [OF nat_lepoll_var]) done lemma not_Finite_imp_exist: "\Finite(A) \ \x. x\A" apply (subgoal_tac "A\0") apply (auto simp add: Finite_0) done lemma Inter_Diff_var_iff: "Finite(A) \ b\(\(RepFun(var-A, B))) \ (\x\var-A. b\B(x))" apply (subgoal_tac "\x. x\var-A", auto) apply (subgoal_tac "\Finite (var-A) ") apply (drule not_Finite_imp_exist, auto) apply (cut_tac var_not_Finite) apply (erule swap) apply (rule_tac B = A in Diff_Finite, auto) done lemma Inter_var_DiffD: "\b\\(RepFun(var-A, B)); Finite(A); x\var-A\ \ b\B(x)" by (simp add: Inter_Diff_var_iff) (* \Finite(A); (\x\var-A. b\B(x))\ \ b\\(RepFun(var-A, B)) *) lemmas Inter_var_DiffI = Inter_Diff_var_iff [THEN iffD2] declare Inter_var_DiffI [intro!] lemma Acts_subset_Int_Pow_simp [simp]: "Acts(F)<= A \ Pow(state*state) \ Acts(F)<=A" by (insert Acts_type [of F], auto) lemma setsum_nsetsum_eq: "\Finite(A); \x\A. g(x)\nat\ \ setsum(\x. $#(g(x)), A) = $# nsetsum(\x. g(x), A)" apply (erule Finite_induct) apply (auto simp add: int_of_add) done lemma nsetsum_cong: "\A=B; \x\A. f(x)=g(x); \x\A. g(x)\nat; Finite(A)\ \ nsetsum(f, A) = nsetsum(g, B)" apply (subgoal_tac "$# nsetsum (f, A) = $# nsetsum (g, B)", simp) apply (simp add: setsum_nsetsum_eq [symmetric] cong: setsum_cong) done end diff --git a/src/ZF/UNITY/ClientImpl.thy b/src/ZF/UNITY/ClientImpl.thy --- a/src/ZF/UNITY/ClientImpl.thy +++ b/src/ZF/UNITY/ClientImpl.thy @@ -1,307 +1,307 @@ (* Title: ZF/UNITY/ClientImpl.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2002 University of Cambridge Distributed Resource Management System: Client Implementation. *) theory ClientImpl imports AllocBase Guar begin abbreviation "ask \ Var(Nil)" (* input history: tokens requested *) abbreviation "giv \ Var([0])" (* output history: tokens granted *) abbreviation "rel \ Var([1])" (* input history: tokens released *) abbreviation "tok \ Var([2])" (* the number of available tokens *) axiomatization where type_assumes: "type_of(ask) = list(tokbag) \ type_of(giv) = list(tokbag) \ type_of(rel) = list(tokbag) \ type_of(tok) = nat" and default_val_assumes: "default_val(ask) = Nil \ default_val(giv) = Nil \ default_val(rel) = Nil \ default_val(tok) = 0" (*Array indexing is translated to list indexing as A[n] \ nth(n-1,A). *) definition (** Release some client_tokens **) "client_rel_act \ {\s,t\ \ state*state. \nrel \ nat. nrel = length(s`rel) \ t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) \ nrel < length(s`giv) \ nth(nrel, s`ask) \ nth(nrel, s`giv)}" (** Choose a new token requirement **) (** Including t=s suppresses fairness, allowing the non-trivial part of the action to be ignored **) definition "client_tok_act \ {\s,t\ \ state*state. t=s | t = s(tok:=succ(s`tok mod NbT))}" definition "client_ask_act \ {\s,t\ \ state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" definition "client_prog \ mk_program({s \ state. s`tok \ NbT \ s`giv = Nil \ s`ask = Nil \ s`rel = Nil}, {client_rel_act, client_tok_act, client_ask_act}, \G \ preserves(lift(rel)) Int preserves(lift(ask)) Int preserves(lift(tok)). Acts(G))" declare type_assumes [simp] default_val_assumes [simp] (* This part should be automated *) lemma ask_value_type [simp,TC]: "s \ state \ s`ask \ list(nat)" unfolding state_def apply (drule_tac a = ask in apply_type, auto) done lemma giv_value_type [simp,TC]: "s \ state \ s`giv \ list(nat)" unfolding state_def apply (drule_tac a = giv in apply_type, auto) done lemma rel_value_type [simp,TC]: "s \ state \ s`rel \ list(nat)" unfolding state_def apply (drule_tac a = rel in apply_type, auto) done lemma tok_value_type [simp,TC]: "s \ state \ s`tok \ nat" unfolding state_def apply (drule_tac a = tok in apply_type, auto) done (** The Client Program **) lemma client_type [simp,TC]: "client_prog \ program" unfolding client_prog_def apply (simp (no_asm)) done declare client_prog_def [THEN def_prg_Init, simp] declare client_prog_def [THEN def_prg_AllowedActs, simp] declare client_prog_def [program] declare client_rel_act_def [THEN def_act_simp, simp] declare client_tok_act_def [THEN def_act_simp, simp] declare client_ask_act_def [THEN def_act_simp, simp] lemma client_prog_ok_iff: "\G \ program. (client_prog ok G) \ (G \ preserves(lift(rel)) \ G \ preserves(lift(ask)) \ G \ preserves(lift(tok)) \ client_prog \ Allowed(G))" by (auto simp add: ok_iff_Allowed client_prog_def [THEN def_prg_Allowed]) lemma client_prog_preserves: "client_prog:(\x \ var-{ask, rel, tok}. preserves(lift(x)))" apply (rule Inter_var_DiffI, force) apply (rule ballI) apply (rule preservesI, safety, auto) done lemma preserves_lift_imp_stable: "G \ preserves(lift(ff)) \ G \ stable({s \ state. P(s`ff)})" apply (drule preserves_imp_stable) apply (simp add: lift_def) done lemma preserves_imp_prefix: "G \ preserves(lift(ff)) \ G \ stable({s \ state. \k, s`ff\ \ prefix(nat)})" by (erule preserves_lift_imp_stable) (*Safety property 1 \ ask, rel are increasing: (24) *) lemma client_prog_Increasing_ask_rel: "client_prog: program guarantees Incr(lift(ask)) \ Incr(lift(rel))" unfolding guar_def apply (auto intro!: increasing_imp_Increasing simp add: client_prog_ok_iff Increasing.increasing_def preserves_imp_prefix) apply (safety, force, force)+ done declare nth_append [simp] append_one_prefix [simp] lemma NbT_pos2: "0 the client never requests too many tokens. With no Substitution Axiom, we must prove the two invariants simultaneously. *) lemma ask_Bounded_lemma: "\client_prog ok G; G \ program\ \ client_prog \ G \ Always({s \ state. s`tok \ NbT} \ {s \ state. \elt \ set_of_list(s`ask). elt \ NbT})" apply (rotate_tac -1) apply (auto simp add: client_prog_ok_iff) apply (rule invariantI [THEN stable_Join_Always2], force) prefer 2 apply (fast intro: stable_Int preserves_lift_imp_stable, safety) apply (auto dest: ActsD) apply (cut_tac NbT_pos) apply (rule NbT_pos2 [THEN mod_less_divisor]) apply (auto dest: ActsD preserves_imp_eq simp add: set_of_list_append) done (* Export version, with no mention of tok in the postcondition, but unfortunately tok must be declared local.*) lemma client_prog_ask_Bounded: "client_prog \ program guarantees Always({s \ state. \elt \ set_of_list(s`ask). elt \ NbT})" apply (rule guaranteesI) apply (erule ask_Bounded_lemma [THEN Always_weaken], auto) done (*** Towards proving the liveness property ***) lemma client_prog_stable_rel_le_giv: "client_prog \ stable({s \ state. \ prefix(nat)})" by (safety, auto) lemma client_prog_Join_Stable_rel_le_giv: "\client_prog \ G \ Incr(lift(giv)); G \ preserves(lift(rel))\ \ client_prog \ G \ Stable({s \ state. \ prefix(nat)})" apply (rule client_prog_stable_rel_le_giv [THEN Increasing_preserves_Stable]) apply (auto simp add: lift_def) done lemma client_prog_Join_Always_rel_le_giv: "\client_prog \ G \ Incr(lift(giv)); G \ preserves(lift(rel))\ \ client_prog \ G \ Always({s \ state. \ prefix(nat)})" by (force intro!: AlwaysI client_prog_Join_Stable_rel_le_giv) lemma def_act_eq: "A \ {\s, t\ \ state*state. P(s, t)} \ A={\s, t\ \ state*state. P(s, t)}" by auto lemma act_subset: "A={\s,t\ \ state*state. P(s, t)} \ A<=state*state" by auto lemma transient_lemma: "client_prog \ transient({s \ state. s`rel = k \ \k, h\ \ strict_prefix(nat) \ \ prefix(nat) \ h pfixGe s`ask})" apply (rule_tac act = client_rel_act in transientI) apply (simp (no_asm) add: client_prog_def [THEN def_prg_Acts]) apply (simp (no_asm) add: client_rel_act_def [THEN def_act_eq, THEN act_subset]) apply (auto simp add: client_prog_def [THEN def_prg_Acts] domain_def) apply (rule ReplaceI) apply (rule_tac x = "x (rel:= x`rel @ [nth (length (x`rel), x`giv) ]) " in exI) apply (auto intro!: state_update_type app_type length_type nth_type, auto) apply (blast intro: lt_trans2 prefix_length_le strict_prefix_length_lt) apply (blast intro: lt_trans2 prefix_length_le strict_prefix_length_lt) apply (simp (no_asm_use) add: gen_prefix_iff_nth) apply (subgoal_tac "h \ list(nat)") apply (simp_all (no_asm_simp) add: prefix_type [THEN subsetD, THEN SigmaD1]) apply (auto simp add: prefix_def Ge_def) apply (drule strict_prefix_length_lt) apply (drule_tac x = "length (x`rel) " in spec) apply auto apply (simp (no_asm_use) add: gen_prefix_iff_nth) apply (auto simp add: id_def lam_def) done lemma strict_prefix_is_prefix: "\xs, ys\ \ strict_prefix(A) \ \xs, ys\ \ prefix(A) \ xs\ys" -apply (unfold strict_prefix_def id_def lam_def) + unfolding strict_prefix_def id_def lam_def apply (auto dest: prefix_type [THEN subsetD]) done lemma induct_lemma: "\client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program\ \ client_prog \ G \ {s \ state. s`rel = k \ \k,h\ \ strict_prefix(nat) \ \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ strict_prefix(nat) \ \ prefix(nat) \ \ prefix(nat) \ h pfixGe s`ask}" apply (rule single_LeadsTo_I) prefer 2 apply simp apply (frule client_prog_Increasing_ask_rel [THEN guaranteesD]) apply (rotate_tac [3] 2) apply (auto simp add: client_prog_ok_iff) apply (rule transient_lemma [THEN Join_transient_I1, THEN transient_imp_leadsTo, THEN leadsTo_imp_LeadsTo, THEN PSP_Stable, THEN LeadsTo_weaken]) apply (rule Stable_Int [THEN Stable_Int, THEN Stable_Int]) apply (erule_tac f = "lift (giv) " and a = "s`giv" in Increasing_imp_Stable) apply (simp (no_asm_simp)) apply (erule_tac f = "lift (ask) " and a = "s`ask" in Increasing_imp_Stable) apply (simp (no_asm_simp)) apply (erule_tac f = "lift (rel) " and a = "s`rel" in Increasing_imp_Stable) apply (simp (no_asm_simp)) apply (erule client_prog_Join_Stable_rel_le_giv, blast, simp_all) prefer 2 apply (blast intro: sym strict_prefix_is_prefix [THEN iffD2] prefix_trans prefix_imp_pfixGe pfixGe_trans) apply (auto intro: strict_prefix_is_prefix [THEN iffD1, THEN conjunct1] prefix_trans) done lemma rel_progress_lemma: "\client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program\ \ client_prog \ G \ {s \ state. \ strict_prefix(nat) \ \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ prefix(nat)}" apply (rule_tac f = "\x \ state. length(h) #- length(x`rel)" in LessThan_induct) apply (auto simp add: vimage_def) prefer 2 apply (force simp add: lam_def) apply (rule single_LeadsTo_I) prefer 2 apply simp apply (subgoal_tac "h \ list(nat)") prefer 2 apply (blast dest: prefix_type [THEN subsetD]) apply (rule induct_lemma [THEN LeadsTo_weaken]) apply (simp add: length_type lam_def) apply (auto intro: strict_prefix_is_prefix [THEN iffD2] dest: common_prefix_linear prefix_type [THEN subsetD]) apply (erule swap) apply (rule imageI) apply (force dest!: simp add: lam_def) apply (simp add: length_type lam_def, clarify) apply (drule strict_prefix_length_lt)+ apply (drule less_imp_succ_add, simp)+ apply clarify apply simp apply (erule diff_le_self [THEN ltD]) done lemma progress_lemma: "\client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program\ \ client_prog \ G \ {s \ state. \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ prefix(nat)}" apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI], assumption) apply (force simp add: client_prog_ok_iff) apply (rule LeadsTo_weaken_L) apply (rule LeadsTo_Un [OF rel_progress_lemma subset_refl [THEN subset_imp_LeadsTo]]) apply (auto intro: strict_prefix_is_prefix [THEN iffD2] dest: common_prefix_linear prefix_type [THEN subsetD]) done (*Progress property: all tokens that are given will be released*) lemma client_prog_progress: "client_prog \ Incr(lift(giv)) guarantees (\h \ list(nat). {s \ state. \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ prefix(nat)})" apply (rule guaranteesI) apply (blast intro: progress_lemma, auto) done lemma client_prog_Allowed: "Allowed(client_prog) = preserves(lift(rel)) \ preserves(lift(ask)) \ preserves(lift(tok))" apply (cut_tac v = "lift (ask)" in preserves_type) apply (auto simp add: Allowed_def client_prog_def [THEN def_prg_Allowed] cons_Int_distrib safety_prop_Acts_iff) done end diff --git a/src/ZF/UNITY/Comp.thy b/src/ZF/UNITY/Comp.thy --- a/src/ZF/UNITY/Comp.thy +++ b/src/ZF/UNITY/Comp.thy @@ -1,342 +1,342 @@ (* Title: ZF/UNITY/Comp.thy Author: Sidi O Ehmety, Computer Laboratory Copyright 1998 University of Cambridge From Chandy and Sanders, "Reasoning About Program Composition", Technical Report 2000-003, University of Florida, 2000. Revised by Sidi Ehmety on January 2001 Added: a strong form of the order relation over component and localize Theory ported from HOL. *) section\Composition\ theory Comp imports Union Increasing begin definition component :: "[i,i]\o" (infixl \component\ 65) where "F component H \ (\G. F \ G = H)" definition strict_component :: "[i,i]\o" (infixl \strict'_component\ 65) where "F strict_component H \ F component H \ F\H" definition (* A stronger form of the component relation *) component_of :: "[i,i]\o" (infixl \component'_of\ 65) where "F component_of H \ (\G. F ok G \ F \ G = H)" definition strict_component_of :: "[i,i]\o" (infixl \strict'_component'_of\ 65) where "F strict_component_of H \ F component_of H \ F\H" definition (* Preserves a state function f, in particular a variable *) preserves :: "(i\i)\i" where "preserves(f) \ {F:program. \z. F: stable({s:state. f(s) = z})}" definition fun_pair :: "[i\i, i \i] \(i\i)" where "fun_pair(f, g) \ \x. " definition localize :: "[i\i, i] \ i" where "localize(f,F) \ mk_program(Init(F), Acts(F), AllowedActs(F) \ (\G\preserves(f). Acts(G)))" (*** component and strict_component relations ***) lemma componentI: "H component F | H component G \ H component (F \ G)" apply (unfold component_def, auto) apply (rule_tac x = "Ga \ G" in exI) apply (rule_tac [2] x = "G \ F" in exI) apply (auto simp add: Join_ac) done lemma component_eq_subset: "G \ program \ (F component G) \ (Init(G) \ Init(F) \ Acts(F) \ Acts(G) \ AllowedActs(G) \ AllowedActs(F))" apply (unfold component_def, auto) apply (rule exI) apply (rule program_equalityI, auto) done lemma component_SKIP [simp]: "F \ program \ SKIP component F" unfolding component_def apply (rule_tac x = F in exI) apply (force intro: Join_SKIP_left) done lemma component_refl [simp]: "F \ program \ F component F" unfolding component_def apply (rule_tac x = F in exI) apply (force intro: Join_SKIP_right) done lemma SKIP_minimal: "F component SKIP \ programify(F) = SKIP" apply (rule program_equalityI) apply (simp_all add: component_eq_subset, blast) done lemma component_Join1: "F component (F \ G)" by (unfold component_def, blast) lemma component_Join2: "G component (F \ G)" unfolding component_def apply (simp (no_asm) add: Join_commute) apply blast done lemma Join_absorb1: "F component G \ F \ G = G" by (auto simp add: component_def Join_left_absorb) lemma Join_absorb2: "G component F \ F \ G = F" by (auto simp add: Join_ac component_def) lemma JOIN_component_iff: "H \ program\(JOIN(I,F) component H) \ (\i \ I. F(i) component H)" apply (case_tac "I=0", force) apply (simp (no_asm_simp) add: component_eq_subset) apply auto apply blast apply (rename_tac "y") apply (drule_tac c = y and A = "AllowedActs (H)" in subsetD) apply (blast elim!: not_emptyE)+ done lemma component_JOIN: "i \ I \ F(i) component (\i \ I. (F(i)))" unfolding component_def apply (blast intro: JOIN_absorb) done lemma component_trans: "\F component G; G component H\ \ F component H" unfolding component_def apply (blast intro: Join_assoc [symmetric]) done lemma component_antisym: "\F \ program; G \ program\ \(F component G \ G component F) \ F = G" apply (simp (no_asm_simp) add: component_eq_subset) apply clarify apply (rule program_equalityI, auto) done lemma Join_component_iff: "H \ program \ ((F \ G) component H) \ (F component H \ G component H)" apply (simp (no_asm_simp) add: component_eq_subset) apply blast done lemma component_constrains: "\F component G; G \ A co B; F \ program\ \ F \ A co B" apply (frule constrainsD2) apply (auto simp add: constrains_def component_eq_subset) done (*** preserves ***) lemma preserves_is_safety_prop [simp]: "safety_prop(preserves(f))" -apply (unfold preserves_def safety_prop_def) + unfolding preserves_def safety_prop_def apply (auto dest: ActsD simp add: stable_def constrains_def) apply (drule_tac c = act in subsetD, auto) done lemma preservesI [rule_format]: "\z. F \ stable({s \ state. f(s) = z}) \ F \ preserves(f)" apply (auto simp add: preserves_def) apply (blast dest: stableD2) done lemma preserves_imp_eq: "\F \ preserves(f); act \ Acts(F); \s,t\ \ act\ \ f(s) = f(t)" -apply (unfold preserves_def stable_def constrains_def) + unfolding preserves_def stable_def constrains_def apply (subgoal_tac "s \ state \ t \ state") prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) done lemma Join_preserves [iff]: "(F \ G \ preserves(v)) \ (programify(F) \ preserves(v) \ programify(G) \ preserves(v))" by (auto simp add: preserves_def INT_iff) lemma JOIN_preserves [iff]: "(JOIN(I,F): preserves(v)) \ (\i \ I. programify(F(i)):preserves(v))" by (auto simp add: JOIN_stable preserves_def INT_iff) lemma SKIP_preserves [iff]: "SKIP \ preserves(v)" by (auto simp add: preserves_def INT_iff) lemma fun_pair_apply [simp]: "fun_pair(f,g,x) = " unfolding fun_pair_def apply (simp (no_asm)) done lemma preserves_fun_pair: "preserves(fun_pair(v,w)) = preserves(v) \ preserves(w)" apply (rule equalityI) apply (auto simp add: preserves_def stable_def constrains_def, blast+) done lemma preserves_fun_pair_iff [iff]: "F \ preserves(fun_pair(v, w)) \ F \ preserves(v) \ preserves(w)" by (simp add: preserves_fun_pair) lemma fun_pair_comp_distrib: "(fun_pair(f, g) comp h)(x) = fun_pair(f comp h, g comp h, x)" by (simp add: fun_pair_def metacomp_def) lemma comp_apply [simp]: "(f comp g)(x) = f(g(x))" by (simp add: metacomp_def) lemma preserves_type: "preserves(v)<=program" by (unfold preserves_def, auto) lemma preserves_into_program [TC]: "F \ preserves(f) \ F \ program" by (blast intro: preserves_type [THEN subsetD]) lemma subset_preserves_comp: "preserves(f) \ preserves(g comp f)" apply (auto simp add: preserves_def stable_def constrains_def) apply (drule_tac x = "f (xb)" in spec) apply (drule_tac x = act in bspec, auto) done lemma imp_preserves_comp: "F \ preserves(f) \ F \ preserves(g comp f)" by (blast intro: subset_preserves_comp [THEN subsetD]) lemma preserves_subset_stable: "preserves(f) \ stable({s \ state. P(f(s))})" apply (auto simp add: preserves_def stable_def constrains_def) apply (rename_tac s' s) apply (subgoal_tac "f (s) = f (s') ") apply (force+) done lemma preserves_imp_stable: "F \ preserves(f) \ F \ stable({s \ state. P(f(s))})" by (blast intro: preserves_subset_stable [THEN subsetD]) lemma preserves_imp_increasing: "\F \ preserves(f); \x \ state. f(x):A\ \ F \ Increasing.increasing(A, r, f)" -apply (unfold Increasing.increasing_def) + unfolding Increasing.increasing_def apply (auto intro: preserves_into_program) apply (rule_tac P = "\x. \k, x\:r" in preserves_imp_stable, auto) done lemma preserves_id_subset_stable: "st_set(A) \ preserves(\x. x) \ stable(A)" apply (unfold preserves_def stable_def constrains_def, auto) apply (drule_tac x = xb in spec) apply (drule_tac x = act in bspec) apply (auto dest: ActsD) done lemma preserves_id_imp_stable: "\F \ preserves(\x. x); st_set(A)\ \ F \ stable(A)" by (blast intro: preserves_id_subset_stable [THEN subsetD]) (** Added by Sidi **) (** component_of **) (* component_of is stronger than component *) lemma component_of_imp_component: "F component_of H \ F component H" apply (unfold component_def component_of_def, blast) done (* component_of satisfies many of component's properties *) lemma component_of_refl [simp]: "F \ program \ F component_of F" unfolding component_of_def apply (rule_tac x = SKIP in exI, auto) done lemma component_of_SKIP [simp]: "F \ program \SKIP component_of F" apply (unfold component_of_def, auto) apply (rule_tac x = F in exI, auto) done lemma component_of_trans: "\F component_of G; G component_of H\ \ F component_of H" unfolding component_of_def apply (blast intro: Join_assoc [symmetric]) done (** localize **) lemma localize_Init_eq [simp]: "Init(localize(v,F)) = Init(F)" by (unfold localize_def, simp) lemma localize_Acts_eq [simp]: "Acts(localize(v,F)) = Acts(F)" by (unfold localize_def, simp) lemma localize_AllowedActs_eq [simp]: "AllowedActs(localize(v,F)) = AllowedActs(F) \ (\G \ preserves(v). Acts(G))" unfolding localize_def apply (rule equalityI) apply (auto dest: Acts_type [THEN subsetD]) done (** Theorems used in ClientImpl **) lemma stable_localTo_stable2: "\F \ stable({s \ state. P(f(s), g(s))}); G \ preserves(f); G \ preserves(g)\ \ F \ G \ stable({s \ state. P(f(s), g(s))})" apply (auto dest: ActsD preserves_into_program simp add: stable_def constrains_def) apply (case_tac "act \ Acts (F) ") apply auto apply (drule preserves_imp_eq) apply (drule_tac [3] preserves_imp_eq, auto) done lemma Increasing_preserves_Stable: "\F \ stable({s \ state. :r}); G \ preserves(f); F \ G \ Increasing(A, r, g); \x \ state. f(x):A \ g(x):A\ \ F \ G \ Stable({s \ state. :r})" apply (auto simp add: stable_def Stable_def Increasing_def Constrains_def all_conj_distrib) apply (simp_all add: constrains_type [THEN subsetD] preserves_type [THEN subsetD]) apply (blast intro: constrains_weaken) (*The G case remains*) apply (auto dest: ActsD simp add: preserves_def stable_def constrains_def ball_conj_distrib all_conj_distrib) (*We have a G-action, so delete assumptions about F-actions*) apply (erule_tac V = "\act \ Acts (F). P (act)" for P in thin_rl) apply (erule_tac V = "\k\A. \act \ Acts (F) . P (k,act)" for P in thin_rl) apply (subgoal_tac "f (x) = f (xa) ") apply (auto dest!: bspec) done (** Lemma used in AllocImpl **) lemma Constrains_UN_left: "\\x \ I. F \ A(x) Co B; F \ program\ \ F:(\x \ I. A(x)) Co B" by (unfold Constrains_def constrains_def, auto) lemma stable_Join_Stable: "\F \ stable({s \ state. P(f(s), g(s))}); \k \ A. F \ G \ Stable({s \ state. P(k, g(s))}); G \ preserves(f); \s \ state. f(s):A\ \ F \ G \ Stable({s \ state. P(f(s), g(s))})" -apply (unfold stable_def Stable_def preserves_def) + unfolding stable_def Stable_def preserves_def apply (rule_tac A = "(\k \ A. {s \ state. f(s)=k} \ {s \ state. P (f (s), g (s))})" in Constrains_weaken_L) prefer 2 apply blast apply (rule Constrains_UN_left, auto) apply (rule_tac A = "{s \ state. f(s)=k} \ {s \ state. P (f (s), g (s))} \ {s \ state. P (k, g (s))}" and A' = "({s \ state. f(s)=k} \ {s \ state. P (f (s), g (s))}) \ {s \ state. P (k, g (s))}" in Constrains_weaken) prefer 2 apply blast prefer 2 apply blast apply (rule Constrains_Int) apply (rule constrains_imp_Constrains) apply (auto simp add: constrains_type [THEN subsetD]) apply (blast intro: constrains_weaken) apply (drule_tac x = k in spec) apply (blast intro: constrains_weaken) done end diff --git a/src/ZF/UNITY/Constrains.thy b/src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy +++ b/src/ZF/UNITY/Constrains.thy @@ -1,507 +1,507 @@ (* Title: ZF/UNITY/Constrains.thy Author: Sidi O Ehmety, Computer Laboratory Copyright 2001 University of Cambridge *) section\Weak Safety Properties\ theory Constrains imports UNITY begin consts traces :: "[i, i] \ i" (* Initial states and program \ (final state, reversed trace to it)... the domain may also be state*list(state) *) inductive domains "traces(init, acts)" <= "(init \ (\act\acts. field(act)))*list(\act\acts. field(act))" intros (*Initial trace is empty*) Init: "s: init \ \ traces(init,acts)" Acts: "\act:acts; \s,evs\ \ traces(init,acts); : act\ \ \ traces(init, acts)" type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1 consts reachable :: "i\i" inductive domains "reachable(F)" \ "Init(F) \ (\act\Acts(F). field(act))" intros Init: "s:Init(F) \ s:reachable(F)" Acts: "\act: Acts(F); s:reachable(F); : act\ \ s':reachable(F)" type_intros UnI1 UnI2 fieldI2 UN_I definition Constrains :: "[i,i] \ i" (infixl \Co\ 60) where "A Co B \ {F:program. F:(reachable(F) \ A) co B}" definition op_Unless :: "[i, i] \ i" (infixl \Unless\ 60) where "A Unless B \ (A-B) Co (A \ B)" definition Stable :: "i \ i" where "Stable(A) \ A Co A" definition (*Always is the weak form of "invariant"*) Always :: "i \ i" where "Always(A) \ initially(A) \ Stable(A)" (*** traces and reachable ***) lemma reachable_type: "reachable(F) \ state" apply (cut_tac F = F in Init_type) apply (cut_tac F = F in Acts_type) apply (cut_tac F = F in reachable.dom_subset, blast) done lemma st_set_reachable: "st_set(reachable(F))" unfolding st_set_def apply (rule reachable_type) done declare st_set_reachable [iff] lemma reachable_Int_state: "reachable(F) \ state = reachable(F)" by (cut_tac reachable_type, auto) declare reachable_Int_state [iff] lemma state_Int_reachable: "state \ reachable(F) = reachable(F)" by (cut_tac reachable_type, auto) declare state_Int_reachable [iff] lemma reachable_equiv_traces: "F \ program \ reachable(F)={s \ state. \evs. \s,evs\:traces(Init(F), Acts(F))}" apply (rule equalityI, safe) apply (blast dest: reachable_type [THEN subsetD]) apply (erule_tac [2] traces.induct) apply (erule reachable.induct) apply (blast intro: reachable.intros traces.intros)+ done lemma Init_into_reachable: "Init(F) \ reachable(F)" by (blast intro: reachable.intros) lemma stable_reachable: "\F \ program; G \ program; Acts(G) \ Acts(F)\ \ G \ stable(reachable(F))" apply (blast intro: stableI constrainsI st_setI reachable_type [THEN subsetD] reachable.intros) done declare stable_reachable [intro!] declare stable_reachable [simp] (*The set of all reachable states is an invariant...*) lemma invariant_reachable: "F \ program \ F \ invariant(reachable(F))" -apply (unfold invariant_def initially_def) + unfolding invariant_def initially_def apply (blast intro: reachable_type [THEN subsetD] reachable.intros) done (*...in fact the strongest invariant!*) lemma invariant_includes_reachable: "F \ invariant(A) \ reachable(F) \ A" apply (cut_tac F = F in Acts_type) apply (cut_tac F = F in Init_type) apply (cut_tac F = F in reachable_type) apply (simp (no_asm_use) add: stable_def constrains_def invariant_def initially_def) apply (rule subsetI) apply (erule reachable.induct) apply (blast intro: reachable.intros)+ done (*** Co ***) lemma constrains_reachable_Int: "F \ B co B'\F:(reachable(F) \ B) co (reachable(F) \ B')" apply (frule constrains_type [THEN subsetD]) apply (frule stable_reachable [OF _ _ subset_refl]) apply (simp_all add: stable_def constrains_Int) done (*Resembles the previous definition of Constrains*) lemma Constrains_eq_constrains: "A Co B = {F \ program. F:(reachable(F) \ A) co (reachable(F) \ B)}" unfolding Constrains_def apply (blast dest: constrains_reachable_Int constrains_type [THEN subsetD] intro: constrains_weaken) done lemmas Constrains_def2 = Constrains_eq_constrains [THEN eq_reflection] lemma constrains_imp_Constrains: "F \ A co A' \ F \ A Co A'" unfolding Constrains_def apply (blast intro: constrains_weaken_L dest: constrainsD2) done lemma ConstrainsI: "\\act s s'. \act \ Acts(F); :act; s \ A\ \ s':A'; F \ program\ \ F \ A Co A'" apply (auto simp add: Constrains_def constrains_def st_set_def) apply (blast dest: reachable_type [THEN subsetD]) done lemma Constrains_type: "A Co B \ program" apply (unfold Constrains_def, blast) done lemma Constrains_empty: "F \ 0 Co B \ F \ program" by (auto dest: Constrains_type [THEN subsetD] intro: constrains_imp_Constrains) declare Constrains_empty [iff] lemma Constrains_state: "F \ A Co state \ F \ program" unfolding Constrains_def apply (auto dest: Constrains_type [THEN subsetD] intro: constrains_imp_Constrains) done declare Constrains_state [iff] lemma Constrains_weaken_R: "\F \ A Co A'; A'<=B'\ \ F \ A Co B'" unfolding Constrains_def2 apply (blast intro: constrains_weaken_R) done lemma Constrains_weaken_L: "\F \ A Co A'; B<=A\ \ F \ B Co A'" unfolding Constrains_def2 apply (blast intro: constrains_weaken_L st_set_subset) done lemma Constrains_weaken: "\F \ A Co A'; B<=A; A'<=B'\ \ F \ B Co B'" unfolding Constrains_def2 apply (blast intro: constrains_weaken st_set_subset) done (** Union **) lemma Constrains_Un: "\F \ A Co A'; F \ B Co B'\ \ F \ (A \ B) Co (A' \ B')" apply (unfold Constrains_def2, auto) apply (simp add: Int_Un_distrib) apply (blast intro: constrains_Un) done lemma Constrains_UN: "\(\i. i \ I\F \ A(i) Co A'(i)); F \ program\ \ F:(\i \ I. A(i)) Co (\i \ I. A'(i))" by (auto intro: constrains_UN simp del: UN_simps simp add: Constrains_def2 Int_UN_distrib) (** Intersection **) lemma Constrains_Int: "\F \ A Co A'; F \ B Co B'\\ F:(A \ B) Co (A' \ B')" unfolding Constrains_def apply (subgoal_tac "reachable (F) \ (A \ B) = (reachable (F) \ A) \ (reachable (F) \ B) ") apply (auto intro: constrains_Int) done lemma Constrains_INT: "\(\i. i \ I \F \ A(i) Co A'(i)); F \ program\ \ F:(\i \ I. A(i)) Co (\i \ I. A'(i))" apply (simp (no_asm_simp) del: INT_simps add: Constrains_def INT_extend_simps) apply (rule constrains_INT) apply (auto simp add: Constrains_def) done lemma Constrains_imp_subset: "F \ A Co A' \ reachable(F) \ A \ A'" unfolding Constrains_def apply (blast dest: constrains_imp_subset) done lemma Constrains_trans: "\F \ A Co B; F \ B Co C\ \ F \ A Co C" unfolding Constrains_def2 apply (blast intro: constrains_trans constrains_weaken) done lemma Constrains_cancel: "\F \ A Co (A' \ B); F \ B Co B'\ \ F \ A Co (A' \ B')" unfolding Constrains_def2 apply (simp (no_asm_use) add: Int_Un_distrib) apply (blast intro: constrains_cancel) done (*** Stable ***) (* Useful because there's no Stable_weaken. [Tanja Vos] *) lemma stable_imp_Stable: "F \ stable(A) \ F \ Stable(A)" -apply (unfold stable_def Stable_def) + unfolding stable_def Stable_def apply (erule constrains_imp_Constrains) done lemma Stable_eq: "\F \ Stable(A); A = B\ \ F \ Stable(B)" by blast lemma Stable_eq_stable: "F \ Stable(A) \ (F \ stable(reachable(F) \ A))" apply (auto dest: constrainsD2 simp add: Stable_def stable_def Constrains_def2) done lemma StableI: "F \ A Co A \ F \ Stable(A)" by (unfold Stable_def, assumption) lemma StableD: "F \ Stable(A) \ F \ A Co A" by (unfold Stable_def, assumption) lemma Stable_Un: "\F \ Stable(A); F \ Stable(A')\ \ F \ Stable(A \ A')" unfolding Stable_def apply (blast intro: Constrains_Un) done lemma Stable_Int: "\F \ Stable(A); F \ Stable(A')\ \ F \ Stable (A \ A')" unfolding Stable_def apply (blast intro: Constrains_Int) done lemma Stable_Constrains_Un: "\F \ Stable(C); F \ A Co (C \ A')\ \ F \ (C \ A) Co (C \ A')" unfolding Stable_def apply (blast intro: Constrains_Un [THEN Constrains_weaken_R]) done lemma Stable_Constrains_Int: "\F \ Stable(C); F \ (C \ A) Co A'\ \ F \ (C \ A) Co (C \ A')" unfolding Stable_def apply (blast intro: Constrains_Int [THEN Constrains_weaken]) done lemma Stable_UN: "\(\i. i \ I \ F \ Stable(A(i))); F \ program\ \ F \ Stable (\i \ I. A(i))" apply (simp add: Stable_def) apply (blast intro: Constrains_UN) done lemma Stable_INT: "\(\i. i \ I \ F \ Stable(A(i))); F \ program\ \ F \ Stable (\i \ I. A(i))" apply (simp add: Stable_def) apply (blast intro: Constrains_INT) done lemma Stable_reachable: "F \ program \F \ Stable (reachable(F))" apply (simp (no_asm_simp) add: Stable_eq_stable Int_absorb) done lemma Stable_type: "Stable(A) \ program" unfolding Stable_def apply (rule Constrains_type) done (*** The Elimination Theorem. The "free" m has become universally quantified! Should the premise be \m instead of \m ? Would make it harder to use in forward proof. ***) lemma Elimination: "\\m \ M. F \ ({s \ A. x(s) = m}) Co (B(m)); F \ program\ \ F \ ({s \ A. x(s):M}) Co (\m \ M. B(m))" apply (unfold Constrains_def, auto) apply (rule_tac A1 = "reachable (F) \ A" in UNITY.elimination [THEN constrains_weaken_L]) apply (auto intro: constrains_weaken_L) done (* As above, but for the special case of A=state *) lemma Elimination2: "\\m \ M. F \ {s \ state. x(s) = m} Co B(m); F \ program\ \ F \ {s \ state. x(s):M} Co (\m \ M. B(m))" apply (blast intro: Elimination) done (** Unless **) lemma Unless_type: "A Unless B <=program" unfolding op_Unless_def apply (rule Constrains_type) done (*** Specialized laws for handling Always ***) (** Natural deduction rules for "Always A" **) lemma AlwaysI: "\Init(F)<=A; F \ Stable(A)\ \ F \ Always(A)" -apply (unfold Always_def initially_def) + unfolding Always_def initially_def apply (frule Stable_type [THEN subsetD], auto) done lemma AlwaysD: "F \ Always(A) \ Init(F)<=A \ F \ Stable(A)" by (simp add: Always_def initially_def) lemmas AlwaysE = AlwaysD [THEN conjE] lemmas Always_imp_Stable = AlwaysD [THEN conjunct2] (*The set of all reachable states is Always*) lemma Always_includes_reachable: "F \ Always(A) \ reachable(F) \ A" apply (simp (no_asm_use) add: Stable_def Constrains_def constrains_def Always_def initially_def) apply (rule subsetI) apply (erule reachable.induct) apply (blast intro: reachable.intros)+ done lemma invariant_imp_Always: "F \ invariant(A) \ F \ Always(A)" -apply (unfold Always_def invariant_def Stable_def stable_def) + unfolding Always_def invariant_def Stable_def stable_def apply (blast intro: constrains_imp_Constrains) done lemmas Always_reachable = invariant_reachable [THEN invariant_imp_Always] lemma Always_eq_invariant_reachable: "Always(A) = {F \ program. F \ invariant(reachable(F) \ A)}" apply (simp (no_asm) add: Always_def invariant_def Stable_def Constrains_def2 stable_def initially_def) apply (rule equalityI, auto) apply (blast intro: reachable.intros reachable_type) done (*the RHS is the traditional definition of the "always" operator*) lemma Always_eq_includes_reachable: "Always(A) = {F \ program. reachable(F) \ A}" apply (rule equalityI, safe) apply (auto dest: invariant_includes_reachable simp add: subset_Int_iff invariant_reachable Always_eq_invariant_reachable) done lemma Always_type: "Always(A) \ program" by (unfold Always_def initially_def, auto) lemma Always_state_eq: "Always(state) = program" apply (rule equalityI) apply (auto dest: Always_type [THEN subsetD] reachable_type [THEN subsetD] simp add: Always_eq_includes_reachable) done declare Always_state_eq [simp] lemma state_AlwaysI: "F \ program \ F \ Always(state)" by (auto dest: reachable_type [THEN subsetD] simp add: Always_eq_includes_reachable) lemma Always_eq_UN_invariant: "st_set(A) \ Always(A) = (\I \ Pow(A). invariant(I))" apply (simp (no_asm) add: Always_eq_includes_reachable) apply (rule equalityI, auto) apply (blast intro: invariantI rev_subsetD [OF _ Init_into_reachable] rev_subsetD [OF _ invariant_includes_reachable] dest: invariant_type [THEN subsetD])+ done lemma Always_weaken: "\F \ Always(A); A \ B\ \ F \ Always(B)" by (auto simp add: Always_eq_includes_reachable) (*** "Co" rules involving Always ***) lemmas Int_absorb2 = subset_Int_iff [unfolded iff_def, THEN conjunct1, THEN mp] lemma Always_Constrains_pre: "F \ Always(I) \ (F:(I \ A) Co A') \ (F \ A Co A')" apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_def Int_assoc [symmetric]) done lemma Always_Constrains_post: "F \ Always(I) \ (F \ A Co (I \ A')) \(F \ A Co A')" apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_eq_constrains Int_assoc [symmetric]) done lemma Always_ConstrainsI: "\F \ Always(I); F \ (I \ A) Co A'\ \ F \ A Co A'" by (blast intro: Always_Constrains_pre [THEN iffD1]) (* \F \ Always(I); F \ A Co A'\ \ F \ A Co (I \ A') *) lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2] (*The analogous proof of Always_LeadsTo_weaken doesn't terminate*) lemma Always_Constrains_weaken: "\F \ Always(C); F \ A Co A'; C \ B<=A; C \ A'<=B'\\F \ B Co B'" apply (rule Always_ConstrainsI) apply (drule_tac [2] Always_ConstrainsD, simp_all) apply (blast intro: Constrains_weaken) done (** Conjoining Always properties **) lemma Always_Int_distrib: "Always(A \ B) = Always(A) \ Always(B)" by (auto simp add: Always_eq_includes_reachable) (* the premise i \ I is need since \is formally not defined for I=0 *) lemma Always_INT_distrib: "i \ I\Always(\i \ I. A(i)) = (\i \ I. Always(A(i)))" apply (rule equalityI) apply (auto simp add: Inter_iff Always_eq_includes_reachable) done lemma Always_Int_I: "\F \ Always(A); F \ Always(B)\ \ F \ Always(A \ B)" apply (simp (no_asm_simp) add: Always_Int_distrib) done (*Allows a kind of "implication introduction"*) lemma Always_Diff_Un_eq: "\F \ Always(A)\ \ (F \ Always(C-A \ B)) \ (F \ Always(B))" by (auto simp add: Always_eq_includes_reachable) (*Delete the nearest invariance assumption (which will be the second one used by Always_Int_I) *) lemmas Always_thin = thin_rl [of "F \ Always(A)"] for F A (*To allow expansion of the program's definition when appropriate*) named_theorems program "program definitions" ML \ (*Combines two invariance ASSUMPTIONS into one. USEFUL??*) fun Always_Int_tac ctxt = dresolve_tac ctxt @{thms Always_Int_I} THEN' assume_tac ctxt THEN' eresolve_tac ctxt @{thms Always_thin}; (*Combines a list of invariance THEOREMS into one.*) val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I}); (*proves "co" properties when the program is specified*) fun constrains_tac ctxt = SELECT_GOAL (EVERY [REPEAT (Always_Int_tac ctxt 1), REPEAT (eresolve_tac ctxt @{thms Always_ConstrainsI} 1 ORELSE resolve_tac ctxt [@{thm StableI}, @{thm stableI}, @{thm constrains_imp_Constrains}] 1), resolve_tac ctxt @{thms constrainsI} 1, (* Three subgoals *) rewrite_goal_tac ctxt [@{thm st_set_def}] 3, REPEAT (force_tac ctxt 2), full_simp_tac (ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\program\)) 1, ALLGOALS (clarify_tac ctxt), REPEAT (FIRSTGOAL (eresolve_tac ctxt @{thms disjE})), ALLGOALS (clarify_tac ctxt), REPEAT (FIRSTGOAL (eresolve_tac ctxt @{thms disjE})), ALLGOALS (clarify_tac ctxt), ALLGOALS (asm_full_simp_tac ctxt), ALLGOALS (clarify_tac ctxt)]); (*For proving invariants*) fun always_tac ctxt i = resolve_tac ctxt @{thms AlwaysI} i THEN force_tac ctxt i THEN constrains_tac ctxt i; \ method_setup safety = \ Scan.succeed (SIMPLE_METHOD' o constrains_tac)\ "for proving safety properties" method_setup always = \ Scan.succeed (SIMPLE_METHOD' o always_tac)\ "for proving invariants" end diff --git a/src/ZF/UNITY/FP.thy b/src/ZF/UNITY/FP.thy --- a/src/ZF/UNITY/FP.thy +++ b/src/ZF/UNITY/FP.thy @@ -1,83 +1,83 @@ (* Title: ZF/UNITY/FP.thy Author: Sidi O Ehmety, Computer Laboratory Copyright 2001 University of Cambridge From Misra, "A Logic for Concurrent Programming", 1994 Theory ported from HOL. *) section\Fixed Point of a Program\ theory FP imports UNITY begin definition FP_Orig :: "i\i" where "FP_Orig(F) \ \({A \ Pow(state). \B. F \ stable(A \ B)})" definition FP :: "i\i" where "FP(F) \ {s\state. F \ stable({s})}" lemma FP_Orig_type: "FP_Orig(F) \ state" by (unfold FP_Orig_def, blast) lemma st_set_FP_Orig [iff]: "st_set(FP_Orig(F))" unfolding st_set_def apply (rule FP_Orig_type) done lemma FP_type: "FP(F) \ state" by (unfold FP_def, blast) lemma st_set_FP [iff]: "st_set(FP(F))" unfolding st_set_def apply (rule FP_type) done lemma stable_FP_Orig_Int: "F \ program \ F \ stable(FP_Orig(F) \ B)" apply (simp only: FP_Orig_def stable_def Int_Union2) apply (blast intro: constrains_UN) done lemma FP_Orig_weakest2: "\\B. F \ stable (A \ B); st_set(A)\ \ A \ FP_Orig(F)" by (unfold FP_Orig_def stable_def st_set_def, blast) lemmas FP_Orig_weakest = allI [THEN FP_Orig_weakest2] lemma stable_FP_Int: "F \ program \ F \ stable (FP(F) \ B)" apply (subgoal_tac "FP (F) \ B = (\x\B. FP (F) \ {x}) ") prefer 2 apply blast apply (simp (no_asm_simp) add: Int_cons_right) -apply (unfold FP_def stable_def) + unfolding FP_def stable_def apply (rule constrains_UN) apply (auto simp add: cons_absorb) done lemma FP_subset_FP_Orig: "F \ program \ FP(F) \ FP_Orig(F)" by (rule stable_FP_Int [THEN FP_Orig_weakest], auto) lemma FP_Orig_subset_FP: "F \ program \ FP_Orig(F) \ FP(F)" apply (unfold FP_Orig_def FP_def, clarify) apply (drule_tac x = "{x}" in spec) apply (simp add: Int_cons_right) apply (frule stableD2) apply (auto simp add: cons_absorb st_set_def) done lemma FP_equivalence: "F \ program \ FP(F) = FP_Orig(F)" by (blast intro!: FP_Orig_subset_FP FP_subset_FP_Orig) lemma FP_weakest [rule_format]: "\\B. F \ stable(A \ B); F \ program; st_set(A)\ \ A \ FP(F)" by (simp add: FP_equivalence FP_Orig_weakest) lemma Diff_FP: "\F \ program; st_set(A)\ \ A-FP(F) = (\act \ Acts(F). A - {s \ state. act``{s} \ {s}})" by (unfold FP_def stable_def constrains_def st_set_def, blast) end diff --git a/src/ZF/UNITY/Follows.thy b/src/ZF/UNITY/Follows.thy --- a/src/ZF/UNITY/Follows.thy +++ b/src/ZF/UNITY/Follows.thy @@ -1,475 +1,475 @@ (* Title: ZF/UNITY/Follows.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2002 University of Cambridge Theory ported from HOL. *) section\The "Follows" relation of Charpentier and Sivilotte\ theory Follows imports SubstAx Increasing begin definition Follows :: "[i, i, i\i, i\i] \ i" where "Follows(A, r, f, g) \ Increasing(A, r, g) Int Increasing(A, r,f) Int Always({s \ state. :r}) Int (\k \ A. {s \ state. :r} \w {s \ state. :r})" abbreviation Incr :: "[i\i]\i" where "Incr(f) \ Increasing(list(nat), prefix(nat), f)" abbreviation n_Incr :: "[i\i]\i" where "n_Incr(f) \ Increasing(nat, Le, f)" abbreviation s_Incr :: "[i\i]\i" where "s_Incr(f) \ Increasing(Pow(nat), SetLe(nat), f)" abbreviation m_Incr :: "[i\i]\i" where "m_Incr(f) \ Increasing(Mult(nat), MultLe(nat, Le), f)" abbreviation n_Fols :: "[i\i, i\i]\i" (infixl \n'_Fols\ 65) where "f n_Fols g \ Follows(nat, Le, f, g)" abbreviation Follows' :: "[i\i, i\i, i, i] \ i" (\(_ /Fols _ /Wrt (_ /'/ _))\ [60, 0, 0, 60] 60) where "f Fols g Wrt r/A \ Follows(A,r,f,g)" (*Does this hold for "invariant"?*) lemma Follows_cong: "\A=A'; r=r'; \x. x \ state \ f(x)=f'(x); \x. x \ state \ g(x)=g'(x)\ \ Follows(A, r, f, g) = Follows(A', r', f', g')" by (simp add: Increasing_def Follows_def) lemma subset_Always_comp: "\mono1(A, r, B, s, h); \x \ state. f(x):A \ g(x):A\ \ Always({x \ state. \ r})<=Always({x \ state. <(h comp f)(x), (h comp g)(x)> \ s})" -apply (unfold mono1_def metacomp_def) + unfolding mono1_def metacomp_def apply (auto simp add: Always_eq_includes_reachable) done lemma imp_Always_comp: "\F \ Always({x \ state. \ r}); mono1(A, r, B, s, h); \x \ state. f(x):A \ g(x):A\ \ F \ Always({x \ state. <(h comp f)(x), (h comp g)(x)> \ s})" by (blast intro: subset_Always_comp [THEN subsetD]) lemma imp_Always_comp2: "\F \ Always({x \ state. \ r}); F \ Always({x \ state. \ s}); mono2(A, r, B, s, C, t, h); \x \ state. f1(x):A \ f(x):A \ g1(x):B \ g(x):B\ \ F \ Always({x \ state. \ t})" apply (auto simp add: Always_eq_includes_reachable mono2_def) apply (auto dest!: subsetD) done (* comp LeadsTo *) lemma subset_LeadsTo_comp: "\mono1(A, r, B, s, h); refl(A,r); trans[B](s); \x \ state. f(x):A \ g(x):A\ \ (\j \ A. {s \ state. \ r} \w {s \ state. \ r}) <= (\k \ B. {x \ state. \ s} \w {x \ state. \ s})" apply (unfold mono1_def metacomp_def, clarify) apply (simp_all (no_asm_use) add: INT_iff) apply auto apply (rule single_LeadsTo_I) prefer 2 apply (blast dest: LeadsTo_type [THEN subsetD], auto) apply (rotate_tac 5) apply (drule_tac x = "g (sa) " in bspec) apply (erule_tac [2] LeadsTo_weaken) apply (auto simp add: part_order_def refl_def) apply (rule_tac b = "h (g (sa))" in trans_onD) apply blast apply auto done lemma imp_LeadsTo_comp: "\F:(\j \ A. {s \ state. \ r} \w {s \ state. \ r}); mono1(A, r, B, s, h); refl(A,r); trans[B](s); \x \ state. f(x):A \ g(x):A\ \ F:(\k \ B. {x \ state. \ s} \w {x \ state. \ s})" apply (rule subset_LeadsTo_comp [THEN subsetD], auto) done lemma imp_LeadsTo_comp_right: "\F \ Increasing(B, s, g); \j \ A. F: {s \ state. \ r} \w {s \ state. \ r}; mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t); \x \ state. f1(x):A \ f(x):A \ g(x):B; k \ C\ \ F:{x \ state. \ t} \w {x \ state. \ t}" -apply (unfold mono2_def Increasing_def) + unfolding mono2_def Increasing_def apply (rule single_LeadsTo_I, auto) apply (drule_tac x = "g (sa) " and A = B in bspec) apply auto apply (drule_tac x = "f (sa) " and P = "\j. F \ X(j) \w Y(j)" for X Y in bspec) apply auto apply (rule PSP_Stable [THEN LeadsTo_weaken], blast, blast) apply auto apply (force simp add: part_order_def refl_def) apply (force simp add: part_order_def refl_def) apply (drule_tac x = "f1 (x)" and x1 = "f (sa) " and P2 = "\x y. \u\B. P (x,y,u)" for P in bspec [THEN bspec]) apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "\x y. P (x,y) \ d (x,y) \ t" for P d in bspec [THEN bspec]) apply auto apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD) apply (auto simp add: part_order_def) done lemma imp_LeadsTo_comp_left: "\F \ Increasing(A, r, f); \j \ B. F: {x \ state. \ s} \w {x \ state. \ s}; mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t); \x \ state. f(x):A \ g1(x):B \ g(x):B; k \ C\ \ F:{x \ state. \ t} \w {x \ state. \ t}" -apply (unfold mono2_def Increasing_def) + unfolding mono2_def Increasing_def apply (rule single_LeadsTo_I, auto) apply (drule_tac x = "f (sa) " and P = "\k. F \ Stable (X (k))" for X in bspec) apply auto apply (drule_tac x = "g (sa) " in bspec) apply auto apply (rule PSP_Stable [THEN LeadsTo_weaken], blast, blast) apply auto apply (force simp add: part_order_def refl_def) apply (force simp add: part_order_def refl_def) apply (drule_tac x = "f (x) " and x1 = "f (sa) " in bspec [THEN bspec]) apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "\x y. P (x,y) \ d (x,y) \ t" for P d in bspec [THEN bspec]) apply auto apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD) apply (auto simp add: part_order_def) done (** This general result is used to prove Follows Un, munion, etc. **) lemma imp_LeadsTo_comp2: "\F \ Increasing(A, r, f1) \ Increasing(B, s, g); \j \ A. F: {s \ state. \ r} \w {s \ state. \ r}; \j \ B. F: {x \ state. \ s} \w {x \ state. \ s}; mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t); \x \ state. f(x):A \ g1(x):B \ f1(x):A \g(x):B; k \ C\ \ F:{x \ state. \ t} \w {x \ state. \ t}" apply (rule_tac B = "{x \ state. \ t}" in LeadsTo_Trans) apply (blast intro: imp_LeadsTo_comp_right) apply (blast intro: imp_LeadsTo_comp_left) done (* Follows type *) lemma Follows_type: "Follows(A, r, f, g)<=program" unfolding Follows_def apply (blast dest: Increasing_type [THEN subsetD]) done lemma Follows_into_program [TC]: "F \ Follows(A, r, f, g) \ F \ program" by (blast dest: Follows_type [THEN subsetD]) lemma FollowsD: "F \ Follows(A, r, f, g)\ F \ program \ (\a. a \ A) \ (\x \ state. f(x):A \ g(x):A)" unfolding Follows_def apply (blast dest: IncreasingD) done lemma Follows_constantI: "\F \ program; c \ A; refl(A, r)\ \ F \ Follows(A, r, \x. c, \x. c)" apply (unfold Follows_def, auto) apply (auto simp add: refl_def) done lemma subset_Follows_comp: "\mono1(A, r, B, s, h); refl(A, r); trans[B](s)\ \ Follows(A, r, f, g) \ Follows(B, s, h comp f, h comp g)" apply (unfold Follows_def, clarify) apply (frule_tac f = g in IncreasingD) apply (frule_tac f = f in IncreasingD) apply (rule IntI) apply (rule_tac [2] h = h in imp_LeadsTo_comp) prefer 5 apply assumption apply (auto intro: imp_Increasing_comp imp_Always_comp simp del: INT_simps) done lemma imp_Follows_comp: "\F \ Follows(A, r, f, g); mono1(A, r, B, s, h); refl(A, r); trans[B](s)\ \ F \ Follows(B, s, h comp f, h comp g)" apply (blast intro: subset_Follows_comp [THEN subsetD]) done (* 2-place monotone operation \ this general result is used to prove Follows_Un, Follows_munion *) (* 2-place monotone operation \ this general result is used to prove Follows_Un, Follows_munion *) lemma imp_Follows_comp2: "\F \ Follows(A, r, f1, f); F \ Follows(B, s, g1, g); mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t)\ \ F \ Follows(C, t, \x. h(f1(x), g1(x)), \x. h(f(x), g(x)))" apply (unfold Follows_def, clarify) apply (frule_tac f = g in IncreasingD) apply (frule_tac f = f in IncreasingD) apply (rule IntI, safe) apply (rule_tac [3] h = h in imp_Always_comp2) prefer 5 apply assumption apply (rule_tac [2] h = h in imp_Increasing_comp2) prefer 4 apply assumption apply (rule_tac h = h in imp_Increasing_comp2) prefer 3 apply assumption apply simp_all apply (blast dest!: IncreasingD) apply (rule_tac h = h in imp_LeadsTo_comp2) prefer 4 apply assumption apply auto prefer 3 apply (simp add: mono2_def) apply (blast dest: IncreasingD)+ done lemma Follows_trans: "\F \ Follows(A, r, f, g); F \ Follows(A,r, g, h); trans[A](r)\ \ F \ Follows(A, r, f, h)" apply (frule_tac f = f in FollowsD) apply (frule_tac f = g in FollowsD) apply (simp add: Follows_def) apply (simp add: Always_eq_includes_reachable INT_iff, auto) apply (rule_tac [2] B = "{s \ state. \ r}" in LeadsTo_Trans) apply (rule_tac b = "g (x) " in trans_onD) apply blast+ done (** Destruction rules for Follows **) lemma Follows_imp_Increasing_left: "F \ Follows(A, r, f,g) \ F \ Increasing(A, r, f)" by (unfold Follows_def, blast) lemma Follows_imp_Increasing_right: "F \ Follows(A, r, f,g) \ F \ Increasing(A, r, g)" by (unfold Follows_def, blast) lemma Follows_imp_Always: "F :Follows(A, r, f, g) \ F \ Always({s \ state. \ r})" by (unfold Follows_def, blast) lemma Follows_imp_LeadsTo: "\F \ Follows(A, r, f, g); k \ A\ \ F: {s \ state. \ r } \w {s \ state. \ r}" by (unfold Follows_def, blast) lemma Follows_LeadsTo_pfixLe: "\F \ Follows(list(nat), gen_prefix(nat, Le), f, g); k \ list(nat)\ \ F \ {s \ state. k pfixLe g(s)} \w {s \ state. k pfixLe f(s)}" by (blast intro: Follows_imp_LeadsTo) lemma Follows_LeadsTo_pfixGe: "\F \ Follows(list(nat), gen_prefix(nat, Ge), f, g); k \ list(nat)\ \ F \ {s \ state. k pfixGe g(s)} \w {s \ state. k pfixGe f(s)}" by (blast intro: Follows_imp_LeadsTo) lemma Always_Follows1: "\F \ Always({s \ state. f(s) = g(s)}); F \ Follows(A, r, f, h); \x \ state. g(x):A\ \ F \ Follows(A, r, g, h)" -apply (unfold Follows_def Increasing_def Stable_def) + unfolding Follows_def Increasing_def Stable_def apply (simp add: INT_iff, auto) apply (rule_tac [3] C = "{s \ state. f(s)=g(s)}" and A = "{s \ state. \ r}" and A' = "{s \ state. \ r}" in Always_LeadsTo_weaken) apply (erule_tac A = "{s \ state. \ r}" and A' = "{s \ state. \ r}" in Always_Constrains_weaken) apply auto apply (drule Always_Int_I, assumption) apply (erule_tac A = "{s \ state. f(s)=g(s)} \ {s \ state. \ r}" in Always_weaken) apply auto done lemma Always_Follows2: "\F \ Always({s \ state. g(s) = h(s)}); F \ Follows(A, r, f, g); \x \ state. h(x):A\ \ F \ Follows(A, r, f, h)" -apply (unfold Follows_def Increasing_def Stable_def) + unfolding Follows_def Increasing_def Stable_def apply (simp add: INT_iff, auto) apply (rule_tac [3] C = "{s \ state. g (s) =h (s) }" and A = "{s \ state. \ r}" and A' = "{s \ state. \ r}" in Always_LeadsTo_weaken) apply (erule_tac A = "{s \ state. \ r}" and A' = "{s \ state. \ r}" in Always_Constrains_weaken) apply auto apply (drule Always_Int_I, assumption) apply (erule_tac A = "{s \ state. g(s)=h(s)} \ {s \ state. \ r}" in Always_weaken) apply auto done (** Union properties (with the subset ordering) **) lemma refl_SetLe [simp]: "refl(Pow(A), SetLe(A))" by (unfold refl_def SetLe_def, auto) lemma trans_on_SetLe [simp]: "trans[Pow(A)](SetLe(A))" by (unfold trans_on_def SetLe_def, auto) lemma antisym_SetLe [simp]: "antisym(SetLe(A))" by (unfold antisym_def SetLe_def, auto) lemma part_order_SetLe [simp]: "part_order(Pow(A), SetLe(A))" by (unfold part_order_def, auto) lemma increasing_Un: "\F \ Increasing.increasing(Pow(A), SetLe(A), f); F \ Increasing.increasing(Pow(A), SetLe(A), g)\ \ F \ Increasing.increasing(Pow(A), SetLe(A), \x. f(x) \ g(x))" by (rule_tac h = "(Un)" in imp_increasing_comp2, auto) lemma Increasing_Un: "\F \ Increasing(Pow(A), SetLe(A), f); F \ Increasing(Pow(A), SetLe(A), g)\ \ F \ Increasing(Pow(A), SetLe(A), \x. f(x) \ g(x))" by (rule_tac h = "(Un)" in imp_Increasing_comp2, auto) lemma Always_Un: "\F \ Always({s \ state. f1(s) \ f(s)}); F \ Always({s \ state. g1(s) \ g(s)})\ \ F \ Always({s \ state. f1(s) \ g1(s) \ f(s) \ g(s)})" by (simp add: Always_eq_includes_reachable, blast) lemma Follows_Un: "\F \ Follows(Pow(A), SetLe(A), f1, f); F \ Follows(Pow(A), SetLe(A), g1, g)\ \ F \ Follows(Pow(A), SetLe(A), \s. f1(s) \ g1(s), \s. f(s) \ g(s))" by (rule_tac h = "(Un)" in imp_Follows_comp2, auto) (** Multiset union properties (with the MultLe ordering) **) lemma refl_MultLe [simp]: "refl(Mult(A), MultLe(A,r))" by (unfold MultLe_def refl_def, auto) lemma MultLe_refl1 [simp]: "\multiset(M); mset_of(M)<=A\ \ \M, M\ \ MultLe(A, r)" -apply (unfold MultLe_def id_def lam_def) + unfolding MultLe_def id_def lam_def apply (auto simp add: Mult_iff_multiset) done lemma MultLe_refl2 [simp]: "M \ Mult(A) \ \M, M\ \ MultLe(A, r)" by (unfold MultLe_def id_def lam_def, auto) lemma trans_on_MultLe [simp]: "trans[Mult(A)](MultLe(A,r))" -apply (unfold MultLe_def trans_on_def) + unfolding MultLe_def trans_on_def apply (auto intro: trancl_trans simp add: multirel_def) done lemma MultLe_type: "MultLe(A, r)<= (Mult(A) * Mult(A))" apply (unfold MultLe_def, auto) apply (drule multirel_type [THEN subsetD], auto) done lemma MultLe_trans: "\\M,K\ \ MultLe(A,r); \K,N\ \ MultLe(A,r)\ \ \M,N\ \ MultLe(A,r)" apply (cut_tac A=A in trans_on_MultLe) apply (drule trans_onD, assumption) apply (auto dest: MultLe_type [THEN subsetD]) done lemma part_order_imp_part_ord: "part_order(A, r) \ part_ord(A, r-id(A))" -apply (unfold part_order_def part_ord_def) + unfolding part_order_def part_ord_def apply (simp add: refl_def id_def lam_def irrefl_def, auto) apply (simp (no_asm) add: trans_on_def) apply auto apply (blast dest: trans_onD) apply (simp (no_asm_use) add: antisym_def) apply auto done lemma antisym_MultLe [simp]: "part_order(A, r) \ antisym(MultLe(A,r))" -apply (unfold MultLe_def antisym_def) + unfolding MultLe_def antisym_def apply (drule part_order_imp_part_ord, auto) apply (drule irrefl_on_multirel) apply (frule multirel_type [THEN subsetD]) apply (drule multirel_trans) apply (auto simp add: irrefl_def) done lemma part_order_MultLe [simp]: "part_order(A, r) \ part_order(Mult(A), MultLe(A, r))" apply (frule antisym_MultLe) apply (auto simp add: part_order_def) done lemma empty_le_MultLe [simp]: "\multiset(M); mset_of(M)<= A\ \ \0, M\ \ MultLe(A, r)" unfolding MultLe_def apply (case_tac "M=0") apply (auto simp add: FiniteFun.intros) apply (subgoal_tac "<0 +# 0, 0 +# M> \ multirel (A, r - id (A))") apply (rule_tac [2] one_step_implies_multirel) apply (auto simp add: Mult_iff_multiset) done lemma empty_le_MultLe2 [simp]: "M \ Mult(A) \ \0, M\ \ MultLe(A, r)" by (simp add: Mult_iff_multiset) lemma munion_mono: "\\M, N\ \ MultLe(A, r); \K, L\ \ MultLe(A, r)\ \ \ MultLe(A, r)" unfolding MultLe_def apply (auto intro: munion_multirel_mono1 munion_multirel_mono2 munion_multirel_mono multiset_into_Mult simp add: Mult_iff_multiset) done lemma increasing_munion: "\F \ Increasing.increasing(Mult(A), MultLe(A,r), f); F \ Increasing.increasing(Mult(A), MultLe(A,r), g)\ \ F \ Increasing.increasing(Mult(A),MultLe(A,r), \x. f(x) +# g(x))" by (rule_tac h = munion in imp_increasing_comp2, auto) lemma Increasing_munion: "\F \ Increasing(Mult(A), MultLe(A,r), f); F \ Increasing(Mult(A), MultLe(A,r), g)\ \ F \ Increasing(Mult(A),MultLe(A,r), \x. f(x) +# g(x))" by (rule_tac h = munion in imp_Increasing_comp2, auto) lemma Always_munion: "\F \ Always({s \ state. \ MultLe(A,r)}); F \ Always({s \ state. \ MultLe(A,r)}); \x \ state. f1(x):Mult(A)\f(x):Mult(A) \ g1(x):Mult(A) \ g(x):Mult(A)\ \ F \ Always({s \ state. \ MultLe(A,r)})" apply (rule_tac h = munion in imp_Always_comp2, simp_all) apply (blast intro: munion_mono, simp_all) done lemma Follows_munion: "\F \ Follows(Mult(A), MultLe(A, r), f1, f); F \ Follows(Mult(A), MultLe(A, r), g1, g)\ \ F \ Follows(Mult(A), MultLe(A, r), \s. f1(s) +# g1(s), \s. f(s) +# g(s))" by (rule_tac h = munion in imp_Follows_comp2, auto) (** Used in ClientImp **) lemma Follows_msetsum_UN: "\f. \\i \ I. F \ Follows(Mult(A), MultLe(A, r), f'(i), f(i)); \s. \i \ I. multiset(f'(i, s)) \ mset_of(f'(i, s))<=A \ multiset(f(i, s)) \ mset_of(f(i, s))<=A ; Finite(I); F \ program\ \ F \ Follows(Mult(A), MultLe(A, r), \x. msetsum(\i. f'(i, x), I, A), \x. msetsum(\i. f(i, x), I, A))" apply (erule rev_mp) apply (drule Finite_into_Fin) apply (erule Fin_induct) apply (simp (no_asm_simp)) apply (rule Follows_constantI) apply (simp_all (no_asm_simp) add: FiniteFun.intros) apply auto apply (rule Follows_munion, auto) done end diff --git a/src/ZF/UNITY/GenPrefix.thy b/src/ZF/UNITY/GenPrefix.thy --- a/src/ZF/UNITY/GenPrefix.thy +++ b/src/ZF/UNITY/GenPrefix.thy @@ -1,680 +1,680 @@ (* Title: ZF/UNITY/GenPrefix.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2001 University of Cambridge \xs,ys\:gen_prefix(r) if ys = xs' @ zs where length(xs) = length(xs') and corresponding elements of xs, xs' are pairwise related by r Based on Lex/Prefix *) section\Charpentier's Generalized Prefix Relation\ theory GenPrefix imports ZF begin definition (*really belongs in ZF/Trancl*) part_order :: "[i, i] \ o" where "part_order(A, r) \ refl(A,r) \ trans[A](r) \ antisym(r)" consts gen_prefix :: "[i, i] \ i" inductive (* Parameter A is the domain of zs's elements *) domains "gen_prefix(A, r)" \ "list(A)*list(A)" intros Nil: "<[],[]>:gen_prefix(A, r)" prepend: "\\xs,ys\:gen_prefix(A, r); \x,y\:r; x \ A; y \ A\ \ : gen_prefix(A, r)" append: "\\xs,ys\:gen_prefix(A, r); zs:list(A)\ \ :gen_prefix(A, r)" type_intros app_type list.Nil list.Cons definition prefix :: "i\i" where "prefix(A) \ gen_prefix(A, id(A))" definition strict_prefix :: "i\i" where "strict_prefix(A) \ prefix(A) - id(list(A))" (* less or equal and greater or equal over prefixes *) abbreviation pfixLe :: "[i, i] \ o" (infixl \pfixLe\ 50) where "xs pfixLe ys \ \xs, ys\:gen_prefix(nat, Le)" abbreviation pfixGe :: "[i, i] \ o" (infixl \pfixGe\ 50) where "xs pfixGe ys \ \xs, ys\:gen_prefix(nat, Ge)" lemma reflD: "\refl(A, r); x \ A\ \ \x,x\:r" apply (unfold refl_def, auto) done (*** preliminary lemmas ***) lemma Nil_gen_prefix: "xs \ list(A) \ <[], xs> \ gen_prefix(A, r)" by (drule gen_prefix.append [OF gen_prefix.Nil], simp) declare Nil_gen_prefix [simp] lemma gen_prefix_length_le: "\xs,ys\ \ gen_prefix(A, r) \ length(xs) \ length(ys)" apply (erule gen_prefix.induct) apply (subgoal_tac [3] "ys \ list (A) ") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] intro: le_trans simp add: length_app) done lemma Cons_gen_prefix_aux: "\ \ gen_prefix(A, r)\ \ (\x xs. x \ A \ xs'= Cons(x,xs) \ (\y ys. y \ A \ ys' = Cons(y,ys) \ \x,y\:r \ \xs, ys\ \ gen_prefix(A, r)))" apply (erule gen_prefix.induct) prefer 3 apply (force intro: gen_prefix.append, auto) done lemma Cons_gen_prefixE: "\ \ gen_prefix(A, r); \y ys. \zs = Cons(y, ys); y \ A; x \ A; \x,y\:r; \xs,ys\ \ gen_prefix(A, r)\ \ P\ \ P" apply (frule gen_prefix.dom_subset [THEN subsetD], auto) apply (blast dest: Cons_gen_prefix_aux) done declare Cons_gen_prefixE [elim!] lemma Cons_gen_prefix_Cons: "( \ gen_prefix(A, r)) \ (x \ A \ y \ A \ \x,y\:r \ \xs,ys\ \ gen_prefix(A, r))" apply (auto intro: gen_prefix.prepend) done declare Cons_gen_prefix_Cons [iff] (** Monotonicity of gen_prefix **) lemma gen_prefix_mono2: "r<=s \ gen_prefix(A, r) \ gen_prefix(A, s)" apply clarify apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) apply (erule rev_mp) apply (erule gen_prefix.induct) apply (auto intro: gen_prefix.append) done lemma gen_prefix_mono1: "A<=B \gen_prefix(A, r) \ gen_prefix(B, r)" apply clarify apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) apply (erule rev_mp) apply (erule_tac P = "y \ list (A) " in rev_mp) apply (erule_tac P = "xa \ list (A) " in rev_mp) apply (erule gen_prefix.induct) apply (simp (no_asm_simp)) apply clarify apply (erule ConsE)+ apply (auto dest: gen_prefix.dom_subset [THEN subsetD] intro: gen_prefix.append list_mono [THEN subsetD]) done lemma gen_prefix_mono: "\A \ B; r \ s\ \ gen_prefix(A, r) \ gen_prefix(B, s)" apply (rule subset_trans) apply (rule gen_prefix_mono1) apply (rule_tac [2] gen_prefix_mono2, auto) done (*** gen_prefix order ***) (* reflexivity *) lemma refl_gen_prefix: "refl(A, r) \ refl(list(A), gen_prefix(A, r))" apply (unfold refl_def, auto) apply (induct_tac "x", auto) done declare refl_gen_prefix [THEN reflD, simp] (* Transitivity *) (* A lemma for proving gen_prefix_trans_comp *) lemma append_gen_prefix [rule_format (no_asm)]: "xs \ list(A) \ \zs. \ gen_prefix(A, r) \ \xs, zs\: gen_prefix(A, r)" apply (erule list.induct) apply (auto dest: gen_prefix.dom_subset [THEN subsetD]) done (* Lemma proving transitivity and more*) lemma gen_prefix_trans_comp [rule_format (no_asm)]: "\x, y\: gen_prefix(A, r) \ (\z \ list(A). \y,z\ \ gen_prefix(A, s)\\x, z\ \ gen_prefix(A, s O r))" apply (erule gen_prefix.induct) apply (auto elim: ConsE simp add: Nil_gen_prefix) apply (subgoal_tac "ys \ list (A) ") prefer 2 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) apply (drule_tac xs = ys and r = s in append_gen_prefix, auto) done lemma trans_comp_subset: "trans(r) \ r O r \ r" by (auto dest: transD) lemma trans_gen_prefix: "trans(r) \ trans(gen_prefix(A,r))" apply (simp (no_asm) add: trans_def) apply clarify apply (rule trans_comp_subset [THEN gen_prefix_mono2, THEN subsetD], assumption) apply (rule gen_prefix_trans_comp) apply (auto dest: gen_prefix.dom_subset [THEN subsetD]) done lemma trans_on_gen_prefix: "trans(r) \ trans[list(A)](gen_prefix(A, r))" apply (drule_tac A = A in trans_gen_prefix) apply (unfold trans_def trans_on_def, blast) done lemma prefix_gen_prefix_trans: "\\x,y\ \ prefix(A); \y, z\ \ gen_prefix(A, r); r<=A*A\ \ \x, z\ \ gen_prefix(A, r)" unfolding prefix_def apply (rule_tac P = "\r. \x,z\ \ gen_prefix (A, r) " in right_comp_id [THEN subst]) apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ done lemma gen_prefix_prefix_trans: "\\x,y\ \ gen_prefix(A,r); \y, z\ \ prefix(A); r<=A*A\ \ \x, z\ \ gen_prefix(A, r)" unfolding prefix_def apply (rule_tac P = "\r. \x,z\ \ gen_prefix (A, r) " in left_comp_id [THEN subst]) apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ done (** Antisymmetry **) lemma nat_le_lemma [rule_format]: "n \ nat \ \b \ nat. n #+ b \ n \ b = 0" by (induct_tac "n", auto) lemma antisym_gen_prefix: "antisym(r) \ antisym(gen_prefix(A, r))" apply (simp (no_asm) add: antisym_def) apply (rule impI [THEN allI, THEN allI]) apply (erule gen_prefix.induct, blast) apply (simp add: antisym_def, blast) txt\append case is hardest\ apply clarify apply (subgoal_tac "length (zs) = 0") apply (subgoal_tac "ys \ list (A) ") prefer 2 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) apply (drule_tac psi = " \ gen_prefix (A,r) " in asm_rl) apply simp apply (subgoal_tac "length (ys @ zs) = length (ys) #+ length (zs) \ys \ list (A) \xs \ list (A) ") prefer 2 apply (blast intro: length_app dest: gen_prefix.dom_subset [THEN subsetD]) apply (drule gen_prefix_length_le)+ apply clarify apply simp apply (drule_tac j = "length (xs) " in le_trans) apply blast apply (auto intro: nat_le_lemma) done (*** recursion equations ***) lemma gen_prefix_Nil: "xs \ list(A) \ \ gen_prefix(A,r) \ (xs = [])" by (induct_tac "xs", auto) declare gen_prefix_Nil [simp] lemma same_gen_prefix_gen_prefix: "\refl(A, r); xs \ list(A)\ \ : gen_prefix(A, r) \ \ys,zs\ \ gen_prefix(A, r)" unfolding refl_def apply (induct_tac "xs") apply (simp_all (no_asm_simp)) done declare same_gen_prefix_gen_prefix [simp] lemma gen_prefix_Cons: "\xs \ list(A); ys \ list(A); y \ A\ \ \ gen_prefix(A,r) \ (xs=[] | (\z zs. xs=Cons(z,zs) \ z \ A \ \z,y\:r \ \zs,ys\ \ gen_prefix(A,r)))" apply (induct_tac "xs", auto) done lemma gen_prefix_take_append: "\refl(A,r); \xs,ys\ \ gen_prefix(A, r); zs \ list(A)\ \ \ gen_prefix(A, r)" apply (erule gen_prefix.induct) apply (simp (no_asm_simp)) apply (frule_tac [!] gen_prefix.dom_subset [THEN subsetD], auto) apply (frule gen_prefix_length_le) apply (subgoal_tac "take (length (xs), ys) \ list (A) ") apply (simp_all (no_asm_simp) add: diff_is_0_iff [THEN iffD2] take_type) done lemma gen_prefix_append_both: "\refl(A, r); \xs,ys\ \ gen_prefix(A,r); length(xs) = length(ys); zs \ list(A)\ \ \ gen_prefix(A, r)" apply (drule_tac zs = zs in gen_prefix_take_append, assumption+) apply (subgoal_tac "take (length (xs), ys) =ys") apply (auto intro!: take_all dest: gen_prefix.dom_subset [THEN subsetD]) done (*NOT suitable for rewriting since [y] has the form y#ys*) lemma append_cons_conv: "xs \ list(A) \ xs @ Cons(y, ys) = (xs @ [y]) @ ys" by (auto simp add: app_assoc) lemma append_one_gen_prefix_lemma [rule_format]: "\\xs,ys\ \ gen_prefix(A, r); refl(A, r)\ \ length(xs) < length(ys) \ \ gen_prefix(A, r)" apply (erule gen_prefix.induct, blast) apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) apply (simp_all add: length_type) (* Append case is hardest *) apply (frule gen_prefix_length_le [THEN le_iff [THEN iffD1]]) apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) apply (subgoal_tac "length (xs) :nat\length (ys) :nat \length (zs) :nat") prefer 2 apply (blast intro: length_type, clarify) apply (simp_all add: nth_append length_type length_app) apply (rule conjI) apply (blast intro: gen_prefix.append) apply (erule_tac V = "length (xs) < length (ys) \u" for u in thin_rl) apply (erule_tac a = zs in list.cases, auto) apply (rule_tac P1 = "\x. :w" for u v w in nat_diff_split [THEN iffD2]) apply auto apply (simplesubst append_cons_conv) apply (rule_tac [2] gen_prefix.append) apply (auto elim: ConsE simp add: gen_prefix_append_both) done lemma append_one_gen_prefix: "\\xs,ys\: gen_prefix(A, r); length(xs) < length(ys); refl(A, r)\ \ \ gen_prefix(A, r)" apply (blast intro: append_one_gen_prefix_lemma) done (** Proving the equivalence with Charpentier's definition **) lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \ list(A) \ \ys \ list(A). \i \ nat. i < length(xs) \ \xs, ys\: gen_prefix(A, r) \ :r" apply (induct_tac "xs", simp, clarify) apply simp apply (erule natE, auto) done lemma gen_prefix_imp_nth: "\\xs,ys\ \ gen_prefix(A,r); i < length(xs)\ \ :r" apply (cut_tac A = A in gen_prefix.dom_subset) apply (rule gen_prefix_imp_nth_lemma) apply (auto simp add: lt_nat_in_nat) done lemma nth_imp_gen_prefix [rule_format]: "xs \ list(A) \ \ys \ list(A). length(xs) \ length(ys) \ (\i. i < length(xs) \ :r) \ \xs, ys\ \ gen_prefix(A, r)" apply (induct_tac "xs") apply (simp_all (no_asm_simp)) apply clarify apply (erule_tac a = ys in list.cases, simp) apply (force intro!: nat_0_le simp add: lt_nat_in_nat) done lemma gen_prefix_iff_nth: "(\xs,ys\ \ gen_prefix(A,r)) \ (xs \ list(A) \ ys \ list(A) \ length(xs) \ length(ys) \ (\i. i < length(xs) \ : r))" apply (rule iffI) apply (frule gen_prefix.dom_subset [THEN subsetD]) apply (frule gen_prefix_length_le, auto) apply (rule_tac [2] nth_imp_gen_prefix) apply (drule gen_prefix_imp_nth) apply (auto simp add: lt_nat_in_nat) done (** prefix is a partial order: **) lemma refl_prefix: "refl(list(A), prefix(A))" unfolding prefix_def apply (rule refl_gen_prefix) apply (auto simp add: refl_def) done declare refl_prefix [THEN reflD, simp] lemma trans_prefix: "trans(prefix(A))" unfolding prefix_def apply (rule trans_gen_prefix) apply (auto simp add: trans_def) done lemmas prefix_trans = trans_prefix [THEN transD] lemma trans_on_prefix: "trans[list(A)](prefix(A))" unfolding prefix_def apply (rule trans_on_gen_prefix) apply (auto simp add: trans_def) done lemmas prefix_trans_on = trans_on_prefix [THEN trans_onD] (* Monotonicity of "set" operator WRT prefix *) lemma set_of_list_prefix_mono: "\xs,ys\ \ prefix(A) \ set_of_list(xs) \ set_of_list(ys)" unfolding prefix_def apply (erule gen_prefix.induct) apply (subgoal_tac [3] "xs \ list (A) \ys \ list (A) ") prefer 4 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) apply (auto simp add: set_of_list_append) done (** recursion equations **) lemma Nil_prefix: "xs \ list(A) \ <[],xs> \ prefix(A)" unfolding prefix_def apply (simp (no_asm_simp) add: Nil_gen_prefix) done declare Nil_prefix [simp] lemma prefix_Nil: " \ prefix(A) \ (xs = [])" apply (unfold prefix_def, auto) apply (frule gen_prefix.dom_subset [THEN subsetD]) apply (drule_tac psi = " \ gen_prefix (A, id (A))" in asm_rl) apply (simp add: gen_prefix_Nil) done declare prefix_Nil [iff] lemma Cons_prefix_Cons: " \ prefix(A) \ (x=y \ \xs,ys\ \ prefix(A) \ y \ A)" apply (unfold prefix_def, auto) done declare Cons_prefix_Cons [iff] lemma same_prefix_prefix: "xs \ list(A)\ \ prefix(A) \ (\ys,zs\ \ prefix(A))" unfolding prefix_def apply (subgoal_tac "refl (A,id (A))") apply (simp (no_asm_simp)) apply (auto simp add: refl_def) done declare same_prefix_prefix [simp] lemma same_prefix_prefix_Nil: "xs \ list(A) \ \ prefix(A) \ ( \ prefix(A))" apply (rule_tac P = "\x. \u, x\:v \ w(x)" for u v w in app_right_Nil [THEN subst]) apply (rule_tac [2] same_prefix_prefix, auto) done declare same_prefix_prefix_Nil [simp] lemma prefix_appendI: "\\xs,ys\ \ prefix(A); zs \ list(A)\ \ \ prefix(A)" unfolding prefix_def apply (erule gen_prefix.append, assumption) done declare prefix_appendI [simp] lemma prefix_Cons: "\xs \ list(A); ys \ list(A); y \ A\ \ \ prefix(A) \ (xs=[] | (\zs. xs=Cons(y,zs) \ \zs,ys\ \ prefix(A)))" unfolding prefix_def apply (auto simp add: gen_prefix_Cons) done lemma append_one_prefix: "\\xs,ys\ \ prefix(A); length(xs) < length(ys)\ \ \ prefix(A)" unfolding prefix_def apply (subgoal_tac "refl (A, id (A))") apply (simp (no_asm_simp) add: append_one_gen_prefix) apply (auto simp add: refl_def) done lemma prefix_length_le: "\xs,ys\ \ prefix(A) \ length(xs) \ length(ys)" unfolding prefix_def apply (blast dest: gen_prefix_length_le) done lemma prefix_type: "prefix(A)<=list(A)*list(A)" unfolding prefix_def apply (blast intro!: gen_prefix.dom_subset) done lemma strict_prefix_type: "strict_prefix(A) \ list(A)*list(A)" unfolding strict_prefix_def apply (blast intro!: prefix_type [THEN subsetD]) done lemma strict_prefix_length_lt_aux: "\xs,ys\ \ prefix(A) \ xs\ys \ length(xs) < length(ys)" unfolding prefix_def apply (erule gen_prefix.induct, clarify) apply (subgoal_tac [!] "ys \ list(A) \ xs \ list(A)") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: length_type) apply (subgoal_tac "length (zs) =0") apply (drule_tac [2] not_lt_imp_le) apply (rule_tac [5] j = "length (ys) " in lt_trans2) apply auto done lemma strict_prefix_length_lt: "\xs,ys\:strict_prefix(A) \ length(xs) < length(ys)" unfolding strict_prefix_def apply (rule strict_prefix_length_lt_aux [THEN mp]) apply (auto dest: prefix_type [THEN subsetD]) done (*Equivalence to the definition used in Lex/Prefix.thy*) lemma prefix_iff: "\xs,zs\ \ prefix(A) \ (\ys \ list(A). zs = xs@ys) \ xs \ list(A)" unfolding prefix_def apply (auto simp add: gen_prefix_iff_nth lt_nat_in_nat nth_append nth_type app_type length_app) apply (subgoal_tac "drop (length (xs), zs) \ list (A) ") apply (rule_tac x = "drop (length (xs), zs) " in bexI) apply safe prefer 2 apply (simp add: length_type drop_type) apply (rule nth_equalityI) apply (simp_all (no_asm_simp) add: nth_append app_type drop_type length_app length_drop) apply (rule nat_diff_split [THEN iffD2], simp_all, clarify) apply (drule_tac i = "length (zs) " in leI) apply (force simp add: le_subset_iff, safe) apply (subgoal_tac "length (xs) #+ (i #- length (xs)) = i") apply (subst nth_drop) apply (simp_all (no_asm_simp) add: leI split: nat_diff_split) done lemma prefix_snoc: "\xs \ list(A); ys \ list(A); y \ A\ \ \ prefix(A) \ (xs = ys@[y] | \xs,ys\ \ prefix(A))" apply (simp (no_asm) add: prefix_iff) apply (rule iffI, clarify) apply (erule_tac xs = ysa in rev_list_elim, simp) apply (simp add: app_type app_assoc [symmetric]) apply (auto simp add: app_assoc app_type) done declare prefix_snoc [simp] lemma prefix_append_iff [rule_format]: "zs \ list(A) \ \xs \ list(A). \ys \ list(A). ( \ prefix(A)) \ (\xs,ys\ \ prefix(A) | (\us. xs = ys@us \ \us,zs\ \ prefix(A)))" apply (erule list_append_induct, force, clarify) apply (rule iffI) apply (simp add: add: app_assoc [symmetric]) apply (erule disjE) apply (rule disjI2) apply (rule_tac x = "y @ [x]" in exI) apply (simp add: add: app_assoc [symmetric], force+) done (*Although the prefix ordering is not linear, the prefixes of a list are linearly ordered.*) lemma common_prefix_linear_lemma [rule_format]: "\zs \ list(A); xs \ list(A); ys \ list(A)\ \ \xs, zs\ \ prefix(A) \ \ys,zs\ \ prefix(A) \\xs,ys\ \ prefix(A) | \ys,xs\ \ prefix(A)" apply (erule list_append_induct, auto) done lemma common_prefix_linear: "\\xs, zs\ \ prefix(A); \ys,zs\ \ prefix(A)\ \ \xs,ys\ \ prefix(A) | \ys,xs\ \ prefix(A)" apply (cut_tac prefix_type) apply (blast del: disjCI intro: common_prefix_linear_lemma) done (*** pfixLe, pfixGe \ properties inherited from the translations ***) (** pfixLe **) lemma refl_Le: "refl(nat,Le)" apply (unfold refl_def, auto) done declare refl_Le [simp] lemma antisym_Le: "antisym(Le)" unfolding antisym_def apply (auto intro: le_anti_sym) done declare antisym_Le [simp] lemma trans_on_Le: "trans[nat](Le)" apply (unfold trans_on_def, auto) apply (blast intro: le_trans) done declare trans_on_Le [simp] lemma trans_Le: "trans(Le)" apply (unfold trans_def, auto) apply (blast intro: le_trans) done declare trans_Le [simp] lemma part_order_Le: "part_order(nat,Le)" by (unfold part_order_def, auto) declare part_order_Le [simp] lemma pfixLe_refl: "x \ list(nat) \ x pfixLe x" by (blast intro: refl_gen_prefix [THEN reflD] refl_Le) declare pfixLe_refl [simp] lemma pfixLe_trans: "\x pfixLe y; y pfixLe z\ \ x pfixLe z" by (blast intro: trans_gen_prefix [THEN transD] trans_Le) lemma pfixLe_antisym: "\x pfixLe y; y pfixLe x\ \ x = y" by (blast intro: antisym_gen_prefix [THEN antisymE] antisym_Le) lemma prefix_imp_pfixLe: "\xs,ys\:prefix(nat)\ xs pfixLe ys" unfolding prefix_def apply (rule gen_prefix_mono [THEN subsetD], auto) done lemma refl_Ge: "refl(nat, Ge)" by (unfold refl_def Ge_def, auto) declare refl_Ge [iff] lemma antisym_Ge: "antisym(Ge)" -apply (unfold antisym_def Ge_def) + unfolding antisym_def Ge_def apply (auto intro: le_anti_sym) done declare antisym_Ge [iff] lemma trans_Ge: "trans(Ge)" -apply (unfold trans_def Ge_def) + unfolding trans_def Ge_def apply (auto intro: le_trans) done declare trans_Ge [iff] lemma pfixGe_refl: "x \ list(nat) \ x pfixGe x" by (blast intro: refl_gen_prefix [THEN reflD]) declare pfixGe_refl [simp] lemma pfixGe_trans: "\x pfixGe y; y pfixGe z\ \ x pfixGe z" by (blast intro: trans_gen_prefix [THEN transD]) lemma pfixGe_antisym: "\x pfixGe y; y pfixGe x\ \ x = y" by (blast intro: antisym_gen_prefix [THEN antisymE]) lemma prefix_imp_pfixGe: "\xs,ys\:prefix(nat) \ xs pfixGe ys" -apply (unfold prefix_def Ge_def) + unfolding prefix_def Ge_def apply (rule gen_prefix_mono [THEN subsetD], auto) done (* Added by Sidi \ prefix and take *) lemma prefix_imp_take: "\xs, ys\ \ prefix(A) \ xs = take(length(xs), ys)" unfolding prefix_def apply (erule gen_prefix.induct) apply (subgoal_tac [3] "length (xs) :nat") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: length_type) apply (frule gen_prefix.dom_subset [THEN subsetD]) apply (frule gen_prefix_length_le) apply (auto simp add: take_append) apply (subgoal_tac "length (xs) #- length (ys) =0") apply (simp_all (no_asm_simp) add: diff_is_0_iff) done lemma prefix_length_equal: "\\xs,ys\ \ prefix(A); length(xs)=length(ys)\ \ xs = ys" apply (cut_tac A = A in prefix_type) apply (drule subsetD, auto) apply (drule prefix_imp_take) apply (erule trans, simp) done lemma prefix_length_le_equal: "\\xs,ys\ \ prefix(A); length(ys) \ length(xs)\ \ xs = ys" by (blast intro: prefix_length_equal le_anti_sym prefix_length_le) lemma take_prefix [rule_format]: "xs \ list(A) \ \n \ nat. \ prefix(A)" unfolding prefix_def apply (erule list.induct, simp, clarify) apply (erule natE, auto) done lemma prefix_take_iff: "\xs,ys\ \ prefix(A) \ (xs=take(length(xs), ys) \ xs \ list(A) \ ys \ list(A))" apply (rule iffI) apply (frule prefix_type [THEN subsetD]) apply (blast intro: prefix_imp_take, clarify) apply (erule ssubst) apply (blast intro: take_prefix length_type) done lemma prefix_imp_nth: "\\xs,ys\ \ prefix(A); i < length(xs)\ \ nth(i,xs) = nth(i,ys)" by (auto dest!: gen_prefix_imp_nth simp add: prefix_def) lemma nth_imp_prefix: "\xs \ list(A); ys \ list(A); length(xs) \ length(ys); \i. i < length(xs) \ nth(i, xs) = nth(i,ys)\ \ \xs,ys\ \ prefix(A)" apply (auto simp add: prefix_def nth_imp_gen_prefix) apply (auto intro!: nth_imp_gen_prefix simp add: prefix_def) apply (blast intro: nth_type lt_trans2) done lemma length_le_prefix_imp_prefix: "\length(xs) \ length(ys); \xs,zs\ \ prefix(A); \ys,zs\ \ prefix(A)\ \ \xs,ys\ \ prefix(A)" apply (cut_tac A = A in prefix_type) apply (rule nth_imp_prefix, blast, blast) apply assumption apply (rule_tac b = "nth (i,zs)" in trans) apply (blast intro: prefix_imp_nth) apply (blast intro: sym prefix_imp_nth prefix_length_le lt_trans2) done end diff --git a/src/ZF/UNITY/Increasing.thy b/src/ZF/UNITY/Increasing.thy --- a/src/ZF/UNITY/Increasing.thy +++ b/src/ZF/UNITY/Increasing.thy @@ -1,227 +1,227 @@ (* Title: ZF/UNITY/Increasing.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2001 University of Cambridge Increasing's parameters are a state function f, a domain A and an order relation r over the domain A. *) section\Charpentier's "Increasing" Relation\ theory Increasing imports Constrains Monotonicity begin definition increasing :: "[i, i, i\i] \ i" (\increasing[_]'(_, _')\) where "increasing[A](r, f) \ {F \ program. (\k \ A. F \ stable({s \ state. \ r})) \ (\x \ state. f(x):A)}" definition Increasing :: "[i, i, i\i] \ i" (\Increasing[_]'(_, _')\) where "Increasing[A](r, f) \ {F \ program. (\k \ A. F \ Stable({s \ state. \ r})) \ (\x \ state. f(x):A)}" abbreviation (input) IncWrt :: "[i\i, i, i] \ i" (\(_ IncreasingWrt _ '/ _)\ [60, 0, 60] 60) where "f IncreasingWrt r/A \ Increasing[A](r,f)" (** increasing **) lemma increasing_type: "increasing[A](r, f) \ program" by (unfold increasing_def, blast) lemma increasing_into_program: "F \ increasing[A](r, f) \ F \ program" by (unfold increasing_def, blast) lemma increasing_imp_stable: "\F \ increasing[A](r, f); x \ A\ \F \ stable({s \ state. :r})" by (unfold increasing_def, blast) lemma increasingD: "F \ increasing[A](r,f) \ F \ program \ (\a. a \ A) \ (\s \ state. f(s):A)" unfolding increasing_def apply (subgoal_tac "\x. x \ state") apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state) done lemma increasing_constant [simp]: "F \ increasing[A](r, \s. c) \ F \ program \ c \ A" -apply (unfold increasing_def stable_def) + unfolding increasing_def stable_def apply (subgoal_tac "\x. x \ state") apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state) done lemma subset_increasing_comp: "\mono1(A, r, B, s, g); refl(A, r); trans[B](s)\ \ increasing[A](r, f) \ increasing[B](s, g comp f)" apply (unfold increasing_def stable_def part_order_def constrains_def mono1_def metacomp_def, clarify, simp) apply clarify apply (subgoal_tac "xa \ state") prefer 2 apply (blast dest!: ActsD) apply (subgoal_tac ":r") prefer 2 apply (force simp add: refl_def) apply (rotate_tac 5) apply (drule_tac x = "f (xb) " in bspec) apply (rotate_tac [2] -1) apply (drule_tac [2] x = act in bspec, simp_all) apply (drule_tac A = "act``u" and c = xa for u in subsetD, blast) apply (drule_tac x = "f(xa) " and x1 = "f(xb)" in bspec [THEN bspec]) apply (rule_tac [3] b = "g (f (xb))" and A = B in trans_onD) apply simp_all done lemma imp_increasing_comp: "\F \ increasing[A](r, f); mono1(A, r, B, s, g); refl(A, r); trans[B](s)\ \ F \ increasing[B](s, g comp f)" by (rule subset_increasing_comp [THEN subsetD], auto) lemma strict_increasing: "increasing[nat](Le, f) \ increasing[nat](Lt, f)" by (unfold increasing_def Lt_def, auto) lemma strict_gt_increasing: "increasing[nat](Ge, f) \ increasing[nat](Gt, f)" apply (unfold increasing_def Gt_def Ge_def, auto) apply (erule natE) apply (auto simp add: stable_def) done (** Increasing **) lemma increasing_imp_Increasing: "F \ increasing[A](r, f) \ F \ Increasing[A](r, f)" -apply (unfold increasing_def Increasing_def) + unfolding increasing_def Increasing_def apply (auto intro: stable_imp_Stable) done lemma Increasing_type: "Increasing[A](r, f) \ program" by (unfold Increasing_def, auto) lemma Increasing_into_program: "F \ Increasing[A](r, f) \ F \ program" by (unfold Increasing_def, auto) lemma Increasing_imp_Stable: "\F \ Increasing[A](r, f); a \ A\ \ F \ Stable({s \ state. :r})" by (unfold Increasing_def, blast) lemma IncreasingD: "F \ Increasing[A](r, f) \ F \ program \ (\a. a \ A) \ (\s \ state. f(s):A)" unfolding Increasing_def apply (subgoal_tac "\x. x \ state") apply (auto intro: st0_in_state) done lemma Increasing_constant [simp]: "F \ Increasing[A](r, \s. c) \ F \ program \ (c \ A)" apply (subgoal_tac "\x. x \ state") apply (auto dest!: IncreasingD intro: st0_in_state increasing_imp_Increasing) done lemma subset_Increasing_comp: "\mono1(A, r, B, s, g); refl(A, r); trans[B](s)\ \ Increasing[A](r, f) \ Increasing[B](s, g comp f)" apply (unfold Increasing_def Stable_def Constrains_def part_order_def constrains_def mono1_def metacomp_def, safe) apply (simp_all add: ActsD) apply (subgoal_tac "xb \ state \ xa \ state") prefer 2 apply (simp add: ActsD) apply (subgoal_tac ":r") prefer 2 apply (force simp add: refl_def) apply (rotate_tac 5) apply (drule_tac x = "f (xb) " in bspec) apply simp_all apply clarify apply (rotate_tac -2) apply (drule_tac x = act in bspec) apply (drule_tac [2] A = "act``u" and c = xa for u in subsetD, simp_all, blast) apply (drule_tac x = "f(xa)" and x1 = "f(xb)" in bspec [THEN bspec]) apply (rule_tac [3] b = "g (f (xb))" and A = B in trans_onD) apply simp_all done lemma imp_Increasing_comp: "\F \ Increasing[A](r, f); mono1(A, r, B, s, g); refl(A, r); trans[B](s)\ \ F \ Increasing[B](s, g comp f)" apply (rule subset_Increasing_comp [THEN subsetD], auto) done lemma strict_Increasing: "Increasing[nat](Le, f) \ Increasing[nat](Lt, f)" by (unfold Increasing_def Lt_def, auto) lemma strict_gt_Increasing: "Increasing[nat](Ge, f)<= Increasing[nat](Gt, f)" apply (unfold Increasing_def Ge_def Gt_def, auto) apply (erule natE) apply (auto simp add: Stable_def) done (** Two-place monotone operations **) lemma imp_increasing_comp2: "\F \ increasing[A](r, f); F \ increasing[B](s, g); mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\ \ F \ increasing[C](t, \x. h(f(x), g(x)))" apply (unfold increasing_def stable_def part_order_def constrains_def mono2_def, clarify, simp) apply clarify apply (rename_tac xa xb) apply (subgoal_tac "xb \ state \ xa \ state") prefer 2 apply (blast dest!: ActsD) apply (subgoal_tac ":r \ :s") prefer 2 apply (force simp add: refl_def) apply (rotate_tac 6) apply (drule_tac x = "f (xb) " in bspec) apply (rotate_tac [2] 1) apply (drule_tac [2] x = "g (xb) " in bspec) apply simp_all apply (rotate_tac -1) apply (drule_tac x = act in bspec) apply (rotate_tac [2] -3) apply (drule_tac [2] x = act in bspec, simp_all) apply (drule_tac A = "act``u" and c = xa for u in subsetD) apply (drule_tac [2] A = "act``u" and c = xa for u in subsetD, blast, blast) apply (rotate_tac -4) apply (drule_tac x = "f (xa) " and x1 = "f (xb) " in bspec [THEN bspec]) apply (rotate_tac [3] -1) apply (drule_tac [3] x = "g (xa) " and x1 = "g (xb) " in bspec [THEN bspec]) apply simp_all apply (rule_tac b = "h (f (xb), g (xb))" and A = C in trans_onD) apply simp_all done lemma imp_Increasing_comp2: "\F \ Increasing[A](r, f); F \ Increasing[B](s, g); mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\ \ F \ Increasing[C](t, \x. h(f(x), g(x)))" apply (unfold Increasing_def stable_def part_order_def constrains_def mono2_def Stable_def Constrains_def, safe) apply (simp_all add: ActsD) apply (subgoal_tac "xa \ state \ x \ state") prefer 2 apply (blast dest!: ActsD) apply (subgoal_tac ":r \ :s") prefer 2 apply (force simp add: refl_def) apply (rotate_tac 6) apply (drule_tac x = "f (xa) " in bspec) apply (rotate_tac [2] 1) apply (drule_tac [2] x = "g (xa) " in bspec) apply simp_all apply clarify apply (rotate_tac -2) apply (drule_tac x = act in bspec) apply (rotate_tac [2] -3) apply (drule_tac [2] x = act in bspec, simp_all) apply (drule_tac A = "act``u" and c = x for u in subsetD) apply (drule_tac [2] A = "act``u" and c = x for u in subsetD, blast, blast) apply (rotate_tac -9) apply (drule_tac x = "f (x) " and x1 = "f (xa) " in bspec [THEN bspec]) apply (rotate_tac [3] -1) apply (drule_tac [3] x = "g (x) " and x1 = "g (xa) " in bspec [THEN bspec]) apply simp_all apply (rule_tac b = "h (f (xa), g (xa))" and A = C in trans_onD) apply simp_all done end diff --git a/src/ZF/UNITY/Monotonicity.thy b/src/ZF/UNITY/Monotonicity.thy --- a/src/ZF/UNITY/Monotonicity.thy +++ b/src/ZF/UNITY/Monotonicity.thy @@ -1,120 +1,120 @@ (* Title: ZF/UNITY/Monotonicity.thy Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2002 University of Cambridge Monotonicity of an operator (meta-function) with respect to arbitrary set relations. *) section\Monotonicity of an Operator WRT a Relation\ theory Monotonicity imports GenPrefix MultisetSum begin definition mono1 :: "[i, i, i, i, i\i] \ o" where "mono1(A, r, B, s, f) \ (\x \ A. \y \ A. \x,y\ \ r \ \ s) \ (\x \ A. f(x) \ B)" (* monotonicity of a 2-place meta-function f *) definition mono2 :: "[i, i, i, i, i, i, [i,i]\i] \ o" where "mono2(A, r, B, s, C, t, f) \ (\x \ A. \y \ A. \u \ B. \v \ B. \x,y\ \ r \ \u,v\ \ s \ \ t) \ (\x \ A. \y \ B. f(x,y) \ C)" (* Internalized relations on sets and multisets *) definition SetLe :: "i \i" where "SetLe(A) \ {\x,y\ \ Pow(A)*Pow(A). x \ y}" definition MultLe :: "[i,i] \i" where "MultLe(A, r) \ multirel(A, r - id(A)) \ id(Mult(A))" lemma mono1D: "\mono1(A, r, B, s, f); \x, y\ \ r; x \ A; y \ A\ \ \ s" by (unfold mono1_def, auto) lemma mono2D: "\mono2(A, r, B, s, C, t, f); \x, y\ \ r; \u,v\ \ s; x \ A; y \ A; u \ B; v \ B\ \ \ t" by (unfold mono2_def, auto) (** Monotonicity of take **) lemma take_mono_left_lemma: "\i \ j; xs \ list(A); i \ nat; j \ nat\ \ \ prefix(A)" apply (case_tac "length (xs) \ i") apply (subgoal_tac "length (xs) \ j") apply (simp) apply (blast intro: le_trans) apply (drule not_lt_imp_le, auto) apply (case_tac "length (xs) \ j") apply (auto simp add: take_prefix) apply (drule not_lt_imp_le, auto) apply (drule_tac m = i in less_imp_succ_add, auto) apply (subgoal_tac "i #+ k \ length (xs) ") apply (simp add: take_add prefix_iff take_type drop_type) apply (blast intro: leI) done lemma take_mono_left: "\i \ j; xs \ list(A); j \ nat\ \ \ prefix(A)" by (blast intro: le_in_nat take_mono_left_lemma) lemma take_mono_right: "\\xs,ys\ \ prefix(A); i \ nat\ \ \ prefix(A)" by (auto simp add: prefix_iff) lemma take_mono: "\i \ j; \xs, ys\ \ prefix(A); j \ nat\ \ \ prefix(A)" apply (rule_tac b = "take (j, xs) " in prefix_trans) apply (auto dest: prefix_type [THEN subsetD] intro: take_mono_left take_mono_right) done lemma mono_take [iff]: "mono2(nat, Le, list(A), prefix(A), list(A), prefix(A), take)" apply (unfold mono2_def Le_def, auto) apply (blast intro: take_mono) done (** Monotonicity of length **) lemmas length_mono = prefix_length_le lemma mono_length [iff]: "mono1(list(A), prefix(A), nat, Le, length)" unfolding mono1_def apply (auto dest: prefix_length_le simp add: Le_def) done (** Monotonicity of \ **) lemma mono_Un [iff]: "mono2(Pow(A), SetLe(A), Pow(A), SetLe(A), Pow(A), SetLe(A), (Un))" by (unfold mono2_def SetLe_def, auto) (* Monotonicity of multiset union *) lemma mono_munion [iff]: "mono2(Mult(A), MultLe(A,r), Mult(A), MultLe(A, r), Mult(A), MultLe(A, r), munion)" -apply (unfold mono2_def MultLe_def) + unfolding mono2_def MultLe_def apply (auto simp add: Mult_iff_multiset) apply (blast intro: munion_multirel_mono munion_multirel_mono1 munion_multirel_mono2 multiset_into_Mult)+ done lemma mono_succ [iff]: "mono1(nat, Le, nat, Le, succ)" by (unfold mono1_def Le_def, auto) end diff --git a/src/ZF/UNITY/UNITY.thy b/src/ZF/UNITY/UNITY.thy --- a/src/ZF/UNITY/UNITY.thy +++ b/src/ZF/UNITY/UNITY.thy @@ -1,625 +1,625 @@ (* Title: ZF/UNITY/UNITY.thy Author: Sidi O Ehmety, Computer Laboratory Copyright 2001 University of Cambridge *) section \The Basic UNITY Theory\ theory UNITY imports State begin text\The basic UNITY theory (revised version, based upon the "co" operator) From Misra, "A Logic for Concurrent Programming", 1994. This ZF theory was ported from its HOL equivalent.\ definition program :: i where "program \ {: Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)). id(state) \ acts \ id(state) \ allowed}" definition mk_program :: "[i,i,i]\i" where \ \The definition yields a program thanks to the coercions init \ state, acts \ Pow(state*state), etc.\ "mk_program(init, acts, allowed) \ state, cons(id(state), acts \ Pow(state*state)), cons(id(state), allowed \ Pow(state*state))>" definition SKIP :: i (\\\) where "SKIP \ mk_program(state, 0, Pow(state*state))" (* Coercion from anything to program *) definition programify :: "i\i" where "programify(F) \ if F \ program then F else SKIP" definition RawInit :: "i\i" where "RawInit(F) \ fst(F)" definition Init :: "i\i" where "Init(F) \ RawInit(programify(F))" definition RawActs :: "i\i" where "RawActs(F) \ cons(id(state), fst(snd(F)))" definition Acts :: "i\i" where "Acts(F) \ RawActs(programify(F))" definition RawAllowedActs :: "i\i" where "RawAllowedActs(F) \ cons(id(state), snd(snd(F)))" definition AllowedActs :: "i\i" where "AllowedActs(F) \ RawAllowedActs(programify(F))" definition Allowed :: "i \i" where "Allowed(F) \ {G \ program. Acts(G) \ AllowedActs(F)}" definition initially :: "i\i" where "initially(A) \ {F \ program. Init(F)\A}" definition "constrains" :: "[i, i] \ i" (infixl \co\ 60) where "A co B \ {F \ program. (\act \ Acts(F). act``A\B) \ st_set(A)}" \ \the condition \<^term>\st_set(A)\ makes the definition slightly stronger than the HOL one\ definition unless :: "[i, i] \ i" (infixl \unless\ 60) where "A unless B \ (A - B) co (A \ B)" definition stable :: "i\i" where "stable(A) \ A co A" definition strongest_rhs :: "[i, i] \ i" where "strongest_rhs(F, A) \ \({B \ Pow(state). F \ A co B})" definition invariant :: "i \ i" where "invariant(A) \ initially(A) \ stable(A)" (* meta-function composition *) definition metacomp :: "[i\i, i\i] \ (i\i)" (infixl \comp\ 65) where "f comp g \ \x. f(g(x))" definition pg_compl :: "i\i" where "pg_compl(X)\ program - X" text\SKIP\ lemma SKIP_in_program [iff,TC]: "SKIP \ program" by (force simp add: SKIP_def program_def mk_program_def) subsection\The function \<^term>\programify\, the coercion from anything to program\ lemma programify_program [simp]: "F \ program \ programify(F)=F" by (force simp add: programify_def) lemma programify_in_program [iff,TC]: "programify(F) \ program" by (force simp add: programify_def) text\Collapsing rules: to remove programify from expressions\ lemma programify_idem [simp]: "programify(programify(F))=programify(F)" by (force simp add: programify_def) lemma Init_programify [simp]: "Init(programify(F)) = Init(F)" by (simp add: Init_def) lemma Acts_programify [simp]: "Acts(programify(F)) = Acts(F)" by (simp add: Acts_def) lemma AllowedActs_programify [simp]: "AllowedActs(programify(F)) = AllowedActs(F)" by (simp add: AllowedActs_def) subsection\The Inspectors for Programs\ lemma id_in_RawActs: "F \ program \id(state) \ RawActs(F)" by (auto simp add: program_def RawActs_def) lemma id_in_Acts [iff,TC]: "id(state) \ Acts(F)" by (simp add: id_in_RawActs Acts_def) lemma id_in_RawAllowedActs: "F \ program \id(state) \ RawAllowedActs(F)" by (auto simp add: program_def RawAllowedActs_def) lemma id_in_AllowedActs [iff,TC]: "id(state) \ AllowedActs(F)" by (simp add: id_in_RawAllowedActs AllowedActs_def) lemma cons_id_Acts [simp]: "cons(id(state), Acts(F)) = Acts(F)" by (simp add: cons_absorb) lemma cons_id_AllowedActs [simp]: "cons(id(state), AllowedActs(F)) = AllowedActs(F)" by (simp add: cons_absorb) subsection\Types of the Inspectors\ lemma RawInit_type: "F \ program \ RawInit(F)\state" by (auto simp add: program_def RawInit_def) lemma RawActs_type: "F \ program \ RawActs(F)\Pow(state*state)" by (auto simp add: program_def RawActs_def) lemma RawAllowedActs_type: "F \ program \ RawAllowedActs(F)\Pow(state*state)" by (auto simp add: program_def RawAllowedActs_def) lemma Init_type: "Init(F)\state" by (simp add: RawInit_type Init_def) lemmas InitD = Init_type [THEN subsetD] lemma st_set_Init [iff]: "st_set(Init(F))" unfolding st_set_def apply (rule Init_type) done lemma Acts_type: "Acts(F)\Pow(state*state)" by (simp add: RawActs_type Acts_def) lemma AllowedActs_type: "AllowedActs(F) \ Pow(state*state)" by (simp add: RawAllowedActs_type AllowedActs_def) text\Needed in Behaviors\ lemma ActsD: "\act \ Acts(F); \ act\ \ s \ state \ s' \ state" by (blast dest: Acts_type [THEN subsetD]) lemma AllowedActsD: "\act \ AllowedActs(F); \ act\ \ s \ state \ s' \ state" by (blast dest: AllowedActs_type [THEN subsetD]) subsection\Simplification rules involving \<^term>\state\, \<^term>\Init\, \<^term>\Acts\, and \<^term>\AllowedActs\\ text\But are they really needed?\ lemma state_subset_is_Init_iff [iff]: "state \ Init(F) \ Init(F)=state" by (cut_tac F = F in Init_type, auto) lemma Pow_state_times_state_is_subset_Acts_iff [iff]: "Pow(state*state) \ Acts(F) \ Acts(F)=Pow(state*state)" by (cut_tac F = F in Acts_type, auto) lemma Pow_state_times_state_is_subset_AllowedActs_iff [iff]: "Pow(state*state) \ AllowedActs(F) \ AllowedActs(F)=Pow(state*state)" by (cut_tac F = F in AllowedActs_type, auto) subsubsection\Eliminating \\ state\ from expressions\ lemma Init_Int_state [simp]: "Init(F) \ state = Init(F)" by (cut_tac F = F in Init_type, blast) lemma state_Int_Init [simp]: "state \ Init(F) = Init(F)" by (cut_tac F = F in Init_type, blast) lemma Acts_Int_Pow_state_times_state [simp]: "Acts(F) \ Pow(state*state) = Acts(F)" by (cut_tac F = F in Acts_type, blast) lemma state_times_state_Int_Acts [simp]: "Pow(state*state) \ Acts(F) = Acts(F)" by (cut_tac F = F in Acts_type, blast) lemma AllowedActs_Int_Pow_state_times_state [simp]: "AllowedActs(F) \ Pow(state*state) = AllowedActs(F)" by (cut_tac F = F in AllowedActs_type, blast) lemma state_times_state_Int_AllowedActs [simp]: "Pow(state*state) \ AllowedActs(F) = AllowedActs(F)" by (cut_tac F = F in AllowedActs_type, blast) subsubsection\The Operator \<^term>\mk_program\\ lemma mk_program_in_program [iff,TC]: "mk_program(init, acts, allowed) \ program" by (auto simp add: mk_program_def program_def) lemma RawInit_eq [simp]: "RawInit(mk_program(init, acts, allowed)) = init \ state" by (auto simp add: mk_program_def RawInit_def) lemma RawActs_eq [simp]: "RawActs(mk_program(init, acts, allowed)) = cons(id(state), acts \ Pow(state*state))" by (auto simp add: mk_program_def RawActs_def) lemma RawAllowedActs_eq [simp]: "RawAllowedActs(mk_program(init, acts, allowed)) = cons(id(state), allowed \ Pow(state*state))" by (auto simp add: mk_program_def RawAllowedActs_def) lemma Init_eq [simp]: "Init(mk_program(init, acts, allowed)) = init \ state" by (simp add: Init_def) lemma Acts_eq [simp]: "Acts(mk_program(init, acts, allowed)) = cons(id(state), acts \ Pow(state*state))" by (simp add: Acts_def) lemma AllowedActs_eq [simp]: "AllowedActs(mk_program(init, acts, allowed))= cons(id(state), allowed \ Pow(state*state))" by (simp add: AllowedActs_def) text\Init, Acts, and AlowedActs of SKIP\ lemma RawInit_SKIP [simp]: "RawInit(SKIP) = state" by (simp add: SKIP_def) lemma RawAllowedActs_SKIP [simp]: "RawAllowedActs(SKIP) = Pow(state*state)" by (force simp add: SKIP_def) lemma RawActs_SKIP [simp]: "RawActs(SKIP) = {id(state)}" by (force simp add: SKIP_def) lemma Init_SKIP [simp]: "Init(SKIP) = state" by (force simp add: SKIP_def) lemma Acts_SKIP [simp]: "Acts(SKIP) = {id(state)}" by (force simp add: SKIP_def) lemma AllowedActs_SKIP [simp]: "AllowedActs(SKIP) = Pow(state*state)" by (force simp add: SKIP_def) text\Equality of UNITY programs\ lemma raw_surjective_mk_program: "F \ program \ mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F" apply (auto simp add: program_def mk_program_def RawInit_def RawActs_def RawAllowedActs_def, blast+) done lemma surjective_mk_program [simp]: "mk_program(Init(F), Acts(F), AllowedActs(F)) = programify(F)" by (auto simp add: raw_surjective_mk_program Init_def Acts_def AllowedActs_def) lemma program_equalityI: "\Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G); F \ program; G \ program\ \ F = G" apply (subgoal_tac "programify(F) = programify(G)") apply simp apply (simp only: surjective_mk_program [symmetric]) done lemma program_equalityE: "\F = G; \Init(F) = Init(G); Acts(F) = Acts(G); AllowedActs(F) = AllowedActs(G)\ \ P\ \ P" by force lemma program_equality_iff: "\F \ program; G \ program\ \(F=G) \ (Init(F) = Init(G) \ Acts(F) = Acts(G) \ AllowedActs(F) = AllowedActs(G))" by (blast intro: program_equalityI program_equalityE) subsection\These rules allow "lazy" definition expansion\ lemma def_prg_Init: "F \ mk_program (init,acts,allowed) \ Init(F) = init \ state" by auto lemma def_prg_Acts: "F \ mk_program (init,acts,allowed) \ Acts(F) = cons(id(state), acts \ Pow(state*state))" by auto lemma def_prg_AllowedActs: "F \ mk_program (init,acts,allowed) \ AllowedActs(F) = cons(id(state), allowed \ Pow(state*state))" by auto lemma def_prg_simps: "\F \ mk_program (init,acts,allowed)\ \ Init(F) = init \ state \ Acts(F) = cons(id(state), acts \ Pow(state*state)) \ AllowedActs(F) = cons(id(state), allowed \ Pow(state*state))" by auto text\An action is expanded only if a pair of states is being tested against it\ lemma def_act_simp: "\act \ { \ A*B. P(s, s')}\ \ ( \ act) \ ( \ A*B \ P(s, s'))" by auto text\A set is expanded only if an element is being tested against it\ lemma def_set_simp: "A \ B \ (x \ A) \ (x \ B)" by auto subsection\The Constrains Operator\ lemma constrains_type: "A co B \ program" by (force simp add: constrains_def) lemma constrainsI: "\(\act s s'. \act: Acts(F); \ act; s \ A\ \ s' \ A'); F \ program; st_set(A)\ \ F \ A co A'" by (force simp add: constrains_def) lemma constrainsD: "F \ A co B \ \act \ Acts(F). act``A\B" by (force simp add: constrains_def) lemma constrainsD2: "F \ A co B \ F \ program \ st_set(A)" by (force simp add: constrains_def) lemma constrains_empty [iff]: "F \ 0 co B \ F \ program" by (force simp add: constrains_def st_set_def) lemma constrains_empty2 [iff]: "(F \ A co 0) \ (A=0 \ F \ program)" by (force simp add: constrains_def st_set_def) lemma constrains_state [iff]: "(F \ state co B) \ (state\B \ F \ program)" apply (cut_tac F = F in Acts_type) apply (force simp add: constrains_def st_set_def) done lemma constrains_state2 [iff]: "F \ A co state \ (F \ program \ st_set(A))" apply (cut_tac F = F in Acts_type) apply (force simp add: constrains_def st_set_def) done text\monotonic in 2nd argument\ lemma constrains_weaken_R: "\F \ A co A'; A'\B'\ \ F \ A co B'" apply (unfold constrains_def, blast) done text\anti-monotonic in 1st argument\ lemma constrains_weaken_L: "\F \ A co A'; B\A\ \ F \ B co A'" apply (unfold constrains_def st_set_def, blast) done lemma constrains_weaken: "\F \ A co A'; B\A; A'\B'\ \ F \ B co B'" apply (drule constrains_weaken_R) apply (drule_tac [2] constrains_weaken_L, blast+) done subsection\Constrains and Union\ lemma constrains_Un: "\F \ A co A'; F \ B co B'\ \ F \ (A \ B) co (A' \ B')" by (auto simp add: constrains_def st_set_def, force) lemma constrains_UN: "\\i. i \ I \ F \ A(i) co A'(i); F \ program\ \ F \ (\i \ I. A(i)) co (\i \ I. A'(i))" by (force simp add: constrains_def st_set_def) lemma constrains_Un_distrib: "(A \ B) co C = (A co C) \ (B co C)" by (force simp add: constrains_def st_set_def) lemma constrains_UN_distrib: "i \ I \ (\i \ I. A(i)) co B = (\i \ I. A(i) co B)" by (force simp add: constrains_def st_set_def) subsection\Constrains and Intersection\ lemma constrains_Int_distrib: "C co (A \ B) = (C co A) \ (C co B)" by (force simp add: constrains_def st_set_def) lemma constrains_INT_distrib: "x \ I \ A co (\i \ I. B(i)) = (\i \ I. A co B(i))" by (force simp add: constrains_def st_set_def) lemma constrains_Int: "\F \ A co A'; F \ B co B'\ \ F \ (A \ B) co (A' \ B')" by (force simp add: constrains_def st_set_def) lemma constrains_INT [rule_format]: "\\i \ I. F \ A(i) co A'(i); F \ program\ \ F \ (\i \ I. A(i)) co (\i \ I. A'(i))" apply (case_tac "I=0") apply (simp add: Inter_def) apply (erule not_emptyE) apply (auto simp add: constrains_def st_set_def, blast) apply (drule bspec, assumption, force) done (* The rule below simulates the HOL's one for (\z. A i) co (\z. B i) *) lemma constrains_All: "\\z. F:{s \ state. P(s, z)} co {s \ state. Q(s, z)}; F \ program\\ F:{s \ state. \z. P(s, z)} co {s \ state. \z. Q(s, z)}" by (unfold constrains_def, blast) lemma constrains_imp_subset: "\F \ A co A'\ \ A \ A'" by (unfold constrains_def st_set_def, force) text\The reasoning is by subsets since "co" refers to single actions only. So this rule isn't that useful.\ lemma constrains_trans: "\F \ A co B; F \ B co C\ \ F \ A co C" by (unfold constrains_def st_set_def, auto, blast) lemma constrains_cancel: "\F \ A co (A' \ B); F \ B co B'\ \ F \ A co (A' \ B')" apply (drule_tac A = B in constrains_imp_subset) apply (blast intro: constrains_weaken_R) done subsection\The Unless Operator\ lemma unless_type: "A unless B \ program" by (force simp add: unless_def constrains_def) lemma unlessI: "\F \ (A-B) co (A \ B)\ \ F \ A unless B" unfolding unless_def apply (blast dest: constrainsD2) done lemma unlessD: "F :A unless B \ F \ (A-B) co (A \ B)" by (unfold unless_def, auto) subsection\The Operator \<^term>\initially\\ lemma initially_type: "initially(A) \ program" by (unfold initially_def, blast) lemma initiallyI: "\F \ program; Init(F)\A\ \ F \ initially(A)" by (unfold initially_def, blast) lemma initiallyD: "F \ initially(A) \ Init(F)\A" by (unfold initially_def, blast) subsection\The Operator \<^term>\stable\\ lemma stable_type: "stable(A)\program" by (unfold stable_def constrains_def, blast) lemma stableI: "F \ A co A \ F \ stable(A)" by (unfold stable_def, assumption) lemma stableD: "F \ stable(A) \ F \ A co A" by (unfold stable_def, assumption) lemma stableD2: "F \ stable(A) \ F \ program \ st_set(A)" by (unfold stable_def constrains_def, auto) lemma stable_state [simp]: "stable(state) = program" by (auto simp add: stable_def constrains_def dest: Acts_type [THEN subsetD]) lemma stable_unless: "stable(A)= A unless 0" by (auto simp add: unless_def stable_def) subsection\Union and Intersection with \<^term>\stable\\ lemma stable_Un: "\F \ stable(A); F \ stable(A')\ \ F \ stable(A \ A')" unfolding stable_def apply (blast intro: constrains_Un) done lemma stable_UN: "\\i. i\I \ F \ stable(A(i)); F \ program\ \ F \ stable (\i \ I. A(i))" unfolding stable_def apply (blast intro: constrains_UN) done lemma stable_Int: "\F \ stable(A); F \ stable(A')\ \ F \ stable (A \ A')" unfolding stable_def apply (blast intro: constrains_Int) done lemma stable_INT: "\\i. i \ I \ F \ stable(A(i)); F \ program\ \ F \ stable (\i \ I. A(i))" unfolding stable_def apply (blast intro: constrains_INT) done lemma stable_All: "\\z. F \ stable({s \ state. P(s, z)}); F \ program\ \ F \ stable({s \ state. \z. P(s, z)})" unfolding stable_def apply (rule constrains_All, auto) done lemma stable_constrains_Un: "\F \ stable(C); F \ A co (C \ A')\ \ F \ (C \ A) co (C \ A')" apply (unfold stable_def constrains_def st_set_def, auto) apply (blast dest!: bspec) done lemma stable_constrains_Int: "\F \ stable(C); F \ (C \ A) co A'\ \ F \ (C \ A) co (C \ A')" by (unfold stable_def constrains_def st_set_def, blast) (* \F \ stable(C); F \ (C \ A) co A\ \ F \ stable(C \ A) *) lemmas stable_constrains_stable = stable_constrains_Int [THEN stableI] subsection\The Operator \<^term>\invariant\\ lemma invariant_type: "invariant(A) \ program" unfolding invariant_def apply (blast dest: stable_type [THEN subsetD]) done lemma invariantI: "\Init(F)\A; F \ stable(A)\ \ F \ invariant(A)" -apply (unfold invariant_def initially_def) + unfolding invariant_def initially_def apply (frule stable_type [THEN subsetD], auto) done lemma invariantD: "F \ invariant(A) \ Init(F)\A \ F \ stable(A)" by (unfold invariant_def initially_def, auto) lemma invariantD2: "F \ invariant(A) \ F \ program \ st_set(A)" unfolding invariant_def apply (blast dest: stableD2) done text\Could also say \<^term>\invariant(A) \ invariant(B) \ invariant (A \ B)\\ lemma invariant_Int: "\F \ invariant(A); F \ invariant(B)\ \ F \ invariant(A \ B)" -apply (unfold invariant_def initially_def) + unfolding invariant_def initially_def apply (simp add: stable_Int, blast) done subsection\The Elimination Theorem\ (** The "free" m has become universally quantified! Should the premise be \m instead of \m ? Would make it harder to use in forward proof. **) text\The general case is easier to prove than the special case!\ lemma "elimination": "\\m \ M. F \ {s \ A. x(s) = m} co B(m); F \ program\ \ F \ {s \ A. x(s) \ M} co (\m \ M. B(m))" by (auto simp add: constrains_def st_set_def, blast) text\As above, but for the special case of A=state\ lemma elimination2: "\\m \ M. F \ {s \ state. x(s) = m} co B(m); F \ program\ \ F:{s \ state. x(s) \ M} co (\m \ M. B(m))" by (rule UNITY.elimination, auto) subsection\The Operator \<^term>\strongest_rhs\\ lemma constrains_strongest_rhs: "\F \ program; st_set(A)\ \ F \ A co (strongest_rhs(F,A))" by (auto simp add: constrains_def strongest_rhs_def st_set_def dest: Acts_type [THEN subsetD]) lemma strongest_rhs_is_strongest: "\F \ A co B; st_set(B)\ \ strongest_rhs(F,A) \ B" by (auto simp add: constrains_def strongest_rhs_def st_set_def) ML \ fun simp_of_act def = def RS @{thm def_act_simp}; fun simp_of_set def = def RS @{thm def_set_simp}; \ end diff --git a/src/ZF/UNITY/Union.thy b/src/ZF/UNITY/Union.thy --- a/src/ZF/UNITY/Union.thy +++ b/src/ZF/UNITY/Union.thy @@ -1,576 +1,576 @@ (* Title: ZF/UNITY/Union.thy Author: Sidi O Ehmety, Computer Laboratory Copyright 2001 University of Cambridge Unions of programs Partly from Misra's Chapter 5 \ Asynchronous Compositions of Programs Theory ported form HOL.. *) theory Union imports SubstAx FP begin definition (*FIXME: conjoin Init(F) \ Init(G) \ 0 *) ok :: "[i, i] \ o" (infixl \ok\ 65) where "F ok G \ Acts(F) \ AllowedActs(G) \ Acts(G) \ AllowedActs(F)" definition (*FIXME: conjoin (\i \ I. Init(F(i))) \ 0 *) OK :: "[i, i\i] \ o" where "OK(I,F) \ (\i \ I. \j \ I-{i}. Acts(F(i)) \ AllowedActs(F(j)))" definition JOIN :: "[i, i\i] \ i" where "JOIN(I,F) \ if I = 0 then SKIP else mk_program(\i \ I. Init(F(i)), \i \ I. Acts(F(i)), \i \ I. AllowedActs(F(i)))" definition Join :: "[i, i] \ i" (infixl \\\ 65) where "F \ G \ mk_program (Init(F) \ Init(G), Acts(F) \ Acts(G), AllowedActs(F) \ AllowedActs(G))" definition (*Characterizes safety properties. Used with specifying AllowedActs*) safety_prop :: "i \ o" where "safety_prop(X) \ X\program \ SKIP \ X \ (\G \ program. Acts(G) \ (\F \ X. Acts(F)) \ G \ X)" syntax "_JOIN1" :: "[pttrns, i] \ i" (\(3\_./ _)\ 10) "_JOIN" :: "[pttrn, i, i] \ i" (\(3\_ \ _./ _)\ 10) translations "\x \ A. B" == "CONST JOIN(A, (\x. B))" "\x y. B" == "\x. \y. B" "\x. B" == "CONST JOIN(CONST state, (\x. B))" subsection\SKIP\ lemma reachable_SKIP [simp]: "reachable(SKIP) = state" by (force elim: reachable.induct intro: reachable.intros) text\Elimination programify from ok and \\ lemma ok_programify_left [iff]: "programify(F) ok G \ F ok G" by (simp add: ok_def) lemma ok_programify_right [iff]: "F ok programify(G) \ F ok G" by (simp add: ok_def) lemma Join_programify_left [simp]: "programify(F) \ G = F \ G" by (simp add: Join_def) lemma Join_programify_right [simp]: "F \ programify(G) = F \ G" by (simp add: Join_def) subsection\SKIP and safety properties\ lemma SKIP_in_constrains_iff [iff]: "(SKIP \ A co B) \ (A\B \ st_set(A))" by (unfold constrains_def st_set_def, auto) lemma SKIP_in_Constrains_iff [iff]: "(SKIP \ A Co B)\ (state \ A\B)" by (unfold Constrains_def, auto) lemma SKIP_in_stable [iff]: "SKIP \ stable(A) \ st_set(A)" by (auto simp add: stable_def) lemma SKIP_in_Stable [iff]: "SKIP \ Stable(A)" by (unfold Stable_def, auto) subsection\Join and JOIN types\ lemma Join_in_program [iff,TC]: "F \ G \ program" by (unfold Join_def, auto) lemma JOIN_in_program [iff,TC]: "JOIN(I,F) \ program" by (unfold JOIN_def, auto) subsection\Init, Acts, and AllowedActs of Join and JOIN\ lemma Init_Join [simp]: "Init(F \ G) = Init(F) \ Init(G)" by (simp add: Int_assoc Join_def) lemma Acts_Join [simp]: "Acts(F \ G) = Acts(F) \ Acts(G)" by (simp add: Int_Un_distrib2 cons_absorb Join_def) lemma AllowedActs_Join [simp]: "AllowedActs(F \ G) = AllowedActs(F) \ AllowedActs(G)" apply (simp add: Int_assoc cons_absorb Join_def) done subsection\Join's algebraic laws\ lemma Join_commute: "F \ G = G \ F" by (simp add: Join_def Un_commute Int_commute) lemma Join_left_commute: "A \ (B \ C) = B \ (A \ C)" apply (simp add: Join_def Int_Un_distrib2 cons_absorb) apply (simp add: Un_ac Int_ac Int_Un_distrib2 cons_absorb) done lemma Join_assoc: "(F \ G) \ H = F \ (G \ H)" by (simp add: Un_ac Join_def cons_absorb Int_assoc Int_Un_distrib2) subsection\Needed below\ lemma cons_id [simp]: "cons(id(state), Pow(state * state)) = Pow(state*state)" by auto lemma Join_SKIP_left [simp]: "SKIP \ F = programify(F)" -apply (unfold Join_def SKIP_def) + unfolding Join_def SKIP_def apply (auto simp add: Int_absorb cons_eq) done lemma Join_SKIP_right [simp]: "F \ SKIP = programify(F)" apply (subst Join_commute) apply (simp add: Join_SKIP_left) done lemma Join_absorb [simp]: "F \ F = programify(F)" by (rule program_equalityI, auto) lemma Join_left_absorb: "F \ (F \ G) = F \ G" by (simp add: Join_assoc [symmetric]) subsection\Join is an AC-operator\ lemmas Join_ac = Join_assoc Join_left_absorb Join_commute Join_left_commute subsection\Eliminating programify form JOIN and OK expressions\ lemma OK_programify [iff]: "OK(I, \x. programify(F(x))) \ OK(I, F)" by (simp add: OK_def) lemma JOIN_programify [iff]: "JOIN(I, \x. programify(F(x))) = JOIN(I, F)" by (simp add: JOIN_def) subsection\JOIN\ lemma JOIN_empty [simp]: "JOIN(0, F) = SKIP" by (unfold JOIN_def, auto) lemma Init_JOIN [simp]: "Init(\i \ I. F(i)) = (if I=0 then state else (\i \ I. Init(F(i))))" by (simp add: JOIN_def INT_extend_simps del: INT_simps) lemma Acts_JOIN [simp]: "Acts(JOIN(I,F)) = cons(id(state), \i \ I. Acts(F(i)))" unfolding JOIN_def apply (auto simp del: INT_simps UN_simps) apply (rule equalityI) apply (auto dest: Acts_type [THEN subsetD]) done lemma AllowedActs_JOIN [simp]: "AllowedActs(\i \ I. F(i)) = (if I=0 then Pow(state*state) else (\i \ I. AllowedActs(F(i))))" apply (unfold JOIN_def, auto) apply (rule equalityI) apply (auto elim!: not_emptyE dest: AllowedActs_type [THEN subsetD]) done lemma JOIN_cons [simp]: "(\i \ cons(a,I). F(i)) = F(a) \ (\i \ I. F(i))" by (rule program_equalityI, auto) lemma JOIN_cong [cong]: "\I=J; \i. i \ J \ F(i) = G(i)\ \ (\i \ I. F(i)) = (\i \ J. G(i))" by (simp add: JOIN_def) subsection\JOIN laws\ lemma JOIN_absorb: "k \ I \F(k) \ (\i \ I. F(i)) = (\i \ I. F(i))" apply (subst JOIN_cons [symmetric]) apply (auto simp add: cons_absorb) done lemma JOIN_Un: "(\i \ I \ J. F(i)) = ((\i \ I. F(i)) \ (\i \ J. F(i)))" apply (rule program_equalityI) apply (simp_all add: UN_Un INT_Un) apply (simp_all del: INT_simps add: INT_extend_simps, blast) done lemma JOIN_constant: "(\i \ I. c) = (if I=0 then SKIP else programify(c))" by (rule program_equalityI, auto) lemma JOIN_Join_distrib: "(\i \ I. F(i) \ G(i)) = (\i \ I. F(i)) \ (\i \ I. G(i))" apply (rule program_equalityI) apply (simp_all add: INT_Int_distrib, blast) done lemma JOIN_Join_miniscope: "(\i \ I. F(i) \ G) = ((\i \ I. F(i) \ G))" by (simp add: JOIN_Join_distrib JOIN_constant) text\Used to prove guarantees_JOIN_I\ lemma JOIN_Join_diff: "i \ I\F(i) \ JOIN(I - {i}, F) = JOIN(I, F)" apply (rule program_equalityI) apply (auto elim!: not_emptyE) done subsection\Safety: co, stable, FP\ (*Fails if I=0 because it collapses to SKIP \ A co B, i.e. to A\B. So an alternative precondition is A\B, but most proofs using this rule require I to be nonempty for other reasons anyway.*) lemma JOIN_constrains: "i \ I\(\i \ I. F(i)) \ A co B \ (\i \ I. programify(F(i)) \ A co B)" apply (unfold constrains_def JOIN_def st_set_def, auto) prefer 2 apply blast apply (rename_tac j act y z) apply (cut_tac F = "F (j) " in Acts_type) apply (drule_tac x = act in bspec, auto) done lemma Join_constrains [iff]: "(F \ G \ A co B) \ (programify(F) \ A co B \ programify(G) \ A co B)" by (auto simp add: constrains_def) lemma Join_unless [iff]: "(F \ G \ A unless B) \ (programify(F) \ A unless B \ programify(G) \ A unless B)" by (simp add: Join_constrains unless_def) (*Analogous weak versions FAIL; see Misra [1994] 5.4.1, Substitution Axiom. reachable (F \ G) could be much bigger than reachable F, reachable G *) lemma Join_constrains_weaken: "\F \ A co A'; G \ B co B'\ \ F \ G \ (A \ B) co (A' \ B')" apply (subgoal_tac "st_set (A) \ st_set (B) \ F \ program \ G \ program") prefer 2 apply (blast dest: constrainsD2, simp) apply (blast intro: constrains_weaken) done (*If I=0, it degenerates to SKIP \ state co 0, which is false.*) lemma JOIN_constrains_weaken: assumes major: "(\i. i \ I \ F(i) \ A(i) co A'(i))" and minor: "i \ I" shows "(\i \ I. F(i)) \ (\i \ I. A(i)) co (\i \ I. A'(i))" apply (cut_tac minor) apply (simp (no_asm_simp) add: JOIN_constrains) apply clarify apply (rename_tac "j") apply (frule_tac i = j in major) apply (frule constrainsD2, simp) apply (blast intro: constrains_weaken) done lemma JOIN_stable: "(\i \ I. F(i)) \ stable(A) \ ((\i \ I. programify(F(i)) \ stable(A)) \ st_set(A))" apply (auto simp add: stable_def constrains_def JOIN_def) apply (cut_tac F = "F (i) " in Acts_type) apply (drule_tac x = act in bspec, auto) done lemma initially_JOIN_I: assumes major: "(\i. i \ I \F(i) \ initially(A))" and minor: "i \ I" shows "(\i \ I. F(i)) \ initially(A)" apply (cut_tac minor) apply (auto elim!: not_emptyE simp add: Inter_iff initially_def) apply (frule_tac i = x in major) apply (auto simp add: initially_def) done lemma invariant_JOIN_I: assumes major: "(\i. i \ I \ F(i) \ invariant(A))" and minor: "i \ I" shows "(\i \ I. F(i)) \ invariant(A)" apply (cut_tac minor) apply (auto intro!: initially_JOIN_I dest: major simp add: invariant_def JOIN_stable) apply (erule_tac V = "i \ I" in thin_rl) apply (frule major) apply (drule_tac [2] major) apply (auto simp add: invariant_def) apply (frule stableD2, force)+ done lemma Join_stable [iff]: " (F \ G \ stable(A)) \ (programify(F) \ stable(A) \ programify(G) \ stable(A))" by (simp add: stable_def) lemma initially_JoinI [intro!]: "\F \ initially(A); G \ initially(A)\ \ F \ G \ initially(A)" by (unfold initially_def, auto) lemma invariant_JoinI: "\F \ invariant(A); G \ invariant(A)\ \ F \ G \ invariant(A)" apply (subgoal_tac "F \ program\G \ program") prefer 2 apply (blast dest: invariantD2) apply (simp add: invariant_def) apply (auto intro: Join_in_program) done (* Fails if I=0 because \i \ 0. A(i) = 0 *) lemma FP_JOIN: "i \ I \ FP(\i \ I. F(i)) = (\i \ I. FP (programify(F(i))))" by (auto simp add: FP_def Inter_def st_set_def JOIN_stable) subsection\Progress: transient, ensures\ lemma JOIN_transient: "i \ I \ (\i \ I. F(i)) \ transient(A) \ (\i \ I. programify(F(i)) \ transient(A))" apply (auto simp add: transient_def JOIN_def) unfolding st_set_def apply (drule_tac [2] x = act in bspec) apply (auto dest: Acts_type [THEN subsetD]) done lemma Join_transient [iff]: "F \ G \ transient(A) \ (programify(F) \ transient(A) | programify(G) \ transient(A))" apply (auto simp add: transient_def Join_def Int_Un_distrib2) done lemma Join_transient_I1: "F \ transient(A) \ F \ G \ transient(A)" by (simp add: Join_transient transientD2) lemma Join_transient_I2: "G \ transient(A) \ F \ G \ transient(A)" by (simp add: Join_transient transientD2) (*If I=0 it degenerates to (SKIP \ A ensures B) = False, i.e. to \(A\B) *) lemma JOIN_ensures: "i \ I \ (\i \ I. F(i)) \ A ensures B \ ((\i \ I. programify(F(i)) \ (A-B) co (A \ B)) \ (\i \ I. programify(F(i)) \ A ensures B))" by (auto simp add: ensures_def JOIN_constrains JOIN_transient) lemma Join_ensures: "F \ G \ A ensures B \ (programify(F) \ (A-B) co (A \ B) \ programify(G) \ (A-B) co (A \ B) \ (programify(F) \ transient (A-B) | programify(G) \ transient (A-B)))" unfolding ensures_def apply (auto simp add: Join_transient) done lemma stable_Join_constrains: "\F \ stable(A); G \ A co A'\ \ F \ G \ A co A'" -apply (unfold stable_def constrains_def Join_def st_set_def) + unfolding stable_def constrains_def Join_def st_set_def apply (cut_tac F = F in Acts_type) apply (cut_tac F = G in Acts_type, force) done (*Premise for G cannot use Always because F \ Stable A is weaker than G \ stable A *) lemma stable_Join_Always1: "\F \ stable(A); G \ invariant(A)\ \ F \ G \ Always(A)" apply (subgoal_tac "F \ program \ G \ program \ st_set (A) ") prefer 2 apply (blast dest: invariantD2 stableD2) apply (simp add: Always_def invariant_def initially_def Stable_eq_stable) apply (force intro: stable_Int) done (*As above, but exchanging the roles of F and G*) lemma stable_Join_Always2: "\F \ invariant(A); G \ stable(A)\ \ F \ G \ Always(A)" apply (subst Join_commute) apply (blast intro: stable_Join_Always1) done lemma stable_Join_ensures1: "\F \ stable(A); G \ A ensures B\ \ F \ G \ A ensures B" apply (subgoal_tac "F \ program \ G \ program \ st_set (A) ") prefer 2 apply (blast dest: stableD2 ensures_type [THEN subsetD]) apply (simp (no_asm_simp) add: Join_ensures) apply (simp add: stable_def ensures_def) apply (erule constrains_weaken, auto) done (*As above, but exchanging the roles of F and G*) lemma stable_Join_ensures2: "\F \ A ensures B; G \ stable(A)\ \ F \ G \ A ensures B" apply (subst Join_commute) apply (blast intro: stable_Join_ensures1) done subsection\The ok and OK relations\ lemma ok_SKIP1 [iff]: "SKIP ok F" by (auto dest: Acts_type [THEN subsetD] simp add: ok_def) lemma ok_SKIP2 [iff]: "F ok SKIP" by (auto dest: Acts_type [THEN subsetD] simp add: ok_def) lemma ok_Join_commute: "(F ok G \ (F \ G) ok H) \ (G ok H \ F ok (G \ H))" by (auto simp add: ok_def) lemma ok_commute: "(F ok G) \(G ok F)" by (auto simp add: ok_def) lemmas ok_sym = ok_commute [THEN iffD1] lemma ok_iff_OK: "OK({\0,F\,\1,G\,\2,H\}, snd) \ (F ok G \ (F \ G) ok H)" by (simp add: ok_def Join_def OK_def Int_assoc cons_absorb Int_Un_distrib2 Ball_def, safe, force+) lemma ok_Join_iff1 [iff]: "F ok (G \ H) \ (F ok G \ F ok H)" by (auto simp add: ok_def) lemma ok_Join_iff2 [iff]: "(G \ H) ok F \ (G ok F \ H ok F)" by (auto simp add: ok_def) (*useful? Not with the previous two around*) lemma ok_Join_commute_I: "\F ok G; (F \ G) ok H\ \ F ok (G \ H)" by (auto simp add: ok_def) lemma ok_JOIN_iff1 [iff]: "F ok JOIN(I,G) \ (\i \ I. F ok G(i))" by (force dest: Acts_type [THEN subsetD] elim!: not_emptyE simp add: ok_def) lemma ok_JOIN_iff2 [iff]: "JOIN(I,G) ok F \ (\i \ I. G(i) ok F)" apply (auto elim!: not_emptyE simp add: ok_def) apply (blast dest: Acts_type [THEN subsetD]) done lemma OK_iff_ok: "OK(I,F) \ (\i \ I. \j \ I-{i}. F(i) ok (F(j)))" by (auto simp add: ok_def OK_def) lemma OK_imp_ok: "\OK(I,F); i \ I; j \ I; i\j\ \ F(i) ok F(j)" by (auto simp add: OK_iff_ok) lemma OK_0 [iff]: "OK(0,F)" by (simp add: OK_def) lemma OK_cons_iff: "OK(cons(i, I), F) \ (i \ I \ OK(I, F)) | (i\I \ OK(I, F) \ F(i) ok JOIN(I,F))" apply (simp add: OK_iff_ok) apply (blast intro: ok_sym) done subsection\Allowed\ lemma Allowed_SKIP [simp]: "Allowed(SKIP) = program" by (auto dest: Acts_type [THEN subsetD] simp add: Allowed_def) lemma Allowed_Join [simp]: "Allowed(F \ G) = Allowed(programify(F)) \ Allowed(programify(G))" apply (auto simp add: Allowed_def) done lemma Allowed_JOIN [simp]: "i \ I \ Allowed(JOIN(I,F)) = (\i \ I. Allowed(programify(F(i))))" apply (auto simp add: Allowed_def, blast) done lemma ok_iff_Allowed: "F ok G \ (programify(F) \ Allowed(programify(G)) \ programify(G) \ Allowed(programify(F)))" by (simp add: ok_def Allowed_def) lemma OK_iff_Allowed: "OK(I,F) \ (\i \ I. \j \ I-{i}. programify(F(i)) \ Allowed(programify(F(j))))" apply (auto simp add: OK_iff_ok ok_iff_Allowed) done subsection\safety_prop, for reasoning about given instances of "ok"\ lemma safety_prop_Acts_iff: "safety_prop(X) \ (Acts(G) \ cons(id(state), (\F \ X. Acts(F)))) \ (programify(G) \ X)" apply (simp (no_asm_use) add: safety_prop_def) apply clarify apply (case_tac "G \ program", simp_all, blast, safe) prefer 2 apply force apply (force simp add: programify_def) done lemma safety_prop_AllowedActs_iff_Allowed: "safety_prop(X) \ (\G \ X. Acts(G)) \ AllowedActs(F) \ (X \ Allowed(programify(F)))" apply (simp add: Allowed_def safety_prop_Acts_iff [THEN iff_sym] safety_prop_def, blast) done lemma Allowed_eq: "safety_prop(X) \ Allowed(mk_program(init, acts, \F \ X. Acts(F))) = X" apply (subgoal_tac "cons (id (state), \(RepFun (X, Acts)) \ Pow (state * state)) = \(RepFun (X, Acts))") apply (rule_tac [2] equalityI) apply (simp del: UN_simps add: Allowed_def safety_prop_Acts_iff safety_prop_def, auto) apply (force dest: Acts_type [THEN subsetD] simp add: safety_prop_def)+ done lemma def_prg_Allowed: "\F \ mk_program (init, acts, \F \ X. Acts(F)); safety_prop(X)\ \ Allowed(F) = X" by (simp add: Allowed_eq) (*For safety_prop to hold, the property must be satisfiable!*) lemma safety_prop_constrains [iff]: "safety_prop(A co B) \ (A \ B \ st_set(A))" by (simp add: safety_prop_def constrains_def st_set_def, blast) (* To be used with resolution *) lemma safety_prop_constrainsI [iff]: "\A\B; st_set(A)\ \safety_prop(A co B)" by auto lemma safety_prop_stable [iff]: "safety_prop(stable(A)) \ st_set(A)" by (simp add: stable_def) lemma safety_prop_stableI: "st_set(A) \ safety_prop(stable(A))" by auto lemma safety_prop_Int [simp]: "\safety_prop(X) ; safety_prop(Y)\ \ safety_prop(X \ Y)" apply (simp add: safety_prop_def, safe, blast) apply (drule_tac [2] B = "\(RepFun (X \ Y, Acts))" and C = "\(RepFun (Y, Acts))" in subset_trans) apply (drule_tac B = "\(RepFun (X \ Y, Acts))" and C = "\(RepFun (X, Acts))" in subset_trans) apply blast+ done (* If I=0 the conclusion becomes safety_prop(0) which is false *) lemma safety_prop_Inter: assumes major: "(\i. i \ I \safety_prop(X(i)))" and minor: "i \ I" shows "safety_prop(\i \ I. X(i))" apply (simp add: safety_prop_def) apply (cut_tac minor, safe) apply (simp (no_asm_use) add: Inter_iff) apply clarify apply (frule major) apply (drule_tac [2] i = xa in major) apply (frule_tac [4] i = xa in major) apply (auto simp add: safety_prop_def) apply (drule_tac B = "\(RepFun (\(RepFun (I, X)), Acts))" and C = "\(RepFun (X (xa), Acts))" in subset_trans) apply blast+ done lemma def_UNION_ok_iff: "\F \ mk_program(init,acts, \G \ X. Acts(G)); safety_prop(X)\ \ F ok G \ (programify(G) \ X \ acts \ Pow(state*state) \ AllowedActs(G))" unfolding ok_def apply (drule_tac G = G in safety_prop_Acts_iff) apply (cut_tac F = G in AllowedActs_type) apply (cut_tac F = G in Acts_type, auto) done end diff --git a/src/ZF/UNITY/WFair.thy b/src/ZF/UNITY/WFair.thy --- a/src/ZF/UNITY/WFair.thy +++ b/src/ZF/UNITY/WFair.thy @@ -1,707 +1,707 @@ (* Title: ZF/UNITY/WFair.thy Author: Sidi Ehmety, Computer Laboratory Copyright 1998 University of Cambridge *) section\Progress under Weak Fairness\ theory WFair imports UNITY ZFC begin text\This theory defines the operators transient, ensures and leadsTo, assuming weak fairness. From Misra, "A Logic for Concurrent Programming", 1994.\ definition (* This definition specifies weak fairness. The rest of the theory is generic to all forms of fairness.*) transient :: "i\i" where "transient(A) \{F \ program. (\act\Acts(F). A<=domain(act) \ act``A \ state-A) \ st_set(A)}" definition ensures :: "[i,i] \ i" (infixl \ensures\ 60) where "A ensures B \ ((A-B) co (A \ B)) \ transient(A-B)" consts (*LEADS-TO constant for the inductive definition*) leads :: "[i, i]\i" inductive domains "leads(D, F)" \ "Pow(D)*Pow(D)" intros Basis: "\F \ A ensures B; A \ Pow(D); B \ Pow(D)\ \ \A,B\:leads(D, F)" Trans: "\\A,B\ \ leads(D, F); \B,C\ \ leads(D, F)\ \ \A,C\:leads(D, F)" Union: "\S \ Pow({A \ S. \A, B\:leads(D, F)}); B \ Pow(D); S \ Pow(Pow(D))\ \ <\(S),B>:leads(D, F)" monos Pow_mono type_intros Union_Pow_iff [THEN iffD2] UnionI PowI definition (* The Visible version of the LEADS-TO relation*) leadsTo :: "[i, i] \ i" (infixl \\\ 60) where "A \ B \ {F \ program. \A,B\:leads(state, F)}" definition (* wlt(F, B) is the largest set that leads to B*) wlt :: "[i, i] \ i" where "wlt(F, B) \ \({A \ Pow(state). F \ A \ B})" (** Ad-hoc set-theory rules **) lemma Int_Union_Union: "\(B) \ A = (\b \ B. b \ A)" by auto lemma Int_Union_Union2: "A \ \(B) = (\b \ B. A \ b)" by auto (*** transient ***) lemma transient_type: "transient(A)<=program" by (unfold transient_def, auto) lemma transientD2: "F \ transient(A) \ F \ program \ st_set(A)" apply (unfold transient_def, auto) done lemma stable_transient_empty: "\F \ stable(A); F \ transient(A)\ \ A = 0" by (simp add: stable_def constrains_def transient_def, fast) lemma transient_strengthen: "\F \ transient(A); B<=A\ \ F \ transient(B)" apply (simp add: transient_def st_set_def, clarify) apply (blast intro!: rev_bexI) done lemma transientI: "\act \ Acts(F); A \ domain(act); act``A \ state-A; F \ program; st_set(A)\ \ F \ transient(A)" by (simp add: transient_def, blast) lemma transientE: "\F \ transient(A); \act. \act \ Acts(F); A \ domain(act); act``A \ state-A\\P\ \P" by (simp add: transient_def, blast) lemma transient_state: "transient(state) = 0" apply (simp add: transient_def) apply (rule equalityI, auto) apply (cut_tac F = x in Acts_type) apply (simp add: Diff_cancel) apply (auto intro: st0_in_state) done lemma transient_state2: "state<=B \ transient(B) = 0" apply (simp add: transient_def st_set_def) apply (rule equalityI, auto) apply (cut_tac F = x in Acts_type) apply (subgoal_tac "B=state") apply (auto intro: st0_in_state) done lemma transient_empty: "transient(0) = program" by (auto simp add: transient_def) declare transient_empty [simp] transient_state [simp] transient_state2 [simp] (*** ensures ***) lemma ensures_type: "A ensures B <=program" by (simp add: ensures_def constrains_def, auto) lemma ensuresI: "\F:(A-B) co (A \ B); F \ transient(A-B)\\F \ A ensures B" unfolding ensures_def apply (auto simp add: transient_type [THEN subsetD]) done (* Added by Sidi, from Misra's notes, Progress chapter, exercise 4 *) lemma ensuresI2: "\F \ A co A \ B; F \ transient(A)\ \ F \ A ensures B" apply (drule_tac B = "A-B" in constrains_weaken_L) apply (drule_tac [2] B = "A-B" in transient_strengthen) apply (auto simp add: ensures_def transient_type [THEN subsetD]) done lemma ensuresD: "F \ A ensures B \ F:(A-B) co (A \ B) \ F \ transient (A-B)" by (unfold ensures_def, auto) lemma ensures_weaken_R: "\F \ A ensures A'; A'<=B'\ \ F \ A ensures B'" unfolding ensures_def apply (blast intro: transient_strengthen constrains_weaken) done (*The L-version (precondition strengthening) fails, but we have this*) lemma stable_ensures_Int: "\F \ stable(C); F \ A ensures B\ \ F:(C \ A) ensures (C \ B)" unfolding ensures_def apply (simp (no_asm) add: Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric]) apply (blast intro: transient_strengthen stable_constrains_Int constrains_weaken) done lemma stable_transient_ensures: "\F \ stable(A); F \ transient(C); A<=B \ C\ \ F \ A ensures B" apply (frule stable_type [THEN subsetD]) apply (simp add: ensures_def stable_def) apply (blast intro: transient_strengthen constrains_weaken) done lemma ensures_eq: "(A ensures B) = (A unless B) \ transient (A-B)" by (auto simp add: ensures_def unless_def) lemma subset_imp_ensures: "\F \ program; A<=B\ \ F \ A ensures B" by (auto simp add: ensures_def constrains_def transient_def st_set_def) (*** leadsTo ***) lemmas leads_left = leads.dom_subset [THEN subsetD, THEN SigmaD1] lemmas leads_right = leads.dom_subset [THEN subsetD, THEN SigmaD2] lemma leadsTo_type: "A \ B \ program" by (unfold leadsTo_def, auto) lemma leadsToD2: "F \ A \ B \ F \ program \ st_set(A) \ st_set(B)" -apply (unfold leadsTo_def st_set_def) + unfolding leadsTo_def st_set_def apply (blast dest: leads_left leads_right) done lemma leadsTo_Basis: "\F \ A ensures B; st_set(A); st_set(B)\ \ F \ A \ B" -apply (unfold leadsTo_def st_set_def) + unfolding leadsTo_def st_set_def apply (cut_tac ensures_type) apply (auto intro: leads.Basis) done declare leadsTo_Basis [intro] (* Added by Sidi, from Misra's notes, Progress chapter, exercise number 4 *) (* \F \ program; A<=B; st_set(A); st_set(B)\ \ A \ B *) lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis] lemma leadsTo_Trans: "\F \ A \ B; F \ B \ C\\F \ A \ C" unfolding leadsTo_def apply (auto intro: leads.Trans) done (* Better when used in association with leadsTo_weaken_R *) lemma transient_imp_leadsTo: "F \ transient(A) \ F \ A \ (state-A)" unfolding transient_def apply (blast intro: ensuresI [THEN leadsTo_Basis] constrains_weaken transientI) done (*Useful with cancellation, disjunction*) lemma leadsTo_Un_duplicate: "F \ A \ (A' \ A') \ F \ A \ A'" by simp lemma leadsTo_Un_duplicate2: "F \ A \ (A' \ C \ C) \ F \ A \ (A' \ C)" by (simp add: Un_ac) (*The Union introduction rule as we should have liked to state it*) lemma leadsTo_Union: "\\A. A \ S \ F \ A \ B; F \ program; st_set(B)\ \ F \ \(S) \ B" -apply (unfold leadsTo_def st_set_def) + unfolding leadsTo_def st_set_def apply (blast intro: leads.Union dest: leads_left) done lemma leadsTo_Union_Int: "\\A. A \ S \F \ (A \ C) \ B; F \ program; st_set(B)\ \ F \ (\(S)Int C)\ B" -apply (unfold leadsTo_def st_set_def) + unfolding leadsTo_def st_set_def apply (simp only: Int_Union_Union) apply (blast dest: leads_left intro: leads.Union) done lemma leadsTo_UN: "\\i. i \ I \ F \ A(i) \ B; F \ program; st_set(B)\ \ F:(\i \ I. A(i)) \ B" apply (simp add: Int_Union_Union leadsTo_def st_set_def) apply (blast dest: leads_left intro: leads.Union) done (* Binary union introduction rule *) lemma leadsTo_Un: "\F \ A \ C; F \ B \ C\ \ F \ (A \ B) \ C" apply (subst Un_eq_Union) apply (blast intro: leadsTo_Union dest: leadsToD2) done lemma single_leadsTo_I: "\\x. x \ A\ F:{x} \ B; F \ program; st_set(B)\ \ F \ A \ B" apply (rule_tac b = A in UN_singleton [THEN subst]) apply (rule leadsTo_UN, auto) done lemma leadsTo_refl: "\F \ program; st_set(A)\ \ F \ A \ A" by (blast intro: subset_imp_leadsTo) lemma leadsTo_refl_iff: "F \ A \ A \ F \ program \ st_set(A)" by (auto intro: leadsTo_refl dest: leadsToD2) lemma empty_leadsTo: "F \ 0 \ B \ (F \ program \ st_set(B))" by (auto intro: subset_imp_leadsTo dest: leadsToD2) declare empty_leadsTo [iff] lemma leadsTo_state: "F \ A \ state \ (F \ program \ st_set(A))" by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD) declare leadsTo_state [iff] lemma leadsTo_weaken_R: "\F \ A \ A'; A'<=B'; st_set(B')\ \ F \ A \ B'" by (blast dest: leadsToD2 intro: subset_imp_leadsTo leadsTo_Trans) lemma leadsTo_weaken_L: "\F \ A \ A'; B<=A\ \ F \ B \ A'" apply (frule leadsToD2) apply (blast intro: leadsTo_Trans subset_imp_leadsTo st_set_subset) done lemma leadsTo_weaken: "\F \ A \ A'; B<=A; A'<=B'; st_set(B')\\ F \ B \ B'" apply (frule leadsToD2) apply (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans leadsTo_refl) done (* This rule has a nicer conclusion *) lemma transient_imp_leadsTo2: "\F \ transient(A); state-A<=B; st_set(B)\ \ F \ A \ B" apply (frule transientD2) apply (rule leadsTo_weaken_R) apply (auto simp add: transient_imp_leadsTo) done (*Distributes over binary unions*) lemma leadsTo_Un_distrib: "F:(A \ B) \ C \ (F \ A \ C \ F \ B \ C)" by (blast intro: leadsTo_Un leadsTo_weaken_L) lemma leadsTo_UN_distrib: "(F:(\i \ I. A(i)) \ B)\ ((\i \ I. F \ A(i) \ B) \ F \ program \ st_set(B))" apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L) done lemma leadsTo_Union_distrib: "(F \ \(S) \ B) \ (\A \ S. F \ A \ B) \ F \ program \ st_set(B)" by (blast dest: leadsToD2 intro: leadsTo_Union leadsTo_weaken_L) text\Set difference: maybe combine with \leadsTo_weaken_L\??\ lemma leadsTo_Diff: "\F: (A-B) \ C; F \ B \ C; st_set(C)\ \ F \ A \ C" by (blast intro: leadsTo_Un leadsTo_weaken dest: leadsToD2) lemma leadsTo_UN_UN: "\\i. i \ I \ F \ A(i) \ A'(i); F \ program\ \ F: (\i \ I. A(i)) \ (\i \ I. A'(i))" apply (rule leadsTo_Union) apply (auto intro: leadsTo_weaken_R dest: leadsToD2) done (*Binary union version*) lemma leadsTo_Un_Un: "\F \ A \ A'; F \ B \ B'\ \ F \ (A \ B) \ (A' \ B')" apply (subgoal_tac "st_set (A) \ st_set (A') \ st_set (B) \ st_set (B') ") prefer 2 apply (blast dest: leadsToD2) apply (blast intro: leadsTo_Un leadsTo_weaken_R) done (** The cancellation law **) lemma leadsTo_cancel2: "\F \ A \ (A' \ B); F \ B \ B'\ \ F \ A \ (A' \ B')" apply (subgoal_tac "st_set (A) \ st_set (A') \ st_set (B) \ st_set (B') \F \ program") prefer 2 apply (blast dest: leadsToD2) apply (blast intro: leadsTo_Trans leadsTo_Un_Un leadsTo_refl) done lemma leadsTo_cancel_Diff2: "\F \ A \ (A' \ B); F \ (B-A') \ B'\\ F \ A \ (A' \ B')" apply (rule leadsTo_cancel2) prefer 2 apply assumption apply (blast dest: leadsToD2 intro: leadsTo_weaken_R) done lemma leadsTo_cancel1: "\F \ A \ (B \ A'); F \ B \ B'\ \ F \ A \ (B' \ A')" apply (simp add: Un_commute) apply (blast intro!: leadsTo_cancel2) done lemma leadsTo_cancel_Diff1: "\F \ A \ (B \ A'); F: (B-A') \ B'\\ F \ A \ (B' \ A')" apply (rule leadsTo_cancel1) prefer 2 apply assumption apply (blast intro: leadsTo_weaken_R dest: leadsToD2) done (*The INDUCTION rule as we should have liked to state it*) lemma leadsTo_induct: assumes major: "F \ za \ zb" and basis: "\A B. \F \ A ensures B; st_set(A); st_set(B)\ \ P(A,B)" and trans: "\A B C. \F \ A \ B; P(A, B); F \ B \ C; P(B, C)\ \ P(A,C)" and union: "\B S. \\A \ S. F \ A \ B; \A \ S. P(A,B); st_set(B); \A \ S. st_set(A)\ \ P(\(S), B)" shows "P(za, zb)" apply (cut_tac major) apply (unfold leadsTo_def, clarify) apply (erule leads.induct) apply (blast intro: basis [unfolded st_set_def]) apply (blast intro: trans [unfolded leadsTo_def]) apply (force intro: union [unfolded st_set_def leadsTo_def]) done (* Added by Sidi, an induction rule without ensures *) lemma leadsTo_induct2: assumes major: "F \ za \ zb" and basis1: "\A B. \A<=B; st_set(B)\ \ P(A, B)" and basis2: "\A B. \F \ A co A \ B; F \ transient(A); st_set(B)\ \ P(A, B)" and trans: "\A B C. \F \ A \ B; P(A, B); F \ B \ C; P(B, C)\ \ P(A,C)" and union: "\B S. \\A \ S. F \ A \ B; \A \ S. P(A,B); st_set(B); \A \ S. st_set(A)\ \ P(\(S), B)" shows "P(za, zb)" apply (cut_tac major) apply (erule leadsTo_induct) apply (auto intro: trans union) apply (simp add: ensures_def, clarify) apply (frule constrainsD2) apply (drule_tac B' = " (A-B) \ B" in constrains_weaken_R) apply blast apply (frule ensuresI2 [THEN leadsTo_Basis]) apply (drule_tac [4] basis2, simp_all) apply (frule_tac A1 = A and B = B in Int_lower2 [THEN basis1]) apply (subgoal_tac "A=\({A - B, A \ B}) ") prefer 2 apply blast apply (erule ssubst) apply (rule union) apply (auto intro: subset_imp_leadsTo) done (** Variant induction rule: on the preconditions for B **) (*Lemma is the weak version: can't see how to do it in one step*) lemma leadsTo_induct_pre_aux: "\F \ za \ zb; P(zb); \A B. \F \ A ensures B; P(B); st_set(A); st_set(B)\ \ P(A); \S. \\A \ S. P(A); \A \ S. st_set(A)\ \ P(\(S)) \ \ P(za)" txt\by induction on this formula\ apply (subgoal_tac "P (zb) \ P (za) ") txt\now solve first subgoal: this formula is sufficient\ apply (blast intro: leadsTo_refl) apply (erule leadsTo_induct) apply (blast+) done lemma leadsTo_induct_pre: "\F \ za \ zb; P(zb); \A B. \F \ A ensures B; F \ B \ zb; P(B); st_set(A)\ \ P(A); \S. \A \ S. F \ A \ zb \ P(A) \ st_set(A) \ P(\(S)) \ \ P(za)" apply (subgoal_tac " (F \ za \ zb) \ P (za) ") apply (erule conjunct2) apply (frule leadsToD2) apply (erule leadsTo_induct_pre_aux) prefer 3 apply (blast dest: leadsToD2 intro: leadsTo_Union) prefer 2 apply (blast intro: leadsTo_Trans leadsTo_Basis) apply (blast intro: leadsTo_refl) done (** The impossibility law **) lemma leadsTo_empty: "F \ A \ 0 \ A=0" apply (erule leadsTo_induct_pre) apply (auto simp add: ensures_def constrains_def transient_def st_set_def) apply (drule bspec, assumption)+ apply blast done declare leadsTo_empty [simp] subsection\PSP: Progress-Safety-Progress\ text\Special case of PSP: Misra's "stable conjunction"\ lemma psp_stable: "\F \ A \ A'; F \ stable(B)\ \ F:(A \ B) \ (A' \ B)" unfolding stable_def apply (frule leadsToD2) apply (erule leadsTo_induct) prefer 3 apply (blast intro: leadsTo_Union_Int) prefer 2 apply (blast intro: leadsTo_Trans) apply (rule leadsTo_Basis) apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] Int_Un_distrib2 [symmetric]) apply (auto intro: transient_strengthen constrains_Int) done lemma psp_stable2: "\F \ A \ A'; F \ stable(B)\\F: (B \ A) \ (B \ A')" apply (simp (no_asm_simp) add: psp_stable Int_ac) done lemma psp_ensures: "\F \ A ensures A'; F \ B co B'\\ F: (A \ B') ensures ((A' \ B) \ (B' - B))" -apply (unfold ensures_def constrains_def st_set_def) + unfolding ensures_def constrains_def st_set_def (*speeds up the proof*) apply clarify apply (blast intro: transient_strengthen) done lemma psp: "\F \ A \ A'; F \ B co B'; st_set(B')\\ F:(A \ B') \ ((A' \ B) \ (B' - B))" apply (subgoal_tac "F \ program \ st_set (A) \ st_set (A') \ st_set (B) ") prefer 2 apply (blast dest!: constrainsD2 leadsToD2) apply (erule leadsTo_induct) prefer 3 apply (blast intro: leadsTo_Union_Int) txt\Basis case\ apply (blast intro: psp_ensures leadsTo_Basis) txt\Transitivity case has a delicate argument involving "cancellation"\ apply (rule leadsTo_Un_duplicate2) apply (erule leadsTo_cancel_Diff1) apply (simp add: Int_Diff Diff_triv) apply (blast intro: leadsTo_weaken_L dest: constrains_imp_subset) done lemma psp2: "\F \ A \ A'; F \ B co B'; st_set(B')\ \ F \ (B' \ A) \ ((B \ A') \ (B' - B))" by (simp (no_asm_simp) add: psp Int_ac) lemma psp_unless: "\F \ A \ A'; F \ B unless B'; st_set(B); st_set(B')\ \ F \ (A \ B) \ ((A' \ B) \ B')" unfolding unless_def apply (subgoal_tac "st_set (A) \st_set (A') ") prefer 2 apply (blast dest: leadsToD2) apply (drule psp, assumption, blast) apply (blast intro: leadsTo_weaken) done subsection\Proving the induction rules\ (** The most general rule \ r is any wf relation; f is any variant function **) lemma leadsTo_wf_induct_aux: "\wf(r); m \ I; field(r)<=I; F \ program; st_set(B); \m \ I. F \ (A \ f-``{m}) \ ((A \ f-``(converse(r)``{m})) \ B)\ \ F \ (A \ f-``{m}) \ B" apply (erule_tac a = m in wf_induct2, simp_all) apply (subgoal_tac "F \ (A \ (f-`` (converse (r) ``{x}))) \ B") apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate) apply (subst vimage_eq_UN) apply (simp del: UN_simps add: Int_UN_distrib) apply (auto intro: leadsTo_UN simp del: UN_simps simp add: Int_UN_distrib) done (** Meta or object quantifier ? **) lemma leadsTo_wf_induct: "\wf(r); field(r)<=I; A<=f-``I; F \ program; st_set(A); st_set(B); \m \ I. F \ (A \ f-``{m}) \ ((A \ f-``(converse(r)``{m})) \ B)\ \ F \ A \ B" apply (rule_tac b = A in subst) defer 1 apply (rule_tac I = I in leadsTo_UN) apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best) done lemma nat_measure_field: "field(measure(nat, \x. x)) = nat" unfolding field_def apply (simp add: measure_def) apply (rule equalityI, force, clarify) apply (erule_tac V = "x\range (y)" for y in thin_rl) apply (erule nat_induct) apply (rule_tac [2] b = "succ (succ (xa))" in domainI) apply (rule_tac b = "succ (0) " in domainI) apply simp_all done lemma Image_inverse_lessThan: "k measure(A, \x. x) -`` {k} = k" apply (rule equalityI) apply (auto simp add: measure_def) apply (blast intro: ltD) apply (rule vimageI) prefer 2 apply blast apply (simp add: lt_Ord lt_Ord2 Ord_mem_iff_lt) apply (blast intro: lt_trans) done (*Alternative proof is via the lemma F \ (A \ f-`(lessThan m)) \ B*) lemma lessThan_induct: "\A<=f-``nat; F \ program; st_set(A); st_set(B); \m \ nat. F:(A \ f-``{m}) \ ((A \ f -`` m) \ B)\ \ F \ A \ B" apply (rule_tac A1 = nat and f1 = "\x. x" in wf_measure [THEN leadsTo_wf_induct]) apply (simp_all add: nat_measure_field) apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric]) done (*** wlt ****) (*Misra's property W3*) lemma wlt_type: "wlt(F,B) <=state" by (unfold wlt_def, auto) lemma wlt_st_set: "st_set(wlt(F, B))" unfolding st_set_def apply (rule wlt_type) done declare wlt_st_set [iff] lemma wlt_leadsTo_iff: "F \ wlt(F, B) \ B \ (F \ program \ st_set(B))" unfolding wlt_def apply (blast dest: leadsToD2 intro!: leadsTo_Union) done (* \F \ program; st_set(B)\ \ F \ wlt(F, B) \ B *) lemmas wlt_leadsTo = conjI [THEN wlt_leadsTo_iff [THEN iffD2]] lemma leadsTo_subset: "F \ A \ B \ A \ wlt(F, B)" unfolding wlt_def apply (frule leadsToD2) apply (auto simp add: st_set_def) done (*Misra's property W2*) lemma leadsTo_eq_subset_wlt: "F \ A \ B \ (A \ wlt(F,B) \ F \ program \ st_set(B))" apply auto apply (blast dest: leadsToD2 leadsTo_subset intro: leadsTo_weaken_L wlt_leadsTo)+ done (*Misra's property W4*) lemma wlt_increasing: "\F \ program; st_set(B)\ \ B \ wlt(F,B)" apply (rule leadsTo_subset) apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [THEN iff_sym] subset_imp_leadsTo) done (*Used in the Trans case below*) lemma leadsTo_123_aux: "\B \ A2; F \ (A1 - B) co (A1 \ B); F \ (A2 - C) co (A2 \ C)\ \ F \ (A1 \ A2 - C) co (A1 \ A2 \ C)" apply (unfold constrains_def st_set_def, blast) done (*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*) (* slightly different from the HOL one \ B here is bounded *) lemma leadsTo_123: "F \ A \ A' \ \B \ Pow(state). A<=B \ F \ B \ A' \ F \ (B-A') co (B \ A')" apply (frule leadsToD2) apply (erule leadsTo_induct) txt\Basis\ apply (blast dest: ensuresD constrainsD2 st_setD) txt\Trans\ apply clarify apply (rule_tac x = "Ba \ Bb" in bexI) apply (blast intro: leadsTo_123_aux leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate, blast) txt\Union\ apply (clarify dest!: ball_conj_distrib [THEN iffD1]) apply (subgoal_tac "\y. y \ Pi (S, \A. {Ba \ Pow (state) . A<=Ba \ F \ Ba \ B \ F \ Ba - B co Ba \ B}) ") defer 1 apply (rule AC_ball_Pi, safe) apply (rotate_tac 1) apply (drule_tac x = x in bspec, blast, blast) apply (rule_tac x = "\A \ S. y`A" in bexI, safe) apply (rule_tac [3] I1 = S in constrains_UN [THEN constrains_weaken]) apply (rule_tac [2] leadsTo_Union) prefer 5 apply (blast dest!: apply_type, simp_all) apply (force dest!: apply_type)+ done (*Misra's property W5*) lemma wlt_constrains_wlt: "\F \ program; st_set(B)\ \F \ (wlt(F, B) - B) co (wlt(F,B))" apply (cut_tac F = F in wlt_leadsTo [THEN leadsTo_123], assumption, blast) apply clarify apply (subgoal_tac "Ba = wlt (F,B) ") prefer 2 apply (blast dest: leadsTo_eq_subset_wlt [THEN iffD1], clarify) apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]]) done subsection\Completion: Binary and General Finite versions\ lemma completion_aux: "\W = wlt(F, (B' \ C)); F \ A \ (A' \ C); F \ A' co (A' \ C); F \ B \ (B' \ C); F \ B' co (B' \ C)\ \ F \ (A \ B) \ ((A' \ B') \ C)" apply (subgoal_tac "st_set (C) \st_set (W) \st_set (W-C) \st_set (A') \st_set (A) \ st_set (B) \ st_set (B') \ F \ program") prefer 2 apply simp apply (blast dest!: leadsToD2) apply (subgoal_tac "F \ (W-C) co (W \ B' \ C) ") prefer 2 apply (blast intro!: constrains_weaken [OF constrains_Un [OF _ wlt_constrains_wlt]]) apply (subgoal_tac "F \ (W-C) co W") prefer 2 apply (simp add: wlt_increasing [THEN subset_Un_iff2 [THEN iffD1]] Un_assoc) apply (subgoal_tac "F \ (A \ W - C) \ (A' \ W \ C) ") prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken]) (** step 13 **) apply (subgoal_tac "F \ (A' \ W \ C) \ (A' \ B' \ C) ") apply (drule leadsTo_Diff) apply (blast intro: subset_imp_leadsTo dest: leadsToD2 constrainsD2) apply (force simp add: st_set_def) apply (subgoal_tac "A \ B \ A \ W") prefer 2 apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono]) apply (blast intro: leadsTo_Trans subset_imp_leadsTo) txt\last subgoal\ apply (rule_tac leadsTo_Un_duplicate2) apply (rule_tac leadsTo_Un_Un) prefer 2 apply (blast intro: leadsTo_refl) apply (rule_tac A'1 = "B' \ C" in wlt_leadsTo[THEN psp2, THEN leadsTo_weaken]) apply blast+ done lemmas completion = refl [THEN completion_aux] lemma finite_completion_aux: "\I \ Fin(X); F \ program; st_set(C)\ \ (\i \ I. F \ (A(i)) \ (A'(i) \ C)) \ (\i \ I. F \ (A'(i)) co (A'(i) \ C)) \ F \ (\i \ I. A(i)) \ ((\i \ I. A'(i)) \ C)" apply (erule Fin_induct) apply (auto simp add: Inter_0) apply (rule completion) apply (auto simp del: INT_simps simp add: INT_extend_simps) apply (blast intro: constrains_INT) done lemma finite_completion: "\I \ Fin(X); \i. i \ I \ F \ A(i) \ (A'(i) \ C); \i. i \ I \ F \ A'(i) co (A'(i) \ C); F \ program; st_set(C)\ \ F \ (\i \ I. A(i)) \ ((\i \ I. A'(i)) \ C)" by (blast intro: finite_completion_aux [THEN mp, THEN mp]) lemma stable_completion: "\F \ A \ A'; F \ stable(A'); F \ B \ B'; F \ stable(B')\ \ F \ (A \ B) \ (A' \ B')" unfolding stable_def apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+) apply (blast dest: leadsToD2) done lemma finite_stable_completion: "\I \ Fin(X); (\i. i \ I \ F \ A(i) \ A'(i)); (\i. i \ I \ F \ stable(A'(i))); F \ program\ \ F \ (\i \ I. A(i)) \ (\i \ I. A'(i))" unfolding stable_def apply (subgoal_tac "st_set (\i \ I. A' (i))") prefer 2 apply (blast dest: leadsToD2) apply (rule_tac C1 = 0 in finite_completion [THEN leadsTo_weaken_R], auto) done end diff --git a/src/ZF/WF.thy b/src/ZF/WF.thy --- a/src/ZF/WF.thy +++ b/src/ZF/WF.thy @@ -1,372 +1,372 @@ (* Title: ZF/WF.thy Author: Tobias Nipkow and Lawrence C Paulson Copyright 1994 University of Cambridge Derived first for transitive relations, and finally for arbitrary WF relations via wf_trancl and trans_trancl. It is difficult to derive this general case directly, using r^+ instead of r. In is_recfun, the two occurrences of the relation must have the same form. Inserting r^+ in the_recfun or wftrec yields a recursion rule with r^+ -`` {a} instead of r-``{a}. This recursion rule is stronger in principle, but harder to use, especially to prove wfrec_eclose_eq in epsilon.ML. Expanding out the definition of wftrec in wfrec would yield a mess. *) section\Well-Founded Recursion\ theory WF imports Trancl begin definition wf :: "i\o" where (*r is a well-founded relation*) "wf(r) \ \Z. Z=0 | (\x\Z. \y. \y,x\:r \ \ y \ Z)" definition wf_on :: "[i,i]\o" (\wf[_]'(_')\) where (*r is well-founded on A*) "wf_on(A,r) \ wf(r \ A*A)" definition is_recfun :: "[i, i, [i,i]\i, i] \o" where "is_recfun(r,a,H,f) \ (f = (\x\r-``{a}. H(x, restrict(f, r-``{x}))))" definition the_recfun :: "[i, i, [i,i]\i] \i" where "the_recfun(r,a,H) \ (THE f. is_recfun(r,a,H,f))" definition wftrec :: "[i, i, [i,i]\i] \i" where "wftrec(r,a,H) \ H(a, the_recfun(r,a,H))" definition wfrec :: "[i, i, [i,i]\i] \i" where (*public version. Does not require r to be transitive*) "wfrec(r,a,H) \ wftrec(r^+, a, \x f. H(x, restrict(f,r-``{x})))" definition wfrec_on :: "[i, i, i, [i,i]\i] \i" (\wfrec[_]'(_,_,_')\) where "wfrec[A](r,a,H) \ wfrec(r \ A*A, a, H)" subsection\Well-Founded Relations\ subsubsection\Equivalences between \<^term>\wf\ and \<^term>\wf_on\\ lemma wf_imp_wf_on: "wf(r) \ wf[A](r)" by (unfold wf_def wf_on_def, force) lemma wf_on_imp_wf: "\wf[A](r); r \ A*A\ \ wf(r)" by (simp add: wf_on_def subset_Int_iff) lemma wf_on_field_imp_wf: "wf[field(r)](r) \ wf(r)" by (unfold wf_def wf_on_def, fast) lemma wf_iff_wf_on_field: "wf(r) \ wf[field(r)](r)" by (blast intro: wf_imp_wf_on wf_on_field_imp_wf) lemma wf_on_subset_A: "\wf[A](r); B<=A\ \ wf[B](r)" by (unfold wf_on_def wf_def, fast) lemma wf_on_subset_r: "\wf[A](r); s<=r\ \ wf[A](s)" by (unfold wf_on_def wf_def, fast) lemma wf_subset: "\wf(s); r<=s\ \ wf(r)" by (simp add: wf_def, fast) subsubsection\Introduction Rules for \<^term>\wf_on\\ text\If every non-empty subset of \<^term>\A\ has an \<^term>\r\-minimal element then we have \<^term>\wf[A](r)\.\ lemma wf_onI: assumes prem: "\Z u. \Z<=A; u \ Z; \x\Z. \y\Z. \y,x\:r\ \ False" shows "wf[A](r)" -apply (unfold wf_on_def wf_def) + unfolding wf_on_def wf_def apply (rule equals0I [THEN disjCI, THEN allI]) apply (rule_tac Z = Z in prem, blast+) done text\If \<^term>\r\ allows well-founded induction over \<^term>\A\ then we have \<^term>\wf[A](r)\. Premise is equivalent to \<^prop>\\B. \x\A. (\y. \y,x\: r \ y \ B) \ x \ B \ A<=B\\ lemma wf_onI2: assumes prem: "\y B. \\x\A. (\y\A. \y,x\:r \ y \ B) \ x \ B; y \ A\ \ y \ B" shows "wf[A](r)" apply (rule wf_onI) apply (rule_tac c=u in prem [THEN DiffE]) prefer 3 apply blast apply fast+ done subsubsection\Well-founded Induction\ text\Consider the least \<^term>\z\ in \<^term>\domain(r)\ such that \<^term>\P(z)\ does not hold...\ lemma wf_induct_raw: "\wf(r); \x.\\y. \y,x\: r \ P(y)\ \ P(x)\ \ P(a)" unfolding wf_def apply (erule_tac x = "{z \ domain(r). \ P(z)}" in allE) apply blast done lemmas wf_induct = wf_induct_raw [rule_format, consumes 1, case_names step, induct set: wf] text\The form of this rule is designed to match \wfI\\ lemma wf_induct2: "\wf(r); a \ A; field(r)<=A; \x.\x \ A; \y. \y,x\: r \ P(y)\ \ P(x)\ \ P(a)" apply (erule_tac P="a \ A" in rev_mp) apply (erule_tac a=a in wf_induct, blast) done lemma field_Int_square: "field(r \ A*A) \ A" by blast lemma wf_on_induct_raw [consumes 2, induct set: wf_on]: "\wf[A](r); a \ A; \x.\x \ A; \y\A. \y,x\: r \ P(y)\ \ P(x) \ \ P(a)" unfolding wf_on_def apply (erule wf_induct2, assumption) apply (rule field_Int_square, blast) done lemma wf_on_induct [consumes 2, case_names step, induct set: wf_on]: "wf[A](r) \ a \ A \ (\x. x \ A \ (\y. y \ A \ \y, x\ \ r \ P(y)) \ P(x)) \ P(a)" using wf_on_induct_raw [of A r a P] by simp text\If \<^term>\r\ allows well-founded induction then we have \<^term>\wf(r)\.\ lemma wfI: "\field(r)<=A; \y B. \\x\A. (\y\A. \y,x\:r \ y \ B) \ x \ B; y \ A\ \ y \ B\ \ wf(r)" apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf]) apply (rule wf_onI2) prefer 2 apply blast apply blast done subsection\Basic Properties of Well-Founded Relations\ lemma wf_not_refl: "wf(r) \ \a,a\ \ r" by (erule_tac a=a in wf_induct, blast) lemma wf_not_sym [rule_format]: "wf(r) \ \x. \a,x\:r \ \x,a\ \ r" by (erule_tac a=a in wf_induct, blast) (* @{term"\wf(r); \a,x\ \ r; \P \ \x,a\ \ r\ \ P"} *) lemmas wf_asym = wf_not_sym [THEN swap] lemma wf_on_not_refl: "\wf[A](r); a \ A\ \ \a,a\ \ r" by (erule_tac a=a in wf_on_induct, assumption, blast) lemma wf_on_not_sym: "\wf[A](r); a \ A\ \ (\b. b\A \ \a,b\:r \ \b,a\\r)" apply (atomize (full), intro impI) apply (erule_tac a=a in wf_on_induct, assumption, blast) done lemma wf_on_asym: "\wf[A](r); \Z \ \a,b\ \ r; \b,a\ \ r \ Z; \Z \ a \ A; \Z \ b \ A\ \ Z" by (blast dest: wf_on_not_sym) (*Needed to prove well_ordI. Could also reason that wf[A](r) means wf(r \ A*A); thus wf( (r \ A*A)^+ ) and use wf_not_refl *) lemma wf_on_chain3: "\wf[A](r); \a,b\:r; \b,c\:r; \c,a\:r; a \ A; b \ A; c \ A\ \ P" apply (subgoal_tac "\y\A. \z\A. \a,y\:r \ \y,z\:r \ \z,a\:r \ P", blast) apply (erule_tac a=a in wf_on_induct, assumption, blast) done text\transitive closure of a WF relation is WF provided \<^term>\A\ is downward closed\ lemma wf_on_trancl: "\wf[A](r); r-``A \ A\ \ wf[A](r^+)" apply (rule wf_onI2) apply (frule bspec [THEN mp], assumption+) apply (erule_tac a = y in wf_on_induct, assumption) apply (blast elim: tranclE, blast) done lemma wf_trancl: "wf(r) \ wf(r^+)" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A) apply (erule wf_on_trancl) apply blast apply (rule trancl_type [THEN field_rel_subset]) done text\\<^term>\r-``{a}\ is the set of everything under \<^term>\a\ in \<^term>\r\\ lemmas underI = vimage_singleton_iff [THEN iffD2] lemmas underD = vimage_singleton_iff [THEN iffD1] subsection\The Predicate \<^term>\is_recfun\\ lemma is_recfun_type: "is_recfun(r,a,H,f) \ f \ r-``{a} -> range(f)" unfolding is_recfun_def apply (erule ssubst) apply (rule lamI [THEN rangeI, THEN lam_type], assumption) done lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function] lemma apply_recfun: "\is_recfun(r,a,H,f); \x,a\:r\ \ f`x = H(x, restrict(f,r-``{x}))" unfolding is_recfun_def txt\replace f only on the left-hand side\ apply (erule_tac P = "\x. t(x) = u" for t u in ssubst) apply (simp add: underI) done lemma is_recfun_equal [rule_format]: "\wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g)\ \ \x,a\:r \ \x,b\:r \ f`x=g`x" apply (frule_tac f = f in is_recfun_type) apply (frule_tac f = g in is_recfun_type) apply (simp add: is_recfun_def) apply (erule_tac a=x in wf_induct) apply (intro impI) apply (elim ssubst) apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) apply (rule_tac t = "\z. H (x, z)" for x in subst_context) apply (subgoal_tac "\y\r-``{x}. \z. \y,z\:f \ \y,z\:g") apply (blast dest: transD) apply (simp add: apply_iff) apply (blast dest: transD intro: sym) done lemma is_recfun_cut: "\wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g); \b,a\:r\ \ restrict(f, r-``{b}) = g" apply (frule_tac f = f in is_recfun_type) apply (rule fun_extension) apply (blast dest: transD intro: restrict_type2) apply (erule is_recfun_type, simp) apply (blast dest: transD intro: is_recfun_equal) done subsection\Recursion: Main Existence Lemma\ lemma is_recfun_functional: "\wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,a,H,g)\ \ f=g" by (blast intro: fun_extension is_recfun_type is_recfun_equal) lemma the_recfun_eq: "\is_recfun(r,a,H,f); wf(r); trans(r)\ \ the_recfun(r,a,H) = f" unfolding the_recfun_def apply (blast intro: is_recfun_functional) done (*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *) lemma is_the_recfun: "\is_recfun(r,a,H,f); wf(r); trans(r)\ \ is_recfun(r, a, H, the_recfun(r,a,H))" by (simp add: the_recfun_eq) lemma unfold_the_recfun: "\wf(r); trans(r)\ \ is_recfun(r, a, H, the_recfun(r,a,H))" apply (rule_tac a=a in wf_induct, assumption) apply (rename_tac a1) apply (rule_tac f = "\y\r-``{a1}. wftrec (r,y,H)" in is_the_recfun) apply typecheck -apply (unfold is_recfun_def wftrec_def) + unfolding is_recfun_def wftrec_def \ \Applying the substitution: must keep the quantified assumption!\ apply (rule lam_cong [OF refl]) apply (drule underD) apply (fold is_recfun_def) apply (rule_tac t = "\z. H(x, z)" for x in subst_context) apply (rule fun_extension) apply (blast intro: is_recfun_type) apply (rule lam_type [THEN restrict_type2]) apply blast apply (blast dest: transD) apply atomize apply (frule spec [THEN mp], assumption) apply (subgoal_tac "\xa,a1\ \ r") apply (drule_tac x1 = xa in spec [THEN mp], assumption) apply (simp add: vimage_singleton_iff apply_recfun is_recfun_cut) apply (blast dest: transD) done subsection\Unfolding \<^term>\wftrec(r,a,H)\\ lemma the_recfun_cut: "\wf(r); trans(r); \b,a\:r\ \ restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)" by (blast intro: is_recfun_cut unfold_the_recfun) (*NOT SUITABLE FOR REWRITING: it is recursive!*) lemma wftrec: "\wf(r); trans(r)\ \ wftrec(r,a,H) = H(a, \x\r-``{a}. wftrec(r,x,H))" unfolding wftrec_def apply (subst unfold_the_recfun [unfolded is_recfun_def]) apply (simp_all add: vimage_singleton_iff [THEN iff_sym] the_recfun_cut) done subsubsection\Removal of the Premise \<^term>\trans(r)\\ (*NOT SUITABLE FOR REWRITING: it is recursive!*) lemma wfrec: "wf(r) \ wfrec(r,a,H) = H(a, \x\r-``{a}. wfrec(r,x,H))" unfolding wfrec_def apply (erule wf_trancl [THEN wftrec, THEN ssubst]) apply (rule trans_trancl) apply (rule vimage_pair_mono [THEN restrict_lam_eq, THEN subst_context]) apply (erule r_into_trancl) apply (rule subset_refl) done (*This form avoids giant explosions in proofs. NOTE USE OF \ *) lemma def_wfrec: "\\x. h(x)\wfrec(r,x,H); wf(r)\ \ h(a) = H(a, \x\r-``{a}. h(x))" apply simp apply (elim wfrec) done lemma wfrec_type: "\wf(r); a \ A; field(r)<=A; \x u. \x \ A; u \ Pi(r-``{x}, B)\ \ H(x,u) \ B(x) \ \ wfrec(r,a,H) \ B(a)" apply (rule_tac a = a in wf_induct2, assumption+) apply (subst wfrec, assumption) apply (simp add: lam_type underD) done lemma wfrec_on: "\wf[A](r); a \ A\ \ wfrec[A](r,a,H) = H(a, \x\(r-``{a}) \ A. wfrec[A](r,x,H))" -apply (unfold wf_on_def wfrec_on_def) + unfolding wf_on_def wfrec_on_def apply (erule wfrec [THEN trans]) apply (simp add: vimage_Int_square cons_subset_iff) done text\Minimal-element characterization of well-foundedness\ lemma wf_eq_minimal: "wf(r) \ (\Q x. x \ Q \ (\z\Q. \y. \y,z\:r \ y\Q))" by (unfold wf_def, blast) end diff --git a/src/ZF/ZF_Base.thy b/src/ZF/ZF_Base.thy --- a/src/ZF/ZF_Base.thy +++ b/src/ZF/ZF_Base.thy @@ -1,650 +1,650 @@ (* Title: ZF/ZF_Base.thy Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory Copyright 1993 University of Cambridge *) section \Base of Zermelo-Fraenkel Set Theory\ theory ZF_Base imports FOL begin subsection \Signature\ declare [[eta_contract = false]] typedecl i instance i :: "term" .. axiomatization mem :: "[i, i] \ o" (infixl \\\ 50) \ \membership relation\ and zero :: "i" (\0\) \ \the empty set\ and Pow :: "i \ i" \ \power sets\ and Inf :: "i" \ \infinite set\ and Union :: "i \ i" (\\_\ [90] 90) and PrimReplace :: "[i, [i, i] \ o] \ i" abbreviation not_mem :: "[i, i] \ o" (infixl \\\ 50) \ \negated membership relation\ where "x \ y \ \ (x \ y)" subsection \Bounded Quantifiers\ definition Ball :: "[i, i \ o] \ o" where "Ball(A, P) \ \x. x\A \ P(x)" definition Bex :: "[i, i \ o] \ o" where "Bex(A, P) \ \x. x\A \ P(x)" syntax "_Ball" :: "[pttrn, i, o] \ o" (\(3\_\_./ _)\ 10) "_Bex" :: "[pttrn, i, o] \ o" (\(3\_\_./ _)\ 10) translations "\x\A. P" \ "CONST Ball(A, \x. P)" "\x\A. P" \ "CONST Bex(A, \x. P)" subsection \Variations on Replacement\ (* Derived form of replacement, restricting P to its functional part. The resulting set (for functional P) is the same as with PrimReplace, but the rules are simpler. *) definition Replace :: "[i, [i, i] \ o] \ i" where "Replace(A,P) \ PrimReplace(A, \x y. (\!z. P(x,z)) \ P(x,y))" syntax "_Replace" :: "[pttrn, pttrn, i, o] \ i" (\(1{_ ./ _ \ _, _})\) translations "{y. x\A, Q}" \ "CONST Replace(A, \x y. Q)" (* Functional form of replacement -- analgous to ML's map functional *) definition RepFun :: "[i, i \ i] \ i" where "RepFun(A,f) \ {y . x\A, y=f(x)}" syntax "_RepFun" :: "[i, pttrn, i] \ i" (\(1{_ ./ _ \ _})\ [51,0,51]) translations "{b. x\A}" \ "CONST RepFun(A, \x. b)" (* Separation and Pairing can be derived from the Replacement and Powerset Axioms using the following definitions. *) definition Collect :: "[i, i \ o] \ i" where "Collect(A,P) \ {y . x\A, x=y \ P(x)}" syntax "_Collect" :: "[pttrn, i, o] \ i" (\(1{_ \ _ ./ _})\) translations "{x\A. P}" \ "CONST Collect(A, \x. P)" subsection \General union and intersection\ definition Inter :: "i \ i" (\\_\ [90] 90) where "\(A) \ { x\\(A) . \y\A. x\y}" syntax "_UNION" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) "_INTER" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) translations "\x\A. B" == "CONST Union({B. x\A})" "\x\A. B" == "CONST Inter({B. x\A})" subsection \Finite sets and binary operations\ (*Unordered pairs (Upair) express binary union/intersection and cons; set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*) definition Upair :: "[i, i] \ i" where "Upair(a,b) \ {y. x\Pow(Pow(0)), (x=0 \ y=a) | (x=Pow(0) \ y=b)}" definition Subset :: "[i, i] \ o" (infixl \\\ 50) \ \subset relation\ where subset_def: "A \ B \ \x\A. x\B" definition Diff :: "[i, i] \ i" (infixl \-\ 65) \ \set difference\ where "A - B \ { x\A . \(x\B) }" definition Un :: "[i, i] \ i" (infixl \\\ 65) \ \binary union\ where "A \ B \ \(Upair(A,B))" definition Int :: "[i, i] \ i" (infixl \\\ 70) \ \binary intersection\ where "A \ B \ \(Upair(A,B))" definition cons :: "[i, i] \ i" where "cons(a,A) \ Upair(a,a) \ A" definition succ :: "i \ i" where "succ(i) \ cons(i, i)" nonterminal "is" syntax "" :: "i \ is" (\_\) "_Enum" :: "[i, is] \ is" (\_,/ _\) "_Finset" :: "is \ i" (\{(_)}\) translations "{x, xs}" == "CONST cons(x, {xs})" "{x}" == "CONST cons(x, 0)" subsection \Axioms\ (* ZF axioms -- see Suppes p.238 Axioms for Union, Pow and Replace state existence only, uniqueness is derivable using extensionality. *) axiomatization where extension: "A = B \ A \ B \ B \ A" and Union_iff: "A \ \(C) \ (\B\C. A\B)" and Pow_iff: "A \ Pow(B) \ A \ B" and (*We may name this set, though it is not uniquely defined.*) infinity: "0 \ Inf \ (\y\Inf. succ(y) \ Inf)" and (*This formulation facilitates case analysis on A.*) foundation: "A = 0 \ (\x\A. \y\x. y\A)" and (*Schema axiom since predicate P is a higher-order variable*) replacement: "(\x\A. \y z. P(x,y) \ P(x,z) \ y = z) \ b \ PrimReplace(A,P) \ (\x\A. P(x,b))" subsection \Definite descriptions -- via Replace over the set "1"\ definition The :: "(i \ o) \ i" (binder \THE \ 10) where the_def: "The(P) \ \({y . x \ {0}, P(y)})" definition If :: "[o, i, i] \ i" (\(if (_)/ then (_)/ else (_))\ [10] 10) where if_def: "if P then a else b \ THE z. P \ z=a | \P \ z=b" abbreviation (input) old_if :: "[o, i, i] \ i" (\if '(_,_,_')\) where "if(P,a,b) \ If(P,a,b)" subsection \Ordered Pairing\ (* this "symmetric" definition works better than {{a}, {a,b}} *) definition Pair :: "[i, i] \ i" where "Pair(a,b) \ {{a,a}, {a,b}}" definition fst :: "i \ i" where "fst(p) \ THE a. \b. p = Pair(a, b)" definition snd :: "i \ i" where "snd(p) \ THE b. \a. p = Pair(a, b)" definition split :: "[[i, i] \ 'a, i] \ 'a::{}" \ \for pattern-matching\ where "split(c) \ \p. c(fst(p), snd(p))" (* Patterns -- extends pre-defined type "pttrn" used in abstractions *) nonterminal patterns syntax "_pattern" :: "patterns \ pttrn" (\\_\\) "" :: "pttrn \ patterns" (\_\) "_patterns" :: "[pttrn, patterns] \ patterns" (\_,/_\) "_Tuple" :: "[i, is] \ i" (\\(_,/ _)\\) translations "\x, y, z\" == "\x, \y, z\\" "\x, y\" == "CONST Pair(x, y)" "\\x,y,zs\.b" == "CONST split(\x \y,zs\.b)" "\\x,y\.b" == "CONST split(\x y. b)" definition Sigma :: "[i, i \ i] \ i" where "Sigma(A,B) \ \x\A. \y\B(x). {\x,y\}" abbreviation cart_prod :: "[i, i] \ i" (infixr \\\ 80) \ \Cartesian product\ where "A \ B \ Sigma(A, \_. B)" subsection \Relations and Functions\ (*converse of relation r, inverse of function*) definition converse :: "i \ i" where "converse(r) \ {z. w\r, \x y. w=\x,y\ \ z=\y,x\}" definition domain :: "i \ i" where "domain(r) \ {x. w\r, \y. w=\x,y\}" definition range :: "i \ i" where "range(r) \ domain(converse(r))" definition field :: "i \ i" where "field(r) \ domain(r) \ range(r)" definition relation :: "i \ o" \ \recognizes sets of pairs\ where "relation(r) \ \z\r. \x y. z = \x,y\" definition "function" :: "i \ o" \ \recognizes functions; can have non-pairs\ where "function(r) \ \x y. \x,y\ \ r \ (\y'. \x,y'\ \ r \ y = y')" definition Image :: "[i, i] \ i" (infixl \``\ 90) \ \image\ where image_def: "r `` A \ {y \ range(r). \x\A. \x,y\ \ r}" definition vimage :: "[i, i] \ i" (infixl \-``\ 90) \ \inverse image\ where vimage_def: "r -`` A \ converse(r)``A" (* Restrict the relation r to the domain A *) definition restrict :: "[i, i] \ i" where "restrict(r,A) \ {z \ r. \x\A. \y. z = \x,y\}" (* Abstraction, application and Cartesian product of a family of sets *) definition Lambda :: "[i, i \ i] \ i" where lam_def: "Lambda(A,b) \ {\x,b(x)\. x\A}" definition "apply" :: "[i, i] \ i" (infixl \`\ 90) \ \function application\ where "f`a \ \(f``{a})" definition Pi :: "[i, i \ i] \ i" where "Pi(A,B) \ {f\Pow(Sigma(A,B)). A\domain(f) \ function(f)}" abbreviation function_space :: "[i, i] \ i" (infixr \\\ 60) \ \function space\ where "A \ B \ Pi(A, \_. B)" (* binder syntax *) syntax "_PROD" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) "_SUM" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) "_lam" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) translations "\x\A. B" == "CONST Pi(A, \x. B)" "\x\A. B" == "CONST Sigma(A, \x. B)" "\x\A. f" == "CONST Lambda(A, \x. f)" subsection \ASCII syntax\ notation (ASCII) cart_prod (infixr \*\ 80) and Int (infixl \Int\ 70) and Un (infixl \Un\ 65) and function_space (infixr \->\ 60) and Subset (infixl \<=\ 50) and mem (infixl \:\ 50) and not_mem (infixl \\:\ 50) syntax (ASCII) "_Ball" :: "[pttrn, i, o] \ o" (\(3ALL _:_./ _)\ 10) "_Bex" :: "[pttrn, i, o] \ o" (\(3EX _:_./ _)\ 10) "_Collect" :: "[pttrn, i, o] \ i" (\(1{_: _ ./ _})\) "_Replace" :: "[pttrn, pttrn, i, o] \ i" (\(1{_ ./ _: _, _})\) "_RepFun" :: "[i, pttrn, i] \ i" (\(1{_ ./ _: _})\ [51,0,51]) "_UNION" :: "[pttrn, i, i] \ i" (\(3UN _:_./ _)\ 10) "_INTER" :: "[pttrn, i, i] \ i" (\(3INT _:_./ _)\ 10) "_PROD" :: "[pttrn, i, i] \ i" (\(3PROD _:_./ _)\ 10) "_SUM" :: "[pttrn, i, i] \ i" (\(3SUM _:_./ _)\ 10) "_lam" :: "[pttrn, i, i] \ i" (\(3lam _:_./ _)\ 10) "_Tuple" :: "[i, is] \ i" (\<(_,/ _)>\) "_pattern" :: "patterns \ pttrn" (\<_>\) subsection \Substitution\ (*Useful examples: singletonI RS subst_elem, subst_elem RSN (2,IntI) *) lemma subst_elem: "\b\A; a=b\ \ a\A" by (erule ssubst, assumption) subsection\Bounded universal quantifier\ lemma ballI [intro!]: "\\x. x\A \ P(x)\ \ \x\A. P(x)" by (simp add: Ball_def) lemmas strip = impI allI ballI lemma bspec [dest?]: "\\x\A. P(x); x: A\ \ P(x)" by (simp add: Ball_def) (*Instantiates x first: better for automatic theorem proving?*) lemma rev_ballE [elim]: "\\x\A. P(x); x\A \ Q; P(x) \ Q\ \ Q" by (simp add: Ball_def, blast) lemma ballE: "\\x\A. P(x); P(x) \ Q; x\A \ Q\ \ Q" by blast (*Used in the datatype package*) lemma rev_bspec: "\x: A; \x\A. P(x)\ \ P(x)" by (simp add: Ball_def) (*Trival rewrite rule; @{term"(\x\A.P)<->P"} holds only if A is nonempty!*) lemma ball_triv [simp]: "(\x\A. P) <-> ((\x. x\A) \ P)" by (simp add: Ball_def) (*Congruence rule for rewriting*) lemma ball_cong [cong]: "\A=A'; \x. x\A' \ P(x) <-> P'(x)\ \ (\x\A. P(x)) <-> (\x\A'. P'(x))" by (simp add: Ball_def) lemma atomize_ball: "(\x. x \ A \ P(x)) \ Trueprop (\x\A. P(x))" by (simp only: Ball_def atomize_all atomize_imp) lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball subsection\Bounded existential quantifier\ lemma bexI [intro]: "\P(x); x: A\ \ \x\A. P(x)" by (simp add: Bex_def, blast) (*The best argument order when there is only one @{term"x\A"}*) lemma rev_bexI: "\x\A; P(x)\ \ \x\A. P(x)" by blast (*Not of the general form for such rules. The existential quanitifer becomes universal. *) lemma bexCI: "\\x\A. \P(x) \ P(a); a: A\ \ \x\A. P(x)" by blast lemma bexE [elim!]: "\\x\A. P(x); \x. \x\A; P(x)\ \ Q\ \ Q" by (simp add: Bex_def, blast) (*We do not even have @{term"(\x\A. True) <-> True"} unless @{term"A" is nonempty\*) lemma bex_triv [simp]: "(\x\A. P) <-> ((\x. x\A) \ P)" by (simp add: Bex_def) lemma bex_cong [cong]: "\A=A'; \x. x\A' \ P(x) <-> P'(x)\ \ (\x\A. P(x)) <-> (\x\A'. P'(x))" by (simp add: Bex_def cong: conj_cong) subsection\Rules for subsets\ lemma subsetI [intro!]: "(\x. x\A \ x\B) \ A \ B" by (simp add: subset_def) (*Rule in Modus Ponens style [was called subsetE] *) lemma subsetD [elim]: "\A \ B; c\A\ \ c\B" unfolding subset_def apply (erule bspec, assumption) done (*Classical elimination rule*) lemma subsetCE [elim]: "\A \ B; c\A \ P; c\B \ P\ \ P" by (simp add: subset_def, blast) (*Sometimes useful with premises in this order*) lemma rev_subsetD: "\c\A; A<=B\ \ c\B" by blast lemma contra_subsetD: "\A \ B; c \ B\ \ c \ A" by blast lemma rev_contra_subsetD: "\c \ B; A \ B\ \ c \ A" by blast lemma subset_refl [simp]: "A \ A" by blast lemma subset_trans: "\A<=B; B<=C\ \ A<=C" by blast (*Useful for proving A<=B by rewriting in some cases*) lemma subset_iff: "A<=B <-> (\x. x\A \ x\B)" -apply (unfold subset_def Ball_def) + unfolding subset_def Ball_def apply (rule iff_refl) done text\For calculations\ declare subsetD [trans] rev_subsetD [trans] subset_trans [trans] subsection\Rules for equality\ (*Anti-symmetry of the subset relation*) lemma equalityI [intro]: "\A \ B; B \ A\ \ A = B" by (rule extension [THEN iffD2], rule conjI) lemma equality_iffI: "(\x. x\A <-> x\B) \ A = B" by (rule equalityI, blast+) lemmas equalityD1 = extension [THEN iffD1, THEN conjunct1] lemmas equalityD2 = extension [THEN iffD1, THEN conjunct2] lemma equalityE: "\A = B; \A<=B; B<=A\ \ P\ \ P" by (blast dest: equalityD1 equalityD2) lemma equalityCE: "\A = B; \c\A; c\B\ \ P; \c\A; c\B\ \ P\ \ P" by (erule equalityE, blast) lemma equality_iffD: "A = B \ (\x. x \ A <-> x \ B)" by auto subsection\Rules for Replace -- the derived form of replacement\ lemma Replace_iff: "b \ {y. x\A, P(x,y)} <-> (\x\A. P(x,b) \ (\y. P(x,y) \ y=b))" unfolding Replace_def apply (rule replacement [THEN iff_trans], blast+) done (*Introduction; there must be a unique y such that P(x,y), namely y=b. *) lemma ReplaceI [intro]: "\P(x,b); x: A; \y. P(x,y) \ y=b\ \ b \ {y. x\A, P(x,y)}" by (rule Replace_iff [THEN iffD2], blast) (*Elimination; may asssume there is a unique y such that P(x,y), namely y=b. *) lemma ReplaceE: "\b \ {y. x\A, P(x,y)}; \x. \x: A; P(x,b); \y. P(x,y)\y=b\ \ R \ \ R" by (rule Replace_iff [THEN iffD1, THEN bexE], simp+) (*As above but without the (generally useless) 3rd assumption*) lemma ReplaceE2 [elim!]: "\b \ {y. x\A, P(x,y)}; \x. \x: A; P(x,b)\ \ R \ \ R" by (erule ReplaceE, blast) lemma Replace_cong [cong]: "\A=B; \x y. x\B \ P(x,y) <-> Q(x,y)\ \ Replace(A,P) = Replace(B,Q)" apply (rule equality_iffI) apply (simp add: Replace_iff) done subsection\Rules for RepFun\ lemma RepFunI: "a \ A \ f(a) \ {f(x). x\A}" by (simp add: RepFun_def Replace_iff, blast) (*Useful for coinduction proofs*) lemma RepFun_eqI [intro]: "\b=f(a); a \ A\ \ b \ {f(x). x\A}" apply (erule ssubst) apply (erule RepFunI) done lemma RepFunE [elim!]: "\b \ {f(x). x\A}; \x.\x\A; b=f(x)\ \ P\ \ P" by (simp add: RepFun_def Replace_iff, blast) lemma RepFun_cong [cong]: "\A=B; \x. x\B \ f(x)=g(x)\ \ RepFun(A,f) = RepFun(B,g)" by (simp add: RepFun_def) lemma RepFun_iff [simp]: "b \ {f(x). x\A} <-> (\x\A. b=f(x))" by (unfold Bex_def, blast) lemma triv_RepFun [simp]: "{x. x\A} = A" by blast subsection\Rules for Collect -- forming a subset by separation\ (*Separation is derivable from Replacement*) lemma separation [simp]: "a \ {x\A. P(x)} <-> a\A \ P(a)" by (unfold Collect_def, blast) lemma CollectI [intro!]: "\a\A; P(a)\ \ a \ {x\A. P(x)}" by simp lemma CollectE [elim!]: "\a \ {x\A. P(x)}; \a\A; P(a)\ \ R\ \ R" by simp lemma CollectD1: "a \ {x\A. P(x)} \ a\A" by (erule CollectE, assumption) lemma CollectD2: "a \ {x\A. P(x)} \ P(a)" by (erule CollectE, assumption) lemma Collect_cong [cong]: "\A=B; \x. x\B \ P(x) <-> Q(x)\ \ Collect(A, \x. P(x)) = Collect(B, \x. Q(x))" by (simp add: Collect_def) subsection\Rules for Unions\ declare Union_iff [simp] (*The order of the premises presupposes that C is rigid; A may be flexible*) lemma UnionI [intro]: "\B: C; A: B\ \ A: \(C)" by (simp, blast) lemma UnionE [elim!]: "\A \ \(C); \B.\A: B; B: C\ \ R\ \ R" by (simp, blast) subsection\Rules for Unions of families\ (* @{term"\x\A. B(x)"} abbreviates @{term"\({B(x). x\A})"} *) lemma UN_iff [simp]: "b \ (\x\A. B(x)) <-> (\x\A. b \ B(x))" by (simp add: Bex_def, blast) (*The order of the premises presupposes that A is rigid; b may be flexible*) lemma UN_I: "\a: A; b: B(a)\ \ b: (\x\A. B(x))" by (simp, blast) lemma UN_E [elim!]: "\b \ (\x\A. B(x)); \x.\x: A; b: B(x)\ \ R\ \ R" by blast lemma UN_cong: "\A=B; \x. x\B \ C(x)=D(x)\ \ (\x\A. C(x)) = (\x\B. D(x))" by simp (*No "Addcongs [UN_cong]" because @{term\} is a combination of constants*) (* UN_E appears before UnionE so that it is tried first, to avoid expensive calls to hyp_subst_tac. Cannot include UN_I as it is unsafe: would enlarge the search space.*) subsection\Rules for the empty set\ (*The set @{term"{x\0. False}"} is empty; by foundation it equals 0 See Suppes, page 21.*) lemma not_mem_empty [simp]: "a \ 0" apply (cut_tac foundation) apply (best dest: equalityD2) done lemmas emptyE [elim!] = not_mem_empty [THEN notE] lemma empty_subsetI [simp]: "0 \ A" by blast lemma equals0I: "\\y. y\A \ False\ \ A=0" by blast lemma equals0D [dest]: "A=0 \ a \ A" by blast declare sym [THEN equals0D, dest] lemma not_emptyI: "a\A \ A \ 0" by blast lemma not_emptyE: "\A \ 0; \x. x\A \ R\ \ R" by blast subsection\Rules for Inter\ (*Not obviously useful for proving InterI, InterD, InterE*) lemma Inter_iff: "A \ \(C) <-> (\x\C. A: x) \ C\0" by (simp add: Inter_def Ball_def, blast) (* Intersection is well-behaved only if the family is non-empty! *) lemma InterI [intro!]: "\\x. x: C \ A: x; C\0\ \ A \ \(C)" by (simp add: Inter_iff) (*A "destruct" rule -- every B in C contains A as an element, but A\B can hold when B\C does not! This rule is analogous to "spec". *) lemma InterD [elim, Pure.elim]: "\A \ \(C); B \ C\ \ A \ B" by (unfold Inter_def, blast) (*"Classical" elimination rule -- does not require exhibiting @{term"B\C"} *) lemma InterE [elim]: "\A \ \(C); B\C \ R; A\B \ R\ \ R" by (simp add: Inter_def, blast) subsection\Rules for Intersections of families\ (* @{term"\x\A. B(x)"} abbreviates @{term"\({B(x). x\A})"} *) lemma INT_iff: "b \ (\x\A. B(x)) <-> (\x\A. b \ B(x)) \ A\0" by (force simp add: Inter_def) lemma INT_I: "\\x. x: A \ b: B(x); A\0\ \ b: (\x\A. B(x))" by blast lemma INT_E: "\b \ (\x\A. B(x)); a: A\ \ b \ B(a)" by blast lemma INT_cong: "\A=B; \x. x\B \ C(x)=D(x)\ \ (\x\A. C(x)) = (\x\B. D(x))" by simp (*No "Addcongs [INT_cong]" because @{term\} is a combination of constants*) subsection\Rules for Powersets\ lemma PowI: "A \ B \ A \ Pow(B)" by (erule Pow_iff [THEN iffD2]) lemma PowD: "A \ Pow(B) \ A<=B" by (erule Pow_iff [THEN iffD1]) declare Pow_iff [iff] lemmas Pow_bottom = empty_subsetI [THEN PowI] \ \\<^term>\0 \ Pow(B)\\ lemmas Pow_top = subset_refl [THEN PowI] \ \\<^term>\A \ Pow(A)\\ subsection\Cantor's Theorem: There is no surjection from a set to its powerset.\ (*The search is undirected. Allowing redundant introduction rules may make it diverge. Variable b represents ANY map, such as (lam x\A.b(x)): A->Pow(A). *) lemma cantor: "\S \ Pow(A). \x\A. b(x) \ S" by (best elim!: equalityCE del: ReplaceI RepFun_eqI) end diff --git a/src/ZF/Zorn.thy b/src/ZF/Zorn.thy --- a/src/ZF/Zorn.thy +++ b/src/ZF/Zorn.thy @@ -1,517 +1,517 @@ (* Title: ZF/Zorn.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section\Zorn's Lemma\ theory Zorn imports OrderArith AC Inductive begin text\Based upon the unpublished article ``Towards the Mechanization of the Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.\ definition Subset_rel :: "i\i" where "Subset_rel(A) \ {z \ A*A . \x y. z=\x,y\ \ x<=y \ x\y}" definition chain :: "i\i" where "chain(A) \ {F \ Pow(A). \X\F. \Y\F. X<=Y | Y<=X}" definition super :: "[i,i]\i" where "super(A,c) \ {d \ chain(A). c<=d \ c\d}" definition maxchain :: "i\i" where "maxchain(A) \ {c \ chain(A). super(A,c)=0}" definition increasing :: "i\i" where "increasing(A) \ {f \ Pow(A)->Pow(A). \x. x<=A \ x<=f`x}" text\Lemma for the inductive definition below\ lemma Union_in_Pow: "Y \ Pow(Pow(A)) \ \(Y) \ Pow(A)" by blast text\We could make the inductive definition conditional on \<^term>\next \ increasing(S)\ but instead we make this a side-condition of an introduction rule. Thus the induction rule lets us assume that condition! Many inductive proofs are therefore unconditional.\ consts "TFin" :: "[i,i]\i" inductive domains "TFin(S,next)" \ "Pow(S)" intros nextI: "\x \ TFin(S,next); next \ increasing(S)\ \ next`x \ TFin(S,next)" Pow_UnionI: "Y \ Pow(TFin(S,next)) \ \(Y) \ TFin(S,next)" monos Pow_mono con_defs increasing_def type_intros CollectD1 [THEN apply_funtype] Union_in_Pow subsection\Mathematical Preamble\ lemma Union_lemma0: "(\x\C. x<=A | B<=x) \ \(C)<=A | B<=\(C)" by blast lemma Inter_lemma0: "\c \ C; \x\C. A<=x | x<=B\ \ A \ \(C) | \(C) \ B" by blast subsection\The Transfinite Construction\ lemma increasingD1: "f \ increasing(A) \ f \ Pow(A)->Pow(A)" unfolding increasing_def apply (erule CollectD1) done lemma increasingD2: "\f \ increasing(A); x<=A\ \ x \ f`x" by (unfold increasing_def, blast) lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI] lemmas TFin_is_subset = TFin.dom_subset [THEN subsetD, THEN PowD] text\Structural induction on \<^term>\TFin(S,next)\\ lemma TFin_induct: "\n \ TFin(S,next); \x. \x \ TFin(S,next); P(x); next \ increasing(S)\ \ P(next`x); \Y. \Y \ TFin(S,next); \y\Y. P(y)\ \ P(\(Y)) \ \ P(n)" by (erule TFin.induct, blast+) subsection\Some Properties of the Transfinite Construction\ lemmas increasing_trans = subset_trans [OF _ increasingD2, OF _ _ TFin_is_subset] text\Lemma 1 of section 3.1\ lemma TFin_linear_lemma1: "\n \ TFin(S,next); m \ TFin(S,next); \x \ TFin(S,next) . x<=m \ x=m | next`x<=m\ \ n<=m | next`m<=n" apply (erule TFin_induct) apply (erule_tac [2] Union_lemma0) (*or just Blast_tac*) (*downgrade subsetI from intro! to intro*) apply (blast dest: increasing_trans) done text\Lemma 2 of section 3.2. Interesting in its own right! Requires \<^term>\next \ increasing(S)\ in the second induction step.\ lemma TFin_linear_lemma2: "\m \ TFin(S,next); next \ increasing(S)\ \ \n \ TFin(S,next). n<=m \ n=m | next`n \ m" apply (erule TFin_induct) apply (rule impI [THEN ballI]) txt\case split using \TFin_linear_lemma1\\ apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], assumption+) apply (blast del: subsetI intro: increasing_trans subsetI, blast) txt\second induction step\ apply (rule impI [THEN ballI]) apply (rule Union_lemma0 [THEN disjE]) apply (erule_tac [3] disjI2) prefer 2 apply blast apply (rule ballI) apply (drule bspec, assumption) apply (drule subsetD, assumption) apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], assumption+, blast) apply (erule increasingD2 [THEN subset_trans, THEN disjI1]) apply (blast dest: TFin_is_subset)+ done text\a more convenient form for Lemma 2\ lemma TFin_subsetD: "\n<=m; m \ TFin(S,next); n \ TFin(S,next); next \ increasing(S)\ \ n=m | next`n \ m" by (blast dest: TFin_linear_lemma2 [rule_format]) text\Consequences from section 3.3 -- Property 3.2, the ordering is total\ lemma TFin_subset_linear: "\m \ TFin(S,next); n \ TFin(S,next); next \ increasing(S)\ \ n \ m | m<=n" apply (rule disjE) apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) apply (assumption+, erule disjI2) apply (blast del: subsetI intro: subsetI increasingD2 [THEN subset_trans] TFin_is_subset) done text\Lemma 3 of section 3.3\ lemma equal_next_upper: "\n \ TFin(S,next); m \ TFin(S,next); m = next`m\ \ n \ m" apply (erule TFin_induct) apply (drule TFin_subsetD) apply (assumption+, force, blast) done text\Property 3.3 of section 3.3\ lemma equal_next_Union: "\m \ TFin(S,next); next \ increasing(S)\ \ m = next`m <-> m = \(TFin(S,next))" apply (rule iffI) apply (rule Union_upper [THEN equalityI]) apply (rule_tac [2] equal_next_upper [THEN Union_least]) apply (assumption+) apply (erule ssubst) apply (rule increasingD2 [THEN equalityI], assumption) apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.nextI TFin_is_subset)+ done subsection\Hausdorff's Theorem: Every Set Contains a Maximal Chain\ text\NOTE: We assume the partial ordering is \\\, the subset relation!\ text\* Defining the "next" operation for Hausdorff's Theorem *\ lemma chain_subset_Pow: "chain(A) \ Pow(A)" unfolding chain_def apply (rule Collect_subset) done lemma super_subset_chain: "super(A,c) \ chain(A)" unfolding super_def apply (rule Collect_subset) done lemma maxchain_subset_chain: "maxchain(A) \ chain(A)" unfolding maxchain_def apply (rule Collect_subset) done lemma choice_super: "\ch \ (\X \ Pow(chain(S)) - {0}. X); X \ chain(S); X \ maxchain(S)\ \ ch ` super(S,X) \ super(S,X)" apply (erule apply_type) apply (unfold super_def maxchain_def, blast) done lemma choice_not_equals: "\ch \ (\X \ Pow(chain(S)) - {0}. X); X \ chain(S); X \ maxchain(S)\ \ ch ` super(S,X) \ X" apply (rule notI) apply (drule choice_super, assumption, assumption) apply (simp add: super_def) done text\This justifies Definition 4.4\ lemma Hausdorff_next_exists: "ch \ (\X \ Pow(chain(S))-{0}. X) \ \next \ increasing(S). \X \ Pow(S). next`X = if(X \ chain(S)-maxchain(S), ch`super(S,X), X)" apply (rule_tac x="\X\Pow(S). if X \ chain(S) - maxchain(S) then ch ` super(S, X) else X" in bexI) apply force unfolding increasing_def apply (rule CollectI) apply (rule lam_type) apply (simp (no_asm_simp)) apply (blast dest: super_subset_chain [THEN subsetD] chain_subset_Pow [THEN subsetD] choice_super) txt\Now, verify that it increases\ apply (simp (no_asm_simp) add: Pow_iff subset_refl) apply safe apply (drule choice_super) apply (assumption+) apply (simp add: super_def, blast) done text\Lemma 4\ lemma TFin_chain_lemma4: "\c \ TFin(S,next); ch \ (\X \ Pow(chain(S))-{0}. X); next \ increasing(S); \X \ Pow(S). next`X = if(X \ chain(S)-maxchain(S), ch`super(S,X), X)\ \ c \ chain(S)" apply (erule TFin_induct) apply (simp (no_asm_simp) add: chain_subset_Pow [THEN subsetD, THEN PowD] choice_super [THEN super_subset_chain [THEN subsetD]]) unfolding chain_def apply (rule CollectI, blast, safe) apply (rule_tac m1=B and n1=Ba in TFin_subset_linear [THEN disjE], fast+) txt\\Blast_tac's\ slow\ done theorem Hausdorff: "\c. c \ maxchain(S)" apply (rule AC_Pi_Pow [THEN exE]) apply (rule Hausdorff_next_exists [THEN bexE], assumption) apply (rename_tac ch "next") apply (subgoal_tac "\(TFin (S,next)) \ chain (S) ") prefer 2 apply (blast intro!: TFin_chain_lemma4 subset_refl [THEN TFin_UnionI]) apply (rule_tac x = "\(TFin (S,next))" in exI) apply (rule classical) apply (subgoal_tac "next ` Union(TFin (S,next)) = \(TFin (S,next))") apply (rule_tac [2] equal_next_Union [THEN iffD2, symmetric]) apply (rule_tac [2] subset_refl [THEN TFin_UnionI]) prefer 2 apply assumption apply (rule_tac [2] refl) apply (simp add: subset_refl [THEN TFin_UnionI, THEN TFin.dom_subset [THEN subsetD, THEN PowD]]) apply (erule choice_not_equals [THEN notE]) apply (assumption+) done subsection\Zorn's Lemma: If All Chains in S Have Upper Bounds In S, then S contains a Maximal Element\ text\Used in the proof of Zorn's Lemma\ lemma chain_extend: "\c \ chain(A); z \ A; \x \ c. x<=z\ \ cons(z,c) \ chain(A)" by (unfold chain_def, blast) lemma Zorn: "\c \ chain(S). \(c) \ S \ \y \ S. \z \ S. y<=z \ y=z" apply (rule Hausdorff [THEN exE]) apply (simp add: maxchain_def) apply (rename_tac c) apply (rule_tac x = "\(c)" in bexI) prefer 2 apply blast apply safe apply (rename_tac z) apply (rule classical) apply (subgoal_tac "cons (z,c) \ super (S,c) ") apply (blast elim: equalityE) apply (unfold super_def, safe) apply (fast elim: chain_extend) apply (fast elim: equalityE) done text \Alternative version of Zorn's Lemma\ theorem Zorn2: "\c \ chain(S). \y \ S. \x \ c. x \ y \ \y \ S. \z \ S. y<=z \ y=z" apply (cut_tac Hausdorff maxchain_subset_chain) apply (erule exE) apply (drule subsetD, assumption) apply (drule bspec, assumption, erule bexE) apply (rule_tac x = y in bexI) prefer 2 apply assumption apply clarify apply rule apply assumption apply rule apply (rule ccontr) apply (frule_tac z=z in chain_extend) apply (assumption, blast) -apply (unfold maxchain_def super_def) + unfolding maxchain_def super_def apply (blast elim!: equalityCE) done subsection\Zermelo's Theorem: Every Set can be Well-Ordered\ text\Lemma 5\ lemma TFin_well_lemma5: "\n \ TFin(S,next); Z \ TFin(S,next); z:Z; \ \(Z) \ Z\ \ \m \ Z. n \ m" apply (erule TFin_induct) prefer 2 apply blast txt\second induction step is easy\ apply (rule ballI) apply (rule bspec [THEN TFin_subsetD, THEN disjE], auto) apply (subgoal_tac "m = \(Z) ") apply blast+ done text\Well-ordering of \<^term>\TFin(S,next)\\ lemma well_ord_TFin_lemma: "\Z \ TFin(S,next); z \ Z\ \ \(Z) \ Z" apply (rule classical) apply (subgoal_tac "Z = {\(TFin (S,next))}") apply (simp (no_asm_simp) add: Inter_singleton) apply (erule equal_singleton) apply (rule Union_upper [THEN equalityI]) apply (rule_tac [2] subset_refl [THEN TFin_UnionI, THEN TFin_well_lemma5, THEN bspec], blast+) done text\This theorem just packages the previous result\ lemma well_ord_TFin: "next \ increasing(S) \ well_ord(TFin(S,next), Subset_rel(TFin(S,next)))" apply (rule well_ordI) -apply (unfold Subset_rel_def linear_def) + unfolding Subset_rel_def linear_def txt\Prove the well-foundedness goal\ apply (rule wf_onI) apply (frule well_ord_TFin_lemma, assumption) apply (drule_tac x = "\(Z) " in bspec, assumption) apply blast txt\Now prove the linearity goal\ apply (intro ballI) apply (case_tac "x=y") apply blast txt\The \<^term>\x\y\ case remains\ apply (rule_tac n1=x and m1=y in TFin_subset_linear [THEN disjE], assumption+, blast+) done text\* Defining the "next" operation for Zermelo's Theorem *\ lemma choice_Diff: "\ch \ (\X \ Pow(S) - {0}. X); X \ S; X\S\ \ ch ` (S-X) \ S-X" apply (erule apply_type) apply (blast elim!: equalityE) done text\This justifies Definition 6.1\ lemma Zermelo_next_exists: "ch \ (\X \ Pow(S)-{0}. X) \ \next \ increasing(S). \X \ Pow(S). next`X = (if X=S then S else cons(ch`(S-X), X))" apply (rule_tac x="\X\Pow(S). if X=S then S else cons(ch`(S-X), X)" in bexI) apply force unfolding increasing_def apply (rule CollectI) apply (rule lam_type) txt\Type checking is surprisingly hard!\ apply (simp (no_asm_simp) add: Pow_iff cons_subset_iff subset_refl) apply (blast intro!: choice_Diff [THEN DiffD1]) txt\Verify that it increases\ apply (intro allI impI) apply (simp add: Pow_iff subset_consI subset_refl) done text\The construction of the injection\ lemma choice_imp_injection: "\ch \ (\X \ Pow(S)-{0}. X); next \ increasing(S); \X \ Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X))\ \ (\ x \ S. \({y \ TFin(S,next). x \ y})) \ inj(S, TFin(S,next) - {S})" apply (rule_tac d = "\y. ch` (S-y) " in lam_injective) apply (rule DiffI) apply (rule Collect_subset [THEN TFin_UnionI]) apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE) apply (subgoal_tac "x \ \({y \ TFin (S,next) . x \ y}) ") prefer 2 apply (blast elim: equalityE) apply (subgoal_tac "\({y \ TFin (S,next) . x \ y}) \ S") prefer 2 apply (blast elim: equalityE) txt\For proving \x \ next`\(...)\. Abrial and Laffitte's justification appears to be faulty.\ apply (subgoal_tac "\ next ` Union({y \ TFin (S,next) . x \ y}) \ \({y \ TFin (S,next) . x \ y}) ") prefer 2 apply (simp del: Union_iff add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset] Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2]) apply (subgoal_tac "x \ next ` Union({y \ TFin (S,next) . x \ y}) ") prefer 2 apply (blast intro!: Collect_subset [THEN TFin_UnionI] TFin.nextI) txt\End of the lemmas!\ apply (simp add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset]) done text\The wellordering theorem\ theorem AC_well_ord: "\r. well_ord(S,r)" apply (rule AC_Pi_Pow [THEN exE]) apply (rule Zermelo_next_exists [THEN bexE], assumption) apply (rule exI) apply (rule well_ord_rvimage) apply (erule_tac [2] well_ord_TFin) apply (rule choice_imp_injection [THEN inj_weaken_type], blast+) done subsection \Zorn's Lemma for Partial Orders\ text \Reimported from HOL by Clemens Ballarin.\ definition Chain :: "i \ i" where "Chain(r) = {A \ Pow(field(r)). \a\A. \b\A. \a, b\ \ r | \b, a\ \ r}" lemma mono_Chain: "r \ s \ Chain(r) \ Chain(s)" unfolding Chain_def by blast theorem Zorn_po: assumes po: "Partial_order(r)" and u: "\C\Chain(r). \u\field(r). \a\C. \a, u\ \ r" shows "\m\field(r). \a\field(r). \m, a\ \ r \ a = m" proof - have "Preorder(r)" using po by (simp add: partial_order_on_def) \ \Mirror r in the set of subsets below (wrt r) elements of A (?).\ let ?B = "\x\field(r). r -`` {x}" let ?S = "?B `` field(r)" have "\C\chain(?S). \U\?S. \A\C. A \ U" proof (clarsimp simp: chain_def Subset_rel_def bex_image_simp) fix C assume 1: "C \ ?S" and 2: "\A\C. \B\C. A \ B | B \ A" let ?A = "{x \ field(r). \M\C. M = ?B`x}" have "C = ?B `` ?A" using 1 apply (auto simp: image_def) apply rule apply rule apply (drule subsetD) apply assumption apply (erule CollectE) apply rule apply assumption apply (erule bexE) apply rule prefer 2 apply assumption apply rule apply (erule lamE) apply simp apply assumption apply (thin_tac "C \ X" for X) apply (fast elim: lamE) done have "?A \ Chain(r)" proof (simp add: Chain_def subsetI, intro conjI ballI impI) fix a b assume "a \ field(r)" "r -`` {a} \ C" "b \ field(r)" "r -`` {b} \ C" hence "r -`` {a} \ r -`` {b} | r -`` {b} \ r -`` {a}" using 2 by auto then show "\a, b\ \ r | \b, a\ \ r" using \Preorder(r)\ \a \ field(r)\ \b \ field(r)\ by (simp add: subset_vimage1_vimage1_iff) qed then obtain u where uA: "u \ field(r)" "\a\?A. \a, u\ \ r" using u apply auto apply (drule bspec) apply assumption apply auto done have "\A\C. A \ r -`` {u}" proof (auto intro!: vimageI) fix a B assume aB: "B \ C" "a \ B" with 1 obtain x where "x \ field(r)" "B = r -`` {x}" apply - apply (drule subsetD) apply assumption apply (erule imageE) apply (erule lamE) apply simp done then show "\a, u\ \ r" using uA aB \Preorder(r)\ by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+ qed then show "\U\field(r). \A\C. A \ r -`` {U}" using \u \ field(r)\ .. qed from Zorn2 [OF this] obtain m B where "m \ field(r)" "B = r -`` {m}" "\x\field(r). B \ r -`` {x} \ B = r -`` {x}" by (auto elim!: lamE simp: ball_image_simp) then have "\a\field(r). \m, a\ \ r \ a = m" using po \Preorder(r)\ \m \ field(r)\ by (auto simp: subset_vimage1_vimage1_iff Partial_order_eq_vimage1_vimage1_iff) then show ?thesis using \m \ field(r)\ by blast qed end diff --git a/src/ZF/ex/CoUnit.thy b/src/ZF/ex/CoUnit.thy --- a/src/ZF/ex/CoUnit.thy +++ b/src/ZF/ex/CoUnit.thy @@ -1,102 +1,102 @@ (* Title: ZF/ex/CoUnit.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Trivial codatatype definitions, one of which goes wrong!\ theory CoUnit imports ZF begin text \ See discussion in: L C Paulson. A Concrete Final Coalgebra Theorem for ZF Set Theory. Report 334, Cambridge University Computer Laboratory. 1994. \bigskip This degenerate definition does not work well because the one constructor's definition is trivial! The same thing occurs with Aczel's Special Final Coalgebra Theorem. \ consts counit :: i codatatype "counit" = Con ("x \ counit") inductive_cases ConE: "Con(x) \ counit" \ \USELESS because folding on \<^term>\Con(xa) \ xa\ fails.\ lemma Con_iff: "Con(x) = Con(y) \ x = y" \ \Proving freeness results.\ by (auto elim!: counit.free_elims) lemma counit_eq_univ: "counit = quniv(0)" \ \Should be a singleton, not everything!\ apply (rule counit.dom_subset [THEN equalityI]) apply (rule subsetI) apply (erule counit.coinduct) apply (rule subset_refl) - apply (unfold counit.con_defs) + unfolding counit.con_defs apply fast done text \ \medskip A similar example, but the constructor is non-degenerate and it works! The resulting set is a singleton. \ consts counit2 :: i codatatype "counit2" = Con2 ("x \ counit2", "y \ counit2") inductive_cases Con2E: "Con2(x, y) \ counit2" lemma Con2_iff: "Con2(x, y) = Con2(x', y') \ x = x' \ y = y'" \ \Proving freeness results.\ by (fast elim!: counit2.free_elims) lemma Con2_bnd_mono: "bnd_mono(univ(0), \x. Con2(x, x))" - apply (unfold counit2.con_defs) + unfolding counit2.con_defs apply (rule bnd_monoI) apply (assumption | rule subset_refl QPair_subset_univ QPair_mono)+ done lemma lfp_Con2_in_counit2: "lfp(univ(0), \x. Con2(x,x)) \ counit2" apply (rule singletonI [THEN counit2.coinduct]) apply (rule qunivI [THEN singleton_subsetI]) apply (rule subset_trans [OF lfp_subset empty_subsetI [THEN univ_mono]]) apply (fast intro!: Con2_bnd_mono [THEN lfp_unfold]) done lemma counit2_Int_Vset_subset [rule_format]: "Ord(i) \ \x y. x \ counit2 \ y \ counit2 \ x \ Vset(i) \ y" \ \Lemma for proving finality.\ apply (erule trans_induct) apply (tactic "safe_tac (put_claset subset_cs \<^context>)") apply (erule counit2.cases) apply (erule counit2.cases) - apply (unfold counit2.con_defs) + unfolding counit2.con_defs apply (tactic \fast_tac (put_claset subset_cs \<^context> addSIs [@{thm QPair_Int_Vset_subset_UN} RS @{thm subset_trans}, @{thm QPair_mono}] addSEs [@{thm Ord_in_Ord}, @{thm Pair_inject}]) 1\) done lemma counit2_implies_equal: "\x \ counit2; y \ counit2\ \ x = y" apply (rule equalityI) apply (assumption | rule conjI counit2_Int_Vset_subset [THEN Int_Vset_subset])+ done lemma counit2_eq_univ: "counit2 = {lfp(univ(0), \x. Con2(x,x))}" apply (rule equalityI) apply (rule_tac [2] lfp_Con2_in_counit2 [THEN singleton_subsetI]) apply (rule subsetI) apply (drule lfp_Con2_in_counit2 [THEN counit2_implies_equal]) apply (erule subst) apply (rule singletonI) done end diff --git a/src/ZF/ex/Commutation.thy b/src/ZF/ex/Commutation.thy --- a/src/ZF/ex/Commutation.thy +++ b/src/ZF/ex/Commutation.thy @@ -1,142 +1,142 @@ (* Title: ZF/ex/Commutation.thy Author: Tobias Nipkow \ Sidi Ould Ehmety Copyright 1995 TU Muenchen Commutation theory for proving the Church Rosser theorem. *) theory Commutation imports ZF begin definition square :: "[i, i, i, i] \ o" where "square(r,s,t,u) \ (\a b. \a,b\ \ r \ (\c. \a, c\ \ s \ (\x. \b,x\ \ t \ \c,x\ \ u)))" definition commute :: "[i, i] \ o" where "commute(r,s) \ square(r,s,s,r)" definition diamond :: "i\o" where "diamond(r) \ commute(r, r)" definition strip :: "i\o" where "strip(r) \ commute(r^*, r)" definition Church_Rosser :: "i \ o" where "Church_Rosser(r) \ (\x y. \x,y\ \ (r \ converse(r))^* \ (\z. \x,z\ \ r^* \ \y,z\ \ r^*))" definition confluent :: "i\o" where "confluent(r) \ diamond(r^*)" lemma square_sym: "square(r,s,t,u) \ square(s,r,u,t)" unfolding square_def by blast lemma square_subset: "\square(r,s,t,u); t \ t'\ \ square(r,s,t',u)" unfolding square_def by blast lemma square_rtrancl: "square(r,s,s,t) \ field(s)<=field(t) \ square(r^*,s,s,t^*)" apply (unfold square_def, clarify) apply (erule rtrancl_induct) apply (blast intro: rtrancl_refl) apply (blast intro: rtrancl_into_rtrancl) done (* A special case of square_rtrancl_on *) lemma diamond_strip: "diamond(r) \ strip(r)" -apply (unfold diamond_def commute_def strip_def) + unfolding diamond_def commute_def strip_def apply (rule square_rtrancl, simp_all) done (*** commute ***) lemma commute_sym: "commute(r,s) \ commute(s,r)" unfolding commute_def by (blast intro: square_sym) lemma commute_rtrancl: "commute(r,s) \ field(r)=field(s) \ commute(r^*,s^*)" unfolding commute_def apply (rule square_rtrancl) apply (rule square_sym [THEN square_rtrancl, THEN square_sym]) apply (simp_all add: rtrancl_field) done lemma confluentD: "confluent(r) \ diamond(r^*)" by (simp add: confluent_def) lemma strip_confluent: "strip(r) \ confluent(r)" -apply (unfold strip_def confluent_def diamond_def) + unfolding strip_def confluent_def diamond_def apply (drule commute_rtrancl) apply (simp_all add: rtrancl_field) done lemma commute_Un: "\commute(r,t); commute(s,t)\ \ commute(r \ s, t)" unfolding commute_def square_def by blast lemma diamond_Un: "\diamond(r); diamond(s); commute(r, s)\ \ diamond(r \ s)" unfolding diamond_def by (blast intro: commute_Un commute_sym) lemma diamond_confluent: "diamond(r) \ confluent(r)" -apply (unfold diamond_def confluent_def) + unfolding diamond_def confluent_def apply (erule commute_rtrancl, simp) done lemma confluent_Un: "\confluent(r); confluent(s); commute(r^*, s^*); relation(r); relation(s)\ \ confluent(r \ s)" unfolding confluent_def apply (rule rtrancl_Un_rtrancl [THEN subst], auto) apply (blast dest: diamond_Un intro: diamond_confluent [THEN confluentD]) done lemma diamond_to_confluence: "\diamond(r); s \ r; r<= s^*\ \ confluent(s)" apply (drule rtrancl_subset [symmetric], assumption) apply (simp_all add: confluent_def) apply (blast intro: diamond_confluent [THEN confluentD]) done (*** Church_Rosser ***) lemma Church_Rosser1: "Church_Rosser(r) \ confluent(r)" apply (unfold confluent_def Church_Rosser_def square_def commute_def diamond_def, auto) apply (drule converseI) apply (simp (no_asm_use) add: rtrancl_converse [symmetric]) apply (drule_tac x = b in spec) apply (drule_tac x1 = c in spec [THEN mp]) apply (rule_tac b = a in rtrancl_trans) apply (blast intro: rtrancl_mono [THEN subsetD])+ done lemma Church_Rosser2: "confluent(r) \ Church_Rosser(r)" apply (unfold confluent_def Church_Rosser_def square_def commute_def diamond_def, auto) apply (frule fieldI1) apply (simp add: rtrancl_field) apply (erule rtrancl_induct, auto) apply (blast intro: rtrancl_refl) apply (blast del: rtrancl_refl intro: r_into_rtrancl rtrancl_trans)+ done lemma Church_Rosser: "Church_Rosser(r) \ confluent(r)" by (blast intro: Church_Rosser1 Church_Rosser2) end diff --git a/src/ZF/ex/Group.thy b/src/ZF/ex/Group.thy --- a/src/ZF/ex/Group.thy +++ b/src/ZF/ex/Group.thy @@ -1,1231 +1,1231 @@ (* Title: ZF/ex/Group.thy *) section \Groups\ theory Group imports ZF begin text\Based on work by Clemens Ballarin, Florian Kammueller, L C Paulson and Markus Wenzel.\ subsection \Monoids\ (*First, we must simulate a record declaration: record monoid = carrier :: i mult :: "[i,i] \ i" (infixl "\\" 70) one :: i ("\\") *) definition carrier :: "i \ i" where "carrier(M) \ fst(M)" definition mmult :: "[i, i, i] \ i" (infixl \\\\ 70) where "mmult(M,x,y) \ fst(snd(M)) ` \x,y\" definition one :: "i \ i" (\\\\) where "one(M) \ fst(snd(snd(M)))" definition update_carrier :: "[i,i] \ i" where "update_carrier(M,A) \ " definition m_inv :: "i \ i \ i" (\inv\ _\ [81] 80) where "inv\<^bsub>G\<^esub> x \ (THE y. y \ carrier(G) \ y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub> \ x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub>)" locale monoid = fixes G (structure) assumes m_closed [intro, simp]: "\x \ carrier(G); y \ carrier(G)\ \ x \ y \ carrier(G)" and m_assoc: "\x \ carrier(G); y \ carrier(G); z \ carrier(G)\ \ (x \ y) \ z = x \ (y \ z)" and one_closed [intro, simp]: "\ \ carrier(G)" and l_one [simp]: "x \ carrier(G) \ \ \ x = x" and r_one [simp]: "x \ carrier(G) \ x \ \ = x" text\Simulating the record\ lemma carrier_eq [simp]: "carrier(\A,Z\) = A" by (simp add: carrier_def) lemma mult_eq [simp]: "mmult(, x, y) = M ` \x,y\" by (simp add: mmult_def) lemma one_eq [simp]: "one() = I" by (simp add: one_def) lemma update_carrier_eq [simp]: "update_carrier(\A,Z\,B) = \B,Z\" by (simp add: update_carrier_def) lemma carrier_update_carrier [simp]: "carrier(update_carrier(M,B)) = B" by (simp add: update_carrier_def) lemma mult_update_carrier [simp]: "mmult(update_carrier(M,B),x,y) = mmult(M,x,y)" by (simp add: update_carrier_def mmult_def) lemma one_update_carrier [simp]: "one(update_carrier(M,B)) = one(M)" by (simp add: update_carrier_def one_def) lemma (in monoid) inv_unique: assumes eq: "y \ x = \" "x \ y' = \" and G: "x \ carrier(G)" "y \ carrier(G)" "y' \ carrier(G)" shows "y = y'" proof - from G eq have "y = y \ (x \ y')" by simp also from G have "... = (y \ x) \ y'" by (simp add: m_assoc) also from G eq have "... = y'" by simp finally show ?thesis . qed text \ A group is a monoid all of whose elements are invertible. \ locale group = monoid + assumes inv_ex: "\x. x \ carrier(G) \ \y \ carrier(G). y \ x = \ \ x \ y = \" lemma (in group) is_group [simp]: "group(G)" by (rule group_axioms) theorem groupI: fixes G (structure) assumes m_closed [simp]: "\x y. \x \ carrier(G); y \ carrier(G)\ \ x \ y \ carrier(G)" and one_closed [simp]: "\ \ carrier(G)" and m_assoc: "\x y z. \x \ carrier(G); y \ carrier(G); z \ carrier(G)\ \ (x \ y) \ z = x \ (y \ z)" and l_one [simp]: "\x. x \ carrier(G) \ \ \ x = x" and l_inv_ex: "\x. x \ carrier(G) \ \y \ carrier(G). y \ x = \" shows "group(G)" proof - have l_cancel [simp]: "\x y z. \x \ carrier(G); y \ carrier(G); z \ carrier(G)\ \ (x \ y = x \ z) \ (y = z)" proof fix x y z assume G: "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" { assume eq: "x \ y = x \ z" with G l_inv_ex obtain x_inv where xG: "x_inv \ carrier(G)" and l_inv: "x_inv \ x = \" by fast from G eq xG have "(x_inv \ x) \ y = (x_inv \ x) \ z" by (simp add: m_assoc) with G show "y = z" by (simp add: l_inv) next assume eq: "y = z" with G show "x \ y = x \ z" by simp } qed have r_one: "\x. x \ carrier(G) \ x \ \ = x" proof - fix x assume x: "x \ carrier(G)" with l_inv_ex obtain x_inv where xG: "x_inv \ carrier(G)" and l_inv: "x_inv \ x = \" by fast from x xG have "x_inv \ (x \ \) = x_inv \ x" by (simp add: m_assoc [symmetric] l_inv) with x xG show "x \ \ = x" by simp qed have inv_ex: "\x. x \ carrier(G) \ \y \ carrier(G). y \ x = \ \ x \ y = \" proof - fix x assume x: "x \ carrier(G)" with l_inv_ex obtain y where y: "y \ carrier(G)" and l_inv: "y \ x = \" by fast from x y have "y \ (x \ y) = y \ \" by (simp add: m_assoc [symmetric] l_inv r_one) with x y have r_inv: "x \ y = \" by simp from x y show "\y \ carrier(G). y \ x = \ \ x \ y = \" by (fast intro: l_inv r_inv) qed show ?thesis by (blast intro: group.intro monoid.intro group_axioms.intro assms r_one inv_ex) qed lemma (in group) inv [simp]: "x \ carrier(G) \ inv x \ carrier(G) \ inv x \ x = \ \ x \ inv x = \" apply (frule inv_ex) - apply (unfold Bex_def m_inv_def) + unfolding Bex_def m_inv_def apply (erule exE) apply (rule theI) apply (rule ex1I, assumption) apply (blast intro: inv_unique) done lemma (in group) inv_closed [intro!]: "x \ carrier(G) \ inv x \ carrier(G)" by simp lemma (in group) l_inv: "x \ carrier(G) \ inv x \ x = \" by simp lemma (in group) r_inv: "x \ carrier(G) \ x \ inv x = \" by simp subsection \Cancellation Laws and Basic Properties\ lemma (in group) l_cancel [simp]: assumes "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" shows "(x \ y = x \ z) \ (y = z)" proof assume eq: "x \ y = x \ z" hence "(inv x \ x) \ y = (inv x \ x) \ z" by (simp only: m_assoc inv_closed assms) thus "y = z" by (simp add: assms) next assume eq: "y = z" then show "x \ y = x \ z" by simp qed lemma (in group) r_cancel [simp]: assumes "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" shows "(y \ x = z \ x) \ (y = z)" proof assume eq: "y \ x = z \ x" then have "y \ (x \ inv x) = z \ (x \ inv x)" by (simp only: m_assoc [symmetric] inv_closed assms) thus "y = z" by (simp add: assms) next assume eq: "y = z" thus "y \ x = z \ x" by simp qed lemma (in group) inv_comm: assumes "x \ y = \" and G: "x \ carrier(G)" "y \ carrier(G)" shows "y \ x = \" proof - from G have "x \ y \ x = x \ \" by (auto simp add: assms) with G show ?thesis by (simp del: r_one add: m_assoc) qed lemma (in group) inv_equality: "\y \ x = \; x \ carrier(G); y \ carrier(G)\ \ inv x = y" apply (simp add: m_inv_def) apply (rule the_equality) apply (simp add: inv_comm [of y x]) apply (rule r_cancel [THEN iffD1], auto) done lemma (in group) inv_one [simp]: "inv \ = \" by (auto intro: inv_equality) lemma (in group) inv_inv [simp]: "x \ carrier(G) \ inv (inv x) = x" by (auto intro: inv_equality) text\This proof is by cancellation\ lemma (in group) inv_mult_group: "\x \ carrier(G); y \ carrier(G)\ \ inv (x \ y) = inv y \ inv x" proof - assume G: "x \ carrier(G)" "y \ carrier(G)" then have "inv (x \ y) \ (x \ y) = (inv y \ inv x) \ (x \ y)" by (simp add: m_assoc l_inv) (simp add: m_assoc [symmetric] l_inv) with G show ?thesis by (simp_all del: inv add: inv_closed) qed subsection \Substructures\ locale subgroup = fixes H and G (structure) assumes subset: "H \ carrier(G)" and m_closed [intro, simp]: "\x \ H; y \ H\ \ x \ y \ H" and one_closed [simp]: "\ \ H" and m_inv_closed [intro,simp]: "x \ H \ inv x \ H" lemma (in subgroup) mem_carrier [simp]: "x \ H \ x \ carrier(G)" using subset by blast lemma subgroup_imp_subset: "subgroup(H,G) \ H \ carrier(G)" by (rule subgroup.subset) lemma (in subgroup) group_axiomsI [intro]: assumes "group(G)" shows "group_axioms (update_carrier(G,H))" proof - interpret group G by fact show ?thesis by (force intro: group_axioms.intro l_inv r_inv) qed lemma (in subgroup) is_group [intro]: assumes "group(G)" shows "group (update_carrier(G,H))" proof - interpret group G by fact show ?thesis by (rule groupI) (auto intro: m_assoc l_inv mem_carrier) qed text \ Since \<^term>\H\ is nonempty, it contains some element \<^term>\x\. Since it is closed under inverse, it contains \inv x\. Since it is closed under product, it contains \x \ inv x = \\. \ text \ Since \<^term>\H\ is nonempty, it contains some element \<^term>\x\. Since it is closed under inverse, it contains \inv x\. Since it is closed under product, it contains \x \ inv x = \\. \ lemma (in group) one_in_subset: "\H \ carrier(G); H \ 0; \a \ H. inv a \ H; \a\H. \b\H. a \ b \ H\ \ \ \ H" by (force simp add: l_inv) text \A characterization of subgroups: closed, non-empty subset.\ declare monoid.one_closed [simp] group.inv_closed [simp] monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp] lemma subgroup_nonempty: "\ subgroup(0,G)" by (blast dest: subgroup.one_closed) subsection \Direct Products\ definition DirProdGroup :: "[i,i] \ i" (infixr \\\ 80) where "G \ H \ carrier(H), (\<\g,h\, > \ (carrier(G) \ carrier(H)) \ (carrier(G) \ carrier(H)). \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>), <\\<^bsub>G\<^esub>, \\<^bsub>H\<^esub>>, 0>" lemma DirProdGroup_group: assumes "group(G)" and "group(H)" shows "group (G \ H)" proof - interpret G: group G by fact interpret H: group H by fact show ?thesis by (force intro!: groupI G.m_assoc H.m_assoc G.l_inv H.l_inv simp add: DirProdGroup_def) qed lemma carrier_DirProdGroup [simp]: "carrier (G \ H) = carrier(G) \ carrier(H)" by (simp add: DirProdGroup_def) lemma one_DirProdGroup [simp]: "\\<^bsub>G \ H\<^esub> = <\\<^bsub>G\<^esub>, \\<^bsub>H\<^esub>>" by (simp add: DirProdGroup_def) lemma mult_DirProdGroup [simp]: "\g \ carrier(G); h \ carrier(H); g' \ carrier(G); h' \ carrier(H)\ \ \g, h\ \\<^bsub>G \ H\<^esub> = \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>" by (simp add: DirProdGroup_def) lemma inv_DirProdGroup [simp]: assumes "group(G)" and "group(H)" assumes g: "g \ carrier(G)" and h: "h \ carrier(H)" shows "inv \<^bsub>G \ H\<^esub> \g, h\ = G\<^esub> g, inv\<^bsub>H\<^esub> h>" apply (rule group.inv_equality [OF DirProdGroup_group]) apply (simp_all add: assms group.l_inv) done subsection \Isomorphisms\ definition hom :: "[i,i] \ i" where "hom(G,H) \ {h \ carrier(G) -> carrier(H). (\x \ carrier(G). \y \ carrier(G). h ` (x \\<^bsub>G\<^esub> y) = (h ` x) \\<^bsub>H\<^esub> (h ` y))}" lemma hom_mult: "\h \ hom(G,H); x \ carrier(G); y \ carrier(G)\ \ h ` (x \\<^bsub>G\<^esub> y) = h ` x \\<^bsub>H\<^esub> h ` y" by (simp add: hom_def) lemma hom_closed: "\h \ hom(G,H); x \ carrier(G)\ \ h ` x \ carrier(H)" by (auto simp add: hom_def) lemma (in group) hom_compose: "\h \ hom(G,H); i \ hom(H,I)\ \ i O h \ hom(G,I)" by (force simp add: hom_def comp_fun) lemma hom_is_fun: "h \ hom(G,H) \ h \ carrier(G) -> carrier(H)" by (simp add: hom_def) subsection \Isomorphisms\ definition iso :: "[i,i] \ i" (infixr \\\ 60) where "G \ H \ hom(G,H) \ bij(carrier(G), carrier(H))" lemma (in group) iso_refl: "id(carrier(G)) \ G \ G" by (simp add: iso_def hom_def id_type id_bij) lemma (in group) iso_sym: "h \ G \ H \ converse(h) \ H \ G" apply (simp add: iso_def bij_converse_bij, clarify) apply (subgoal_tac "converse(h) \ carrier(H) \ carrier(G)") prefer 2 apply (simp add: bij_converse_bij bij_is_fun) apply (auto intro: left_inverse_eq [of _ "carrier(G)" "carrier(H)"] simp add: hom_def bij_is_inj right_inverse_bij) done lemma (in group) iso_trans: "\h \ G \ H; i \ H \ I\ \ i O h \ G \ I" by (auto simp add: iso_def hom_compose comp_bij) lemma DirProdGroup_commute_iso: assumes "group(G)" and "group(H)" shows "(\\x,y\ \ carrier(G \ H). \y,x\) \ (G \ H) \ (H \ G)" proof - interpret group G by fact interpret group H by fact show ?thesis by (auto simp add: iso_def hom_def inj_def surj_def bij_def) qed lemma DirProdGroup_assoc_iso: assumes "group(G)" and "group(H)" and "group(I)" shows "(\<\x,y\,z> \ carrier((G \ H) \ I). y,z\>) \ ((G \ H) \ I) \ (G \ (H \ I))" proof - interpret group G by fact interpret group H by fact interpret group I by fact show ?thesis by (auto intro: lam_type simp add: iso_def hom_def inj_def surj_def bij_def) qed text\Basis for homomorphism proofs: we assume two groups \<^term>\G\ and \<^term>\H\, with a homomorphism \<^term>\h\ between them\ locale group_hom = G: group G + H: group H for G (structure) and H (structure) and h + assumes homh: "h \ hom(G,H)" notes hom_mult [simp] = hom_mult [OF homh] and hom_closed [simp] = hom_closed [OF homh] and hom_is_fun [simp] = hom_is_fun [OF homh] lemma (in group_hom) one_closed [simp]: "h ` \ \ carrier(H)" by simp lemma (in group_hom) hom_one [simp]: "h ` \ = \\<^bsub>H\<^esub>" proof - have "h ` \ \\<^bsub>H\<^esub> \\<^bsub>H\<^esub> = (h ` \) \\<^bsub>H\<^esub> (h ` \)" by (simp add: hom_mult [symmetric] del: hom_mult) then show ?thesis by (simp del: H.r_one) qed lemma (in group_hom) inv_closed [simp]: "x \ carrier(G) \ h ` (inv x) \ carrier(H)" by simp lemma (in group_hom) hom_inv [simp]: "x \ carrier(G) \ h ` (inv x) = inv\<^bsub>H\<^esub> (h ` x)" proof - assume x: "x \ carrier(G)" then have "h ` x \\<^bsub>H\<^esub> h ` (inv x) = \\<^bsub>H\<^esub>" by (simp add: hom_mult [symmetric] G.r_inv del: hom_mult) also from x have "... = h ` x \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h ` x)" by (simp add: hom_mult [symmetric] H.r_inv del: hom_mult) finally have "h ` x \\<^bsub>H\<^esub> h ` (inv x) = h ` x \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h ` x)" . with x show ?thesis by (simp del: H.inv) qed subsection \Commutative Structures\ text \ Naming convention: multiplicative structures that are commutative are called \emph{commutative}, additive structures are called \emph{Abelian}. \ subsection \Definition\ locale comm_monoid = monoid + assumes m_comm: "\x \ carrier(G); y \ carrier(G)\ \ x \ y = y \ x" lemma (in comm_monoid) m_lcomm: "\x \ carrier(G); y \ carrier(G); z \ carrier(G)\ \ x \ (y \ z) = y \ (x \ z)" proof - assume xyz: "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" from xyz have "x \ (y \ z) = (x \ y) \ z" by (simp add: m_assoc) also from xyz have "... = (y \ x) \ z" by (simp add: m_comm) also from xyz have "... = y \ (x \ z)" by (simp add: m_assoc) finally show ?thesis . qed lemmas (in comm_monoid) m_ac = m_assoc m_comm m_lcomm locale comm_group = comm_monoid + group lemma (in comm_group) inv_mult: "\x \ carrier(G); y \ carrier(G)\ \ inv (x \ y) = inv x \ inv y" by (simp add: m_ac inv_mult_group) lemma (in group) subgroup_self: "subgroup (carrier(G),G)" by (simp add: subgroup_def) lemma (in group) subgroup_imp_group: "subgroup(H,G) \ group (update_carrier(G,H))" by (simp add: subgroup.is_group) lemma (in group) subgroupI: assumes subset: "H \ carrier(G)" and non_empty: "H \ 0" and "\a. a \ H \ inv a \ H" and "\a b. \a \ H; b \ H\ \ a \ b \ H" shows "subgroup(H,G)" proof (simp add: subgroup_def assms) show "\ \ H" by (rule one_in_subset) (auto simp only: assms) qed subsection \Bijections of a Set, Permutation Groups, Automorphism Groups\ definition BijGroup :: "i\i" where "BijGroup(S) \ \g,f\ \ bij(S,S) \ bij(S,S). g O f, id(S), 0>" subsection \Bijections Form a Group\ theorem group_BijGroup: "group(BijGroup(S))" apply (simp add: BijGroup_def) apply (rule groupI) apply (simp_all add: id_bij comp_bij comp_assoc) apply (simp add: id_bij bij_is_fun left_comp_id [of _ S S] fun_is_rel) apply (blast intro: left_comp_inverse bij_is_inj bij_converse_bij) done subsection\Automorphisms Form a Group\ lemma Bij_Inv_mem: "\f \ bij(S,S); x \ S\ \ converse(f) ` x \ S" by (blast intro: apply_funtype bij_is_fun bij_converse_bij) lemma inv_BijGroup: "f \ bij(S,S) \ m_inv (BijGroup(S), f) = converse(f)" apply (rule group.inv_equality) apply (rule group_BijGroup) apply (simp_all add: BijGroup_def bij_converse_bij left_comp_inverse [OF bij_is_inj]) done lemma iso_is_bij: "h \ G \ H \ h \ bij(carrier(G), carrier(H))" by (simp add: iso_def) definition auto :: "i\i" where "auto(G) \ iso(G,G)" definition AutoGroup :: "i\i" where "AutoGroup(G) \ update_carrier(BijGroup(carrier(G)), auto(G))" lemma (in group) id_in_auto: "id(carrier(G)) \ auto(G)" by (simp add: iso_refl auto_def) lemma (in group) subgroup_auto: "subgroup (auto(G)) (BijGroup (carrier(G)))" proof (rule subgroup.intro) show "auto(G) \ carrier (BijGroup (carrier(G)))" by (auto simp add: auto_def BijGroup_def iso_def) next fix x y assume "x \ auto(G)" "y \ auto(G)" thus "x \\<^bsub>BijGroup (carrier(G))\<^esub> y \ auto(G)" by (auto simp add: BijGroup_def auto_def iso_def bij_is_fun group.hom_compose comp_bij) next show "\\<^bsub>BijGroup (carrier(G))\<^esub> \ auto(G)" by (simp add: BijGroup_def id_in_auto) next fix x assume "x \ auto(G)" thus "inv\<^bsub>BijGroup (carrier(G))\<^esub> x \ auto(G)" by (simp add: auto_def inv_BijGroup iso_is_bij iso_sym) qed theorem (in group) AutoGroup: "group (AutoGroup(G))" by (simp add: AutoGroup_def subgroup.is_group subgroup_auto group_BijGroup) subsection\Cosets and Quotient Groups\ definition r_coset :: "[i,i,i] \ i" (infixl \#>\\ 60) where "H #>\<^bsub>G\<^esub> a \ \h\H. {h \\<^bsub>G\<^esub> a}" definition l_coset :: "[i,i,i] \ i" (infixl \<#\\ 60) where "a <#\<^bsub>G\<^esub> H \ \h\H. {a \\<^bsub>G\<^esub> h}" definition RCOSETS :: "[i,i] \ i" (\rcosets\ _\ [81] 80) where "rcosets\<^bsub>G\<^esub> H \ \a\carrier(G). {H #>\<^bsub>G\<^esub> a}" definition set_mult :: "[i,i,i] \ i" (infixl \<#>\\ 60) where "H <#>\<^bsub>G\<^esub> K \ \h\H. \k\K. {h \\<^bsub>G\<^esub> k}" definition SET_INV :: "[i,i] \ i" (\set'_inv\ _\ [81] 80) where "set_inv\<^bsub>G\<^esub> H \ \h\H. {inv\<^bsub>G\<^esub> h}" locale normal = subgroup + group + assumes coset_eq: "(\x \ carrier(G). H #> x = x <# H)" notation normal (infixl \\\ 60) subsection \Basic Properties of Cosets\ lemma (in group) coset_mult_assoc: "\M \ carrier(G); g \ carrier(G); h \ carrier(G)\ \ (M #> g) #> h = M #> (g \ h)" by (force simp add: r_coset_def m_assoc) lemma (in group) coset_mult_one [simp]: "M \ carrier(G) \ M #> \ = M" by (force simp add: r_coset_def) lemma (in group) solve_equation: "\subgroup(H,G); x \ H; y \ H\ \ \h\H. y = h \ x" apply (rule bexI [of _ "y \ (inv x)"]) apply (auto simp add: subgroup.m_closed subgroup.m_inv_closed m_assoc subgroup.subset [THEN subsetD]) done lemma (in group) repr_independence: "\y \ H #> x; x \ carrier(G); subgroup(H,G)\ \ H #> x = H #> y" by (auto simp add: r_coset_def m_assoc [symmetric] subgroup.subset [THEN subsetD] subgroup.m_closed solve_equation) lemma (in group) coset_join2: "\x \ carrier(G); subgroup(H,G); x\H\ \ H #> x = H" \ \Alternative proof is to put \<^term>\x=\\ in \repr_independence\.\ by (force simp add: subgroup.m_closed r_coset_def solve_equation) lemma (in group) r_coset_subset_G: "\H \ carrier(G); x \ carrier(G)\ \ H #> x \ carrier(G)" by (auto simp add: r_coset_def) lemma (in group) rcosI: "\h \ H; H \ carrier(G); x \ carrier(G)\ \ h \ x \ H #> x" by (auto simp add: r_coset_def) lemma (in group) rcosetsI: "\H \ carrier(G); x \ carrier(G)\ \ H #> x \ rcosets H" by (auto simp add: RCOSETS_def) text\Really needed?\ lemma (in group) transpose_inv: "\x \ y = z; x \ carrier(G); y \ carrier(G); z \ carrier(G)\ \ (inv x) \ z = y" by (force simp add: m_assoc [symmetric]) subsection \Normal subgroups\ lemma normal_imp_subgroup: "H \ G \ subgroup(H,G)" by (simp add: normal_def subgroup_def) lemma (in group) normalI: "subgroup(H,G) \ (\x \ carrier(G). H #> x = x <# H) \ H \ G" by (simp add: normal_def normal_axioms_def) lemma (in normal) inv_op_closed1: "\x \ carrier(G); h \ H\ \ (inv x) \ h \ x \ H" apply (insert coset_eq) apply (auto simp add: l_coset_def r_coset_def) apply (drule bspec, assumption) apply (drule equalityD1 [THEN subsetD], blast, clarify) apply (simp add: m_assoc) apply (simp add: m_assoc [symmetric]) done lemma (in normal) inv_op_closed2: "\x \ carrier(G); h \ H\ \ x \ h \ (inv x) \ H" apply (subgoal_tac "inv (inv x) \ h \ (inv x) \ H") apply simp apply (blast intro: inv_op_closed1) done text\Alternative characterization of normal subgroups\ lemma (in group) normal_inv_iff: "(N \ G) \ (subgroup(N,G) \ (\x \ carrier(G). \h \ N. x \ h \ (inv x) \ N))" (is "_ \ ?rhs") proof assume N: "N \ G" show ?rhs by (blast intro: N normal.inv_op_closed2 normal_imp_subgroup) next assume ?rhs hence sg: "subgroup(N,G)" and closed: "\x. x\carrier(G) \ \h\N. x \ h \ inv x \ N" by auto hence sb: "N \ carrier(G)" by (simp add: subgroup.subset) show "N \ G" proof (intro normalI [OF sg], simp add: l_coset_def r_coset_def, clarify) fix x assume x: "x \ carrier(G)" show "(\h\N. {h \ x}) = (\h\N. {x \ h})" proof show "(\h\N. {h \ x}) \ (\h\N. {x \ h})" proof clarify fix n assume n: "n \ N" show "n \ x \ (\h\N. {x \ h})" proof (rule UN_I) from closed [of "inv x"] show "inv x \ n \ x \ N" by (simp add: x n) show "n \ x \ {x \ (inv x \ n \ x)}" by (simp add: x n m_assoc [symmetric] sb [THEN subsetD]) qed qed next show "(\h\N. {x \ h}) \ (\h\N. {h \ x})" proof clarify fix n assume n: "n \ N" show "x \ n \ (\h\N. {h \ x})" proof (rule UN_I) show "x \ n \ inv x \ N" by (simp add: x n closed) show "x \ n \ {x \ n \ inv x \ x}" by (simp add: x n m_assoc sb [THEN subsetD]) qed qed qed qed qed subsection\More Properties of Cosets\ lemma (in group) l_coset_subset_G: "\H \ carrier(G); x \ carrier(G)\ \ x <# H \ carrier(G)" by (auto simp add: l_coset_def subsetD) lemma (in group) l_coset_swap: "\y \ x <# H; x \ carrier(G); subgroup(H,G)\ \ x \ y <# H" proof (simp add: l_coset_def) assume "\h\H. y = x \ h" and x: "x \ carrier(G)" and sb: "subgroup(H,G)" then obtain h' where h': "h' \ H \ x \ h' = y" by blast show "\h\H. x = y \ h" proof show "x = y \ inv h'" using h' x sb by (auto simp add: m_assoc subgroup.subset [THEN subsetD]) show "inv h' \ H" using h' sb by (auto simp add: subgroup.subset [THEN subsetD] subgroup.m_inv_closed) qed qed lemma (in group) l_coset_carrier: "\y \ x <# H; x \ carrier(G); subgroup(H,G)\ \ y \ carrier(G)" by (auto simp add: l_coset_def m_assoc subgroup.subset [THEN subsetD] subgroup.m_closed) lemma (in group) l_repr_imp_subset: assumes y: "y \ x <# H" and x: "x \ carrier(G)" and sb: "subgroup(H,G)" shows "y <# H \ x <# H" proof - from y obtain h' where "h' \ H" "x \ h' = y" by (auto simp add: l_coset_def) thus ?thesis using x sb by (auto simp add: l_coset_def m_assoc subgroup.subset [THEN subsetD] subgroup.m_closed) qed lemma (in group) l_repr_independence: assumes y: "y \ x <# H" and x: "x \ carrier(G)" and sb: "subgroup(H,G)" shows "x <# H = y <# H" proof show "x <# H \ y <# H" by (rule l_repr_imp_subset, (blast intro: l_coset_swap l_coset_carrier y x sb)+) show "y <# H \ x <# H" by (rule l_repr_imp_subset [OF y x sb]) qed lemma (in group) setmult_subset_G: "\H \ carrier(G); K \ carrier(G)\ \ H <#> K \ carrier(G)" by (auto simp add: set_mult_def subsetD) lemma (in group) subgroup_mult_id: "subgroup(H,G) \ H <#> H = H" apply (rule equalityI) apply (auto simp add: subgroup.m_closed set_mult_def Sigma_def image_def) apply (rule_tac x = x in bexI) apply (rule bexI [of _ "\"]) apply (auto simp add: subgroup.one_closed subgroup.subset [THEN subsetD]) done subsubsection \Set of inverses of an \r_coset\.\ lemma (in normal) rcos_inv: assumes x: "x \ carrier(G)" shows "set_inv (H #> x) = H #> (inv x)" proof (simp add: r_coset_def SET_INV_def x inv_mult_group, safe intro!: equalityI) fix h assume h: "h \ H" { show "inv x \ inv h \ (\j\H. {j \ inv x})" proof (rule UN_I) show "inv x \ inv h \ x \ H" by (simp add: inv_op_closed1 h x) show "inv x \ inv h \ {inv x \ inv h \ x \ inv x}" by (simp add: h x m_assoc) qed next show "h \ inv x \ (\j\H. {inv x \ inv j})" proof (rule UN_I) show "x \ inv h \ inv x \ H" by (simp add: inv_op_closed2 h x) show "h \ inv x \ {inv x \ inv (x \ inv h \ inv x)}" by (simp add: h x m_assoc [symmetric] inv_mult_group) qed } qed subsubsection \Theorems for \<#>\ with \#>\ or \<#\.\ lemma (in group) setmult_rcos_assoc: "\H \ carrier(G); K \ carrier(G); x \ carrier(G)\ \ H <#> (K #> x) = (H <#> K) #> x" by (force simp add: r_coset_def set_mult_def m_assoc) lemma (in group) rcos_assoc_lcos: "\H \ carrier(G); K \ carrier(G); x \ carrier(G)\ \ (H #> x) <#> K = H <#> (x <# K)" by (force simp add: r_coset_def l_coset_def set_mult_def m_assoc) lemma (in normal) rcos_mult_step1: "\x \ carrier(G); y \ carrier(G)\ \ (H #> x) <#> (H #> y) = (H <#> (x <# H)) #> y" by (simp add: setmult_rcos_assoc subset r_coset_subset_G l_coset_subset_G rcos_assoc_lcos) lemma (in normal) rcos_mult_step2: "\x \ carrier(G); y \ carrier(G)\ \ (H <#> (x <# H)) #> y = (H <#> (H #> x)) #> y" by (insert coset_eq, simp add: normal_def) lemma (in normal) rcos_mult_step3: "\x \ carrier(G); y \ carrier(G)\ \ (H <#> (H #> x)) #> y = H #> (x \ y)" by (simp add: setmult_rcos_assoc coset_mult_assoc subgroup_mult_id subset normal_axioms normal.axioms) lemma (in normal) rcos_sum: "\x \ carrier(G); y \ carrier(G)\ \ (H #> x) <#> (H #> y) = H #> (x \ y)" by (simp add: rcos_mult_step1 rcos_mult_step2 rcos_mult_step3) lemma (in normal) rcosets_mult_eq: "M \ rcosets H \ H <#> M = M" \ \generalizes \subgroup_mult_id\\ by (auto simp add: RCOSETS_def subset setmult_rcos_assoc subgroup_mult_id normal_axioms normal.axioms) subsubsection\Two distinct right cosets are disjoint\ definition r_congruent :: "[i,i] \ i" (\rcong\ _\ [60] 60) where "rcong\<^bsub>G\<^esub> H \ {\x,y\ \ carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \\<^bsub>G\<^esub> y \ H}" lemma (in subgroup) equiv_rcong: assumes "group(G)" shows "equiv (carrier(G), rcong H)" proof - interpret group G by fact show ?thesis proof (simp add: equiv_def, intro conjI) show "rcong H \ carrier(G) \ carrier(G)" by (auto simp add: r_congruent_def) next show "refl (carrier(G), rcong H)" by (auto simp add: r_congruent_def refl_def) next show "sym (rcong H)" proof (simp add: r_congruent_def sym_def, clarify) fix x y assume [simp]: "x \ carrier(G)" "y \ carrier(G)" and "inv x \ y \ H" hence "inv (inv x \ y) \ H" by simp thus "inv y \ x \ H" by (simp add: inv_mult_group) qed next show "trans (rcong H)" proof (simp add: r_congruent_def trans_def, clarify) fix x y z assume [simp]: "x \ carrier(G)" "y \ carrier(G)" "z \ carrier(G)" and "inv x \ y \ H" and "inv y \ z \ H" hence "(inv x \ y) \ (inv y \ z) \ H" by simp hence "inv x \ (y \ inv y) \ z \ H" by (simp add: m_assoc del: inv) thus "inv x \ z \ H" by simp qed qed qed text\Equivalence classes of \rcong\ correspond to left cosets. Was there a mistake in the definitions? I'd have expected them to correspond to right cosets.\ lemma (in subgroup) l_coset_eq_rcong: assumes "group(G)" assumes a: "a \ carrier(G)" shows "a <# H = (rcong H) `` {a}" proof - interpret group G by fact show ?thesis by (force simp add: r_congruent_def l_coset_def m_assoc [symmetric] a Collect_image_eq) qed lemma (in group) rcos_equation: assumes "subgroup(H, G)" shows "\ha \ a = h \ b; a \ carrier(G); b \ carrier(G); h \ H; ha \ H; hb \ H\ \ hb \ a \ (\h\H. {h \ b})" (is "PROP ?P") proof - interpret subgroup H G by fact show "PROP ?P" apply (rule UN_I [of "hb \ ((inv ha) \ h)"], simp) apply (simp add: m_assoc transpose_inv) done qed lemma (in group) rcos_disjoint: assumes "subgroup(H, G)" shows "\a \ rcosets H; b \ rcosets H; a\b\ \ a \ b = 0" (is "PROP ?P") proof - interpret subgroup H G by fact show "PROP ?P" apply (simp add: RCOSETS_def r_coset_def) apply (blast intro: rcos_equation assms sym) done qed subsection \Order of a Group and Lagrange's Theorem\ definition order :: "i \ i" where "order(S) \ |carrier(S)|" lemma (in group) rcos_self: assumes "subgroup(H, G)" shows "x \ carrier(G) \ x \ H #> x" (is "PROP ?P") proof - interpret subgroup H G by fact show "PROP ?P" apply (simp add: r_coset_def) apply (rule_tac x="\" in bexI) apply (auto) done qed lemma (in group) rcosets_part_G: assumes "subgroup(H, G)" shows "\(rcosets H) = carrier(G)" proof - interpret subgroup H G by fact show ?thesis apply (rule equalityI) apply (force simp add: RCOSETS_def r_coset_def) apply (auto simp add: RCOSETS_def intro: rcos_self assms) done qed lemma (in group) cosets_finite: "\c \ rcosets H; H \ carrier(G); Finite (carrier(G))\ \ Finite(c)" apply (auto simp add: RCOSETS_def) apply (simp add: r_coset_subset_G [THEN subset_Finite]) done text\More general than the HOL version, which also requires \<^term>\G\ to be finite.\ lemma (in group) card_cosets_equal: assumes H: "H \ carrier(G)" shows "c \ rcosets H \ |c| = |H|" proof (simp add: RCOSETS_def, clarify) fix a assume a: "a \ carrier(G)" show "|H #> a| = |H|" proof (rule eqpollI [THEN cardinal_cong]) show "H #> a \ H" proof (simp add: lepoll_def, intro exI) show "(\y \ H#>a. y \ inv a) \ inj(H #> a, H)" by (auto intro: lam_type simp add: inj_def r_coset_def m_assoc subsetD [OF H] a) qed show "H \ H #> a" proof (simp add: lepoll_def, intro exI) show "(\y\ H. y \ a) \ inj(H, H #> a)" by (auto intro: lam_type simp add: inj_def r_coset_def subsetD [OF H] a) qed qed qed lemma (in group) rcosets_subset_PowG: "subgroup(H,G) \ rcosets H \ Pow(carrier(G))" apply (simp add: RCOSETS_def) apply (blast dest: r_coset_subset_G subgroup.subset) done theorem (in group) lagrange: "\Finite(carrier(G)); subgroup(H,G)\ \ |rcosets H| #* |H| = order(G)" apply (simp (no_asm_simp) add: order_def rcosets_part_G [symmetric]) apply (subst mult_commute) apply (rule card_partition) apply (simp add: rcosets_subset_PowG [THEN subset_Finite]) apply (simp add: rcosets_part_G) apply (simp add: card_cosets_equal [OF subgroup.subset]) apply (simp add: rcos_disjoint) done subsection \Quotient Groups: Factorization of a Group\ definition FactGroup :: "[i,i] \ i" (infixl \Mod\ 65) where \ \Actually defined for groups rather than monoids\ "G Mod H \ G\<^esub> H, \\K1,K2\ \ (rcosets\<^bsub>G\<^esub> H) \ (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>" lemma (in normal) setmult_closed: "\K1 \ rcosets H; K2 \ rcosets H\ \ K1 <#> K2 \ rcosets H" by (auto simp add: rcos_sum RCOSETS_def) lemma (in normal) setinv_closed: "K \ rcosets H \ set_inv K \ rcosets H" by (auto simp add: rcos_inv RCOSETS_def) lemma (in normal) rcosets_assoc: "\M1 \ rcosets H; M2 \ rcosets H; M3 \ rcosets H\ \ M1 <#> M2 <#> M3 = M1 <#> (M2 <#> M3)" by (auto simp add: RCOSETS_def rcos_sum m_assoc) lemma (in subgroup) subgroup_in_rcosets: assumes "group(G)" shows "H \ rcosets H" proof - interpret group G by fact have "H #> \ = H" using _ subgroup_axioms by (rule coset_join2) simp_all then show ?thesis by (auto simp add: RCOSETS_def intro: sym) qed lemma (in normal) rcosets_inv_mult_group_eq: "M \ rcosets H \ set_inv M <#> M = H" by (auto simp add: RCOSETS_def rcos_inv rcos_sum subgroup.subset normal_axioms normal.axioms) theorem (in normal) factorgroup_is_group: "group (G Mod H)" apply (simp add: FactGroup_def) apply (rule groupI) apply (simp add: setmult_closed) apply (simp add: normal_imp_subgroup subgroup_in_rcosets) apply (simp add: setmult_closed rcosets_assoc) apply (simp add: normal_imp_subgroup subgroup_in_rcosets rcosets_mult_eq) apply (auto dest: rcosets_inv_mult_group_eq simp add: setinv_closed) done lemma (in normal) inv_FactGroup: "X \ carrier (G Mod H) \ inv\<^bsub>G Mod H\<^esub> X = set_inv X" apply (rule group.inv_equality [OF factorgroup_is_group]) apply (simp_all add: FactGroup_def setinv_closed rcosets_inv_mult_group_eq) done text\The coset map is a homomorphism from \<^term>\G\ to the quotient group \<^term>\G Mod H\\ lemma (in normal) r_coset_hom_Mod: "(\a \ carrier(G). H #> a) \ hom(G, G Mod H)" by (auto simp add: FactGroup_def RCOSETS_def hom_def rcos_sum intro: lam_type) subsection\The First Isomorphism Theorem\ text\The quotient by the kernel of a homomorphism is isomorphic to the range of that homomorphism.\ definition kernel :: "[i,i,i] \ i" where \ \the kernel of a homomorphism\ "kernel(G,H,h) \ {x \ carrier(G). h ` x = \\<^bsub>H\<^esub>}" lemma (in group_hom) subgroup_kernel: "subgroup (kernel(G,H,h), G)" apply (rule subgroup.intro) apply (auto simp add: kernel_def group.intro) done text\The kernel of a homomorphism is a normal subgroup\ lemma (in group_hom) normal_kernel: "(kernel(G,H,h)) \ G" apply (simp add: group.normal_inv_iff subgroup_kernel group.intro) apply (simp add: kernel_def) done lemma (in group_hom) FactGroup_nonempty: assumes X: "X \ carrier (G Mod kernel(G,H,h))" shows "X \ 0" proof - from X obtain g where "g \ carrier(G)" and "X = kernel(G,H,h) #> g" by (auto simp add: FactGroup_def RCOSETS_def) thus ?thesis by (auto simp add: kernel_def r_coset_def image_def intro: hom_one) qed lemma (in group_hom) FactGroup_contents_mem: assumes X: "X \ carrier (G Mod (kernel(G,H,h)))" shows "contents (h``X) \ carrier(H)" proof - from X obtain g where g: "g \ carrier(G)" and "X = kernel(G,H,h) #> g" by (auto simp add: FactGroup_def RCOSETS_def) hence "h `` X = {h ` g}" by (auto simp add: kernel_def r_coset_def image_UN image_eq_UN [OF hom_is_fun] g) thus ?thesis by (auto simp add: g) qed lemma mult_FactGroup: "\X \ carrier(G Mod H); X' \ carrier(G Mod H)\ \ X \\<^bsub>(G Mod H)\<^esub> X' = X <#>\<^bsub>G\<^esub> X'" by (simp add: FactGroup_def) lemma (in normal) FactGroup_m_closed: "\X \ carrier(G Mod H); X' \ carrier(G Mod H)\ \ X <#>\<^bsub>G\<^esub> X' \ carrier(G Mod H)" by (simp add: FactGroup_def setmult_closed) lemma (in group_hom) FactGroup_hom: "(\X \ carrier(G Mod (kernel(G,H,h))). contents (h``X)) \ hom (G Mod (kernel(G,H,h)), H)" proof (simp add: hom_def FactGroup_contents_mem lam_type mult_FactGroup normal.FactGroup_m_closed [OF normal_kernel], intro ballI) fix X and X' assume X: "X \ carrier (G Mod kernel(G,H,h))" and X': "X' \ carrier (G Mod kernel(G,H,h))" then obtain g and g' where "g \ carrier(G)" and "g' \ carrier(G)" and "X = kernel(G,H,h) #> g" and "X' = kernel(G,H,h) #> g'" by (auto simp add: FactGroup_def RCOSETS_def) hence all: "\x\X. h ` x = h ` g" "\x\X'. h ` x = h ` g'" and Xsub: "X \ carrier(G)" and X'sub: "X' \ carrier(G)" by (force simp add: kernel_def r_coset_def image_def)+ hence "h `` (X <#> X') = {h ` g \\<^bsub>H\<^esub> h ` g'}" using X X' by (auto dest!: FactGroup_nonempty simp add: set_mult_def image_eq_UN [OF hom_is_fun] image_UN subsetD [OF Xsub] subsetD [OF X'sub]) thus "contents (h `` (X <#> X')) = contents (h `` X) \\<^bsub>H\<^esub> contents (h `` X')" by (simp add: all image_eq_UN [OF hom_is_fun] FactGroup_nonempty X X' Xsub X'sub) qed text\Lemma for the following injectivity result\ lemma (in group_hom) FactGroup_subset: "\g \ carrier(G); g' \ carrier(G); h ` g = h ` g'\ \ kernel(G,H,h) #> g \ kernel(G,H,h) #> g'" apply (clarsimp simp add: kernel_def r_coset_def image_def) apply (rename_tac y) apply (rule_tac x="y \ g \ inv g'" in bexI) apply (simp_all add: G.m_assoc) done lemma (in group_hom) FactGroup_inj: "(\X\carrier (G Mod kernel(G,H,h)). contents (h `` X)) \ inj(carrier (G Mod kernel(G,H,h)), carrier(H))" proof (simp add: inj_def FactGroup_contents_mem lam_type, clarify) fix X and X' assume X: "X \ carrier (G Mod kernel(G,H,h))" and X': "X' \ carrier (G Mod kernel(G,H,h))" then obtain g and g' where gX: "g \ carrier(G)" "g' \ carrier(G)" "X = kernel(G,H,h) #> g" "X' = kernel(G,H,h) #> g'" by (auto simp add: FactGroup_def RCOSETS_def) hence all: "\x\X. h ` x = h ` g" "\x\X'. h ` x = h ` g'" and Xsub: "X \ carrier(G)" and X'sub: "X' \ carrier(G)" by (force simp add: kernel_def r_coset_def image_def)+ assume "contents (h `` X) = contents (h `` X')" hence h: "h ` g = h ` g'" by (simp add: all image_eq_UN [OF hom_is_fun] FactGroup_nonempty X X' Xsub X'sub) show "X=X'" by (rule equalityI) (simp_all add: FactGroup_subset h gX) qed lemma (in group_hom) kernel_rcoset_subset: assumes g: "g \ carrier(G)" shows "kernel(G,H,h) #> g \ carrier (G)" by (auto simp add: g kernel_def r_coset_def) text\If the homomorphism \<^term>\h\ is onto \<^term>\H\, then so is the homomorphism from the quotient group\ lemma (in group_hom) FactGroup_surj: assumes h: "h \ surj(carrier(G), carrier(H))" shows "(\X\carrier (G Mod kernel(G,H,h)). contents (h `` X)) \ surj(carrier (G Mod kernel(G,H,h)), carrier(H))" proof (simp add: surj_def FactGroup_contents_mem lam_type, clarify) fix y assume y: "y \ carrier(H)" with h obtain g where g: "g \ carrier(G)" "h ` g = y" by (auto simp add: surj_def) hence "(\x\kernel(G,H,h) #> g. {h ` x}) = {y}" by (auto simp add: y kernel_def r_coset_def) with g show "\x\carrier(G Mod kernel(G, H, h)). contents(h `` x) = y" \ \The witness is \<^term>\kernel(G,H,h) #> g\\ by (force simp add: FactGroup_def RCOSETS_def image_eq_UN [OF hom_is_fun] kernel_rcoset_subset) qed text\If \<^term>\h\ is a homomorphism from \<^term>\G\ onto \<^term>\H\, then the quotient group \<^term>\G Mod (kernel(G,H,h))\ is isomorphic to \<^term>\H\.\ theorem (in group_hom) FactGroup_iso: "h \ surj(carrier(G), carrier(H)) \ (\X\carrier (G Mod kernel(G,H,h)). contents (h``X)) \ (G Mod (kernel(G,H,h))) \ H" by (simp add: iso_def FactGroup_hom FactGroup_inj bij_def FactGroup_surj) end diff --git a/src/ZF/ex/LList.thy b/src/ZF/ex/LList.thy --- a/src/ZF/ex/LList.thy +++ b/src/ZF/ex/LList.thy @@ -1,262 +1,262 @@ (* Title: ZF/ex/LList.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Codatatype definition of Lazy Lists. Equality for llist(A) as a greatest fixed point Functions for Lazy Lists STILL NEEDS: co_recursion for defining lconst, flip, etc. a typing rule for it, based on some notion of "productivity..." *) theory LList imports ZF begin consts llist :: "i\i" codatatype "llist(A)" = LNil | LCons ("a \ A", "l \ llist(A)") (*Coinductive definition of equality*) consts lleq :: "i\i" (*Previously used <*> in the domain and variant pairs as elements. But standard pairs work just as well. To use variant pairs, must change prefix a q/Q to the Sigma, Pair and converse rules.*) coinductive domains "lleq(A)" \ "llist(A) * llist(A)" intros LNil: "\LNil, LNil\ \ lleq(A)" LCons: "\a \ A; \ lleq(A)\ \ \ lleq(A)" type_intros llist.intros (*Lazy list functions; flip is not definitional!*) definition lconst :: "i \ i" where "lconst(a) \ lfp(univ(a), \l. LCons(a,l))" axiomatization flip :: "i \ i" where flip_LNil: "flip(LNil) = LNil" and flip_LCons: "\x \ bool; l \ llist(bool)\ \ flip(LCons(x,l)) = LCons(not(x), flip(l))" (*These commands cause classical reasoning to regard the subset relation as primitive, not reducing it to membership*) declare subsetI [rule del] subsetCE [rule del] declare subset_refl [intro!] cons_subsetI [intro!] subset_consI [intro!] Union_least [intro!] UN_least [intro!] Un_least [intro!] Inter_greatest [intro!] Int_greatest [intro!] RepFun_subset [intro!] Un_upper1 [intro!] Un_upper2 [intro!] Int_lower1 [intro!] Int_lower2 [intro!] (*An elimination rule, for type-checking*) inductive_cases LConsE: "LCons(a,l) \ llist(A)" (*Proving freeness results*) lemma LCons_iff: "LCons(a,l)=LCons(a',l') \ a=a' \ l=l'" by auto lemma LNil_LCons_iff: "\ LNil=LCons(a,l)" by auto (* lemma llist_unfold: "llist(A) = {0} <+> (A <*> llist(A))"; let open llist val rew = rewrite_rule con_defs in by (fast_tac (claset() addSIs (subsetI ::map rew intros) addEs [rew elim]) 1) end done *) (*** Lemmas to justify using "llist" in other recursive type definitions ***) lemma llist_mono: "A \ B \ llist(A) \ llist(B)" -apply (unfold llist.defs ) + unfolding llist.defs apply (rule gfp_mono) apply (rule llist.bnd_mono) apply (assumption | rule quniv_mono basic_monos)+ done (** Closure of quniv(A) under llist -- why so complex? Its a gfp... **) declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!] QPair_subset_univ [intro!] empty_subsetI [intro!] one_in_quniv [THEN qunivD, intro!] declare qunivD [dest!] declare Ord_in_Ord [elim!] lemma llist_quniv_lemma: "Ord(i) \ l \ llist(quniv(A)) \ l \ Vset(i) \ univ(eclose(A))" proof (induct i arbitrary: l rule: trans_induct) case (step i l) show ?case using \l \ llist(quniv(A))\ proof (cases l rule: llist.cases) case LNil thus ?thesis by (simp add: QInl_def QInr_def llist.con_defs) next case (LCons a l) thus ?thesis using step.hyps proof (simp add: QInl_def QInr_def llist.con_defs) show "<1; > \ Vset(i) \ univ(eclose(A))" using LCons \Ord(i)\ by (fast intro: step Ord_trans Int_lower1 [THEN subset_trans]) qed qed qed lemma llist_quniv: "llist(quniv(A)) \ quniv(A)" apply (rule qunivI [THEN subsetI]) apply (rule Int_Vset_subset) apply (assumption | rule llist_quniv_lemma)+ done lemmas llist_subset_quniv = subset_trans [OF llist_mono llist_quniv] (*** Lazy List Equality: lleq ***) declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!] QPair_mono [intro!] declare Ord_in_Ord [elim!] (*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*) lemma lleq_Int_Vset_subset: "Ord(i) \ \ lleq(A) \ l \ Vset(i) \ l'" proof (induct i arbitrary: l l' rule: trans_induct) case (step i l l') show ?case using \\l, l'\ \ lleq(A)\ proof (cases rule: lleq.cases) case LNil thus ?thesis by (auto simp add: QInr_def llist.con_defs) next case (LCons a l l') thus ?thesis using step.hyps by (auto simp add: QInr_def llist.con_defs intro: Ord_trans) qed qed (*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*) lemma lleq_symmetric: " \ lleq(A) \ \ lleq(A)" apply (erule lleq.coinduct [OF converseI]) apply (rule lleq.dom_subset [THEN converse_type], safe) apply (erule lleq.cases, blast+) done lemma lleq_implies_equal: " \ lleq(A) \ l=l'" apply (rule equalityI) apply (assumption | rule lleq_Int_Vset_subset [THEN Int_Vset_subset] | erule lleq_symmetric)+ done lemma equal_llist_implies_leq: "\l=l'; l \ llist(A)\ \ \ lleq(A)" apply (rule_tac X = "{\l,l\. l \ llist (A) }" in lleq.coinduct) apply blast apply safe apply (erule_tac a=la in llist.cases, fast+) done (*** Lazy List Functions ***) (*Examples of coinduction for type-checking and to prove llist equations*) (*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***) lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), \l. LCons(a,l))" -apply (unfold llist.con_defs ) + unfolding llist.con_defs apply (rule bnd_monoI) apply (blast intro: A_subset_univ QInr_subset_univ) apply (blast intro: subset_refl QInr_mono QPair_mono) done (* lconst(a) = LCons(a,lconst(a)) *) lemmas lconst = def_lfp_unfold [OF lconst_def lconst_fun_bnd_mono] lemmas lconst_subset = lconst_def [THEN def_lfp_subset] lemmas member_subset_Union_eclose = arg_into_eclose [THEN Union_upper] lemma lconst_in_quniv: "a \ A \ lconst(a) \ quniv(A)" apply (rule lconst_subset [THEN subset_trans, THEN qunivI]) apply (erule arg_into_eclose [THEN eclose_subset, THEN univ_mono]) done lemma lconst_type: "a \ A \ lconst(a): llist(A)" apply (rule singletonI [THEN llist.coinduct]) apply (erule lconst_in_quniv [THEN singleton_subsetI]) apply (fast intro!: lconst) done (*** flip --- equations merely assumed; certain consequences proved ***) declare flip_LNil [simp] flip_LCons [simp] not_type [simp] lemma bool_Int_subset_univ: "b \ bool \ b \ X \ univ(eclose(A))" by (fast intro: Int_lower1 [THEN subset_trans] elim!: boolE) declare not_type [intro!] declare bool_Int_subset_univ [intro] (*Reasoning borrowed from lleq.ML; a similar proof works for all "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*) lemma flip_llist_quniv_lemma: "Ord(i) \ l \ llist(bool) \ flip(l) \ Vset(i) \ univ(eclose(bool))" proof (induct i arbitrary: l rule: trans_induct) case (step i l) show ?case using \l \ llist(bool)\ proof (cases l rule: llist.cases) case LNil thus ?thesis by (simp, simp add: QInl_def QInr_def llist.con_defs) next case (LCons a l) thus ?thesis using step.hyps proof (simp, simp add: QInl_def QInr_def llist.con_defs) show "<1; > \ Vset(i) \ univ(eclose(bool))" using LCons step.hyps by (auto intro: Ord_trans) qed qed qed lemma flip_in_quniv: "l \ llist(bool) \ flip(l) \ quniv(bool)" by (rule flip_llist_quniv_lemma [THEN Int_Vset_subset, THEN qunivI], assumption+) lemma flip_type: "l \ llist(bool) \ flip(l): llist(bool)" apply (rule_tac X = "{flip (l) . l \ llist (bool) }" in llist.coinduct) apply blast apply (fast intro!: flip_in_quniv) apply (erule RepFunE) apply (erule_tac a=la in llist.cases, auto) done lemma flip_flip: "l \ llist(bool) \ flip(flip(l)) = l" apply (rule_tac X1 = "{ . l \ llist (bool) }" in lleq.coinduct [THEN lleq_implies_equal]) apply blast apply (fast intro!: flip_type) apply (erule RepFunE) apply (erule_tac a=la in llist.cases) apply (simp_all add: flip_type not_not) done end diff --git a/src/ZF/ex/Ramsey.thy b/src/ZF/ex/Ramsey.thy --- a/src/ZF/ex/Ramsey.thy +++ b/src/ZF/ex/Ramsey.thy @@ -1,197 +1,197 @@ (* Title: ZF/ex/Ramsey.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Ramsey's Theorem (finite exponent 2 version) Based upon the article D Basin and M Kaufmann, The Boyer-Moore Prover and Nuprl: An Experimental Comparison. In G Huet and G Plotkin, editors, Logical Frameworks. (CUP, 1991), pages 89-119 See also M Kaufmann, An example in NQTHM: Ramsey's Theorem Internal Note, Computational Logic, Inc., Austin, Texas 78703 Available from the author: kaufmann@cli.com This function compute Ramsey numbers according to the proof given below (which, does not constrain the base case values at all. fun ram 0 j = 1 | ram i 0 = 1 | ram i j = ram (i-1) j + ram i (j-1) *) theory Ramsey imports ZF begin definition Symmetric :: "i\o" where "Symmetric(E) \ (\x y. \x,y\:E \ \y,x\:E)" definition Atleast :: "[i,i]\o" where \ \not really necessary: ZF defines cardinality\ "Atleast(n,S) \ (\f. f \ inj(n,S))" definition Clique :: "[i,i,i]\o" where "Clique(C,V,E) \ (C \ V) \ (\x \ C. \y \ C. x\y \ \x,y\ \ E)" definition Indept :: "[i,i,i]\o" where "Indept(I,V,E) \ (I \ V) \ (\x \ I. \y \ I. x\y \ \x,y\ \ E)" definition Ramsey :: "[i,i,i]\o" where "Ramsey(n,i,j) \ \V E. Symmetric(E) \ Atleast(n,V) \ (\C. Clique(C,V,E) \ Atleast(i,C)) | (\I. Indept(I,V,E) \ Atleast(j,I))" (*** Cliques and Independent sets ***) lemma Clique0 [intro]: "Clique(0,V,E)" by (unfold Clique_def, blast) lemma Clique_superset: "\Clique(C,V',E); V'<=V\ \ Clique(C,V,E)" by (unfold Clique_def, blast) lemma Indept0 [intro]: "Indept(0,V,E)" by (unfold Indept_def, blast) lemma Indept_superset: "\Indept(I,V',E); V'<=V\ \ Indept(I,V,E)" by (unfold Indept_def, blast) (*** Atleast ***) lemma Atleast0 [intro]: "Atleast(0,A)" by (unfold Atleast_def inj_def Pi_def function_def, blast) lemma Atleast_succD: "Atleast(succ(m),A) \ \x \ A. Atleast(m, A-{x})" unfolding Atleast_def apply (blast dest: inj_is_fun [THEN apply_type] inj_succ_restrict) done lemma Atleast_superset: "\Atleast(n,A); A \ B\ \ Atleast(n,B)" by (unfold Atleast_def, blast intro: inj_weaken_type) lemma Atleast_succI: "\Atleast(m,B); b\ B\ \ Atleast(succ(m), cons(b,B))" -apply (unfold Atleast_def succ_def) + unfolding Atleast_def succ_def apply (blast intro: inj_extend elim: mem_irrefl) done lemma Atleast_Diff_succI: "\Atleast(m, B-{x}); x \ B\ \ Atleast(succ(m), B)" by (blast intro: Atleast_succI [THEN Atleast_superset]) (*** Main Cardinality Lemma ***) (*The #-succ(0) strengthens the original theorem statement, but precisely the same proof could be used\*) lemma pigeon2 [rule_format]: "m \ nat \ \n \ nat. \A B. Atleast((m#+n) #- succ(0), A \ B) \ Atleast(m,A) | Atleast(n,B)" apply (induct_tac "m") apply (blast intro!: Atleast0, simp) apply (rule ballI) apply (rename_tac m' n) (*simplifier does NOT preserve bound names!*) apply (induct_tac "n", auto) apply (erule Atleast_succD [THEN bexE]) apply (rename_tac n' A B z) apply (erule UnE) (**case z \ B. Instantiate the '\A B' induction hypothesis. **) apply (drule_tac [2] x1 = A and x = "B-{z}" in spec [THEN spec]) apply (erule_tac [2] mp [THEN disjE]) (*cases Atleast(succ(m1),A) and Atleast(succ(k),B)*) apply (erule_tac [3] asm_rl notE Atleast_Diff_succI)+ (*proving the condition*) prefer 2 apply (blast intro: Atleast_superset) (**case z \ A. Instantiate the '\n \ nat. \A B' induction hypothesis. **) apply (drule_tac x2="succ(n')" and x1="A-{z}" and x=B in bspec [THEN spec, THEN spec]) apply (erule nat_succI) apply (erule mp [THEN disjE]) (*cases Atleast(succ(m1),A) and Atleast(succ(k),B)*) apply (erule_tac [2] asm_rl Atleast_Diff_succI notE)+ (*proving the condition*) apply simp apply (blast intro: Atleast_superset) done (**** Ramsey's Theorem ****) (** Base cases of induction; they now admit ANY Ramsey number **) lemma Ramsey0j: "Ramsey(n,0,j)" by (unfold Ramsey_def, blast) lemma Ramseyi0: "Ramsey(n,i,0)" by (unfold Ramsey_def, blast) (** Lemmas for induction step **) (*The use of succ(m) here, rather than #-succ(0), simplifies the proof of Ramsey_step_lemma.*) lemma Atleast_partition: "\Atleast(m #+ n, A); m \ nat; n \ nat\ \ Atleast(succ(m), {x \ A. \P(x)}) | Atleast(n, {x \ A. P(x)})" apply (rule nat_succI [THEN pigeon2], assumption+) apply (rule Atleast_superset, auto) done (*For the Atleast part, proves \(a \ I) from the second premise!*) lemma Indept_succ: "\Indept(I, {z \ V-{a}. \a,z\ \ E}, E); Symmetric(E); a \ V; Atleast(j,I)\ \ Indept(cons(a,I), V, E) \ Atleast(succ(j), cons(a,I))" -apply (unfold Symmetric_def Indept_def) + unfolding Symmetric_def Indept_def apply (blast intro!: Atleast_succI) done lemma Clique_succ: "\Clique(C, {z \ V-{a}. \a,z\:E}, E); Symmetric(E); a \ V; Atleast(j,C)\ \ Clique(cons(a,C), V, E) \ Atleast(succ(j), cons(a,C))" -apply (unfold Symmetric_def Clique_def) + unfolding Symmetric_def Clique_def apply (blast intro!: Atleast_succI) done (** Induction step **) (*Published proofs gloss over the need for Ramsey numbers to be POSITIVE.*) lemma Ramsey_step_lemma: "\Ramsey(succ(m), succ(i), j); Ramsey(n, i, succ(j)); m \ nat; n \ nat\ \ Ramsey(succ(m#+n), succ(i), succ(j))" apply (unfold Ramsey_def, clarify) apply (erule Atleast_succD [THEN bexE]) apply (erule_tac P1 = "\z.\x,z\:E" in Atleast_partition [THEN disjE], assumption+) (*case m*) apply (fast dest!: Indept_succ elim: Clique_superset) (*case n*) apply (fast dest!: Clique_succ elim: Indept_superset) done (** The actual proof **) (*Again, the induction requires Ramsey numbers to be positive.*) lemma ramsey_lemma: "i \ nat \ \j \ nat. \n \ nat. Ramsey(succ(n), i, j)" apply (induct_tac "i") apply (blast intro!: Ramsey0j) apply (rule ballI) apply (induct_tac "j") apply (blast intro!: Ramseyi0) apply (blast intro!: add_type Ramsey_step_lemma) done (*Final statement in a tidy form, without succ(...) *) lemma ramsey: "\i \ nat; j \ nat\ \ \n \ nat. Ramsey(n,i,j)" by (blast dest: ramsey_lemma) end diff --git a/src/ZF/func.thy b/src/ZF/func.thy --- a/src/ZF/func.thy +++ b/src/ZF/func.thy @@ -1,611 +1,611 @@ (* Title: ZF/func.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge *) section\Functions, Function Spaces, Lambda-Abstraction\ theory func imports equalities Sum begin subsection\The Pi Operator: Dependent Function Space\ lemma subset_Sigma_imp_relation: "r \ Sigma(A,B) \ relation(r)" by (simp add: relation_def, blast) lemma relation_converse_converse [simp]: "relation(r) \ converse(converse(r)) = r" by (simp add: relation_def, blast) lemma relation_restrict [simp]: "relation(restrict(r,A))" by (simp add: restrict_def relation_def, blast) lemma Pi_iff: "f \ Pi(A,B) \ function(f) \ f<=Sigma(A,B) \ A<=domain(f)" by (unfold Pi_def, blast) (*For upward compatibility with the former definition*) lemma Pi_iff_old: "f \ Pi(A,B) \ f<=Sigma(A,B) \ (\x\A. \!y. \x,y\: f)" by (unfold Pi_def function_def, blast) lemma fun_is_function: "f \ Pi(A,B) \ function(f)" by (simp only: Pi_iff) lemma function_imp_Pi: "\function(f); relation(f)\ \ f \ domain(f) -> range(f)" by (simp add: Pi_iff relation_def, blast) lemma functionI: "\\x y y'. \\x,y\:r; :r\ \ y=y'\ \ function(r)" by (simp add: function_def, blast) (*Functions are relations*) lemma fun_is_rel: "f \ Pi(A,B) \ f \ Sigma(A,B)" by (unfold Pi_def, blast) lemma Pi_cong: "\A=A'; \x. x \ A' \ B(x)=B'(x)\ \ Pi(A,B) = Pi(A',B')" by (simp add: Pi_def cong add: Sigma_cong) (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause flex-flex pairs and the "Check your prover" error. Most Sigmas and Pis are abbreviated as * or -> *) (*Weakening one function type to another; see also Pi_type*) lemma fun_weaken_type: "\f \ A->B; B<=D\ \ f \ A->D" by (unfold Pi_def, best) subsection\Function Application\ lemma apply_equality2: "\\a,b\: f; \a,c\: f; f \ Pi(A,B)\ \ b=c" by (unfold Pi_def function_def, blast) lemma function_apply_equality: "\\a,b\: f; function(f)\ \ f`a = b" by (unfold apply_def function_def, blast) lemma apply_equality: "\\a,b\: f; f \ Pi(A,B)\ \ f`a = b" unfolding Pi_def apply (blast intro: function_apply_equality) done (*Applying a function outside its domain yields 0*) lemma apply_0: "a \ domain(f) \ f`a = 0" by (unfold apply_def, blast) lemma Pi_memberD: "\f \ Pi(A,B); c \ f\ \ \x\A. c = " apply (frule fun_is_rel) apply (blast dest: apply_equality) done lemma function_apply_Pair: "\function(f); a \ domain(f)\ \ : f" apply (simp add: function_def, clarify) apply (subgoal_tac "f`a = y", blast) apply (simp add: apply_def, blast) done lemma apply_Pair: "\f \ Pi(A,B); a \ A\ \ : f" apply (simp add: Pi_iff) apply (blast intro: function_apply_Pair) done (*Conclusion is flexible -- use rule_tac or else apply_funtype below!*) lemma apply_type [TC]: "\f \ Pi(A,B); a \ A\ \ f`a \ B(a)" by (blast intro: apply_Pair dest: fun_is_rel) (*This version is acceptable to the simplifier*) lemma apply_funtype: "\f \ A->B; a \ A\ \ f`a \ B" by (blast dest: apply_type) lemma apply_iff: "f \ Pi(A,B) \ \a,b\: f \ a \ A \ f`a = b" apply (frule fun_is_rel) apply (blast intro!: apply_Pair apply_equality) done (*Refining one Pi type to another*) lemma Pi_type: "\f \ Pi(A,C); \x. x \ A \ f`x \ B(x)\ \ f \ Pi(A,B)" apply (simp only: Pi_iff) apply (blast dest: function_apply_equality) done (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*) lemma Pi_Collect_iff: "(f \ Pi(A, \x. {y \ B(x). P(x,y)})) \ f \ Pi(A,B) \ (\x\A. P(x, f`x))" by (blast intro: Pi_type dest: apply_type) lemma Pi_weaken_type: "\f \ Pi(A,B); \x. x \ A \ B(x)<=C(x)\ \ f \ Pi(A,C)" by (blast intro: Pi_type dest: apply_type) (** Elimination of membership in a function **) lemma domain_type: "\\a,b\ \ f; f \ Pi(A,B)\ \ a \ A" by (blast dest: fun_is_rel) lemma range_type: "\\a,b\ \ f; f \ Pi(A,B)\ \ b \ B(a)" by (blast dest: fun_is_rel) lemma Pair_mem_PiD: "\\a,b\: f; f \ Pi(A,B)\ \ a \ A \ b \ B(a) \ f`a = b" by (blast intro: domain_type range_type apply_equality) subsection\Lambda Abstraction\ lemma lamI: "a \ A \ \ (\x\A. b(x))" unfolding lam_def apply (erule RepFunI) done lemma lamE: "\p: (\x\A. b(x)); \x.\x \ A; p=\ \ P \ \ P" by (simp add: lam_def, blast) lemma lamD: "\\a,c\: (\x\A. b(x))\ \ c = b(a)" by (simp add: lam_def) lemma lam_type [TC]: "\\x. x \ A \ b(x): B(x)\ \ (\x\A. b(x)) \ Pi(A,B)" by (simp add: lam_def Pi_def function_def, blast) lemma lam_funtype: "(\x\A. b(x)) \ A -> {b(x). x \ A}" by (blast intro: lam_type) lemma function_lam: "function (\x\A. b(x))" by (simp add: function_def lam_def) lemma relation_lam: "relation (\x\A. b(x))" by (simp add: relation_def lam_def) lemma beta_if [simp]: "(\x\A. b(x)) ` a = (if a \ A then b(a) else 0)" by (simp add: apply_def lam_def, blast) lemma beta: "a \ A \ (\x\A. b(x)) ` a = b(a)" by (simp add: apply_def lam_def, blast) lemma lam_empty [simp]: "(\x\0. b(x)) = 0" by (simp add: lam_def) lemma domain_lam [simp]: "domain(Lambda(A,b)) = A" by (simp add: lam_def, blast) (*congruence rule for lambda abstraction*) lemma lam_cong [cong]: "\A=A'; \x. x \ A' \ b(x)=b'(x)\ \ Lambda(A,b) = Lambda(A',b')" by (simp only: lam_def cong add: RepFun_cong) lemma lam_theI: "(\x. x \ A \ \!y. Q(x,y)) \ \f. \x\A. Q(x, f`x)" apply (rule_tac x = "\x\A. THE y. Q (x,y)" in exI) apply simp apply (blast intro: theI) done lemma lam_eqE: "\(\x\A. f(x)) = (\x\A. g(x)); a \ A\ \ f(a)=g(a)" by (fast intro!: lamI elim: equalityE lamE) (*Empty function spaces*) lemma Pi_empty1 [simp]: "Pi(0,A) = {0}" by (unfold Pi_def function_def, blast) (*The singleton function*) lemma singleton_fun [simp]: "{\a,b\} \ {a} -> {b}" by (unfold Pi_def function_def, blast) lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)" by (unfold Pi_def function_def, force) lemma fun_space_empty_iff [iff]: "(A->X)=0 \ X=0 \ (A \ 0)" apply auto apply (fast intro!: equals0I intro: lam_type) done subsection\Extensionality\ (*Semi-extensionality!*) lemma fun_subset: "\f \ Pi(A,B); g \ Pi(C,D); A<=C; \x. x \ A \ f`x = g`x\ \ f<=g" by (force dest: Pi_memberD intro: apply_Pair) lemma fun_extension: "\f \ Pi(A,B); g \ Pi(A,D); \x. x \ A \ f`x = g`x\ \ f=g" by (blast del: subsetI intro: subset_refl sym fun_subset) lemma eta [simp]: "f \ Pi(A,B) \ (\x\A. f`x) = f" apply (rule fun_extension) apply (auto simp add: lam_type apply_type beta) done lemma fun_extension_iff: "\f \ Pi(A,B); g \ Pi(A,C)\ \ (\a\A. f`a = g`a) \ f=g" by (blast intro: fun_extension) (*thm by Mark Staples, proof by lcp*) lemma fun_subset_eq: "\f \ Pi(A,B); g \ Pi(A,C)\ \ f \ g \ (f = g)" by (blast dest: apply_Pair intro: fun_extension apply_equality [symmetric]) (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*) lemma Pi_lamE: assumes major: "f \ Pi(A,B)" and minor: "\b. \\x\A. b(x):B(x); f = (\x\A. b(x))\ \ P" shows "P" apply (rule minor) apply (rule_tac [2] eta [symmetric]) apply (blast intro: major apply_type)+ done subsection\Images of Functions\ lemma image_lam: "C \ A \ (\x\A. b(x)) `` C = {b(x). x \ C}" by (unfold lam_def, blast) lemma Repfun_function_if: "function(f) \ {f`x. x \ C} = (if C \ domain(f) then f``C else cons(0,f``C))" apply simp apply (intro conjI impI) apply (blast dest: function_apply_equality intro: function_apply_Pair) apply (rule equalityI) apply (blast intro!: function_apply_Pair apply_0) apply (blast dest: function_apply_equality intro: apply_0 [symmetric]) done (*For this lemma and the next, the right-hand side could equivalently be written \x\C. {f`x} *) lemma image_function: "\function(f); C \ domain(f)\ \ f``C = {f`x. x \ C}" by (simp add: Repfun_function_if) lemma image_fun: "\f \ Pi(A,B); C \ A\ \ f``C = {f`x. x \ C}" apply (simp add: Pi_iff) apply (blast intro: image_function) done lemma image_eq_UN: assumes f: "f \ Pi(A,B)" "C \ A" shows "f``C = (\x\C. {f ` x})" by (auto simp add: image_fun [OF f]) lemma Pi_image_cons: "\f \ Pi(A,B); x \ A\ \ f `` cons(x,y) = cons(f`x, f``y)" by (blast dest: apply_equality apply_Pair) subsection\Properties of \<^term>\restrict(f,A)\\ lemma restrict_subset: "restrict(f,A) \ f" by (unfold restrict_def, blast) lemma function_restrictI: "function(f) \ function(restrict(f,A))" by (unfold restrict_def function_def, blast) lemma restrict_type2: "\f \ Pi(C,B); A<=C\ \ restrict(f,A) \ Pi(A,B)" by (simp add: Pi_iff function_def restrict_def, blast) lemma restrict: "restrict(f,A) ` a = (if a \ A then f`a else 0)" by (simp add: apply_def restrict_def, blast) lemma restrict_empty [simp]: "restrict(f,0) = 0" by (unfold restrict_def, simp) lemma restrict_iff: "z \ restrict(r,A) \ z \ r \ (\x\A. \y. z = \x, y\)" by (simp add: restrict_def) lemma restrict_restrict [simp]: "restrict(restrict(r,A),B) = restrict(r, A \ B)" by (unfold restrict_def, blast) lemma domain_restrict [simp]: "domain(restrict(f,C)) = domain(f) \ C" unfolding restrict_def apply (auto simp add: domain_def) done lemma restrict_idem: "f \ Sigma(A,B) \ restrict(f,A) = f" by (simp add: restrict_def, blast) (*converse probably holds too*) lemma domain_restrict_idem: "\domain(r) \ A; relation(r)\ \ restrict(r,A) = r" by (simp add: restrict_def relation_def, blast) lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A \ C" -apply (unfold restrict_def lam_def) + unfolding restrict_def lam_def apply (rule equalityI) apply (auto simp add: domain_iff) done lemma restrict_if [simp]: "restrict(f,A) ` a = (if a \ A then f`a else 0)" by (simp add: restrict apply_0) lemma restrict_lam_eq: "A<=C \ restrict(\x\C. b(x), A) = (\x\A. b(x))" by (unfold restrict_def lam_def, auto) lemma fun_cons_restrict_eq: "f \ cons(a, b) -> B \ f = cons(, restrict(f, b))" apply (rule equalityI) prefer 2 apply (blast intro: apply_Pair restrict_subset [THEN subsetD]) apply (auto dest!: Pi_memberD simp add: restrict_def lam_def) done subsection\Unions of Functions\ (** The Union of a set of COMPATIBLE functions is a function **) lemma function_Union: "\\x\S. function(x); \x\S. \y\S. x<=y | y<=x\ \ function(\(S))" by (unfold function_def, blast) lemma fun_Union: "\\f\S. \C D. f \ C->D; \f\S. \y\S. f<=y | y<=f\ \ \(S) \ domain(\(S)) -> range(\(S))" unfolding Pi_def apply (blast intro!: rel_Union function_Union) done lemma gen_relation_Union: "(\f. f\F \ relation(f)) \ relation(\(F))" by (simp add: relation_def) (** The Union of 2 disjoint functions is a function **) lemmas Un_rls = Un_subset_iff SUM_Un_distrib1 prod_Un_distrib2 subset_trans [OF _ Un_upper1] subset_trans [OF _ Un_upper2] lemma fun_disjoint_Un: "\f \ A->B; g \ C->D; A \ C = 0\ \ (f \ g) \ (A \ C) -> (B \ D)" (*Prove the product and domain subgoals using distributive laws*) apply (simp add: Pi_iff extension Un_rls) apply (unfold function_def, blast) done lemma fun_disjoint_apply1: "a \ domain(g) \ (f \ g)`a = f`a" by (simp add: apply_def, blast) lemma fun_disjoint_apply2: "c \ domain(f) \ (f \ g)`c = g`c" by (simp add: apply_def, blast) subsection\Domain and Range of a Function or Relation\ lemma domain_of_fun: "f \ Pi(A,B) \ domain(f)=A" by (unfold Pi_def, blast) lemma apply_rangeI: "\f \ Pi(A,B); a \ A\ \ f`a \ range(f)" by (erule apply_Pair [THEN rangeI], assumption) lemma range_of_fun: "f \ Pi(A,B) \ f \ A->range(f)" by (blast intro: Pi_type apply_rangeI) subsection\Extensions of Functions\ lemma fun_extend: "\f \ A->B; c\A\ \ cons(\c,b\,f) \ cons(c,A) -> cons(b,B)" apply (frule singleton_fun [THEN fun_disjoint_Un], blast) apply (simp add: cons_eq) done lemma fun_extend3: "\f \ A->B; c\A; b \ B\ \ cons(\c,b\,f) \ cons(c,A) -> B" by (blast intro: fun_extend [THEN fun_weaken_type]) lemma extend_apply: "c \ domain(f) \ cons(\c,b\,f)`a = (if a=c then b else f`a)" by (auto simp add: apply_def) lemma fun_extend_apply [simp]: "\f \ A->B; c\A\ \ cons(\c,b\,f)`a = (if a=c then b else f`a)" apply (rule extend_apply) apply (simp add: Pi_def, blast) done lemmas singleton_apply = apply_equality [OF singletonI singleton_fun, simp] (*For Finite.ML. Inclusion of right into left is easy*) lemma cons_fun_eq: "c \ A \ cons(c,A) -> B = (\f \ A->B. \b\B. {cons(\c,b\, f)})" apply (rule equalityI) apply (safe elim!: fun_extend3) (*Inclusion of left into right*) apply (subgoal_tac "restrict (x, A) \ A -> B") prefer 2 apply (blast intro: restrict_type2) apply (rule UN_I, assumption) apply (rule apply_funtype [THEN UN_I]) apply assumption apply (rule consI1) apply (simp (no_asm)) apply (rule fun_extension) apply assumption apply (blast intro: fun_extend) apply (erule consE, simp_all) done lemma succ_fun_eq: "succ(n) -> B = (\f \ n->B. \b\B. {cons(\n,b\, f)})" by (simp add: succ_def mem_not_refl cons_fun_eq) subsection\Function Updates\ definition update :: "[i,i,i] \ i" where "update(f,a,b) \ \x\cons(a, domain(f)). if(x=a, b, f`x)" nonterminal updbinds and updbind syntax (* Let expressions *) "_updbind" :: "[i, i] \ updbind" (\(2_ :=/ _)\) "" :: "updbind \ updbinds" (\_\) "_updbinds" :: "[updbind, updbinds] \ updbinds" (\_,/ _\) "_Update" :: "[i, updbinds] \ i" (\_/'((_)')\ [900,0] 900) translations "_Update (f, _updbinds(b,bs))" == "_Update (_Update(f,b), bs)" "f(x:=y)" == "CONST update(f,x,y)" lemma update_apply [simp]: "f(x:=y) ` z = (if z=x then y else f`z)" apply (simp add: update_def) apply (case_tac "z \ domain(f)") apply (simp_all add: apply_0) done lemma update_idem: "\f`x = y; f \ Pi(A,B); x \ A\ \ f(x:=y) = f" unfolding update_def apply (simp add: domain_of_fun cons_absorb) apply (rule fun_extension) apply (best intro: apply_type if_type lam_type, assumption, simp) done (* \f \ Pi(A, B); x \ A\ \ f(x := f`x) = f *) declare refl [THEN update_idem, simp] lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))" by (unfold update_def, simp) lemma update_type: "\f \ Pi(A,B); x \ A; y \ B(x)\ \ f(x:=y) \ Pi(A, B)" unfolding update_def apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type) done subsection\Monotonicity Theorems\ subsubsection\Replacement in its Various Forms\ (*Not easy to express monotonicity in P, since any "bigger" predicate would have to be single-valued*) lemma Replace_mono: "A<=B \ Replace(A,P) \ Replace(B,P)" by (blast elim!: ReplaceE) lemma RepFun_mono: "A<=B \ {f(x). x \ A} \ {f(x). x \ B}" by blast lemma Pow_mono: "A<=B \ Pow(A) \ Pow(B)" by blast lemma Union_mono: "A<=B \ \(A) \ \(B)" by blast lemma UN_mono: "\A<=C; \x. x \ A \ B(x)<=D(x)\ \ (\x\A. B(x)) \ (\x\C. D(x))" by blast (*Intersection is ANTI-monotonic. There are TWO premises! *) lemma Inter_anti_mono: "\A<=B; A\0\ \ \(B) \ \(A)" by blast lemma cons_mono: "C<=D \ cons(a,C) \ cons(a,D)" by blast lemma Un_mono: "\A<=C; B<=D\ \ A \ B \ C \ D" by blast lemma Int_mono: "\A<=C; B<=D\ \ A \ B \ C \ D" by blast lemma Diff_mono: "\A<=C; D<=B\ \ A-B \ C-D" by blast subsubsection\Standard Products, Sums and Function Spaces\ lemma Sigma_mono [rule_format]: "\A<=C; \x. x \ A \ B(x) \ D(x)\ \ Sigma(A,B) \ Sigma(C,D)" by blast lemma sum_mono: "\A<=C; B<=D\ \ A+B \ C+D" by (unfold sum_def, blast) (*Note that B->A and C->A are typically disjoint!*) lemma Pi_mono: "B<=C \ A->B \ A->C" by (blast intro: lam_type elim: Pi_lamE) lemma lam_mono: "A<=B \ Lambda(A,c) \ Lambda(B,c)" unfolding lam_def apply (erule RepFun_mono) done subsubsection\Converse, Domain, Range, Field\ lemma converse_mono: "r<=s \ converse(r) \ converse(s)" by blast lemma domain_mono: "r<=s \ domain(r)<=domain(s)" by blast lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset] lemma range_mono: "r<=s \ range(r)<=range(s)" by blast lemmas range_rel_subset = subset_trans [OF range_mono range_subset] lemma field_mono: "r<=s \ field(r)<=field(s)" by blast lemma field_rel_subset: "r \ A*A \ field(r) \ A" by (erule field_mono [THEN subset_trans], blast) subsubsection\Images\ lemma image_pair_mono: "\\x y. \x,y\:r \ \x,y\:s; A<=B\ \ r``A \ s``B" by blast lemma vimage_pair_mono: "\\x y. \x,y\:r \ \x,y\:s; A<=B\ \ r-``A \ s-``B" by blast lemma image_mono: "\r<=s; A<=B\ \ r``A \ s``B" by blast lemma vimage_mono: "\r<=s; A<=B\ \ r-``A \ s-``B" by blast lemma Collect_mono: "\A<=B; \x. x \ A \ P(x) \ Q(x)\ \ Collect(A,P) \ Collect(B,Q)" by blast (*Used in intr_elim.ML and in individual datatype definitions*) lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono Part_mono in_mono (* Useful with simp; contributed by Clemens Ballarin. *) lemma bex_image_simp: "\f \ Pi(X, Y); A \ X\ \ (\x\f``A. P(x)) \ (\x\A. P(f`x))" apply safe apply rule prefer 2 apply assumption apply (simp add: apply_equality) apply (blast intro: apply_Pair) done lemma ball_image_simp: "\f \ Pi(X, Y); A \ X\ \ (\x\f``A. P(x)) \ (\x\A. P(f`x))" apply safe apply (blast intro: apply_Pair) apply (drule bspec) apply assumption apply (simp add: apply_equality) done end