# HG changeset patch # User paulson # Date 1664293895 -3600 # Node ID e44d861316485dd381d9b2afe07a9000b2b38781 # Parent f2094906e4916b460cded1a5a7938d24ff6327a1 Removal of obsolete ASCII syntax diff -r f2094906e491 -r e44d86131648 src/ZF/AC.thy --- a/src/ZF/AC.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,10 +9,10 @@ text\This definition comes from Halmos (1960), page 59.\ axiomatization where - AC: "[| a \ A; !!x. x \ A ==> (\y. y \ B(x)) |] ==> \z. z \ Pi(A,B)" + AC: "\a \ A; \x. x \ A \ (\y. y \ B(x))\ \ \z. z \ Pi(A,B)" (*The same as AC, but no premise @{term"a \ A"}*) -lemma AC_Pi: "[| !!x. x \ A ==> (\y. y \ B(x)) |] ==> \z. z \ Pi(A,B)" +lemma AC_Pi: "\\x. x \ A \ (\y. y \ B(x))\ \ \z. z \ Pi(A,B)" apply (case_tac "A=0") apply (simp add: Pi_empty1) (*The non-trivial case*) @@ -20,7 +20,7 @@ done (*Using dtac, this has the advantage of DELETING the universal quantifier*) -lemma AC_ball_Pi: "\x \ A. \y. y \ B(x) ==> \y. y \ Pi(A,B)" +lemma AC_ball_Pi: "\x \ A. \y. y \ B(x) \ \y. y \ Pi(A,B)" apply (rule AC_Pi) apply (erule bspec, assumption) done @@ -31,15 +31,15 @@ done lemma AC_func: - "[| !!x. x \ A ==> (\y. y \ x) |] ==> \f \ A->\(A). \x \ A. f`x \ x" + "\\x. x \ A \ (\y. y \ x)\ \ \f \ A->\(A). \x \ A. f`x \ x" apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE]) prefer 2 apply (blast dest: apply_type intro: Pi_type, blast) done -lemma non_empty_family: "[| 0 \ A; x \ A |] ==> \y. y \ x" +lemma non_empty_family: "\0 \ A; x \ A\ \ \y. y \ x" by (subgoal_tac "x \ 0", blast+) -lemma AC_func0: "0 \ A ==> \f \ A->\(A). \x \ A. f`x \ x" +lemma AC_func0: "0 \ A \ \f \ A->\(A). \x \ A. f`x \ x" apply (rule AC_func) apply (simp_all add: non_empty_family) done @@ -51,7 +51,7 @@ apply (erule_tac [2] fun_weaken_type, blast+) done -lemma AC_Pi0: "0 \ A ==> \f. f \ (\x \ A. x)" +lemma AC_Pi0: "0 \ A \ \f. f \ (\x \ A. x)" apply (rule AC_Pi) apply (simp_all add: non_empty_family) done diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC15_WO6.thy --- a/src/ZF/AC/AC15_WO6.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC15_WO6.thy Tue Sep 27 16:51:35 2022 +0100 @@ -4,18 +4,18 @@ The proofs needed to state that AC10, ..., AC15 are equivalent to the rest. We need the following: -WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6 +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 +AC10(succ(n)) \ AC13(n) \ AC14 \ AC15 or -AC1 ==> AC13(1); AC13(m) ==> AC13(n) ==> AC14 ==> AC15 (m\n) +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 +Moreover we don't need to prove AC13(1) \ AC1 and AC11 \ AC14 as Rubin & Rubin do. *) @@ -31,7 +31,7 @@ (* - ex_fun_AC13_AC15 *) (* ********************************************************************** *) -lemma lepoll_Sigma: "A\0 ==> B \ A*B" +lemma lepoll_Sigma: "A\0 \ B \ A*B" apply (unfold lepoll_def) apply (erule not_emptyE) apply (rule_tac x = "\z \ B. " in exI) @@ -39,24 +39,24 @@ done lemma cons_times_nat_not_Finite: - "0\A ==> \B \ {cons(0,x*nat). x \ A}. ~Finite(B)" + "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" +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" + "\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) & + \ \B \ A. \! u. u \ f`cons(0, B*nat) & u \ cons(0, B*nat) & 0 \ u & 2 \ u & u \ n" apply (unfold sets_of_size_between_def) apply (rule ballI) @@ -64,7 +64,7 @@ apply (blast dest: lemma1 intro!: lemma2, blast) done -lemma lemma4: "[| A \ i; Ord(i) |] ==> {P(a). a \ A} \ i" +lemma lemma4: "\A \ i; Ord(i)\ \ {P(a). a \ A} \ i" apply (unfold lepoll_def) apply (erule exE) apply (rule_tac x = "\x \ RepFun(A,P). \ j. \a\A. x=P(a) & f`a=j" @@ -81,7 +81,7 @@ done lemma lemma5_1: - "[| B \ A; 2 \ u(B) |] ==> (\x \ A. {fst(x). x \ u(x)-{0}})`B \ 0" + "\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 @@ -89,24 +89,24 @@ done lemma lemma5_2: - "[| B \ A; u(B) \ cons(0, B*nat) |] - ==> (\x \ A. {fst(x). x \ u(x)-{0}})`B \ B" + "\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" + "\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}. + "\\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" + 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) @@ -116,39 +116,39 @@ (* ********************************************************************** *) (* ********************************************************************** *) -(* AC10(n) ==> AC11 *) +(* AC10(n) \ AC11 *) (* ********************************************************************** *) -theorem AC10_AC11: "[| n \ nat; 1\n; AC10(n) |] ==> AC11" +theorem AC10_AC11: "\n \ nat; 1\n; AC10(n)\ \ AC11" by (unfold AC10_def AC11_def, blast) (* ********************************************************************** *) -(* AC11 ==> AC12 *) +(* AC11 \ AC12 *) (* ********************************************************************** *) -theorem AC11_AC12: "AC11 ==> AC12" +theorem AC11_AC12: "AC11 \ AC12" by (unfold AC10_def AC11_def AC11_def AC12_def, blast) (* ********************************************************************** *) -(* AC12 ==> AC15 *) +(* AC12 \ AC15 *) (* ********************************************************************** *) -theorem AC12_AC15: "AC12 ==> AC15" +theorem AC12_AC15: "AC12 \ AC15" apply (unfold AC12_def AC15_def) apply (blast del: ballI intro!: cons_times_nat_not_Finite ex_fun_AC13_AC15) done (* ********************************************************************** *) -(* AC15 ==> WO6 *) +(* AC15 \ WO6 *) (* ********************************************************************** *) -lemma OUN_eq_UN: "Ord(x) ==> (\aa \ x. F(a))" +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" + \ (\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) @@ -158,7 +158,7 @@ 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" + \ \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) @@ -167,7 +167,7 @@ (*but can't use del: DiffE despite the obvious conflict*) done -theorem AC15_WO6: "AC15 ==> WO6" +theorem AC15_WO6: "AC15 \ WO6" apply (unfold AC15_def WO6_def) apply (rule allI) apply (erule_tac x = "Pow (A) -{0}" in allE) @@ -188,14 +188,14 @@ (* ********************************************************************** *) (* ********************************************************************** *) -(* AC10(n) ==> AC13(n-1) if 2\n *) +(* 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)" +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) @@ -208,10 +208,10 @@ (* ********************************************************************** *) (* ********************************************************************** *) -(* AC1 ==> AC13(1) *) +(* AC1 \ AC13(1) *) (* ********************************************************************** *) -lemma AC1_AC13: "AC1 ==> AC13(1)" +lemma AC1_AC13: "AC1 \ AC13(1)" apply (unfold AC1_def AC13_def) apply (rule allI) apply (erule allE) @@ -223,10 +223,10 @@ done (* ********************************************************************** *) -(* AC13(m) ==> AC13(n) for m \ n *) +(* AC13(m) \ AC13(n) for m \ n *) (* ********************************************************************** *) -lemma AC13_mono: "[| m\n; AC13(m) |] ==> AC13(n)" +lemma AC13_mono: "\m\n; AC13(m)\ \ AC13(n)" apply (unfold AC13_def) apply (drule le_imp_lepoll) apply (fast elim!: lepoll_trans) @@ -237,17 +237,17 @@ (* ********************************************************************** *) (* ********************************************************************** *) -(* AC13(n) ==> AC14 if 1 \ n *) +(* AC13(n) \ AC14 if 1 \ n *) (* ********************************************************************** *) -theorem AC13_AC14: "[| n \ nat; 1\n; AC13(n) |] ==> AC14" +theorem AC13_AC14: "\n \ nat; 1\n; AC13(n)\ \ AC14" by (unfold AC13_def AC14_def, auto) (* ********************************************************************** *) -(* AC14 ==> AC15 *) +(* AC14 \ AC15 *) (* ********************************************************************** *) -theorem AC14_AC15: "AC14 ==> AC15" +theorem AC14_AC15: "AC14 \ AC15" by (unfold AC13_def AC14_def AC15_def, fast) (* ********************************************************************** *) @@ -255,15 +255,15 @@ (* ********************************************************************** *) (* ********************************************************************** *) -(* AC13(1) ==> AC1 *) +(* AC13(1) \ AC1 *) (* ********************************************************************** *) -lemma lemma_aux: "[| A\0; A \ 1 |] ==> \a. A={a}" +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)" + \ (\x \ A. THE y. f(x)={y}) \ (\X \ A. X)" apply (rule lam_type) apply (drule bspec, assumption) apply (elim conjE) @@ -271,16 +271,16 @@ apply (simp add: the_equality) done -theorem AC13_AC1: "AC13(1) ==> AC1" +theorem AC13_AC1: "AC13(1) \ AC1" apply (unfold AC13_def AC1_def) apply (fast elim!: AC13_AC1_lemma) done (* ********************************************************************** *) -(* AC11 ==> AC14 *) +(* AC11 \ AC14 *) (* ********************************************************************** *) -theorem AC11_AC14: "AC11 ==> AC14" +theorem AC11_AC14: "AC11 \ AC14" apply (unfold AC11_def AC14_def) apply (fast intro!: AC10_AC13) done diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC16_WO4.thy --- a/src/ZF/AC/AC16_WO4.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC16_WO4.thy Tue Sep 27 16:51:35 2022 +0100 @@ -1,7 +1,7 @@ (* Title: ZF/AC/AC16_WO4.thy Author: Krzysztof Grabczewski -The proof of AC16(n, k) ==> WO4(n-k) +The proof of AC16(n, k) \ WO4(n-k) Tidied (using locales) by LCP *) @@ -15,8 +15,8 @@ (* ********************************************************************** *) lemma lemma1: - "[| Finite(A); 0 nat |] - ==> \a f. Ord(a) & domain(f) = a & + "\Finite(A); 0 nat\ + \ \a f. Ord(a) & domain(f) = a & (\bb m)" apply (unfold Finite_def) apply (erule bexE) @@ -35,10 +35,10 @@ (* The case of infinite set *) (* ********************************************************************** *) -(* well_ord(x,r) ==> well_ord({{y,z}. y \ x}, Something(x,z)) **) +(* 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" +lemma lepoll_trans1: "\A \ B; \ A \ C\ \ \ B \ C" by (blast intro: lepoll_trans) (* ********************************************************************** *) @@ -47,7 +47,7 @@ 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)" +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]) @@ -60,7 +60,7 @@ lepoll_paired [THEN lepoll_Finite]) done -lemma infinite_Un: "~Finite(B) ==> ~Finite(A \ B)" +lemma infinite_Un: "\Finite(B) \ \Finite(A \ B)" by (blast intro: subset_Finite) (* ********************************************************************** *) @@ -77,8 +77,8 @@ (*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)" + "\\(\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*) @@ -86,7 +86,7 @@ apply force done -lemma succ_not_lepoll_imp_eqpoll: "[| ~A \ B; A \ B |] ==> succ(A) \ B" +lemma succ_not_lepoll_imp_eqpoll: "\\A \ B; A \ B\ \ succ(A) \ B" apply (unfold lepoll_def eqpoll_def bij_def surj_def) apply (fast elim!: succ_not_lepoll_lemma inj_is_fun) done @@ -100,16 +100,16 @@ 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)" + "\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))" + "\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)" + "\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]) @@ -121,10 +121,10 @@ OF eqpoll_sym [THEN eqpoll_imp_lepoll] subset_imp_lepoll]) done -lemma cons_eqE: "[| cons(x,a) = cons(y,a); x \ a |] ==> x = y " +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" +lemma eq_imp_Int_eq: "A = B \ A \ C = B \ C" by blast (* ********************************************************************** *) @@ -132,8 +132,8 @@ (* ********************************************************************** *) 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" + "\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 @@ -150,8 +150,8 @@ done lemma eqpoll_sum_imp_Diff_lepoll: - "[| A \ succ(k #+ m); B \ A; succ(k) \ B; k \ nat; m \ nat |] - ==> A-B \ m" + "\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 @@ -161,8 +161,8 @@ (* ********************************************************************** *) 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" + "\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) @@ -177,8 +177,8 @@ done lemma eqpoll_sum_imp_Diff_eqpoll: - "[| A \ succ(k #+ m); B \ A; succ(k) \ B; k \ nat; m \ nat |] - ==> A-B \ m" + "\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 @@ -192,24 +192,24 @@ by (fast dest!: lepoll_0_is_0 intro!: lepoll_refl) lemma subsets_lepoll_succ: - "n \ nat ==> {z \ Pow(y). z \ succ(n)} = + "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" + "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}" + 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" @@ -218,8 +218,8 @@ and lnat[iff]: "l \ nat" and mnat[iff]: "m \ nat" and mpos[iff]: "0 {v \ Pow(x). v \ m}" + and Infinite[iff]: "\ Finite(y)" + and noLepoll: "\ y \ {v \ Pow(x). v \ m}" begin lemma knat [iff]: "k \ nat" @@ -231,7 +231,7 @@ (* where a is certain k-2 element subset of y *) (* ********************************************************************** *) -lemma Diff_Finite_eqpoll: "[| l \ a; a \ y |] ==> y - a \ 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) @@ -252,20 +252,20 @@ 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)" + "\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) 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" +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" + "\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) @@ -275,18 +275,18 @@ 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)" + "\\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}" + "\\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" @@ -307,18 +307,18 @@ (*relies on the disjointness of x, y*) -lemma x_imp_not_y [dest]: "a \ x ==> a \ 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)" + "w \ x \ y \ w \ (x - {u}) = w - cons(u, w \ y)" by blast lemma w_Int_eqpoll_m: - "[| w \ {v \ s(u). a \ v}; + "\w \ {v \ s(u). a \ v}; l \ a; u \ x; - \v \ s(u). succ(l) \ v \ y |] - ==> w \ (x - {u}) \ m" + \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] @@ -336,22 +336,22 @@ (* to {v \ Pow(x). v \ n-k} *) (* ********************************************************************** *) -lemma eqpoll_m_not_empty: "a \ m ==> a \ 0" +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" + "\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]) apply (unfold 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}" + "\\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) @@ -373,14 +373,14 @@ (* ********************************************************************** *) lemma well_ord_subsets_eqpoll_n: - "n \ nat ==> \S. well_ord({z \ Pow(y) . z \ succ(n)}, S)" + "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)" + "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) @@ -410,7 +410,7 @@ (* ********************************************************************** *) lemma unique_superset_in_MM: - "v \ LL ==> \! w. w \ MM & v \ w" + "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]) @@ -426,11 +426,11 @@ (* The union of appropriate values is the whole x *) (* ********************************************************************** *) -lemma Int_in_LL: "w \ MM ==> w \ y \ LL" +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" + "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) @@ -440,13 +440,13 @@ 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" + "v \ LL \ (THE x. x \ MM & v \ x) \ x \ y" apply (drule unique_superset1) apply (unfold MM_def) apply (fast dest!: unique_superset1 "includes" [THEN subsetD]) done -lemma GG_subset: "v \ LL ==> GG ` v \ x" +lemma GG_subset: "v \ LL \ GG ` v \ x" apply (unfold GG_def) apply (frule the_in_MM_subset) apply (frule in_LL_eq_Int) @@ -461,7 +461,7 @@ apply (rule Infinite) done -lemma ex_subset_eqpoll_n: "n \ nat ==> \z. z \ y & n \ z" +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]) @@ -469,7 +469,7 @@ done -lemma exists_proper_in_s: "u \ x ==> \v \ s(u). succ(k) \ v \ y" +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) @@ -480,12 +480,12 @@ 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" +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" +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) @@ -495,7 +495,7 @@ apply (fast elim!: Int_in_LL)+ done -lemma OUN_eq_x: "well_ord(LL,S) ==> +lemma OUN_eq_x: "well_ord(LL,S) \ (\b MM ==> w \ succ(k #+ m)" +lemma in_MM_eqpoll_n: "w \ MM \ w \ succ(k #+ m)" apply (unfold MM_def) apply (fast dest: "includes" [THEN subsetD]) done -lemma in_LL_eqpoll_n: "w \ LL ==> succ(k) \ w" +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)" +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) ==> + "well_ord(LL,S) \ \b m" apply (unfold GG_def) apply (rule oallI) @@ -558,11 +558,11 @@ (* ********************************************************************** *) -(* The main theorem AC16(n, k) ==> WO4(n-k) *) +(* 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)" + "\AC_Equiv.AC16(k #+ m, k); 0 < k; 0 < m; k \ nat; m \ nat\ \ WO4(m)" apply (unfold AC_Equiv.AC16_def WO4_def) apply (rule allI) apply (case_tac "Finite (A)") diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC16_lemmas.thy --- a/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 16:51:35 2022 +0100 @@ -8,7 +8,7 @@ imports AC_Equiv Hartog Cardinal_aux begin -lemma cons_Diff_eq: "a\A ==> cons(a,A)-{a}=A" +lemma cons_Diff_eq: "a\A \ cons(a,A)-{a}=A" by fast lemma nat_1_lepoll_iff: "1\X \ (\x. x \ X)" @@ -28,7 +28,7 @@ apply (fast intro!: singleton_eqpoll_1) done -lemma cons_eqpoll_succ: "[| x\n; y\x |] ==> cons(y,x)\succ(n)" +lemma cons_eqpoll_succ: "\x\n; y\x\ \ cons(y,x)\succ(n)" apply (unfold succ_def) apply (fast elim!: cons_eqpoll_cong mem_irrefl) done @@ -60,7 +60,7 @@ done lemma InfCard_Least_in: - "[| InfCard(x); y \ x; y \ succ(z) |] ==> (\ i. i \ y) \ y" + "\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 @@ -69,8 +69,8 @@ done lemma subsets_lepoll_lemma1: - "[| InfCard(x); n \ nat |] - ==> {y \ Pow(x). y\succ(succ(n))} \ x*{y \ Pow(x). y\succ(n)}" + "\InfCard(x); n \ nat\ + \ {y \ Pow(x). y\succ(succ(n))} \ x*{y \ Pow(x). y\succ(n)}" apply (unfold lepoll_def) apply (rule_tac x = "\y \ {y \ Pow(x) . y\succ (succ (n))}. <\ i. i \ y, y-{\ i. i \ y}>" in exI) @@ -79,7 +79,7 @@ apply (simp, blast intro: InfCard_Least_in) done -lemma set_of_Ord_succ_Union: "(\y \ z. Ord(y)) ==> z \ succ(\(z))" +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) @@ -87,11 +87,11 @@ apply (blast dest: le_imp_subset elim: leE ltE)+ done -lemma subset_not_mem: "j \ i ==> i \ j" +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" + "(\y. y \ z \ Ord(y)) \ succ(\(z)) \ z" apply (rule set_of_Ord_succ_Union [THEN subset_not_mem], blast) done @@ -99,14 +99,14 @@ "\(cons(succ(\(z)),z)) = succ(\(z))" by fast -lemma Un_Ord_disj: "[| Ord(i); Ord(j) |] ==> i \ j = i | i \ j = j" +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})" +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" + "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) @@ -126,11 +126,11 @@ apply (drule Un_Ord_disj, assumption, auto) done -lemma Union_in: "[| \x \ z. Ord(x); z\n; z\0; n \ nat |] ==> \(z) \ z" +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" + "\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) @@ -142,8 +142,8 @@ done lemma succ_lepoll_succ_succ: - "[| InfCard(x); n \ nat |] - ==> {y \ Pow(x). y\succ(n)} \ {y \ Pow(x). y\succ(succ(n))}" + "\InfCard(x); n \ nat\ + \ {y \ Pow(x). y\succ(n)} \ {y \ Pow(x). y\succ(succ(n))}" apply (unfold lepoll_def) apply (rule_tac x = "\z \ {y\Pow(x). y\succ(n)}. cons(succ(\(z)), z)" in exI) @@ -159,7 +159,7 @@ done lemma subsets_eqpoll_X: - "[| InfCard(X); n \ nat |] ==> {Y \ Pow(X). Y\succ(n)} \ 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) @@ -172,16 +172,16 @@ done lemma image_vimage_eq: - "[| f \ surj(A,B); y \ B |] ==> f``(converse(f)``y) = y" + "\f \ surj(A,B); y \ B\ \ f``(converse(f)``y) = y" apply (unfold 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" +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}" + "A\B \ {Y \ Pow(A). Y\n}\{Y \ Pow(B). Y\n}" apply (unfold eqpoll_def) apply (erule exE) apply (rule_tac x = "\X \ {Y \ Pow (A) . \f. f \ bij (Y, n) }. f``X" in exI) @@ -196,23 +196,23 @@ apply (fast elim!: bij_is_surj [THEN image_vimage_eq]) done -lemma WO2_imp_ex_Card: "WO2 ==> \a. Card(a) & X\a" +lemma WO2_imp_ex_Card: "WO2 \ \a. Card(a) & X\a" apply (unfold 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)" +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)" +lemma infinite_Card_is_InfCard: "\\Finite(X); Card(X)\ \ InfCard(X)" apply (unfold 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" +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) @@ -220,12 +220,12 @@ 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" +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" + "\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) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC17_AC1.thy --- a/src/ZF/AC/AC17_AC1.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC17_AC1.thy Tue Sep 27 16:51:35 2022 +0100 @@ -15,26 +15,26 @@ (** 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)" +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" +lemma AC0_AC1: "AC0 \ AC1" apply (unfold AC0_def AC1_def) apply (blast intro: AC0_AC1_lemma) done -lemma AC1_AC0: "AC1 ==> AC0" +lemma AC1_AC0: "AC1 \ AC0" by (unfold AC0_def AC1_def, blast) -(**** The proof of AC1 ==> AC17 ****) +(**** The proof of AC1 \ AC17 ****) -lemma AC1_AC17_lemma: "f \ (\X \ Pow(A) - {0}. X) ==> f \ (Pow(A) - {0} -> A)" +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" +lemma AC1_AC17: "AC1 \ AC17" apply (unfold AC1_def AC17_def) apply (rule allI) apply (rule ballI) @@ -48,17 +48,17 @@ done -(**** The proof of AC17 ==> AC1 ****) +(**** 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}. + "\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)" + 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] @@ -70,7 +70,7 @@ (* *********************************************************************** *) lemma not_AC1_imp_ex: - "~AC1 ==> \A. \f \ Pow(A)-{0} -> A. \u \ Pow(A)-{0}. f`u \ u" + "\AC1 \ \A. \f \ Pow(A)-{0} -> A. \u \ Pow(A)-{0}. f`u \ u" apply (unfold AC1_def) apply (erule swap) apply (rule allI) @@ -80,11 +80,11 @@ done lemma AC17_AC1_aux1: - "[| \f \ Pow(x) - {0} -> x. \u \ Pow(x) - {0}. f`u\u; + "\\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" + 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]) @@ -96,22 +96,22 @@ done lemma AC17_AC1_aux2: - "~ (\f \ Pow(x)-{0}->x. x - F(f) = 0) - ==> (\f \ Pow(x)-{0}->x . x - F(f)) + "\ (\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}" + "\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)" + \ \f \ F. f`Q(f) \ Q(f)" by simp -lemma AC17_AC1: "AC17 ==> AC1" +lemma AC17_AC1: "AC17 \ AC1" apply (unfold AC17_def) apply (rule classical) apply (erule not_AC1_imp_ex [THEN exE]) @@ -133,25 +133,25 @@ (* ********************************************************************** - AC1 ==> AC2 ==> AC1 - AC1 ==> AC4 ==> AC3 ==> AC1 - AC4 ==> AC5 ==> AC4 + AC1 \ AC2 \ AC1 + AC1 \ AC4 \ AC3 \ AC1 + AC4 \ AC5 \ AC4 AC1 \ AC6 ************************************************************************* *) (* ********************************************************************** *) -(* AC1 ==> AC2 *) +(* AC1 \ AC2 *) (* ********************************************************************** *) lemma AC1_AC2_aux1: - "[| f:(\X \ A. X); B \ A; 0\A |] ==> {f`B} \ B \ {f`C. C \ A}" + "\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" + "\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" +lemma AC1_AC2: "AC1 \ AC2" apply (unfold AC1_def AC2_def) apply (rule allI) apply (rule impI) @@ -163,14 +163,14 @@ (* ********************************************************************** *) -(* AC2 ==> AC1 *) +(* AC2 \ AC1 *) (* ********************************************************************** *) -lemma AC2_AC1_aux1: "0\A ==> 0 \ {B*{B}. B \ A}" +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" +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) @@ -178,13 +178,13 @@ 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)" + \ (\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" +lemma AC2_AC1: "AC2 \ AC1" apply (unfold AC1_def AC2_def pairwise_disjoint_def) apply (intro allI impI) apply (elim allE impE) @@ -194,13 +194,13 @@ (* ********************************************************************** *) -(* AC1 ==> AC4 *) +(* AC1 \ AC4 *) (* ********************************************************************** *) lemma empty_notin_images: "0 \ {R``{x}. x \ domain(R)}" by blast -lemma AC1_AC4: "AC1 ==> AC4" +lemma AC1_AC4: "AC1 \ AC4" apply (unfold AC1_def AC4_def) apply (intro allI impI) apply (drule spec, drule mp [OF _ empty_notin_images]) @@ -209,19 +209,19 @@ (* ********************************************************************** *) -(* AC4 ==> AC3 *) +(* AC4 \ AC3 *) (* ********************************************************************** *) -lemma AC4_AC3_aux1: "f \ A->B ==> (\z \ A. {z}*f`z) \ A*\(B)" +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)" +lemma AC4_AC3_aux3: "x \ A \ (\z \ A. {z}*f(z))``{x} = f(x)" by fast -lemma AC4_AC3: "AC4 ==> AC3" +lemma AC4_AC3: "AC4 \ AC3" apply (unfold AC3_def AC4_def) apply (intro allI ballI) apply (elim allE impE) @@ -230,25 +230,25 @@ done (* ********************************************************************** *) -(* AC3 ==> AC1 *) +(* AC3 \ AC1 *) (* ********************************************************************** *) lemma AC3_AC1_lemma: - "b\A ==> (\x \ {a \ A. id(A)`a\b}. id(A)`x) = (\x \ A. x)" + "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" +lemma AC3_AC1: "AC3 \ AC1" apply (unfold AC1_def AC3_def) apply (fast intro!: id_type elim: AC3_AC1_lemma [THEN subst]) done (* ********************************************************************** *) -(* AC4 ==> AC5 *) +(* AC4 \ AC5 *) (* ********************************************************************** *) -lemma AC4_AC5: "AC4 ==> AC5" +lemma AC4_AC5: "AC4 \ AC5" apply (unfold range_def AC4_def AC5_def) apply (intro allI ballI) apply (elim allE impE) @@ -262,27 +262,27 @@ (* ********************************************************************** *) -(* AC5 ==> AC4, Rubin & Rubin, p. 11 *) +(* AC5 \ AC4, Rubin & Rubin, p. 11 *) (* ********************************************************************** *) -lemma AC5_AC4_aux1: "R \ A*B ==> (\x \ R. fst(x)) \ R -> A" +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)" +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)" +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})" +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" +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) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC18_AC19.thy --- a/src/ZF/AC/AC18_AC19.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC18_AC19.thy Tue Sep 27 16:51:35 2022 +0100 @@ -1,7 +1,7 @@ (* Title: ZF/AC/AC18_AC19.thy Author: Krzysztof Grabczewski -The proof of AC1 ==> AC18 ==> AC19 ==> AC1 +The proof of AC1 \ AC18 \ AC19 \ AC1 *) theory AC18_AC19 @@ -10,21 +10,21 @@ definition uu :: "i => i" where - "uu(a) == {c \ {0}. c \ a}" + "uu(a) \ {c \ {0}. c \ a}" (* ********************************************************************** *) -(* AC1 ==> AC18 *) +(* AC1 \ AC18 *) (* ********************************************************************** *) lemma PROD_subsets: - "[| f \ (\b \ {P(a). a \ A}. b); \a \ A. P(a)<=Q(a) |] - ==> (\a \ A. f`P(a)) \ (\a \ A. Q(a))" + "\f \ (\b \ {P(a). a \ A}. b); \a \ A. P(a)<=Q(a)\ + \ (\a \ A. f`P(a)) \ (\a \ A. Q(a))" by (rule lam_type, drule apply_type, auto) lemma lemma_AC18: - "[| \A. 0 \ A \ (\f. f \ (\X \ A. X)); A \ 0 |] - ==> (\a \ A. \b \ B(a). X(a, b)) \ + "\\A. 0 \ A \ (\f. f \ (\X \ A. X)); A \ 0\ + \ (\a \ A. \b \ B(a). X(a, b)) \ (\f \ \a \ A. B(a). \a \ A. X(a, f`a))" apply (rule subsetI) apply (erule_tac x = "{{b \ B (a) . x \ X (a,b) }. a \ A}" in allE) @@ -35,14 +35,14 @@ apply (simp, fast elim!: not_emptyE dest: apply_type [OF _ RepFunI]) done -lemma AC1_AC18: "AC1 ==> PROP AC18" +lemma AC1_AC18: "AC1 \ PROP AC18" apply (unfold AC1_def) apply (rule AC18.intro) apply (fast elim!: lemma_AC18 apply_type intro!: equalityI INT_I UN_I) done (* ********************************************************************** *) -(* AC18 ==> AC19 *) +(* AC18 \ AC19 *) (* ********************************************************************** *) theorem (in AC18) AC19 @@ -52,28 +52,28 @@ done (* ********************************************************************** *) -(* AC19 ==> AC1 *) +(* AC19 \ AC1 *) (* ********************************************************************** *) lemma RepRep_conj: - "[| A \ 0; 0 \ A |] ==> {uu(a). a \ A} \ 0 & 0 \ {uu(a). a \ A}" + "\A \ 0; 0 \ A\ \ {uu(a). a \ A} \ 0 & 0 \ {uu(a). a \ A}" apply (unfold uu_def, auto) apply (blast dest!: sym [THEN RepFun_eq_0_iff [THEN iffD1]]) done -lemma lemma1_1: "[|c \ a; x = c \ {0}; x \ a |] ==> x - {0} \ a" +lemma lemma1_1: "\c \ a; x = c \ {0}; x \ a\ \ x - {0} \ a" apply clarify apply (rule subst_elem, assumption) apply (fast elim: notE subst_elem) done lemma lemma1_2: - "[| f`(uu(a)) \ a; f \ (\B \ {uu(a). a \ A}. B); a \ A |] - ==> f`(uu(a))-{0} \ a" + "\f`(uu(a)) \ a; f \ (\B \ {uu(a). a \ A}. B); a \ A\ + \ f`(uu(a))-{0} \ a" apply (unfold uu_def, fast elim!: lemma1_1 dest!: apply_type) done -lemma lemma1: "\f. f \ (\B \ {uu(a). a \ A}. B) ==> \f. f \ (\B \ A. B)" +lemma lemma1: "\f. f \ (\B \ {uu(a). a \ A}. B) \ \f. f \ (\B \ A. B)" apply (erule exE) apply (rule_tac x = "\a\A. if (f` (uu(a)) \ a, f` (uu(a)), f` (uu(a))-{0})" in exI) @@ -81,16 +81,16 @@ apply (simp add: lemma1_2) done -lemma lemma2_1: "a\0 ==> 0 \ (\b \ uu(a). b)" +lemma lemma2_1: "a\0 \ 0 \ (\b \ uu(a). b)" by (unfold uu_def, auto) -lemma lemma2: "[| A\0; 0\A |] ==> (\x \ {uu(a). a \ A}. \b \ x. b) \ 0" +lemma lemma2: "\A\0; 0\A\ \ (\x \ {uu(a). a \ A}. \b \ x. b) \ 0" apply (erule not_emptyE) apply (rule_tac a = 0 in not_emptyI) apply (fast intro!: lemma2_1) done -lemma AC19_AC1: "AC19 ==> AC1" +lemma AC19_AC1: "AC19 \ AC1" apply (unfold AC19_def AC1_def, clarify) apply (case_tac "A=0", force) apply (erule_tac x = "{uu (a) . a \ A}" in allE) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC7_AC9.thy --- a/src/ZF/AC/AC7_AC9.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC7_AC9.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,16 +10,16 @@ begin (* ********************************************************************** *) -(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1 *) +(* 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" +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. + "C \ A \ (\g \ (nat->\(A))*C. (\n \ nat. if(n=0, snd(g), fst(g)`(n #- 1)))) \ inj((nat->\(A))*C, (nat->\(A)) ) " apply (unfold inj_def) @@ -31,7 +31,7 @@ done lemma Sigma_fun_space_eqpoll: - "[| C \ A; 0\A |] ==> (nat->\(A)) * C \ (nat->\(A))" + "\C \ A; 0\A\ \ (nat->\(A)) * C \ (nat->\(A))" apply (rule eqpollI) apply (simp add: lepoll_def) apply (fast intro!: inj_lemma) @@ -41,34 +41,34 @@ (* ********************************************************************** *) -(* AC6 ==> AC7 *) +(* AC6 \ AC7 *) (* ********************************************************************** *) -lemma AC6_AC7: "AC6 ==> AC7" +lemma AC6_AC7: "AC6 \ AC7" by (unfold AC6_def AC7_def, blast) (* ********************************************************************** *) -(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8 *) +(* 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)" +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)" + "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" + "(\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}" +lemma AC7_AC6_lemma2: "0 \ A \ 0 \ {(nat -> \(A)) * C. C \ A}" by (blast dest: Sigma_fun_space_not0) -lemma AC7_AC6: "AC7 ==> AC6" +lemma AC7_AC6: "AC7 \ AC6" apply (unfold AC6_def AC7_def) apply (rule allI) apply (rule impI) @@ -82,38 +82,38 @@ (* ********************************************************************** *) -(* AC1 ==> AC8 *) +(* AC1 \ AC8 *) (* ********************************************************************** *) lemma AC1_AC8_lemma1: "\B \ A. \B1 B2. B= & B1 \ B2 - ==> 0 \ { bij(fst(B),snd(B)). B \ A }" + \ 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)" + "\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" +lemma AC1_AC8: "AC1 \ AC8" apply (unfold AC1_def AC8_def) apply (fast dest: AC1_AC8_lemma1 AC1_AC8_lemma2) done (* ********************************************************************** *) -(* AC8 ==> AC9 *) +(* AC8 \ AC9 *) (* - this proof replaces the following two from Rubin & Rubin: *) -(* AC8 ==> AC1 and AC1 ==> AC9 *) +(* AC8 \ AC1 and AC1 \ AC9 *) (* ********************************************************************** *) lemma AC8_AC9_lemma: "\B1 \ A. \B2 \ A. B1 \ B2 - ==> \B \ A*A. \B1 B2. B= & B1 \ B2" + \ \B \ A*A. \B1 B2. B= & B1 \ B2" by fast -lemma AC8_AC9: "AC8 ==> AC9" +lemma AC8_AC9: "AC8 \ AC9" apply (unfold AC8_def AC9_def) apply (intro allI impI) apply (erule allE) @@ -122,7 +122,7 @@ (* ********************************************************************** *) -(* AC9 ==> AC1 *) +(* 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 *) @@ -134,15 +134,15 @@ prod_commute_eqpoll) lemma nat_lepoll_lemma: - "[|0 \ A; B \ A|] ==> nat \ ((nat \ \(A)) \ B) \ nat" + "\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; + "\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" + 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] @@ -152,13 +152,13 @@ "\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` \ bij(B1, B2) - ==> (\B \ A. snd(fst((f`)`0))) \ (\X \ A. X)" + \ (\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" +lemma AC9_AC1: "AC9 \ AC1" apply (unfold AC1_def AC9_def) apply (intro allI impI) apply (erule allE) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/AC_Equiv.thy --- a/src/ZF/AC/AC_Equiv.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/AC_Equiv.thy Tue Sep 27 16:51:35 2022 +0100 @@ -18,106 +18,106 @@ (* Well Ordering Theorems *) definition - "WO1 == \A. \R. well_ord(A,R)" + "WO1 \ \A. \R. well_ord(A,R)" definition - "WO2 == \A. \a. Ord(a) & A\a" + "WO2 \ \A. \a. Ord(a) & A\a" definition - "WO3 == \A. \a. Ord(a) & (\b. b \ a & A\b)" + "WO3 \ \A. \a. Ord(a) & (\b. b \ a & A\b)" definition - "WO4(m) == \A. \a f. Ord(a) & domain(f)=a & + "WO4(m) \ \A. \a f. Ord(a) & domain(f)=a & (\bb m)" definition - "WO5 == \m \ nat. 1\m & WO4(m)" + "WO5 \ \m \ nat. 1\m & WO4(m)" definition - "WO6 == \A. \m \ nat. 1\m & (\a f. Ord(a) & domain(f)=a + "WO6 \ \A. \m \ nat. 1\m & (\a f. Ord(a) & domain(f)=a & (\bb m))" definition - "WO7 == \A. Finite(A) \ (\R. well_ord(A,R) \ well_ord(A,converse(R)))" + "WO7 \ \A. Finite(A) \ (\R. well_ord(A,R) \ well_ord(A,converse(R)))" definition - "WO8 == \A. (\f. f \ (\X \ A. X)) \ (\R. well_ord(A,R))" + "WO8 \ \A. (\f. f \ (\X \ A. X)) \ (\R. well_ord(A,R))" definition (* Auxiliary concepts needed below *) pairwise_disjoint :: "i => o" where - "pairwise_disjoint(A) == \A1 \ A. \A2 \ A. A1 \ A2 \ 0 \ A1=A2" + "pairwise_disjoint(A) \ \A1 \ A. \A2 \ A. A1 \ A2 \ 0 \ A1=A2" definition sets_of_size_between :: "[i, i, i] => o" where - "sets_of_size_between(A,m,n) == \B \ A. m \ B & B \ n" + "sets_of_size_between(A,m,n) \ \B \ A. m \ B & B \ n" (* Axioms of Choice *) definition - "AC0 == \A. \f. f \ (\X \ Pow(A)-{0}. X)" + "AC0 \ \A. \f. f \ (\X \ Pow(A)-{0}. X)" definition - "AC1 == \A. 0\A \ (\f. f \ (\X \ A. X))" + "AC1 \ \A. 0\A \ (\f. f \ (\X \ A. X))" definition - "AC2 == \A. 0\A & pairwise_disjoint(A) + "AC2 \ \A. 0\A & pairwise_disjoint(A) \ (\C. \B \ A. \y. B \ C = {y})" definition - "AC3 == \A B. \f \ A->B. \g. g \ (\x \ {a \ A. f`a\0}. f`x)" + "AC3 \ \A B. \f \ A->B. \g. g \ (\x \ {a \ A. f`a\0}. f`x)" definition - "AC4 == \R A B. (R \ A*B \ (\f. f \ (\x \ domain(R). R``{x})))" + "AC4 \ \R A B. (R \ A*B \ (\f. f \ (\x \ domain(R). R``{x})))" definition - "AC5 == \A B. \f \ A->B. \g \ range(f)->A. \x \ domain(g). f`(g`x) = x" + "AC5 \ \A B. \f \ A->B. \g \ range(f)->A. \x \ domain(g). f`(g`x) = x" definition - "AC6 == \A. 0\A \ (\B \ A. B)\0" + "AC6 \ \A. 0\A \ (\B \ A. B)\0" definition - "AC7 == \A. 0\A & (\B1 \ A. \B2 \ A. B1\B2) \ (\B \ A. B) \ 0" + "AC7 \ \A. 0\A & (\B1 \ A. \B2 \ A. B1\B2) \ (\B \ A. B) \ 0" definition - "AC8 == \A. (\B \ A. \B1 B2. B= & B1\B2) + "AC8 \ \A. (\B \ A. \B1 B2. B= & B1\B2) \ (\f. \B \ A. f`B \ bij(fst(B),snd(B)))" definition - "AC9 == \A. (\B1 \ A. \B2 \ A. B1\B2) \ + "AC9 \ \A. (\B1 \ A. \B2 \ A. B1\B2) \ (\f. \B1 \ A. \B2 \ A. f` \ bij(B1,B2))" definition - "AC10(n) == \A. (\B \ A. ~Finite(B)) \ + "AC10(n) \ \A. (\B \ A. \Finite(B)) \ (\f. \B \ A. (pairwise_disjoint(f`B) & sets_of_size_between(f`B, 2, succ(n)) & \(f`B)=B))" definition - "AC11 == \n \ nat. 1\n & AC10(n)" + "AC11 \ \n \ nat. 1\n & AC10(n)" definition - "AC12 == \A. (\B \ A. ~Finite(B)) \ + "AC12 \ \A. (\B \ A. \Finite(B)) \ (\n \ nat. 1\n & (\f. \B \ A. (pairwise_disjoint(f`B) & sets_of_size_between(f`B, 2, succ(n)) & \(f`B)=B)))" definition - "AC13(m) == \A. 0\A \ (\f. \B \ A. f`B\0 & f`B \ B & f`B \ m)" + "AC13(m) \ \A. 0\A \ (\f. \B \ A. f`B\0 & f`B \ B & f`B \ m)" definition - "AC14 == \m \ nat. 1\m & AC13(m)" + "AC14 \ \m \ nat. 1\m & AC13(m)" definition - "AC15 == \A. 0\A \ + "AC15 \ \A. 0\A \ (\m \ nat. 1\m & (\f. \B \ A. f`B\0 & f`B \ B & f`B \ m))" definition - "AC16(n, k) == - \A. ~Finite(A) \ + "AC16(n, k) \ + \A. \Finite(A) \ (\T. T \ {X \ Pow(A). X\succ(n)} & (\X \ {X \ Pow(A). X\succ(k)}. \! Y. Y \ T & X \ Y))" definition - "AC17 == \A. \g \ (Pow(A)-{0} -> A) -> Pow(A)-{0}. + "AC17 \ \A. \g \ (Pow(A)-{0} -> A) -> Pow(A)-{0}. \f \ Pow(A)-{0} -> A. f`(g`f) \ g`f" locale AC18 = @@ -127,7 +127,7 @@ \ \AC18 cannot be expressed within the object-logic\ definition - "AC19 == \A. A\0 & 0\A \ ((\a \ A. \b \ a. b) = + "AC19 \ \A. A\0 & 0\A \ ((\a \ A. \b \ a. b) = (\f \ (\B \ A. B). \a \ A. f`a))" @@ -148,7 +148,7 @@ (* used only in Hartog.ML *) lemma ordertype_Int: - "well_ord(A,r) ==> ordertype(A, r \ A*A) = ordertype(A,r)" + "well_ord(A,r) \ ordertype(A, r \ A*A) = ordertype(A,r)" apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst]) apply (erule id_bij [THEN bij_ordertype_vimage]) done @@ -159,29 +159,29 @@ done lemma inj_strengthen_type: - "[| f \ inj(A, B); !!a. a \ A ==> f`a \ C |] ==> f \ inj(A,C)" + "\f \ inj(A, B); \a. a \ A \ f`a \ C\ \ f \ inj(A,C)" by (unfold inj_def, blast intro: Pi_type) (* ********************************************************************** *) (* Another elimination rule for \! *) (* ********************************************************************** *) -lemma ex1_two_eq: "[| \! x. P(x); P(x); P(y) |] ==> x=y" +lemma ex1_two_eq: "\\! x. P(x); P(x); P(y)\ \ x=y" by blast (* ********************************************************************** *) -(* Lemmas used in the proofs like WO? ==> AC? *) +(* Lemmas used in the proofs like WO? \ AC? *) (* ********************************************************************** *) lemma first_in_B: - "[| well_ord(\(A),r); 0 \ A; B \ A |] ==> (THE b. first(b,B,r)) \ B" + "\well_ord(\(A),r); 0 \ A; B \ A\ \ (THE b. first(b,B,r)) \ B" by (blast dest!: well_ord_imp_ex1_first [THEN theI, THEN first_def [THEN def_imp_iff, THEN iffD1]]) -lemma ex_choice_fun: "[| well_ord(\(A), R); 0 \ A |] ==> \f. f \ (\X \ A. X)" +lemma ex_choice_fun: "\well_ord(\(A), R); 0 \ A\ \ \f. f \ (\X \ A. X)" by (fast elim!: first_in_B intro!: lam_type) -lemma ex_choice_fun_Pow: "well_ord(A, R) ==> \f. f \ (\X \ Pow(A)-{0}. X)" +lemma ex_choice_fun_Pow: "well_ord(A, R) \ \f. f \ (\X \ Pow(A)-{0}. X)" by (fast elim!: well_ord_subset [THEN ex_choice_fun]) @@ -193,7 +193,7 @@ (*Using AC we could trivially prove, for all u, domain(u) \ u*) lemma lepoll_m_imp_domain_lepoll_m: - "[| m \ nat; u \ m |] ==> domain(u) \ m" + "\m \ nat; u \ m\ \ domain(u) \ m" apply (unfold lepoll_def) apply (erule exE) apply (rule_tac x = "\x \ domain(u). \ i. \y. \ u & f` = i" @@ -208,7 +208,7 @@ done lemma rel_domain_ex1: - "[| succ(m) \ domain(r); r \ succ(m); m \ nat |] ==> function(r)" + "\succ(m) \ domain(r); r \ succ(m); m \ nat\ \ function(r)" apply (unfold function_def, safe) apply (rule ccontr) apply (fast elim!: lepoll_trans [THEN succ_lepoll_natE] @@ -217,8 +217,8 @@ done lemma rel_is_fun: - "[| succ(m) \ domain(r); r \ succ(m); m \ nat; - r \ A*B; A=domain(r) |] ==> r \ A->B" + "\succ(m) \ domain(r); r \ succ(m); m \ nat; + r \ A*B; A=domain(r)\ \ r \ A->B" by (simp add: Pi_iff rel_domain_ex1) end diff -r f2094906e491 -r e44d86131648 src/ZF/AC/Cardinal_aux.thy --- a/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 16:51:35 2022 +0100 @@ -6,7 +6,7 @@ theory Cardinal_aux imports AC_Equiv begin -lemma Diff_lepoll: "[| A \ succ(m); B \ A; B\0 |] ==> A-B \ m" +lemma Diff_lepoll: "\A \ succ(m); B \ A; B\0\ \ A-B \ m" apply (rule not_emptyE, assumption) apply (blast intro: lepoll_trans [OF subset_imp_lepoll Diff_sing_lepoll]) done @@ -20,14 +20,14 @@ (* j=|A| *) lemma lepoll_imp_ex_le_eqpoll: - "[| A \ i; Ord(i) |] ==> \j. j \ i & A \ j" + "\A \ i; Ord(i)\ \ \j. j \ i & A \ j" by (blast intro!: lepoll_cardinal_le well_ord_Memrel well_ord_cardinal_eqpoll [THEN eqpoll_sym] dest: lepoll_well_ord) (* j=|A| *) lemma lesspoll_imp_ex_lt_eqpoll: - "[| A \ i; Ord(i) |] ==> \j. j j" + "\A \ i; Ord(i)\ \ \j. j j" by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE) lemma Un_eqpoll_Inf_Ord: @@ -70,7 +70,7 @@ (*Finally we reach this result. Surely there's a simpler proof?*) lemma Un_lepoll_Inf_Ord: - "[| A \ i; B \ i; ~Finite(i); Ord(i) |] ==> A \ B \ i" + "\A \ i; B \ i; \Finite(i); Ord(i)\ \ A \ B \ i" apply (rule_tac A1 = i and C1 = i in ex_eqpoll_disjoint [THEN exE]) apply (erule conjE) apply (drule lepoll_trans) @@ -79,7 +79,7 @@ apply (blast intro: eqpoll_refl Un_eqpoll_Inf_Ord eqpoll_imp_lepoll) done -lemma Least_in_Ord: "[| P(i); i \ j; Ord(j) |] ==> (\ i. P(i)) \ j" +lemma Least_in_Ord: "\P(i); i \ j; Ord(j)\ \ (\ i. P(i)) \ j" apply (erule Least_le [THEN leE]) apply (erule Ord_in_Ord, assumption) apply (erule ltE) @@ -88,8 +88,8 @@ done lemma Diff_first_lepoll: - "[| well_ord(x,r); y \ x; y \ succ(n); n \ nat |] - ==> y - {THE b. first(b,y,r)} \ n" + "\well_ord(x,r); y \ x; y \ succ(n); n \ nat\ + \ y - {THE b. first(b,y,r)} \ n" apply (case_tac "y=0", simp add: empty_lepollI) apply (fast intro!: Diff_sing_lepoll the_first_in) done @@ -98,7 +98,7 @@ "(\x \ X. P(x)) \ (\x \ X. P(x)-Q(x)) \ (\x \ X. Q(x))" by blast -lemma UN_sing_lepoll: "Ord(a) ==> (\x \ a. {P(x)}) \ a" +lemma UN_sing_lepoll: "Ord(a) \ (\x \ a. {P(x)}) \ a" apply (unfold lepoll_def) apply (rule_tac x = "\z \ (\x \ a. {P (x) }) . (\ i. P (i) =z) " in exI) apply (rule_tac d = "%z. P (z) " in lam_injective) @@ -107,8 +107,8 @@ done lemma UN_fun_lepoll_lemma [rule_format]: - "[| well_ord(T, R); ~Finite(a); Ord(a); n \ nat |] - ==> \f. (\b \ a. f`b \ n & f`b \ T) \ (\b \ a. f`b) \ a" + "\well_ord(T, R); \Finite(a); Ord(a); n \ nat\ + \ \f. (\b \ a. f`b \ n & f`b \ T) \ (\b \ a. f`b) \ a" apply (induct_tac "n") apply (rule allI) apply (rule impI) @@ -126,21 +126,21 @@ done lemma UN_fun_lepoll: - "[| \b \ a. f`b \ n & f`b \ T; well_ord(T, R); - ~Finite(a); Ord(a); n \ nat |] ==> (\b \ a. f`b) \ a" + "\\b \ a. f`b \ n & f`b \ T; well_ord(T, R); + \Finite(a); Ord(a); n \ nat\ \ (\b \ a. f`b) \ a" by (blast intro: UN_fun_lepoll_lemma) lemma UN_lepoll: - "[| \b \ a. F(b) \ n & F(b) \ T; well_ord(T, R); - ~Finite(a); Ord(a); n \ nat |] - ==> (\b \ a. F(b)) \ a" + "\\b \ a. F(b) \ n & F(b) \ T; well_ord(T, R); + \Finite(a); Ord(a); n \ nat\ + \ (\b \ a. F(b)) \ a" apply (rule rev_mp) apply (rule_tac f="\b \ a. F (b)" in UN_fun_lepoll) apply auto done lemma UN_eq_UN_Diffs: - "Ord(a) ==> (\b \ a. F(b)) = (\b \ a. F(b) - (\c \ b. F(c)))" + "Ord(a) \ (\b \ a. F(b)) = (\b \ a. F(b) - (\c \ b. F(c)))" apply (rule equalityI) prefer 2 apply fast apply (rule subsetI) @@ -153,7 +153,7 @@ done lemma lepoll_imp_eqpoll_subset: - "a \ X ==> \Y. Y \ X & a \ Y" + "a \ X \ \Y. Y \ X & a \ Y" apply (unfold lepoll_def eqpoll_def, clarify) apply (blast intro: restrict_bij dest: inj_is_fun [THEN fun_is_rel, THEN image_subset]) @@ -164,7 +164,7 @@ (* ********************************************************************** *) lemma Diff_lesspoll_eqpoll_Card_lemma: - "[| A\a; ~Finite(a); Card(a); B \ a; A-B \ a |] ==> P" + "\A\a; \Finite(a); Card(a); B \ a; A-B \ a\ \ P" apply (elim lesspoll_imp_ex_lt_eqpoll [THEN exE] Card_is_Ord conjE) apply (frule_tac j=xa in Un_upper1_le [OF lt_Ord lt_Ord], assumption) apply (frule_tac j=xa in Un_upper2_le [OF lt_Ord lt_Ord], assumption) @@ -186,7 +186,7 @@ done lemma Diff_lesspoll_eqpoll_Card: - "[| A \ a; ~Finite(a); Card(a); B \ a |] ==> A - B \ a" + "\A \ a; \Finite(a); Card(a); B \ a\ \ A - B \ a" apply (rule ccontr) apply (rule Diff_lesspoll_eqpoll_Card_lemma, (assumption+)) apply (blast intro: lesspoll_def [THEN def_imp_iff, THEN iffD2] diff -r f2094906e491 -r e44d86131648 src/ZF/AC/DC.thy --- a/src/ZF/AC/DC.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/DC.thy Tue Sep 27 16:51:35 2022 +0100 @@ -8,7 +8,7 @@ imports AC_Equiv Hartog Cardinal_aux begin -lemma RepFun_lepoll: "Ord(a) ==> {P(b). b \ a} \ a" +lemma RepFun_lepoll: "Ord(a) \ {P(b). b \ a} \ a" apply (unfold 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) @@ -18,7 +18,7 @@ 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" +lemma image_Ord_lepoll: "\f \ X->Y; Ord(X)\ \ f``X \ X" apply (unfold 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) @@ -29,70 +29,70 @@ done lemma range_subset_domain: - "[| R \ X*X; !!g. g \ X ==> \u. \ R |] - ==> range(R) \ domain(R)" + "\R \ X*X; \g. g \ X \ \u. \ R\ + \ range(R) \ domain(R)" by blast -lemma cons_fun_type: "g \ n->X ==> cons(, g) \ succ(n) -> cons(x, X)" +lemma cons_fun_type: "g \ n->X \ cons(, g) \ succ(n) -> cons(x, X)" apply (unfold succ_def) apply (fast intro!: fun_extend elim!: mem_irrefl) done lemma cons_fun_type2: - "[| g \ n->X; x \ X |] ==> cons(, g) \ succ(n) -> X" + "\g \ n->X; x \ X\ \ cons(, g) \ succ(n) -> X" by (erule cons_absorb [THEN subst], erule cons_fun_type) -lemma cons_image_n: "n \ nat ==> cons(, g)``n = g``n" +lemma cons_image_n: "n \ nat \ cons(, g)``n = g``n" by (fast elim!: mem_irrefl) -lemma cons_val_n: "g \ n->X ==> cons(, g)`n = x" +lemma cons_val_n: "g \ n->X \ cons(, g)`n = x" by (fast intro!: apply_equality elim!: cons_fun_type) -lemma cons_image_k: "k \ n ==> cons(, g)``k = g``k" +lemma cons_image_k: "k \ n \ cons(, g)``k = g``k" by (fast elim: mem_asym) -lemma cons_val_k: "[| k \ n; g \ n->X |] ==> cons(, g)`k = g`k" +lemma cons_val_k: "\k \ n; g \ n->X\ \ cons(, 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(, f)) = succ(x)" +lemma domain_cons_eq_succ: "domain(f)=x \ domain(cons(, f)) = succ(x)" by (simp add: domain_cons succ_def) -lemma restrict_cons_eq: "g \ n->X ==> restrict(cons(, g), n) = g" +lemma restrict_cons_eq: "g \ n->X \ restrict(cons(, 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)" +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" + "\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" +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)" +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 & + "DC(a) \ \X R. R \ Pow(X)*X & (\Y \ Pow(X). Y \ a \ (\x \ X. \ R)) \ (\f \ a->X. \b \ R)" definition DC0 :: o where - "DC0 == \A B R. R \ A*B & R\0 & range(R) \ domain(R) + "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) == + "ff(b, X, Q, R) \ transrec(b, %c r. THE x. first(x, {x \ X. \ R}, Q))" @@ -101,13 +101,13 @@ assumes all_ex: "\Y \ Pow(X). Y \ nat \ (\x \ X. \ R)" - defines XX_def: "XX == (\n \ nat. {f \ n->X. \k \ n. \ R})" - and RR_def: "RR == {:XX*XX. domain(z2)=succ(domain(z1)) + defines XX_def: "XX \ (\n \ nat. {f \ n->X. \k \ n. \ R})" + and RR_def: "RR \ {:XX*XX. domain(z2)=succ(domain(z1)) & restrict(z2, domain(z1)) = z1}" begin (* ********************************************************************** *) -(* DC ==> DC(omega) *) +(* DC \ DC(omega) *) (* *) (* The scheme of the proof: *) (* *) @@ -169,8 +169,8 @@ done lemma lemma2: - "[| \n \ nat. \ RR; f \ nat -> XX; n \ nat |] - ==> \k \ nat. f`succ(n) \ k -> X & n \ k + "\\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]) @@ -198,8 +198,8 @@ done lemma lemma3_1: - "[| \n \ nat. \ RR; f \ nat -> XX; m \ nat |] - ==> {f`succ(x)`x. x \ m} = {f`succ(m)`x. x \ m}" + "\\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) @@ -221,8 +221,8 @@ done lemma lemma3: - "[| \n \ nat. \ RR; f \ nat -> XX; m \ nat |] - ==> (\x \ nat. f`succ(x)`x) `` m = f`succ(m)``m" + "\\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]) @@ -234,7 +234,7 @@ end -theorem DC0_imp_DC_nat: "DC0 ==> DC(nat)" +theorem DC0_imp_DC_nat: "DC0 \ DC(nat)" apply (unfold DC_def DC0_def, clarify) apply (elim allE) apply (erule impE) @@ -256,7 +256,7 @@ (* ************************************************************************ - DC(omega) ==> DC + DC(omega) \ DC The scheme of the proof: @@ -269,7 +269,7 @@ RR = {: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)) & + (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) & (\f \ z1. restrict(g, domain(f)) = f)) & z2={<0,x>})} @@ -287,7 +287,7 @@ ************************************************************************* *) lemma singleton_in_funs: - "x \ X ==> {<0,x>} \ + "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] @@ -297,16 +297,16 @@ locale imp_DC0 = fixes XX and RR and x and R and f and allRR - defines XX_def: "XX == (\n \ nat. + defines XX_def: "XX \ (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R})" and RR_def: - "RR == {:Fin(XX)*XX. + "RR \ {: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)) + | (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) & (\f \ z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}" and allRR_def: - "allRR == \b \b \ {\Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) & (\f \ z1. domain(f)) = b @@ -314,8 +314,8 @@ begin lemma lemma4: - "[| range(R) \ domain(R); x \ domain(R) |] - ==> RR \ Pow(XX)*XX & + "\range(R) \ domain(R); x \ domain(R)\ + \ RR \ Pow(XX)*XX & (\Y \ Pow(XX). Y \ nat \ (\x \ XX. :RR))" apply (rule conjI) apply (force dest!: FinD [THEN PowI] simp add: RR_def) @@ -332,34 +332,34 @@ 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))" + "\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)" + "\(\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" + "\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)" + "\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)" + "\domain(f`n) = succ(k); f \ nat -> XX; n \ nat\ + \ f`n \ succ(k) -> domain(R)" apply (unfold 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" + "\h \ nat -> XX; domain(h`n) = succ(k); n \ nat\ + \ \i \ k. \ R" apply (unfold XX_def) apply (drule apply_type, assumption) apply (elim UN_E CollectE) @@ -367,8 +367,8 @@ done lemma restrict_cons_eq_restrict: - "[| restrict(h, domain(u))=u; h \ n->X; domain(u) \ n |] - ==> restrict(cons(, h), domain(u)) = u" + "\restrict(h, domain(u))=u; h \ n->X; domain(u) \ n\ + \ restrict(cons(, h), domain(u)) = u" apply (unfold restrict_def) apply (simp add: restrict_def Pi_iff) apply (erule sym [THEN trans, symmetric]) @@ -376,11 +376,11 @@ done lemma all_in_image_restrict_eq: - "[| \x \ f``n. restrict(f`n, domain(x))=x; + "\\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" + (\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+) @@ -390,16 +390,16 @@ done lemma simplify_recursion: - "[| \b \ RR; - f \ nat -> XX; range(R) \ domain(R); x \ domain(R)|] - ==> allRR" + "\\b \ RR; + f \ nat -> XX; range(R) \ domain(R); x \ domain(R)\ + \ allRR" apply (unfold 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 \~ *) +(*prevent simplification of \\ to \\ *) apply (simp only: separation split) apply (drule_tac x="succ(xa)" in ospec, blast intro: ltI) apply (elim conjE disjE) @@ -435,8 +435,8 @@ lemma lemma2: - "[| allRR; f \ nat->XX; range(R) \ domain(R); x \ domain(R); n \ nat |] - ==> f`n \ succ(n) -> domain(R) & (\i \ n. :R)" + "\allRR; f \ nat->XX; range(R) \ domain(R); x \ domain(R); n \ nat\ + \ f`n \ succ(n) -> domain(R) & (\i \ n. :R)" apply (unfold allRR_def) apply (drule ospec) apply (erule ltI [OF _ Ord_nat]) @@ -448,8 +448,8 @@ done lemma lemma3: - "[| allRR; f \ nat->XX; n\nat; range(R) \ domain(R); x \ domain(R) |] - ==> f`n`n = f`succ(n)`n" + "\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+) apply (unfold allRR_def) apply (drule ospec) @@ -462,7 +462,7 @@ end -theorem DC_nat_imp_DC0: "DC(nat) ==> DC0" +theorem DC_nat_imp_DC0: "DC(nat) \ DC0" apply (unfold DC_def DC0_def) apply (intro allI impI) apply (erule asm_rl conjE ex_in_domain [THEN exE] allE)+ @@ -479,13 +479,13 @@ done (* ********************************************************************** *) -(* \K. Card(K) \ DC(K) ==> WO3 *) +(* \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)" + "\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+) @@ -493,17 +493,17 @@ apply (atomize, blast dest: not_sym) done -lemma value_in_image: "[| f \ X->Y; A \ X; a \ A |] ==> f`a \ f``A" +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" +lemma lesspoll_lemma: "\\ A \ B; C \ B\ \ A - C \ 0" apply (unfold 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" +theorem DC_WO3: "(\K. Card(K) \ DC(K)) \ WO3" apply (unfold DC_def WO3_def) apply (rule allI) apply (case_tac "A \ Hartog (A)") @@ -527,17 +527,17 @@ done (* ********************************************************************** *) -(* WO1 ==> \K. Card(K) \ DC(K) *) +(* 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" + "\\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" + "\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+) @@ -549,16 +549,16 @@ by (fast intro!: lam_type RepFunI) lemma lemmaX: - "[| \Y \ Pow(X). Y \ K \ (\x \ X. \ R); - b \ K; Z \ Pow(X); Z \ K |] - ==> {x \ X. \ R} \ 0" + "\\Y \ Pow(X). Y \ K \ (\x \ X. \ R); + b \ K; Z \ Pow(X); Z \ K\ + \ {x \ X. \ R} \ 0" by blast lemma WO1_DC_lemma: - "[| Card(K); well_ord(X,Q); - \Y \ Pow(X). Y \ K \ (\x \ X. \ R); b \ K |] - ==> ff(b, X, Q, R) \ {x \ X. <(\c \ b. ff(c, X, Q, R))``b, x> \ R}" + "\Card(K); well_ord(X,Q); + \Y \ Pow(X). Y \ K \ (\x \ 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+) @@ -571,7 +571,7 @@ apply (blast intro: lesspoll_trans1 in_Card_imp_lesspoll RepFun_lepoll) done -theorem WO1_DC_Card: "WO1 ==> \K. Card(K) \ DC(K)" +theorem WO1_DC_Card: "WO1 \ \K. Card(K) \ DC(K)" apply (unfold DC_def WO1_def) apply (rule allI impI)+ apply (erule allE exE conjE)+ diff -r f2094906e491 -r e44d86131648 src/ZF/AC/HH.thy --- a/src/ZF/AC/HH.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/HH.thy Tue Sep 27 16:51:35 2022 +0100 @@ -2,9 +2,9 @@ Author: Krzysztof Grabczewski Some properties of the recursive definition of HH used in the proofs of - AC17 ==> AC1 - AC1 ==> WO2 - AC15 ==> WO6 + AC17 \ AC1 + AC1 \ WO2 + AC15 \ WO6 *) theory HH @@ -13,7 +13,7 @@ definition HH :: "[i, i, i] => i" where - "HH(f,x,a) == transrec(a, %b r. let z = x - (\c \ b. r`c) + "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\ @@ -29,26 +29,26 @@ done lemma subset_imp_Diff_eq: - "B \ A ==> X-(\a \ A. P(a)) = X-(\a \ A-B. P(a))-(\b \ B. P(b))" + "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 | bc \ a-b; b \ c=b | bA ==> P(y) = {x}) ==> x - (\y \ A. P(y)) = x" +lemma Diff_UN_eq_self: "(\y. 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)" + \ 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}" +lemma HH_is_x_gt_too: "\HH(f,x,b)={x}; b \ HH(f,x,a)={x}" apply (rule_tac P = "b Pow(x)-{0}; b HH(f,x,b) \ Pow(x)-{0}" + "\HH(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) @@ -72,7 +72,7 @@ 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}" + "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]) @@ -81,7 +81,7 @@ done lemma HH_eq_arg_lt: - "[| HH(f,x,v)=HH(f,x,w); HH(f,x,v) \ Pow(x)-{0}; v \ w |] ==> P" + "\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) @@ -89,7 +89,7 @@ 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" + "\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) @@ -98,7 +98,7 @@ done lemma HH_subset_x_imp_lepoll: - "[| HH(f, x, i) \ Pow(x)-{0}; Ord(i) |] ==> i \ Pow(x)-{0}" + "\HH(f, x, i) \ Pow(x)-{0}; Ord(i)\ \ i \ Pow(x)-{0}" apply (unfold lepoll_def inj_def) apply (rule_tac x = "\j \ i. HH (f, x, j) " in exI) apply (simp (no_asm_simp)) @@ -120,13 +120,13 @@ 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}" + "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\ +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)) @@ -138,30 +138,30 @@ 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)) + \ (\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})" + "\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" +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}" + \ 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" + \ 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) @@ -176,14 +176,14 @@ 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)))" + "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}" + "\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) @@ -192,9 +192,9 @@ 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)) + "\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})" apply (unfold bij_def) apply (fast intro!: lam_Least_HH_inj lam_surj_sing f_sing_imp_HH_sing) @@ -202,7 +202,7 @@ lemma lam_singI: "f \ (\X \ Pow(x)-{0}. F(X)) - ==> (\X \ Pow(x)-{0}. {f`X}) \ (\X \ Pow(x)-{0}. {{z}. z \ 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) @@ -213,7 +213,7 @@ lam_sing_bij [THEN bij_converse_bij]] -subsection\The proof of AC1 ==> WO2\ +subsection\The proof of AC1 \ WO2\ (*Establishing the existence of a bijection, namely converse @@ -225,7 +225,7 @@ lemma bijection: "f \ (\X \ Pow(x) - {0}. X) - ==> \g. g \ bij(x, \ i. HH(\X \ Pow(x)-{0}. {f`X}, x, i) = {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) @@ -234,7 +234,7 @@ apply (fast intro: Pi_weaken_type) done -lemma AC1_WO2: "AC1 ==> WO2" +lemma AC1_WO2: "AC1 \ WO2" apply (unfold AC1_def WO2_def eqpoll_def) apply (intro allI) apply (drule_tac x = "Pow(A) - {0}" in spec) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/Hartog.thy --- a/src/ZF/AC/Hartog.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/Hartog.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,16 +10,16 @@ definition Hartog :: "i => i" where - "Hartog(X) == \ i. ~ i \ X" + "Hartog(X) \ \ i. \ i \ X" -lemma Ords_in_set: "\a. Ord(a) \ a \ X ==> P" +lemma Ords_in_set: "\a. Ord(a) \ a \ X \ P" apply (rule_tac X = "{y \ X. Ord (y) }" in ON_class [elim_format]) apply fast done lemma Ord_lepoll_imp_ex_well_ord: - "[| Ord(a); a \ X |] - ==> \Y. Y \ X & (\R. well_ord(Y,R) & ordertype(Y,R)=a)" + "\Ord(a); a \ X\ + \ \Y. Y \ X & (\R. well_ord(Y,R) & ordertype(Y,R)=a)" apply (unfold lepoll_def) apply (erule exE) apply (intro exI conjI) @@ -36,14 +36,14 @@ done lemma Ord_lepoll_imp_eq_ordertype: - "[| Ord(a); a \ X |] ==> \Y. Y \ X & (\R. R \ X*X & ordertype(Y,R)=a)" + "\Ord(a); a \ X\ \ \Y. Y \ X & (\R. R \ X*X & ordertype(Y,R)=a)" apply (drule Ord_lepoll_imp_ex_well_ord, assumption, clarify) apply (intro exI conjI) apply (erule_tac [3] ordertype_Int, auto) done lemma Ords_lepoll_set_lemma: - "(\a. Ord(a) \ a \ X) ==> + "(\a. Ord(a) \ a \ X) \ \a. Ord(a) \ a \ {b. Z \ Pow(X)*Pow(X*X), \Y R. Z= & ordertype(Y,R)=b}" apply (intro allI impI) @@ -51,15 +51,15 @@ apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) done -lemma Ords_lepoll_set: "\a. Ord(a) \ a \ X ==> P" +lemma Ords_lepoll_set: "\a. Ord(a) \ a \ X \ P" by (erule Ords_lepoll_set_lemma [THEN Ords_in_set]) -lemma ex_Ord_not_lepoll: "\a. Ord(a) & ~a \ X" +lemma ex_Ord_not_lepoll: "\a. Ord(a) & \a \ X" apply (rule ccontr) apply (best intro: Ords_lepoll_set) done -lemma not_Hartog_lepoll_self: "~ Hartog(A) \ A" +lemma not_Hartog_lepoll_self: "\ Hartog(A) \ A" apply (unfold Hartog_def) apply (rule ex_Ord_not_lepoll [THEN exE]) apply (rule LeastI, auto) @@ -70,10 +70,10 @@ lemma Ord_Hartog: "Ord(Hartog(A))" by (unfold Hartog_def, rule Ord_Least) -lemma less_HartogE1: "[| i < Hartog(A); ~ i \ A |] ==> P" +lemma less_HartogE1: "\i < Hartog(A); \ i \ A\ \ P" by (unfold Hartog_def, fast elim: less_LeastE) -lemma less_HartogE: "[| i < Hartog(A); i \ Hartog(A) |] ==> P" +lemma less_HartogE: "\i < Hartog(A); i \ Hartog(A)\ \ P" by (blast intro: less_HartogE1 eqpoll_sym eqpoll_imp_lepoll lepoll_trans [THEN Hartog_lepoll_selfE]) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/WO1_AC.thy --- a/src/ZF/AC/WO1_AC.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/WO1_AC.thy Tue Sep 27 16:51:35 2022 +0100 @@ -1,7 +1,7 @@ (* Title: ZF/AC/WO1_AC.thy Author: Krzysztof Grabczewski -The proofs of WO1 ==> AC1 and WO1 ==> AC10(n) for n >= 1 +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. @@ -29,17 +29,17 @@ begin (* ********************************************************************** *) -(* WO1 ==> AC1 *) +(* WO1 \ AC1 *) (* ********************************************************************** *) -theorem WO1_AC1: "WO1 ==> AC1" +theorem WO1_AC1: "WO1 \ AC1" by (unfold AC1_def WO1_def, fast elim!: ex_choice_fun) (* ********************************************************************** *) -(* WO1 ==> AC10(n) (n >= 1) *) +(* WO1 \ AC10(n) (n >= 1) *) (* ********************************************************************** *) -lemma lemma1: "[| WO1; \B \ A. \C \ D(B). P(C,B) |] ==> \f. \B \ A. P(f`B,B)" +lemma lemma1: "\WO1; \B \ A. \C \ D(B). P(C,B)\ \ \f. \B \ A. P(f`B,B)" apply (unfold WO1_def) apply (erule_tac x = "\({{C \ D (B) . P (C,B) }. B \ A}) " in allE) apply (erule exE, drule ex_choice_fun, fast) @@ -48,7 +48,7 @@ apply (simp, blast dest!: apply_type [OF _ RepFunI]) done -lemma lemma2_1: "[| ~Finite(B); WO1 |] ==> |B| + |B| \ B" +lemma lemma2_1: "\\Finite(B); WO1\ \ |B| + |B| \ B" apply (unfold WO1_def) apply (rule eqpoll_trans) prefer 2 apply (fast elim!: well_ord_cardinal_eqpoll) @@ -61,19 +61,19 @@ done lemma lemma2_2: - "f \ bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i \ D} \ Pow(Pow(B))" + "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})" + "f \ bij(D+D, B) \ pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i \ D})" apply (unfold 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))" + "\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] @@ -82,14 +82,14 @@ done lemma lemma2_5: - "f \ bij(D+D, B) ==> \({{f`Inl(i), f`Inr(i)}. i \ D})=B" + "f \ bij(D+D, B) \ \({{f`Inl(i), f`Inr(i)}. i \ D})=B" apply (unfold 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) & + "\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]], @@ -97,7 +97,7 @@ apply (blast intro!: lemma2_2 lemma2_3 lemma2_4 lemma2_5) done -theorem WO1_AC10: "[| WO1; 1\n |] ==> AC10(n)" +theorem WO1_AC10: "\WO1; 1\n\ \ AC10(n)" apply (unfold AC10_def) apply (fast intro!: lemma1 elim!: lemma2) done diff -r f2094906e491 -r e44d86131648 src/ZF/AC/WO1_WO7.thy --- a/src/ZF/AC/WO1_WO7.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/WO1_WO7.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,8 +13,8 @@ begin definition - "LEMMA == - \X. ~Finite(X) \ (\R. well_ord(X,R) & ~well_ord(X,converse(R)))" + "LEMMA \ + \X. \Finite(X) \ (\R. well_ord(X,R) & \well_ord(X,converse(R)))" (* ********************************************************************** *) (* It is easy to see that WO7 is equivalent to (**) *) @@ -29,7 +29,7 @@ (* It is also easy to show that LEMMA implies WO1. *) (* ********************************************************************** *) -lemma LEMMA_imp_WO1: "LEMMA ==> WO1" +lemma LEMMA_imp_WO1: "LEMMA \ WO1" apply (unfold WO1_def LEMMA_def Finite_def eqpoll_def) apply (blast intro!: well_ord_rvimage [OF bij_is_inj nat_implies_well_ord]) done @@ -48,7 +48,7 @@ (* ********************************************************************** *) lemma converse_Memrel_not_wf_on: - "[| Ord(a); ~Finite(a) |] ==> ~wf[a](converse(Memrel(a)))" + "\Ord(a); \Finite(a)\ \ \wf[a](converse(Memrel(a)))" apply (unfold wf_on_def wf_def) apply (drule nat_le_infinite_Ord [THEN le_imp_subset], assumption) apply (rule notI) @@ -56,7 +56,7 @@ done lemma converse_Memrel_not_well_ord: - "[| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))" + "\Ord(a); \Finite(a)\ \ \well_ord(a,converse(Memrel(a)))" apply (unfold well_ord_def) apply (blast dest: converse_Memrel_not_wf_on) done @@ -69,15 +69,15 @@ 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))))" + "\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" +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] @@ -96,12 +96,12 @@ (* The proof of WO8 \ WO1 (Rubin & Rubin p. 6) *) (* ********************************************************************** *) -lemma WO1_WO8: "WO1 ==> WO8" +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" +(* The implication "WO8 \ WO1": a faithful image of Rubin & Rubin's proof*) +lemma WO8_WO1: "WO8 \ WO1" apply (unfold WO1_def WO8_def) apply (rule allI) apply (erule_tac x = "{{x}. x \ A}" in allE) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/WO2_AC16.thy --- a/src/ZF/AC/WO2_AC16.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/WO2_AC16.thy Tue Sep 27 16:51:35 2022 +0100 @@ -1,7 +1,7 @@ (* Title: ZF/AC/WO2_AC16.thy Author: Krzysztof Grabczewski -The proof of WO2 ==> AC16(k #+ m, k) +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. @@ -15,15 +15,15 @@ theory WO2_AC16 imports AC_Equiv AC16_lemmas Cardinal_aux begin -(**** A recursive definition used in the proof of WO2 ==> AC16 ****) +(**** A recursive definition used in the proof of WO2 \ AC16 ****) definition recfunAC16 :: "[i,i,i,i] => i" where - "recfunAC16(f,h,i,a) == + "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))))})" + (\b f`i \ (\t \ r. \ h`b \ t))))})" (* ********************************************************************** *) (* Basic properties of recfunAC16 *) @@ -38,12 +38,12 @@ else recfunAC16(f,h,i,a) \ {f ` (\ j. h ` i \ f ` j & (\b f`j - \ (\t \ recfunAC16(f,h,i,a). ~ h`b \ t))))})" + \ (\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 recfunAC16(f,h,i,a) = (\j B(g,r); Ord(i) |] - ==> j transrec2(j, 0, B) \ transrec2(i, 0, B)" + "\\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)) @@ -63,8 +63,8 @@ done lemma transrec2_mono: - "[| !!g r. r \ B(g,r); j\i |] - ==> transrec2(j, 0, B) \ transrec2(i, 0, B)" + "\\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 ) @@ -75,7 +75,7 @@ (* ********************************************************************** *) lemma recfunAC16_mono: - "i\j ==> recfunAC16(f, g, i, a) \ recfunAC16(f, g, j, a)" + "i\j \ recfunAC16(f, g, i, a) \ recfunAC16(f, g, j, a)" apply (unfold recfunAC16_def) apply (rule transrec2_mono, auto) done @@ -86,31 +86,31 @@ lemma lemma3_1: - "[| \yzY \ F(y). f(z)<=Y) \ (\! Y. Y \ F(y) & f(z)<=Y); + "\\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" + V \ 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); + "\\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" + V \ 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 & + "\\y X & (\xY \ F(y). h(x) \ Y) \ (\! Y. Y \ F(y) & h(x) \ Y)); - x < a |] - ==> \yzY \ F(y). h(z) \ 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) @@ -118,11 +118,11 @@ done lemma lemma5: - "[| \y X & + "\\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 & + x < a; Limit(x); \i j. i\j \ F(i) \ F(j)\ + \ (\x X & (\xax \ \x x) \ (\! Y. Y \ (\x Y))" apply (rule conjI) @@ -146,7 +146,7 @@ (* 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 + 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). *) @@ -157,7 +157,7 @@ lemma dbl_Diff_eqpoll_Card: - "[| A\a; Card(a); ~Finite(a); B\a; C\a |] ==> A - B - C\a" + "\A\a; Card(a); \Finite(a); B\a; C\a\ \ A - B - C\a" by (blast intro: Diff_lesspoll_eqpoll_Card) (* ********************************************************************** *) @@ -166,7 +166,7 @@ lemma Finite_lesspoll_infinite_Ord: - "[| Finite(X); ~Finite(a); Ord(a) |] ==> X\a" + "\Finite(X); \Finite(a); Ord(a)\ \ X\a" apply (unfold lesspoll_def) apply (rule conjI) apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption) @@ -177,9 +177,9 @@ done lemma Union_lesspoll: - "[| \x \ X. x \ n & x \ T; well_ord(T, R); X \ b; - b nat |] - ==> \(X)\a" + "\\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) @@ -203,20 +203,20 @@ lemma Un_sing_eq_cons: "A \ {a} = cons(a, A)" by fast -lemma Un_lepoll_succ: "A \ B ==> A \ {a} \ succ(B)" +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) - (\b F(a) \ X - (\b X" +lemma Diff_UN_succ_subset: "Ord(a) \ F(a) \ X - (\b X" by blast lemma recfunAC16_Diff_lepoll_1: "Ord(x) - ==> recfunAC16(f, g, x, a) - (\i 1" + \ recfunAC16(f, g, x, a) - (\i 1" apply (erule Ord_cases) apply (simp add: recfunAC16_0 empty_subsetI [THEN subset_imp_lepoll]) (*Limit case*) @@ -231,14 +231,14 @@ done lemma in_Least_Diff: - "[| z \ F(x); Ord(x) |] - ==> z \ F(\ i. z \ F(i)) - (\j<(\ i. z \ F(i)). F(j))" + "\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 (F(b) - (\c(\ i. w \ F(i)) = (\ i. z \ F(i)); + w \ (\i (\i + \ \b (F(b) - (\c (F(b) - (\c 1; a \ A; b \ A |] ==> a=b" +lemma two_in_lepoll_1: "\A \ 1; a \ A; b \ A\ \ a=b" by (fast dest!: lepoll_1_is_sing) lemma UN_lepoll_index: - "[| \ij 1; Limit(a) |] - ==> (\x a" + "\\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) apply (unfold inj_def) @@ -271,7 +271,7 @@ done -lemma recfunAC16_lepoll_index: "Ord(y) ==> recfunAC16(f, h, y, a) \ y" +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) @@ -286,9 +286,9 @@ lemma Union_recfunAC16_lesspoll: - "[| recfunAC16(f,g,y,a) \ {X \ Pow(A). X\n}; - A\a; y nat |] - ==> \(recfunAC16(f,g,y,a))\a" + "\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) @@ -298,11 +298,11 @@ done lemma dbl_Diff_eqpoll: - "[| recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; - Card(a); ~ Finite(a); A\a; + "\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" + h \ 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) @@ -318,9 +318,9 @@ 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}" +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]) @@ -330,14 +330,14 @@ (* ********************************************************************** *) lemma lemma6: - "[| \yx Q(x,y)); succ(j) F(j)<=X & (\x Q(x,j))" + "\\yx 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))" + "\\x Q(x,j); succ(j) + \ P(j,j) \ (\xj | P(x,j) \ Q(x,j))" by (fast elim!: leE) (* ********************************************************************** *) @@ -346,7 +346,7 @@ (* ********************************************************************** *) lemma ex_subset_eqpoll: - "[| A\a; ~ Finite(a); Ord(a); m \ nat |] ==> \X \ Pow(A). X\m" + "\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) @@ -354,12 +354,12 @@ apply (fast elim!: eqpoll_sym) done -lemma subset_Un_disjoint: "[| A \ B \ C; A \ C = 0 |] ==> A \ B" +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" + "\X \ Pow(A - \(B) -C); T \ B; F \ T\ \ F \ X = 0" by blast @@ -369,7 +369,7 @@ lemma subset_imp_eq_lemma: - "m \ nat ==> \A B. A \ B & m \ A & B \ m \ A=B" + "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) @@ -385,13 +385,13 @@ done -lemma subset_imp_eq: "[| A \ B; m \ A; B \ m; m \ nat |] ==> A=B" +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" + "\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) @@ -401,14 +401,14 @@ lemma ex_next_set: - "[| recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; - Card(a); ~ Finite(a); A\a; + "\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 & + \ (\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))" + (\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) @@ -427,15 +427,15 @@ (* ********************************************************************** *) lemma ex_next_Ord: - "[| recfunAC16(f, h, y, a) \ {X \ Pow(A) . X\succ(k #+ m)}; - Card(a); ~ Finite(a); A\a; + "\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 & + \ (\Y \ recfunAC16(f, h, y, a). h`y \ Y)\ + \ \c f`c & (\b f`c \ - (\T \ recfunAC16(f, h, y, a). ~ h`b \ T))" + (\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) @@ -450,10 +450,10 @@ (* ********************************************************************** *) lemma lemma8: - "[| \xxa \ F(j). P(x, xa)) + "\\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 & + 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) @@ -468,10 +468,10 @@ (* ********************************************************************** *) lemma main_induct: - "[| b < a; f \ bij(a, {Y \ Pow(A) . Y\succ(k #+ m)}); + "\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)} & + \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) @@ -507,11 +507,11 @@ (* ********************************************************************** *) lemma lemma_simp_induct: - "[| \b. b F(b) \ S & (\xY \ F(b). f`x \ Y)) + "\\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 & + \i j. i\j \ F(i) \ F(j)\ + \ (\j S & (\x \ f``a. \! Y. Y \ (\j Y)" apply (rule conjI) apply (rule subsetI) @@ -546,7 +546,7 @@ (* The target theorem *) (* ********************************************************************** *) -theorem WO2_AC16: "[| WO2; 0 nat; m \ nat |] ==> AC16(k #+ m,k)" +theorem WO2_AC16: "\WO2; 0 nat; m \ nat\ \ AC16(k #+ m,k)" apply (unfold AC16_def) apply (rule allI) apply (rule impI) diff -r f2094906e491 -r e44d86131648 src/ZF/AC/WO6_WO1.thy --- a/src/ZF/AC/WO6_WO1.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/AC/WO6_WO1.thy Tue Sep 27 16:51:35 2022 +0100 @@ -2,11 +2,11 @@ Author: Krzysztof Grabczewski Proofs needed to state that formulations WO1,...,WO6 are all equivalent. -The only hard one is WO6 ==> WO1. +The only hard one is WO6 \ WO1. -Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear" +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. +They refer reader to a book by Gödel to see the proof WO1 \ WO2. Fortunately order types made this proof also very easy. *) @@ -17,18 +17,18 @@ (* Auxiliary definitions used in proof *) definition NN :: "i => i" where - "NN(y) == {m \ nat. \a. \f. Ord(a) & domain(f)=a & + "NN(y) \ {m \ nat. \a. \f. Ord(a) & domain(f)=a & (\bb m)}" definition uu :: "[i, i, i, i] => i" where - "uu(f, beta, gamma, delta) == (f`beta * f`gamma) \ f`delta" + "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) == + "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 & @@ -37,35 +37,35 @@ definition ww1 :: "[i, i, i] => i" where - "ww1(f,m,b) == f`b - vv1(f,m,b)" + "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 \b \ a++a. if b i" where - "vv2(f,b,g,s) == + "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)" + "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) == + "gg2(f,a,b,s) \ \g \ a++a. if g WO3" +lemma WO2_WO3: "WO2 \ WO3" by (unfold WO2_def WO3_def, fast) (* ********************************************************************** *) -lemma WO3_WO1: "WO3 ==> WO1" +lemma WO3_WO1: "WO3 \ WO1" apply (unfold eqpoll_def WO1_def WO3_def) apply (intro allI) apply (drule_tac x=A in spec) @@ -75,25 +75,25 @@ (* ********************************************************************** *) -lemma WO1_WO2: "WO1 ==> WO2" +lemma WO1_WO2: "WO1 \ WO2" apply (unfold 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}" +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" +lemma surj_imp_eq': "f \ surj(A,B) \ (\a \ A. {f`a}) = B" apply (unfold surj_def) apply (fast elim!: apply_type) done -lemma surj_imp_eq: "[| f \ surj(A,B); Ord(A) |] ==> (\af \ surj(A,B); Ord(A)\ \ (\a WO4(1)" +lemma WO1_WO4: "WO1 \ WO4(1)" apply (unfold WO1_def WO4_def) apply (rule allI) apply (erule_tac x = A in allE) @@ -108,32 +108,32 @@ (* ********************************************************************** *) -lemma WO4_mono: "[| m\n; WO4(m) |] ==> WO4(n)" +lemma WO4_mono: "\m\n; WO4(m)\ \ WO4(n)" apply (unfold WO4_def) apply (blast dest!: spec intro: lepoll_trans [OF _ le_imp_lepoll]) done (* ********************************************************************** *) -lemma WO4_WO5: "[| m \ nat; 1\m; WO4(m) |] ==> WO5" +lemma WO4_WO5: "\m \ nat; 1\m; WO4(m)\ \ WO5" by (unfold WO4_def WO5_def, blast) (* ********************************************************************** *) -lemma WO5_WO6: "WO5 ==> WO6" +lemma WO5_WO6: "WO5 \ WO6" by (unfold WO4_def WO5_def WO6_def, blast) (* ********************************************************************** - The proof of "WO6 ==> WO1". Simplified by L C Paulson. + 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 | (~ kk < i++j; Ord(i); Ord(j)\ + \ k < i | (\ kb m ==> \bgd 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" @@ -163,7 +163,7 @@ 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" +lemma uu_lepoll_m: "\\b m; d \ uu(f,b,g,d) \ m" by (blast intro: uu_subset2 [THEN subset_imp_lepoll] lepoll_trans) (* ********************************************************************** *) @@ -171,7 +171,7 @@ (* ********************************************************************** *) lemma cases: "\bgd m - ==> (\b 0 \ + \ (\b 0 \ (\gd 0 & u(f,b,g,d) \ m)) | (\b 0 & (\gd 0 \ u(f,b,g,d) \ m))" @@ -182,7 +182,7 @@ (* ********************************************************************** *) (* Lemmas used in both cases *) (* ********************************************************************** *) -lemma UN_oadd: "Ord(a) ==> (\bb C(a++b))" +lemma UN_oadd: "Ord(a) \ (\bb C(a++b))" by (blast intro: ltI lt_oadd1 oadd_lt_mono2 dest!: lt_oadd_disj) @@ -197,7 +197,7 @@ (* Case 1: Union of images is the whole "y" *) (* ********************************************************************** *) lemma UN_gg1_eq: - "[| Ord(a); m \ nat |] ==> (\bbOrd(a); m \ nat\ \ (\bb a. \x. Ord(x) & P(a, x)) |] - ==> P(Least_a, \ b. P(Least_a, b))" + "\P(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)+ @@ -222,12 +222,12 @@ domain(uu(f,b,g,d)) \ m"] for f b m lemma gg1_lepoll_m: - "[| Ord(a); m \ nat; + "\Ord(a); m \ nat; \b0 \ (\gd 0 & domain(uu(f,b,g,d)) \ m); - \b succ(m); b gg1(f,a,m)`b \ 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 *) @@ -253,34 +253,34 @@ (* ********************************************************************** *) lemma ex_d_uu_not_empty: - "[| b0; f`g\0; - y*y \ y; (\b \d 0" + "\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" + "\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" +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" + "\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}" +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" + "\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) @@ -293,12 +293,12 @@ done lemma uu_Least_is_fun: - "[| \gd0 \ + "\\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" + f`b\0; 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+) @@ -312,11 +312,11 @@ done lemma vv2_subset: - "[| \gd0 \ + "\\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" + (\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 @@ -325,11 +325,11 @@ (* Case 2: Union of images is the whole "y" *) (* ********************************************************************** *) lemma UN_gg2_eq: - "[| \gd 0 \ + "\\gd 0 \ domain(uu(f,b,g,d)) \ succ(m); \b succ(m); y*y \ y; - (\b nat; s \ f`b; b (\gb nat; s \ f`b; b + \ (\g nat; m\0 |] ==> vv2(f,b,g,s) \ m" +lemma vv2_lepoll: "\m \ nat; m\0\ \ vv2(f,b,g,s) \ m" apply (unfold vv2_def) apply (simp add: empty_lepollI) apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0] @@ -355,8 +355,8 @@ done lemma ww2_lepoll: - "[| \b succ(m); g nat; vv2(f,b,g,d) \ f`g |] - ==> ww2(f,b,g,d) \ m" + "\\b succ(m); g nat; vv2(f,b,g,d) \ f`g\ + \ ww2(f,b,g,d) \ m" apply (unfold ww2_def) apply (case_tac "f`g = 0") apply (simp add: empty_lepollI) @@ -366,11 +366,11 @@ done lemma gg2_lepoll_m: - "[| \gd 0 \ + "\\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" + (\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) @@ -380,7 +380,7 @@ (* ********************************************************************** *) (* lemma ii *) (* ********************************************************************** *) -lemma lemma_ii: "[| succ(m) \ NN(y); y*y \ y; m \ nat; m\0 |] ==> m \ NN(y)" +lemma lemma_ii: "\succ(m) \ NN(y); y*y \ y; m \ nat; m\0\ \ m \ NN(y)" apply (unfold NN_def) apply (elim CollectE exE conjE) apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption) @@ -413,16 +413,16 @@ 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)" + "\\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)" + "\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 @@ -444,34 +444,34 @@ (* y can be well-ordered" *) (* In fact we have to prove *) -(* * WO6 ==> NN(y) \ 0 *) +(* * WO6 \ NN(y) \ 0 *) (* * reverse induction which lets us infer that 1 \ NN(y) *) -(* * 1 \ NN(y) ==> y can be well-ordered *) +(* * 1 \ NN(y) \ y can be well-ordered *) (* ********************************************************************** *) (* ********************************************************************** *) -(* WO6 ==> NN(y) \ 0 *) +(* WO6 \ NN(y) \ 0 *) (* ********************************************************************** *) -lemma WO6_imp_NN_not_empty: "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 *) +(* 1 \ NN(y) \ y can be well-ordered *) (* ********************************************************************** *) lemma lemma1: - "[| (\b y; \b 1; Ord(a) |] ==> \c(\b y; \b 1; Ord(a)\ \ \cb y; \b 1; Ord(a) |] - ==> f` (\ i. f`i = {x}) = {x}" + "\(\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)" +lemma NN_imp_ex_inj: "1 \ NN(y) \ \a f. Ord(a) & f \ inj(y, a)" apply (unfold NN_def) apply (elim CollectE exE conjE) apply (rule_tac x = a in exI) @@ -483,7 +483,7 @@ apply (rule lemma2 [THEN ssubst], assumption+, blast) done -lemma y_well_ord: "[| y*y \ y; 1 \ NN(y) |] ==> \r. well_ord(y, r)" +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 @@ -493,35 +493,35 @@ (* ********************************************************************** *) lemma rev_induct_lemma [rule_format]: - "[| n \ nat; !!m. [| m \ nat; m\0; P(succ(m)) |] ==> P(m) |] - ==> n\0 \ P(n) \ P(1)" + "\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)" + "\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" +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)" +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" *) +(* Main theorem "WO6 \ WO1" *) (* ********************************************************************** *) (* another helpful lemma *) -lemma NN_y_0: "0 \ NN(y) ==> y=0" +lemma NN_y_0: "0 \ NN(y) \ y=0" apply (unfold NN_def) apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst) done -lemma WO6_imp_WO1: "WO6 ==> WO1" +lemma WO6_imp_WO1: "WO6 \ WO1" apply (unfold WO1_def) apply (rule allI) apply (case_tac "A=0") diff -r f2094906e491 -r e44d86131648 src/ZF/Arith.thy --- a/src/ZF/Arith.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Arith.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,11 +17,11 @@ definition pred :: "i=>i" (*inverse of succ*) where - "pred(y) == nat_case(0, %x. x, y)" + "pred(y) \ nat_case(0, %x. x, y)" definition natify :: "i=>i" (*coerces non-nats to nats*) where - "natify == Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a)) + "natify \ Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a)) else 0)" consts @@ -44,44 +44,44 @@ definition add :: "[i,i]=>i" (infixl \#+\ 65) where - "m #+ n == raw_add (natify(m), natify(n))" + "m #+ n \ raw_add (natify(m), natify(n))" definition diff :: "[i,i]=>i" (infixl \#-\ 65) where - "m #- n == raw_diff (natify(m), natify(n))" + "m #- n \ raw_diff (natify(m), natify(n))" definition mult :: "[i,i]=>i" (infixl \#*\ 70) where - "m #* n == raw_mult (natify(m), natify(n))" + "m #* n \ raw_mult (natify(m), natify(n))" definition raw_div :: "[i,i]=>i" where - "raw_div (m, n) == + "raw_div (m, n) \ transrec(m, %j f. if ji" where - "raw_mod (m, n) == + "raw_mod (m, n) \ transrec(m, %j f. if ji" (infixl \div\ 70) where - "m div n == raw_div (natify(m), natify(n))" + "m div n \ raw_div (natify(m), natify(n))" definition mod :: "[i,i]=>i" (infixl \mod\ 70) where - "m mod n == raw_mod (natify(m), natify(n))" + "m mod n \ raw_mod (natify(m), natify(n))" declare rec_type [simp] nat_0_le [simp] -lemma zero_lt_lemma: "[| 0 nat |] ==> \j\nat. k = succ(j)" +lemma zero_lt_lemma: "\0 nat\ \ \j\nat. k = succ(j)" apply (erule rev_mp) apply (induct_tac "k", auto) done -(* @{term"[| 0 < k; k \ nat; !!j. [| j \ nat; k = succ(j) |] ==> Q |] ==> Q"} *) +(* @{term"\0 < k; k \ nat; \j. \j \ nat; k = succ(j)\ \ Q\ \ Q"} *) lemmas zero_lt_natE = zero_lt_lemma [THEN bexE] @@ -96,7 +96,7 @@ lemma natify_0 [simp]: "natify(0) = 0" by (rule natify_def [THEN def_Vrecursor, THEN trans], auto) -lemma natify_non_succ: "\z. x \ succ(z) ==> natify(x) = 0" +lemma natify_non_succ: "\z. x \ succ(z) \ natify(x) = 0" by (rule natify_def [THEN def_Vrecursor, THEN trans], auto) lemma natify_in_nat [iff,TC]: "natify(x) \ nat" @@ -105,12 +105,12 @@ apply (auto simp add: natify_succ natify_non_succ) done -lemma natify_ident [simp]: "n \ nat ==> natify(n) = n" +lemma natify_ident [simp]: "n \ nat \ natify(n) = n" apply (induct_tac "n") apply (auto simp add: natify_succ) done -lemma natify_eqE: "[|natify(x) = y; x \ nat|] ==> x=y" +lemma natify_eqE: "\natify(x) = y; x \ nat\ \ x=y" by auto @@ -165,7 +165,7 @@ (** Addition **) -lemma raw_add_type: "[| m\nat; n\nat |] ==> raw_add (m, n) \ nat" +lemma raw_add_type: "\m\nat; n\nat\ \ raw_add (m, n) \ nat" by (induct_tac "m", auto) lemma add_type [iff,TC]: "m #+ n \ nat" @@ -173,7 +173,7 @@ (** Multiplication **) -lemma raw_mult_type: "[| m\nat; n\nat |] ==> raw_mult (m, n) \ nat" +lemma raw_mult_type: "\m\nat; n\nat\ \ raw_mult (m, n) \ nat" apply (induct_tac "m") apply (simp_all add: raw_add_type) done @@ -184,7 +184,7 @@ (** Difference **) -lemma raw_diff_type: "[| m\nat; n\nat |] ==> raw_diff (m, n) \ nat" +lemma raw_diff_type: "\m\nat; n\nat\ \ raw_diff (m, n) \ nat" by (induct_tac "n", auto) lemma diff_type [iff,TC]: "m #- n \ nat" @@ -208,7 +208,7 @@ lemma diff_0 [simp]: "m #- 0 = natify(m)" by (simp add: diff_def) -lemma diff_le_self: "m\nat ==> (m #- n) \ m" +lemma diff_le_self: "m\nat \ (m #- n) \ m" apply (subgoal_tac " (m #- natify (n)) \ m") apply (rule_tac [2] m = m and n = "natify (n) " in diff_induct) apply (erule_tac [6] leE) @@ -225,7 +225,7 @@ lemma add_succ [simp]: "succ(m) #+ n = succ(m #+ n)" by (simp add: natify_succ add_def) -lemma add_0: "m \ nat ==> 0 #+ m = m" +lemma add_0: "m \ nat \ 0 #+ m = m" by simp (*Associative law for addition*) @@ -250,7 +250,7 @@ apply (auto simp add: natify_succ) done -lemma add_0_right: "m \ nat ==> m #+ 0 = m" +lemma add_0_right: "m \ nat \ m #+ 0 = m" by auto (*Commutative law for addition*) @@ -272,40 +272,40 @@ (*Cancellation law on the left*) lemma raw_add_left_cancel: - "[| raw_add(k, m) = raw_add(k, n); k\nat |] ==> m=n" + "\raw_add(k, m) = raw_add(k, n); k\nat\ \ m=n" apply (erule rev_mp) apply (induct_tac "k", auto) done -lemma add_left_cancel_natify: "k #+ m = k #+ n ==> natify(m) = natify(n)" +lemma add_left_cancel_natify: "k #+ m = k #+ n \ natify(m) = natify(n)" apply (unfold add_def) apply (drule raw_add_left_cancel, auto) done lemma add_left_cancel: - "[| i = j; i #+ m = j #+ n; m\nat; n\nat |] ==> m = n" + "\i = j; i #+ m = j #+ n; m\nat; n\nat\ \ m = n" by (force dest!: add_left_cancel_natify) (*Thanks to Sten Agerholm*) -lemma add_le_elim1_natify: "k#+m \ k#+n ==> natify(m) \ natify(n)" +lemma add_le_elim1_natify: "k#+m \ k#+n \ natify(m) \ natify(n)" apply (rule_tac P = "natify(k) #+m \ natify(k) #+n" in rev_mp) apply (rule_tac [2] n = "natify(k) " in nat_induct) apply auto done -lemma add_le_elim1: "[| k#+m \ k#+n; m \ nat; n \ nat |] ==> m \ n" +lemma add_le_elim1: "\k#+m \ k#+n; m \ nat; n \ nat\ \ m \ n" by (drule add_le_elim1_natify, auto) -lemma add_lt_elim1_natify: "k#+m < k#+n ==> natify(m) < natify(n)" +lemma add_lt_elim1_natify: "k#+m < k#+n \ natify(m) < natify(n)" apply (rule_tac P = "natify(k) #+m < natify(k) #+n" in rev_mp) apply (rule_tac [2] n = "natify(k) " in nat_induct) apply auto done -lemma add_lt_elim1: "[| k#+m < k#+n; m \ nat; n \ nat |] ==> m < n" +lemma add_lt_elim1: "\k#+m < k#+n; m \ nat; n \ nat\ \ m < n" by (drule add_lt_elim1_natify, auto) -lemma zero_less_add: "[| n \ nat; m \ nat |] ==> 0 < m #+ n \ (0n \ nat; m \ nat\ \ 0 < m #+ n \ (0nat, for consider j = omega. Then we can have inat, but natify(j)=0, so the conclusion fails.*) -lemma add_lt_mono1: "[| inat |] ==> i#+k < j#+k" +lemma add_lt_mono1: "\inat\ \ i#+k < j#+k" apply (frule lt_nat_in_nat, assumption) apply (erule succ_lt_induct) apply (simp_all add: leI) done text\strict, in second argument\ -lemma add_lt_mono2: "[| inat |] ==> k#+i < k#+j" +lemma add_lt_mono2: "\inat\ \ k#+i < k#+j" by (simp add: add_commute [of k] add_lt_mono1) text\A [clumsy] way of lifting < monotonicity to \\\ monotonicity\ lemma Ord_lt_mono_imp_le_mono: - assumes lt_mono: "!!i j. [| i f(i) < f(j)" - and ford: "!!i. i:k ==> Ord(f(i))" + assumes lt_mono: "\i j. \i \ f(i) < f(j)" + and ford: "\i. i:k \ Ord(f(i))" and leij: "i \ j" and jink: "j:k" shows "f(i) \ f(j)" @@ -336,29 +336,29 @@ done text\\\\ monotonicity, 1st argument\ -lemma add_le_mono1: "[| i \ j; j\nat |] ==> i#+k \ j#+k" +lemma add_le_mono1: "\i \ j; j\nat\ \ i#+k \ j#+k" apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck) apply (blast intro: add_lt_mono1 add_type [THEN nat_into_Ord])+ done text\\\\ monotonicity, both arguments\ -lemma add_le_mono: "[| i \ j; k \ l; j\nat; l\nat |] ==> i#+k \ j#+l" +lemma add_le_mono: "\i \ j; k \ l; j\nat; l\nat\ \ i#+k \ j#+l" apply (rule add_le_mono1 [THEN le_trans], assumption+) apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+) done text\Combinations of less-than and less-than-or-equals\ -lemma add_lt_le_mono: "[| il; j\nat; l\nat |] ==> i#+k < j#+l" +lemma add_lt_le_mono: "\il; j\nat; l\nat\ \ i#+k < j#+l" apply (rule add_lt_mono1 [THEN lt_trans2], assumption+) apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+) done -lemma add_le_lt_mono: "[| i\j; knat; l\nat |] ==> i#+k < j#+l" +lemma add_le_lt_mono: "\i\j; knat; l\nat\ \ i#+k < j#+l" by (subst add_commute, subst add_commute, erule add_lt_le_mono, assumption+) text\Less-than: in other words, strict in both arguments\ -lemma add_lt_mono: "[| inat; l\nat |] ==> i#+k < j#+l" +lemma add_lt_mono: "\inat; l\nat\ \ i#+k < j#+l" apply (rule add_lt_le_mono) apply (auto intro: leI) done @@ -393,21 +393,21 @@ lemma pred_0 [simp]: "pred(0) = 0" by (simp add: pred_def) -lemma eq_succ_imp_eq_m1: "[|i = succ(j); i\nat|] ==> j = i #- 1 & j \nat" +lemma eq_succ_imp_eq_m1: "\i = succ(j); i\nat\ \ j = i #- 1 & j \nat" by simp lemma pred_Un_distrib: - "[|i\nat; j\nat|] ==> pred(i \ j) = pred(i) \ pred(j)" + "\i\nat; j\nat\ \ pred(i \ j) = pred(i) \ pred(j)" apply (erule_tac n=i in natE, simp) apply (erule_tac n=j in natE, simp) apply (simp add: succ_Un_distrib [symmetric]) done lemma pred_type [TC,simp]: - "i \ nat ==> pred(i) \ nat" + "i \ nat \ pred(i) \ nat" by (simp add: pred_def split: split_nat_case) -lemma nat_diff_pred: "[|i\nat; j\nat|] ==> i #- succ(j) = pred(i #- j)" +lemma nat_diff_pred: "\i\nat; j\nat\ \ i #- succ(j) = pred(i #- j)" apply (rule_tac m=i and n=j in diff_induct) apply (auto simp add: pred_def nat_imp_quasinat split: split_nat_case) done @@ -418,13 +418,13 @@ done lemma nat_diff_Un_distrib: - "[|i\nat; j\nat; k\nat|] ==> (i \ j) #- k = (i#-k) \ (j#-k)" + "\i\nat; j\nat; k\nat\ \ (i \ j) #- k = (i#-k) \ (j#-k)" apply (rule_tac n=k in nat_induct) apply (simp_all add: diff_succ_eq_pred pred_Un_distrib) done lemma diff_Un_distrib: - "[|i\nat; j\nat|] ==> (i \ j) #- k = (i#-k) \ (j#-k)" + "\i\nat; j\nat\ \ (i \ j) #- k = (i#-k) \ (j#-k)" by (insert nat_diff_Un_distrib [of i j "natify(k)"], simp) text\We actually prove \<^term>\i #- j #- k = i #- (j #+ k)\\ @@ -451,10 +451,10 @@ by (simp add: diff_cancel) (*To tidy up the result of a simproc. Only the RHS will be simplified.*) -lemma eq_cong2: "u = u' ==> (t==u) == (t==u')" +lemma eq_cong2: "u = u' \ (t\u) \ (t\u')" by auto -lemma iff_cong2: "u \ u' ==> (t==u) == (t==u')" +lemma iff_cong2: "u \ u' \ (t\u) \ (t\u')" by auto @@ -488,10 +488,10 @@ lemma mult_1_right_natify [simp]: "n #* 1 = natify(n)" by auto -lemma mult_1: "n \ nat ==> 1 #* n = n" +lemma mult_1: "n \ nat \ 1 #* n = n" by simp -lemma mult_1_right: "n \ nat ==> n #* 1 = n" +lemma mult_1_right: "n \ nat \ n #* 1 = n" by simp (*Commutative law for multiplication*) @@ -536,12 +536,12 @@ lemma lt_succ_eq_0_disj: - "[| m\nat; n\nat |] - ==> (m < succ(n)) \ (m = 0 | (\j\nat. m = succ(j) & j < n))" + "\m\nat; n\nat\ + \ (m < succ(n)) \ (m = 0 | (\j\nat. m = succ(j) & j < n))" by (induct_tac "m", auto) lemma less_diff_conv [rule_format]: - "[| j\nat; k\nat |] ==> \i\nat. (i < j #- k) \ (i #+ k < j)" + "\j\nat; k\nat\ \ \i\nat. (i < j #- k) \ (i #+ k < j)" by (erule_tac m = k in diff_induct, auto) lemmas nat_typechecks = rec_type nat_0I nat_1I nat_succI Ord_nat diff -r f2094906e491 -r e44d86131648 src/ZF/ArithSimp.thy --- a/src/ZF/ArithSimp.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ArithSimp.thy Tue Sep 27 16:51:35 2022 +0100 @@ -25,19 +25,19 @@ (*We need m:nat even if we replace the RHS by natify(m), for consider e.g. n=2, m=omega; then n + (m-n) = 2 + (0-2) = 2 \ 0 = natify(m).*) -lemma add_diff_inverse: "[| n \ m; m:nat |] ==> n #+ (m#-n) = m" +lemma add_diff_inverse: "\n \ m; m:nat\ \ n #+ (m#-n) = m" apply (frule lt_nat_in_nat, erule nat_succI) apply (erule rev_mp) apply (rule_tac m = m and n = n in diff_induct, auto) done -lemma add_diff_inverse2: "[| n \ m; m:nat |] ==> (m#-n) #+ n = m" +lemma add_diff_inverse2: "\n \ m; m:nat\ \ (m#-n) #+ n = m" apply (frule lt_nat_in_nat, erule nat_succI) apply (simp (no_asm_simp) add: add_commute add_diff_inverse) done (*Proof is IDENTICAL to that of add_diff_inverse*) -lemma diff_succ: "[| n \ m; m:nat |] ==> succ(m) #- n = succ(m#-n)" +lemma diff_succ: "\n \ m; m:nat\ \ succ(m) #- n = succ(m#-n)" apply (frule lt_nat_in_nat, erule nat_succI) apply (erule rev_mp) apply (rule_tac m = m and n = n in diff_induct) @@ -45,7 +45,7 @@ done lemma zero_less_diff [simp]: - "[| m: nat; n: nat |] ==> 0 < (n #- m) \ mm: nat; n: nat\ \ 0 < (n #- m) \ mRemainder\ (*We need m:nat even with natify*) -lemma div_termination: "[| 0 m; m:nat |] ==> m #- n < m" +lemma div_termination: "\0 m; m:nat\ \ m #- n < m" apply (frule lt_nat_in_nat, erule nat_succI) apply (erule rev_mp) apply (erule rev_mp) @@ -81,7 +81,7 @@ div_termination [THEN ltD] nat_into_Ord not_lt_iff_le [THEN iffD1] -lemma raw_mod_type: "[| m:nat; n:nat |] ==> raw_mod (m, n) \ nat" +lemma raw_mod_type: "\m:nat; n:nat\ \ raw_mod (m, n) \ nat" apply (unfold raw_mod_def) apply (rule Ord_transrec_type) apply (auto simp add: nat_into_Ord [THEN Ord_0_lt_iff]) @@ -109,25 +109,25 @@ apply (simp (no_asm_simp)) done (*NOT for adding to default simpset*) -lemma raw_mod_less: "m raw_mod (m,n) = m" +lemma raw_mod_less: "m raw_mod (m,n) = m" apply (rule raw_mod_def [THEN def_transrec, THEN trans]) apply (simp (no_asm_simp) add: div_termination [THEN ltD]) done -lemma mod_less [simp]: "[| m nat |] ==> m mod n = m" +lemma mod_less [simp]: "\m nat\ \ m mod n = m" apply (frule lt_nat_in_nat, assumption) apply (simp (no_asm_simp) add: mod_def raw_mod_less) done lemma raw_mod_geq: - "[| 0 m; m:nat |] ==> raw_mod (m, n) = raw_mod (m#-n, n)" + "\0 m; m:nat\ \ raw_mod (m, n) = raw_mod (m#-n, n)" apply (frule lt_nat_in_nat, erule nat_succI) apply (rule raw_mod_def [THEN def_transrec, THEN trans]) apply (simp (no_asm_simp) add: div_termination [THEN ltD] not_lt_iff_le [THEN iffD2], blast) done -lemma mod_geq: "[| n \ m; m:nat |] ==> m mod n = (m#-n) mod n" +lemma mod_geq: "\n \ m; m:nat\ \ m mod n = (m#-n) mod n" apply (frule lt_nat_in_nat, erule nat_succI) apply (case_tac "n=0") apply (simp add: DIVISION_BY_ZERO_MOD) @@ -137,7 +137,7 @@ subsection\Division\ -lemma raw_div_type: "[| m:nat; n:nat |] ==> raw_div (m, n) \ nat" +lemma raw_div_type: "\m:nat; n:nat\ \ raw_div (m, n) \ nat" apply (unfold raw_div_def) apply (rule Ord_transrec_type) apply (auto simp add: nat_into_Ord [THEN Ord_0_lt_iff]) @@ -149,17 +149,17 @@ apply (simp (no_asm) add: div_def raw_div_type) done -lemma raw_div_less: "m raw_div (m,n) = 0" +lemma raw_div_less: "m raw_div (m,n) = 0" apply (rule raw_div_def [THEN def_transrec, THEN trans]) apply (simp (no_asm_simp) add: div_termination [THEN ltD]) done -lemma div_less [simp]: "[| m nat |] ==> m div n = 0" +lemma div_less [simp]: "\m nat\ \ m div n = 0" apply (frule lt_nat_in_nat, assumption) apply (simp (no_asm_simp) add: div_def raw_div_less) done -lemma raw_div_geq: "[| 0 m; m:nat |] ==> raw_div(m,n) = succ(raw_div(m#-n, n))" +lemma raw_div_geq: "\0 m; m:nat\ \ raw_div(m,n) = succ(raw_div(m#-n, n))" apply (subgoal_tac "n \ 0") prefer 2 apply blast apply (frule lt_nat_in_nat, erule nat_succI) @@ -168,7 +168,7 @@ done lemma div_geq [simp]: - "[| 0 m; m:nat |] ==> m div n = succ ((m#-n) div n)" + "\0 m; m:nat\ \ m div n = succ ((m#-n) div n)" apply (frule lt_nat_in_nat, erule nat_succI) apply (simp (no_asm_simp) add: div_def raw_div_geq) done @@ -177,7 +177,7 @@ (*A key result*) -lemma mod_div_lemma: "[| m: nat; n: nat |] ==> (m div n)#*n #+ m mod n = m" +lemma mod_div_lemma: "\m: nat; n: nat\ \ (m div n)#*n #+ m mod n = m" apply (case_tac "n=0") apply (simp add: DIVISION_BY_ZERO_MOD) apply (simp add: nat_into_Ord [THEN Ord_0_lt_iff]) @@ -195,7 +195,7 @@ apply (subst mod_div_lemma, auto) done -lemma mod_div_equality: "m: nat ==> (m div n)#*n #+ m mod n = m" +lemma mod_div_equality: "m: nat \ (m div n)#*n #+ m mod n = m" apply (simp (no_asm_simp) add: mod_div_equality_natify) done @@ -205,8 +205,8 @@ text\(mainly for mutilated chess board)\ lemma mod_succ_lemma: - "[| 0 succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))" + "\0 + \ succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))" apply (erule complete_induct) apply (case_tac "succ (x) case succ(x) < n\ @@ -221,7 +221,7 @@ done lemma mod_succ: - "n:nat ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))" + "n:nat \ succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))" apply (case_tac "n=0") apply (simp (no_asm_simp) add: natify_succ DIVISION_BY_ZERO_MOD) apply (subgoal_tac "natify (succ (m)) mod n = (if succ (natify (m) mod n) = n then 0 else succ (natify (m) mod n))") @@ -231,7 +231,7 @@ apply (auto simp del: natify_succ simp add: nat_into_Ord [THEN Ord_0_lt_iff]) done -lemma mod_less_divisor: "[| 0 m mod n < n" +lemma mod_less_divisor: "\0 \ m mod n < n" apply (subgoal_tac "natify (m) mod n < n") apply (rule_tac [2] i = "natify (m) " in complete_induct) apply (case_tac [3] "x k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)" +lemma mod2_cases: "b<2 \ k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)" apply (subgoal_tac "k mod 2: 2") prefer 2 apply (simp add: mod_less_divisor [THEN ltD]) apply (drule ltD, auto) @@ -266,17 +266,17 @@ subsection\Additional theorems about \\\\ -lemma add_le_self: "m:nat ==> m \ (m #+ n)" +lemma add_le_self: "m:nat \ m \ (m #+ n)" apply (simp (no_asm_simp)) done -lemma add_le_self2: "m:nat ==> m \ (n #+ m)" +lemma add_le_self2: "m:nat \ m \ (n #+ m)" apply (simp (no_asm_simp)) done (*** Monotonicity of Multiplication ***) -lemma mult_le_mono1: "[| i \ j; j:nat |] ==> (i#*k) \ (j#*k)" +lemma mult_le_mono1: "\i \ j; j:nat\ \ (i#*k) \ (j#*k)" apply (subgoal_tac "natify (i) #*natify (k) \ j#*natify (k) ") apply (frule_tac [2] lt_nat_in_nat) apply (rule_tac [3] n = "natify (k) " in nat_induct) @@ -284,14 +284,14 @@ done (* @{text"\"} monotonicity, BOTH arguments*) -lemma mult_le_mono: "[| i \ j; k \ l; j:nat; l:nat |] ==> i#*k \ j#*l" +lemma mult_le_mono: "\i \ j; k \ l; j:nat; l:nat\ \ i#*k \ j#*l" apply (rule mult_le_mono1 [THEN le_trans], assumption+) apply (subst mult_commute, subst mult_commute, rule mult_le_mono1, assumption+) done (*strict, in 1st argument; proof is by induction on k>0. I can't see how to relax the typing conditions.*) -lemma mult_lt_mono2: "[| i k#*i < k#*j" +lemma mult_lt_mono2: "\i \ k#*i < k#*j" apply (erule zero_lt_natE) apply (frule_tac [2] lt_nat_in_nat) apply (simp_all (no_asm_simp)) @@ -299,7 +299,7 @@ apply (simp_all (no_asm_simp) add: add_lt_mono) done -lemma mult_lt_mono1: "[| i i#*k < j#*k" +lemma mult_lt_mono1: "\i \ i#*k < j#*k" apply (simp (no_asm_simp) add: mult_lt_mono2 mult_commute [of _ k]) done @@ -326,7 +326,7 @@ done -lemma mult_is_zero: "[|m: nat; n: nat|] ==> (m #* n = 0) \ (m = 0 | n = 0)" +lemma mult_is_zero: "\m: nat; n: nat\ \ (m #* n = 0) \ (m = 0 | n = 0)" apply auto apply (erule natE) apply (erule_tac [2] natE, auto) @@ -342,7 +342,7 @@ subsection\Cancellation Laws for Common Factors in Comparisons\ lemma mult_less_cancel_lemma: - "[| k: nat; m: nat; n: nat |] ==> (m#*k < n#*k) \ (0k: nat; m: nat; n: nat\ \ (m#*k < n#*k) \ (0 nat ==> k #* m \ k \ (0 < k \ natify(m) \ 1)" +lemma mult_le_cancel_le1: "k \ nat \ k #* m \ k \ (0 < k \ natify(m) \ 1)" by (cut_tac k = k and m = m and n = 1 in mult_le_cancel1, auto) -lemma Ord_eq_iff_le: "[| Ord(m); Ord(n) |] ==> m=n \ (m \ n & n \ m)" +lemma Ord_eq_iff_le: "\Ord(m); Ord(n)\ \ m=n \ (m \ n & n \ m)" by (blast intro: le_anti_sym) lemma mult_cancel2_lemma: - "[| k: nat; m: nat; n: nat |] ==> (m#*k = n#*k) \ (m=n | k=0)" + "\k: nat; m: nat; n: nat\ \ (m#*k = n#*k) \ (m=n | k=0)" apply (simp (no_asm_simp) add: Ord_eq_iff_le [of "m#*k"] Ord_eq_iff_le [of m]) apply (auto simp add: Ord_0_lt_iff) done @@ -398,7 +398,7 @@ (** Cancellation law for division **) lemma div_cancel_raw: - "[| 0 (k#*m) div (k#*n) = m div n" + "\0 \ (k#*m) div (k#*n) = m div n" apply (erule_tac i = m in complete_induct) apply (case_tac "x (k#*m) div (k#*n) = m div n" + "\0 < natify(n); 0 < natify(k)\ \ (k#*m) div (k#*n) = m div n" apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)" in div_cancel_raw) apply auto @@ -417,7 +417,7 @@ subsection\More Lemmas about Remainder\ lemma mult_mod_distrib_raw: - "[| k:nat; m:nat; n:nat |] ==> (k#*m) mod (k#*n) = k #* (m mod n)" + "\k:nat; m:nat; n:nat\ \ (k#*m) mod (k#*n) = k #* (m mod n)" apply (case_tac "k=0") apply (simp add: DIVISION_BY_ZERO_MOD) apply (case_tac "n=0") @@ -440,7 +440,7 @@ apply (simp (no_asm) add: mult_commute mod_mult_distrib2) done -lemma mod_add_self2_raw: "n \ nat ==> (m #+ n) mod n = m mod n" +lemma mod_add_self2_raw: "n \ nat \ (m #+ n) mod n = m mod n" apply (subgoal_tac " (n #+ m) mod n = (n #+ m #- n) mod n") apply (simp add: add_commute) apply (subst mod_geq [symmetric], auto) @@ -455,7 +455,7 @@ apply (simp (no_asm_simp) add: add_commute mod_add_self2) done -lemma mod_mult_self1_raw: "k \ nat ==> (m #+ k#*n) mod n = m mod n" +lemma mod_mult_self1_raw: "k \ nat \ (m #+ k#*n) mod n = m mod n" apply (erule nat_induct) apply (simp_all (no_asm_simp) add: add_left_commute [of _ n]) done @@ -470,7 +470,7 @@ done (*Lemma for gcd*) -lemma mult_eq_self_implies_10: "m = m#*n ==> natify(n)=1 | m=0" +lemma mult_eq_self_implies_10: "m = m#*n \ natify(n)=1 | m=0" apply (subgoal_tac "m: nat") prefer 2 apply (erule ssubst) @@ -486,7 +486,7 @@ done lemma less_imp_succ_add [rule_format]: - "[| m \k\nat. n = succ(m#+k)" + "\m \ \k\nat. n = succ(m#+k)" apply (frule lt_nat_in_nat, assumption) apply (erule rev_mp) apply (induct_tac "n") @@ -495,7 +495,7 @@ done lemma less_iff_succ_add: - "[| m: nat; n: nat |] ==> (m (\k\nat. n = succ(m#+k))" + "\m: nat; n: nat\ \ (m (\k\nat. n = succ(m#+k))" by (auto intro: less_imp_succ_add) lemma add_lt_elim2: @@ -510,7 +510,7 @@ subsubsection\More Lemmas About Difference\ lemma diff_is_0_lemma: - "[| m: nat; n: nat |] ==> m #- n = 0 \ m \ n" + "\m: nat; n: nat\ \ m #- n = 0 \ m \ n" apply (rule_tac m = m and n = n in diff_induct, simp_all) done @@ -518,11 +518,11 @@ by (simp add: diff_is_0_lemma [symmetric]) lemma nat_lt_imp_diff_eq_0: - "[| a:nat; b:nat; a a #- b = 0" + "\a:nat; b:nat; a \ a #- b = 0" by (simp add: diff_is_0_iff le_iff) lemma raw_nat_diff_split: - "[| a:nat; b:nat |] ==> + "\a:nat; b:nat\ \ (P(a #- b)) \ ((a < b \P(0)) & (\d\nat. a = b #+ d \ P(d)))" apply (case_tac "a < b") apply (force simp add: nat_lt_imp_diff_eq_0) @@ -540,7 +540,7 @@ text\Difference and less-than\ -lemma diff_lt_imp_lt: "[|(k#-i) < (k#-j); i\nat; j\nat; k\nat|] ==> j(k#-i) < (k#-j); i\nat; j\nat; k\nat\ \ jk; k\nat|] ==> (k#-i) < (k#-j)" +lemma lt_imp_diff_lt: "\jk; k\nat\ \ (k#-i) < (k#-j)" apply (frule le_in_nat, assumption) apply (frule lt_nat_in_nat, assumption) apply (simp split: nat_diff_split, auto) @@ -561,7 +561,7 @@ done -lemma diff_lt_iff_lt: "[|i\k; j\nat; k\nat|] ==> (k#-i) < (k#-j) \ ji\k; j\nat; k\nat\ \ (k#-i) < (k#-j) \ j5 div 2 = \3 and \5 mod 2 = 1; thus \5 = (\3)*2 + 1 *) section\Arithmetic on Binary Integers\ @@ -88,7 +88,7 @@ definition bin_add :: "[i,i]=>i" where - "bin_add(v,w) == bin_adder(v)`w" + "bin_add(v,w) \ bin_adder(v)`w" primrec @@ -144,26 +144,26 @@ (** Type checking **) -lemma integ_of_type [TC]: "w \ bin ==> integ_of(w) \ int" +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" +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" +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" +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" +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" + "v \ bin \ \w\bin. bin_add(v,w) \ bin" apply (unfold bin_add_def) apply (induct_tac "v") apply (rule_tac [3] ballI) @@ -174,7 +174,7 @@ declare bin_add_type [TC] -lemma bin_mult_type [TC]: "[| v \ bin; w \ bin |] ==> bin_mult(v,w) \ bin" +lemma bin_mult_type [TC]: "\v \ bin; w \ bin\ \ bin_mult(v,w) \ bin" by (induct_tac "v", auto) @@ -183,19 +183,19 @@ (*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)" + "\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)" + "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)" + "w \ bin \ integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)" apply (erule bin.induct) apply (auto simp add: zadd_ac elim!: boolE) done @@ -203,7 +203,7 @@ subsubsection\\<^term>\bin_minus\: Unary Negation of Binary Integers\ -lemma integ_of_minus: "w \ bin ==> integ_of(bin_minus(w)) = $- integ_of(w)" +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 @@ -211,18 +211,18 @@ subsubsection\\<^term>\bin_add\: Binary Addition\ -lemma bin_add_Pls [simp]: "w \ bin ==> bin_add(Pls,w) = w" +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" +lemma bin_add_Pls_right: "w \ bin \ bin_add(w,Pls) = w" apply (unfold bin_add_def) apply (erule bin.induct, auto) done -lemma bin_add_Min [simp]: "w \ bin ==> bin_add(Min,w) = bin_pred(w)" +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)" +lemma bin_add_Min_right: "w \ bin \ bin_add(w,Min) = bin_pred(w)" apply (unfold bin_add_def) apply (erule bin.induct, auto) done @@ -234,13 +234,13 @@ by (unfold bin_add_def, simp) lemma bin_add_BIT_BIT [simp]: - "[| w \ bin; y \ bool |] - ==> bin_add(v BIT x, w BIT y) = + "\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 ==> + "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) @@ -250,8 +250,8 @@ (*Subtraction*) lemma diff_integ_of_eq: - "[| v \ bin; w \ bin |] - ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))" + "\v \ bin; w \ bin\ + \ integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))" apply (unfold zdiff_def) apply (simp add: integ_of_add integ_of_minus) done @@ -260,8 +260,8 @@ 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)" + "\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) @@ -294,16 +294,16 @@ (** extra rules for bin_add **) -lemma bin_add_BIT_11: "w \ bin ==> bin_add(v BIT 1, w BIT 1) = +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) = +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)" +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 **) @@ -359,8 +359,8 @@ (** Equals (=) **) lemma eq_integ_of_eq: - "[| v \ bin; w \ bin |] - ==> ((integ_of(v)) = integ_of(w)) \ + "\v \ bin; w \ bin\ + \ ((integ_of(v)) = integ_of(w)) \ iszero (integ_of (bin_add (v, bin_minus(w))))" apply (unfold iszero_def) apply (simp add: zcompare_rls integ_of_add integ_of_minus) @@ -370,14 +370,14 @@ by (unfold iszero_def, simp) -lemma nonzero_integ_of_Min: "~ iszero (integ_of(Min))" +lemma nonzero_integ_of_Min: "\ iszero (integ_of(Min))" apply (unfold 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)))" + "\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 @@ -388,10 +388,10 @@ done lemma iszero_integ_of_0: - "w \ bin ==> iszero (integ_of (w BIT 0)) \ iszero (integ_of(w))" + "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))" +lemma iszero_integ_of_1: "w \ bin \ \ iszero (integ_of (w BIT 1))" by (simp only: iszero_integ_of_BIT, blast) @@ -399,22 +399,22 @@ (** Less-than (<) **) lemma less_integ_of_eq_neg: - "[| v \ bin; w \ bin |] - ==> integ_of(v) $< integ_of(w) + "\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) apply (simp add: integ_of_minus integ_of_add) done -lemma not_neg_integ_of_Pls: "~ znegative (integ_of(Pls))" +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))" + "\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 @@ -430,7 +430,7 @@ (** Less-than-or-equals (<=) **) lemma le_integ_of_eq_not_less: - "(integ_of(x) $\ (integ_of(w))) \ ~ (integ_of(w) $< (integ_of(x)))" + "(integ_of(x) $\ (integ_of(w))) \ \ (integ_of(w) $< (integ_of(x)))" by (simp add: not_zless_iff_zle [THEN iff_sym]) @@ -461,7 +461,7 @@ (*For making a minimal simpset, one must include these default simprules - of thy. Also include simp_thms, or at least (~False)=True*) + 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 @@ -485,25 +485,25 @@ (** 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)" + "\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)" + "\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)" + "\v \ bin; w \ bin\ + \ integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)" apply (unfold 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)) = + "\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) @@ -523,10 +523,10 @@ lemma zdiff_self [simp]: "x $- x = #0" by (simp add: zdiff_def) -lemma znegative_iff_zless_0: "k \ int ==> znegative(k) \ k $< #0" +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)" +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" @@ -535,21 +535,21 @@ 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" +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" +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" +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)" +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 @@ -559,7 +559,7 @@ 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)" +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] @@ -570,7 +570,7 @@ (** 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)" +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) @@ -589,7 +589,7 @@ 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))" + "True \ (integ_of(w) = x) \ (x = integ_of(w))" by auto *) @@ -639,7 +639,7 @@ 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)" +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)" @@ -709,15 +709,15 @@ schematic_goal "y2 $+ ?x42 = y $+ y2" apply simp oops -lemma "oo : int ==> l $+ (l $+ #2) $+ oo = oo" 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 "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 @@ -728,7 +728,7 @@ 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 "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 diff -r f2094906e491 -r e44d86131648 src/ZF/Bool.thy --- a/src/ZF/Bool.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Bool.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,31 +9,31 @@ abbreviation one (\1\) where - "1 == succ(0)" + "1 \ succ(0)" abbreviation two (\2\) where - "2 == succ(1)" + "2 \ succ(1)" text\2 is equal to bool, but is used as a number rather than a type.\ -definition "bool == {0,1}" +definition "bool \ {0,1}" -definition "cond(b,c,d) == if(b=1,c,d)" +definition "cond(b,c,d) \ if(b=1,c,d)" -definition "not(b) == cond(b,0,1)" +definition "not(b) \ cond(b,0,1)" definition "and" :: "[i,i]=>i" (infixl \and\ 70) where - "a and b == cond(a,b,0)" + "a and b \ cond(a,b,0)" definition or :: "[i,i]=>i" (infixl \or\ 65) where - "a or b == cond(a,1,b)" + "a or b \ cond(a,1,b)" definition xor :: "[i,i]=>i" (infixl \xor\ 65) where - "a xor b == cond(a,not(b),b)" + "a xor b \ cond(a,not(b),b)" lemmas bool_defs = bool_def cond_def @@ -52,11 +52,11 @@ lemma one_not_0: "1\0" by (simp add: bool_defs ) -(** 1=0 ==> R **) +(** 1=0 \ R **) lemmas one_neq_0 = one_not_0 [THEN notE] lemma boolE: - "[| c: bool; c=1 ==> P; c=0 ==> P |] ==> P" + "\c: bool; c=1 \ P; c=0 \ P\ \ P" by (simp add: bool_defs, blast) (** cond **) @@ -69,17 +69,17 @@ lemma cond_0 [simp]: "cond(0,c,d) = d" by (simp add: bool_defs ) -lemma cond_type [TC]: "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)" +lemma cond_type [TC]: "\b: bool; c: A(1); d: A(0)\ \ cond(b,c,d): A(b)" by (simp add: bool_defs, blast) (*For Simp_tac and Blast_tac*) -lemma cond_simple_type: "[| b: bool; c: A; d: A |] ==> cond(b,c,d): A" +lemma cond_simple_type: "\b: bool; c: A; d: A\ \ cond(b,c,d): A" by (simp add: bool_defs ) -lemma def_cond_1: "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c" +lemma def_cond_1: "\\b. j(b)\cond(b,c,d)\ \ j(1) = c" by simp -lemma def_cond_0: "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d" +lemma def_cond_0: "\\b. j(b)\cond(b,c,d)\ \ j(0) = d" by simp lemmas not_1 = not_def [THEN def_cond_1, simp] @@ -94,16 +94,16 @@ lemmas xor_1 = xor_def [THEN def_cond_1, simp] lemmas xor_0 = xor_def [THEN def_cond_0, simp] -lemma not_type [TC]: "a:bool ==> not(a) \ bool" +lemma not_type [TC]: "a:bool \ not(a) \ bool" by (simp add: not_def) -lemma and_type [TC]: "[| a:bool; b:bool |] ==> a and b \ bool" +lemma and_type [TC]: "\a:bool; b:bool\ \ a and b \ bool" by (simp add: and_def) -lemma or_type [TC]: "[| a:bool; b:bool |] ==> a or b \ bool" +lemma or_type [TC]: "\a:bool; b:bool\ \ a or b \ bool" by (simp add: or_def) -lemma xor_type [TC]: "[| a:bool; b:bool |] ==> a xor b \ bool" +lemma xor_type [TC]: "\a:bool; b:bool\ \ a xor b \ bool" by (simp add: xor_def) lemmas bool_typechecks = bool_1I bool_0I cond_type not_type and_type @@ -111,49 +111,49 @@ subsection\Laws About 'not'\ -lemma not_not [simp]: "a:bool ==> not(not(a)) = a" +lemma not_not [simp]: "a:bool \ not(not(a)) = a" by (elim boolE, auto) -lemma not_and [simp]: "a:bool ==> not(a and b) = not(a) or not(b)" +lemma not_and [simp]: "a:bool \ not(a and b) = not(a) or not(b)" by (elim boolE, auto) -lemma not_or [simp]: "a:bool ==> not(a or b) = not(a) and not(b)" +lemma not_or [simp]: "a:bool \ not(a or b) = not(a) and not(b)" by (elim boolE, auto) subsection\Laws About 'and'\ -lemma and_absorb [simp]: "a: bool ==> a and a = a" +lemma and_absorb [simp]: "a: bool \ a and a = a" by (elim boolE, auto) -lemma and_commute: "[| a: bool; b:bool |] ==> a and b = b and a" +lemma and_commute: "\a: bool; b:bool\ \ a and b = b and a" by (elim boolE, auto) -lemma and_assoc: "a: bool ==> (a and b) and c = a and (b and c)" +lemma and_assoc: "a: bool \ (a and b) and c = a and (b and c)" by (elim boolE, auto) -lemma and_or_distrib: "[| a: bool; b:bool; c:bool |] ==> +lemma and_or_distrib: "\a: bool; b:bool; c:bool\ \ (a or b) and c = (a and c) or (b and c)" by (elim boolE, auto) subsection\Laws About 'or'\ -lemma or_absorb [simp]: "a: bool ==> a or a = a" +lemma or_absorb [simp]: "a: bool \ a or a = a" by (elim boolE, auto) -lemma or_commute: "[| a: bool; b:bool |] ==> a or b = b or a" +lemma or_commute: "\a: bool; b:bool\ \ a or b = b or a" by (elim boolE, auto) -lemma or_assoc: "a: bool ==> (a or b) or c = a or (b or c)" +lemma or_assoc: "a: bool \ (a or b) or c = a or (b or c)" by (elim boolE, auto) -lemma or_and_distrib: "[| a: bool; b: bool; c: bool |] ==> +lemma or_and_distrib: "\a: bool; b: bool; c: bool\ \ (a and b) or c = (a or c) and (b or c)" by (elim boolE, auto) definition bool_of_o :: "o=>i" where - "bool_of_o(P) == (if P then 1 else 0)" + "bool_of_o(P) \ (if P then 1 else 0)" lemma [simp]: "bool_of_o(True) = 1" by (simp add: bool_of_o_def) @@ -167,7 +167,7 @@ lemma [simp]: "(bool_of_o(P) = 1) \ P" by (simp add: bool_of_o_def) -lemma [simp]: "(bool_of_o(P) = 0) \ ~P" +lemma [simp]: "(bool_of_o(P) = 0) \ \P" by (simp add: bool_of_o_def) end diff -r f2094906e491 -r e44d86131648 src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Cardinal.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,31 +10,31 @@ definition (*least ordinal operator*) Least :: "(i=>o) => i" (binder \\ \ 10) where - "Least(P) == THE i. Ord(i) & P(i) & (\j. j ~P(j))" + "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)" + "A \ B \ \f. f \ bij(A,B)" definition lepoll :: "[i,i] => o" (infixl \\\ 50) where - "A \ B == \f. f \ inj(A,B)" + "A \ B \ \f. f \ inj(A,B)" definition lesspoll :: "[i,i] => o" (infixl \\\ 50) where - "A \ B == A \ B & ~(A \ B)" + "A \ B \ A \ B & \(A \ B)" definition cardinal :: "i=>i" (\|_|\) where - "|A| == (\ i. i \ A)" + "|A| \ (\ i. i \ A)" definition Finite :: "i=>o" where - "Finite(A) == \n\nat. A \ n" + "Finite(A) \ \n\nat. A \ n" definition Card :: "i=>o" where - "Card(i) == (i = |i|)" + "Card(i) \ (i = |i|)" subsection\The Schroeder-Bernstein Theorem\ @@ -47,7 +47,7 @@ lemma Banach_last_equation: "g \ Y->X - ==> g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) = + \ 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]) @@ -55,7 +55,7 @@ done lemma decomposition: - "[| f \ X->Y; g \ Y->X |] ==> + "\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" @@ -67,18 +67,18 @@ done lemma schroeder_bernstein: - "[| f \ inj(X,Y); g \ inj(Y,X) |] ==> \h. h \ bij(X,Y)" + "\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!! *) + is forced by the context\*) done (** Equipollence is an equivalence relation **) -lemma bij_imp_eqpoll: "f \ bij(A,B) ==> A \ B" +lemma bij_imp_eqpoll: "f \ bij(A,B) \ A \ B" apply (unfold eqpoll_def) apply (erule exI) done @@ -86,20 +86,20 @@ (*A \ A*) lemmas eqpoll_refl = id_bij [THEN bij_imp_eqpoll, simp] -lemma eqpoll_sym: "X \ Y ==> Y \ X" +lemma eqpoll_sym: "X \ Y \ Y \ X" apply (unfold eqpoll_def) apply (blast intro: bij_converse_bij) done lemma eqpoll_trans [trans]: - "[| X \ Y; Y \ Z |] ==> X \ Z" + "\X \ Y; Y \ Z\ \ X \ Z" apply (unfold eqpoll_def) apply (blast intro: comp_bij) done (** Le-pollence is a partial ordering **) -lemma subset_imp_lepoll: "X<=Y ==> X \ Y" +lemma subset_imp_lepoll: "X<=Y \ X \ Y" apply (unfold lepoll_def) apply (rule exI) apply (erule id_subset_inj) @@ -109,35 +109,35 @@ lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll] -lemma eqpoll_imp_lepoll: "X \ Y ==> X \ Y" +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" +lemma lepoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" apply (unfold lepoll_def) apply (blast intro: comp_inj) done -lemma eq_lepoll_trans [trans]: "[| X \ Y; Y \ Z |] ==> X \ Z" +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" +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" +lemma eqpollI: "\X \ Y; Y \ X\ \ X \ Y" apply (unfold lepoll_def eqpoll_def) apply (elim exE) apply (rule schroeder_bernstein, assumption+) done lemma eqpollE: - "[| X \ Y; [| X \ Y; Y \ X |] ==> P |] ==> P" + "\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" +lemma lepoll_0_is_0: "A \ 0 \ A = 0" apply (unfold lepoll_def inj_def) apply (blast dest: apply_type) done @@ -149,20 +149,20 @@ by (blast intro: lepoll_0_is_0 lepoll_refl) lemma Un_lepoll_Un: - "[| A \ B; C \ D; B \ D = 0 |] ==> A \ C \ B \ D" + "\A \ B; C \ D; B \ D = 0\ \ A \ C \ B \ D" apply (unfold lepoll_def) apply (blast intro: inj_disjoint_Un) done -(*A \ 0 ==> A=0*) +(*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" + "\A \ B; C \ D; A \ C = 0; B \ D = 0\ + \ A \ C \ B \ D" apply (unfold eqpoll_def) apply (blast intro: bij_disjoint_Un) done @@ -170,16 +170,16 @@ subsection\lesspoll: contributions by Krzysztof Grabczewski\ -lemma lesspoll_not_refl: "~ (i \ i)" +lemma lesspoll_not_refl: "\ (i \ i)" by (simp add: lesspoll_def) -lemma lesspoll_irrefl [elim!]: "i \ i ==> P" +lemma lesspoll_irrefl [elim!]: "i \ i \ P" by (simp add: lesspoll_def) -lemma lesspoll_imp_lepoll: "A \ B ==> A \ B" +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)" +lemma lepoll_well_ord: "\A \ B; well_ord(B,r)\ \ \s. well_ord(A,s)" apply (unfold lepoll_def) apply (blast intro: well_ord_rvimage) done @@ -207,36 +207,36 @@ (** Variations on transitivity **) lemma lesspoll_trans [trans]: - "[| X \ Y; Y \ Z |] ==> X \ Z" + "\X \ Y; Y \ Z\ \ X \ Z" apply (unfold lesspoll_def) apply (blast elim!: eqpollE intro: eqpollI lepoll_trans) done lemma lesspoll_trans1 [trans]: - "[| X \ Y; Y \ Z |] ==> X \ Z" + "\X \ Y; Y \ Z\ \ X \ Z" apply (unfold lesspoll_def) apply (blast elim!: eqpollE intro: eqpollI lepoll_trans) done lemma lesspoll_trans2 [trans]: - "[| X \ Y; Y \ Z |] ==> X \ Z" + "\X \ Y; Y \ Z\ \ X \ Z" apply (unfold lesspoll_def) apply (blast elim!: eqpollE intro: eqpollI lepoll_trans) done lemma eq_lesspoll_trans [trans]: - "[| X \ Y; Y \ Z |] ==> X \ Z" + "\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" + "\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" + "\P(i); Ord(i); \x. x \P(x)\ \ (\ x. P(x)) = i" apply (unfold Least_def) apply (rule the_equality, blast) apply (elim conjE) @@ -254,7 +254,7 @@ case True thus ?thesis . next case False - hence "\x. x \ i \ ~P(x)" using step + hence "\x. x \ i \ \P(x)" using step by blast hence "(\ a. P(a)) = i" using step by (blast intro: Least_equality ltD) @@ -278,7 +278,7 @@ case True thus ?thesis . next case False - hence "\x. x \ i \ ~ (\ a. P(a)) \ i" using step + 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) @@ -291,19 +291,19 @@ qed (*\ really is the smallest*) -lemma less_LeastE: "[| P(i); i < (\ x. P(x)) |] ==> Q" +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))" + "\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" + "\\ (\i. Ord(i) & P(i))\ \ (\ x. P(x)) = 0" apply (unfold Least_def) apply (rule the_0, blast) done @@ -326,12 +326,12 @@ 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))" +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 +(*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|" +lemma cardinal_cong: "X \ Y \ |X| = |Y|" apply (unfold eqpoll_def cardinal_def) apply (rule Least_cong) apply (blast intro: comp_bij bij_converse_bij) @@ -345,7 +345,7 @@ by (best intro: LeastI Ord_ordertype ordermap_bij bij_converse_bij bij_imp_eqpoll r) qed -(* @{term"Ord(A) ==> |A| \ A"} *) +(* @{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|" @@ -362,36 +362,36 @@ qed lemma well_ord_cardinal_eqpoll_iff: - "[| well_ord(X,r); well_ord(Y,s) |] ==> |X| = |Y| \ X \ Y" + "\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" +lemma Ord_cardinal_le: "Ord(i) \ |i| \ i" apply (unfold cardinal_def) apply (erule eqpoll_refl [THEN Least_le]) done -lemma Card_cardinal_eq: "Card(K) ==> |K| = K" +lemma Card_cardinal_eq: "Card(K) \ |K| = K" apply (unfold 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)" +(* 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) apply (subst Least_equality) apply (blast intro: eqpoll_refl)+ done -lemma Card_is_Ord: "Card(i) ==> Ord(i)" +lemma Card_is_Ord: "Card(i) \ Ord(i)" apply (unfold Card_def cardinal_def) apply (erule ssubst) apply (rule Ord_Least) done -lemma Card_cardinal_le: "Card(K) ==> K \ |K|" +lemma Card_cardinal_le: "Card(K) \ K \ |K|" apply (simp (no_asm_simp) add: Card_is_Ord Card_cardinal_eq) done @@ -401,7 +401,7 @@ done text\The cardinals are the initial ordinals.\ -lemma Card_iff_initial: "Card(K) \ Ord(K) & (\j. j ~ j \ K)" +lemma Card_iff_initial: "Card(K) \ Ord(K) & (\j. j \ j \ K)" proof - { fix j assume K: "Card(K)" "j \ K" @@ -416,7 +416,7 @@ by (blast intro: CardI Card_is_Ord) qed -lemma lt_Card_imp_lesspoll: "[| Card(a); i i \ a" +lemma lt_Card_imp_lesspoll: "\Card(a); i \ i \ a" apply (unfold lesspoll_def) apply (drule Card_iff_initial [THEN iffD1]) apply (blast intro!: leI [THEN le_imp_lepoll]) @@ -427,7 +427,7 @@ apply (blast elim!: ltE) done -lemma Card_Un: "[| Card(K); Card(L) |] ==> Card(K \ L)" +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]) @@ -491,19 +491,19 @@ 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" +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" +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)" +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)" +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*) @@ -526,21 +526,21 @@ thus ?thesis by simp qed -lemma lepoll_cardinal_le: "[| A \ i; Ord(i) |] ==> |A| \ i" +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" +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" +lemma lesspoll_imp_eqpoll: "\A \ i; Ord(i)\ \ |A| \ A" apply (unfold lesspoll_def) apply (blast intro: lepoll_Ord_imp_eqpoll) done -lemma cardinal_subset_Ord: "[|A<=i; Ord(i)|] ==> |A| \ i" +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) @@ -549,7 +549,7 @@ subsection\The finite cardinals\ lemma cons_lepoll_consD: - "[| cons(u,A) \ cons(v,B); u\A; v\B |] ==> A \ B" + "\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) @@ -562,13 +562,13 @@ apply blast done -lemma cons_eqpoll_consD: "[| cons(u,A) \ cons(v,B); u\A; v\B |] ==> A \ B" +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" +lemma succ_lepoll_succD: "succ(m) \ succ(n) \ m \ n" apply (unfold succ_def) apply (erule cons_lepoll_consD) apply (rule mem_not_refl)+ @@ -576,7 +576,7 @@ lemma nat_lepoll_imp_le: - "m \ nat ==> n \ nat \ m \ n \ m \ n" + "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 @@ -591,7 +591,7 @@ qed qed -lemma nat_eqpoll_iff: "[| m \ nat; n \ nat |] ==> m \ n \ m = n" +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) @@ -616,11 +616,11 @@ (*Part of Kunen's Lemma 10.6*) -lemma succ_lepoll_natE: "[| succ(n) \ n; n \ nat |] ==> P" +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" + "\n \ nat; nat \ X\ \ \Y. Y \ X & n \ Y" apply (unfold lepoll_def eqpoll_def) apply (fast del: subsetI subsetCE intro!: subset_SIs @@ -649,22 +649,22 @@ qed lemma lesspoll_succ_imp_lepoll: - "[| A \ succ(m); m \ nat |] ==> A \ m" + "\A \ succ(m); m \ nat\ \ A \ m" apply (unfold 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" +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)" +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" +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] @@ -676,7 +676,7 @@ (*This implies Kunen's Lemma 10.6*) lemma lt_not_lepoll: - assumes n: "n nat" shows "~ i \ n" + assumes n: "n nat" shows "\ i \ n" proof - { assume i: "i \ n" have "succ(n) \ i" using n @@ -700,8 +700,8 @@ 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) + 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 @@ -710,7 +710,7 @@ proof - { fix i assume i: "i < nat" "i \ nat" - hence "~ nat \ i" + hence "\ nat \ i" by (simp add: lt_def lt_not_lepoll) hence False using i by (simp add: eqpoll_iff) @@ -721,12 +721,12 @@ qed (*Allows showing that |i| is a limit cardinal*) -lemma nat_le_cardinal: "nat \ i ==> nat \ |i|" +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" +lemma n_lesspoll_nat: "n \ nat \ n \ nat" by (blast intro: Ord_nat Card_nat ltI lt_Card_imp_lesspoll) @@ -735,7 +735,7 @@ (*Congruence law for cons under equipollence*) lemma cons_lepoll_cong: - "[| A \ B; b \ B |] ==> cons(a,A) \ cons(b,B)" + "\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) @@ -745,15 +745,15 @@ done lemma cons_eqpoll_cong: - "[| A \ B; a \ A; b \ B |] ==> cons(a,A) \ cons(b,B)" + "\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" + "\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" + "\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" @@ -766,7 +766,7 @@ apply (simp (no_asm) add: nat_into_Card [THEN Card_cardinal_eq]) done -lemma not_0_is_lepoll_1: "A \ 0 ==> 1 \ A" +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) @@ -774,26 +774,26 @@ done (*Congruence law for succ under equipollence*) -lemma succ_eqpoll_cong: "A \ B ==> succ(A) \ succ(B)" +lemma succ_eqpoll_cong: "A \ B \ succ(A) \ succ(B)" apply (unfold 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" +lemma sum_eqpoll_cong: "\A \ C; B \ D\ \ A+B \ C+D" apply (unfold 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" + "\A \ C; B \ D\ \ A*B \ C*D" apply (unfold eqpoll_def) apply (blast intro!: prod_bij) done lemma inj_disjoint_eqpoll: - "[| f \ inj(A,B); A \ B = 0 |] ==> A \ (B - range(f)) \ B" + "\f \ inj(A,B); A \ B = 0\ \ A \ (B - range(f)) \ B" apply (unfold eqpoll_def) apply (rule exI) apply (rule_tac c = "%x. if x \ A then f`x else x" @@ -814,7 +814,7 @@ 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" + "\a \ A; A \ succ(n)\ \ A - {a} \ n" apply (unfold succ_def) apply (rule cons_lepoll_consD) apply (rule_tac [3] mem_not_refl) @@ -834,12 +834,12 @@ by (blast intro: cons_lepoll_consD mem_irrefl) qed -lemma Diff_sing_eqpoll: "[| a \ A; A \ succ(n) |] ==> A - {a} \ n" +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}" +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) @@ -854,12 +854,12 @@ done lemma well_ord_Un: - "[| well_ord(X,R); well_ord(Y,S) |] ==> \T. well_ord(X \ Y, T)" + "\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" +lemma disj_Un_eqpoll_sum: "A \ B = 0 \ A \ B \ A + B" apply (unfold 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) @@ -869,7 +869,7 @@ subsection \Finite and infinite sets\ -lemma eqpoll_imp_Finite_iff: "A \ B ==> Finite(A) \ Finite(B)" +lemma eqpoll_imp_Finite_iff: "A \ B \ Finite(A) \ Finite(B)" apply (unfold Finite_def) apply (blast intro: eqpoll_trans eqpoll_sym) done @@ -879,7 +879,7 @@ apply (blast intro!: eqpoll_refl nat_0I) done -lemma Finite_cons: "Finite(x) ==> Finite(cons(y,x))" +lemma Finite_cons: "Finite(x) \ Finite(cons(y,x))" apply (unfold Finite_def) apply (case_tac "y \ x") apply (simp add: cons_absorb) @@ -889,7 +889,7 @@ apply (simp (no_asm_simp) add: succ_def cons_eqpoll_cong mem_not_refl) done -lemma Finite_succ: "Finite(x) ==> Finite(succ(x))" +lemma Finite_succ: "Finite(x) \ Finite(succ(x))" apply (unfold succ_def) apply (erule Finite_cons) done @@ -911,7 +911,7 @@ qed lemma lesspoll_nat_is_Finite: - "A \ nat ==> Finite(A)" + "A \ nat \ Finite(A)" apply (unfold Finite_def) apply (blast dest: ltD lesspoll_cardinal_lt lesspoll_imp_eqpoll [THEN eqpoll_sym]) @@ -936,13 +936,13 @@ 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)" +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" + "\Ord(i); \ Finite(i)\ \ nat \ i" apply (unfold Finite_def) apply (erule Ord_nat [THEN [2] Ord_linear2]) prefer 2 apply assumption @@ -950,19 +950,19 @@ done lemma Finite_imp_well_ord: - "Finite(A) ==> \r. well_ord(A,r)" + "Finite(A) \ \r. well_ord(A,r)" apply (unfold 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" +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" +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)" + "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]) @@ -978,10 +978,10 @@ apply (simp add: cons_Diff) done -lemma Finite_Fin: "[| Finite(A); A \ X |] ==> A \ Fin(X)" +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)" +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") @@ -997,18 +997,18 @@ apply (blast intro: bij_converse_bij [THEN bij_is_fun, THEN apply_type]) done -lemma Finite_into_Fin: "Finite(A) ==> A \ Fin(A)" +lemma Finite_into_Fin: "Finite(A) \ A \ Fin(A)" apply (unfold Finite_def) apply (blast intro: Fin_lemma) done -lemma Fin_into_Finite: "A \ Fin(U) ==> Finite(A)" +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)" +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] @@ -1018,7 +1018,7 @@ by (blast intro: subset_Finite Finite_Un) text\The converse must hold too.\ -lemma Finite_Union: "[| \y\X. Finite(y); Finite(X) |] ==> Finite(\(X))" +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) @@ -1027,15 +1027,15 @@ (* 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)" +"\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)" +(*Sidi Ehmety. The contrapositive says \Finite(A) \ \Finite(A-{a}) *) +lemma Diff_sing_Finite: "Finite(A - {a}) \ Finite(A)" apply (unfold Finite_def) apply (case_tac "a \ A") apply (subgoal_tac [2] "A-{a}=A", auto) @@ -1046,8 +1046,8 @@ 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)" + \\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") @@ -1055,12 +1055,12 @@ apply (drule Diff_sing_Finite, auto) done -lemma Finite_RepFun: "Finite(A) ==> Finite(RepFun(A,f))" +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)" + "\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) @@ -1078,15 +1078,15 @@ 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)" + "(\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))" +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)" +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) @@ -1103,7 +1103,7 @@ (*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)))" +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 @@ -1125,15 +1125,15 @@ qed qed -lemma nat_well_ord_converse_Memrel: "n \ nat ==> well_ord(n,converse(Memrel(n)))" +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))" + "\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 @@ -1153,16 +1153,16 @@ qed lemma Finite_well_ord_converse: - "[| Finite(A); well_ord(A,r) |] ==> well_ord(A,converse(r))" + "\Finite(A); well_ord(A,r)\ \ well_ord(A,converse(r))" apply (unfold 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)" +lemma nat_into_Finite: "n \ nat \ Finite(n)" by (auto simp add: Finite_def intro: eqpoll_refl) -lemma nat_not_Finite: "~ Finite(nat)" +lemma nat_not_Finite: "\ Finite(nat)" proof - { fix n assume n: "n \ nat" "nat \ n" diff -r f2094906e491 -r e44d86131648 src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/CardinalArith.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,19 +9,19 @@ definition InfCard :: "i=>o" where - "InfCard(i) == Card(i) & nat \ i" + "InfCard(i) \ Card(i) & nat \ i" definition cmult :: "[i,i]=>i" (infixl \\\ 70) where - "i \ j == |i*j|" + "i \ j \ |i*j|" definition cadd :: "[i,i]=>i" (infixl \\\ 65) where - "i \ j == |i+j|" + "i \ j \ |i+j|" definition csquare_rel :: "i=>i" where - "csquare_rel(K) == + "csquare_rel(K) \ rvimage(K*K, lam :K*K. y, x, y>, rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))" @@ -30,14 +30,14 @@ 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) == + "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) & K \ L. Card(L) & K j \ \A" by blast qed -lemma Card_UN: "(!!x. x \ A ==> Card(K(x))) ==> Card(\x\A. K(x))" +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(\xx. x \ A \ Card(K(x))) \ Card(\x K |] ==> b \ K" +lemma in_Card_imp_lesspoll: "\Card(K); b \ K\ \ b \ K" apply (unfold lesspoll_def) apply (simp add: Card_iff_initial) apply (fast intro!: le_imp_lepoll ltI leI) @@ -130,7 +130,7 @@ apply (rule bij_0_sum) done -lemma cadd_0 [simp]: "Card(K) ==> 0 \ K = K" +lemma cadd_0 [simp]: "Card(K) \ 0 \ K = K" apply (unfold cadd_def) apply (simp add: sum_0_eqpoll [THEN cardinal_cong] Card_cardinal_eq) done @@ -160,7 +160,7 @@ subsubsection\Monotonicity of addition\ lemma sum_lepoll_mono: - "[| A \ C; B \ D |] ==> A + B \ C + D" + "\A \ C; B \ D\ \ A + B \ C + D" apply (unfold lepoll_def) apply (elim exE) apply (rule_tac x = "\z\A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI) @@ -170,7 +170,7 @@ done lemma cadd_le_mono: - "[| K' \ K; L' \ L |] ==> (K' \ L') \ (K \ L)" + "\K' \ K; L' \ L\ \ (K' \ L') \ (K \ L)" apply (unfold cadd_def) apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule well_ord_lepoll_imp_cardinal_le) @@ -292,7 +292,7 @@ apply (rule singleton_prod_bij [THEN bij_converse_bij]) done -lemma cmult_1 [simp]: "Card(K) ==> 1 \ K = K" +lemma cmult_1 [simp]: "Card(K) \ 1 \ K = K" apply (unfold cmult_def succ_def) apply (simp add: prod_singleton_eqpoll [THEN cardinal_cong] Card_cardinal_eq) done @@ -305,7 +305,7 @@ done (*Could probably weaken the premise to well_ord(K,r), or remove using AC*) -lemma cmult_square_le: "Card(K) ==> K \ K \ K" +lemma cmult_square_le: "Card(K) \ K \ K \ K" apply (unfold cmult_def) apply (rule le_trans) apply (rule_tac [2] well_ord_lepoll_imp_cardinal_le) @@ -316,14 +316,14 @@ subsubsection\Multiplication by a non-zero cardinal\ -lemma prod_lepoll_self: "b \ B ==> A \ A*B" +lemma prod_lepoll_self: "b \ B \ A \ A*B" apply (unfold lepoll_def inj_def) apply (rule_tac x = "\x\A. " 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)" + "\Card(K); Ord(L); 0 \ K \ (K \ L)" apply (unfold cmult_def) apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_cardinal_le]) apply assumption @@ -334,7 +334,7 @@ subsubsection\Monotonicity of multiplication\ lemma prod_lepoll_mono: - "[| A \ C; B \ D |] ==> A * B \ C * D" + "\A \ C; B \ D\ \ A * B \ C * D" apply (unfold lepoll_def) apply (elim exE) apply (rule_tac x = "lam :A*B. " in exI) @@ -344,7 +344,7 @@ done lemma cmult_le_mono: - "[| K' \ K; L' \ L |] ==> (K' \ L') \ (K \ L)" + "\K' \ K; L' \ L\ \ (K' \ L') \ (K \ L)" apply (unfold cmult_def) apply (safe dest!: le_subset_iff [THEN iffD1]) apply (rule well_ord_lepoll_imp_cardinal_le) @@ -365,7 +365,7 @@ (*Unconditional version requires AC*) lemma cmult_succ_lemma: - "[| Ord(m); Ord(n) |] ==> succ(m) \ n = n \ (m \ n)" + "\Ord(m); Ord(n)\ \ succ(m) \ n = n \ (m \ n)" apply (unfold cmult_def cadd_def) apply (rule prod_succ_eqpoll [THEN cardinal_cong, THEN trans]) apply (rule cardinal_cong [symmetric]) @@ -373,12 +373,12 @@ apply (blast intro: well_ord_rmult well_ord_Memrel) done -lemma nat_cmult_eq_mult: "[| m \ nat; n \ nat |] ==> m \ n = m#*n" +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" +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: @@ -391,7 +391,7 @@ finally show "B+B \ C*B" . qed -lemma lepoll_imp_sum_lepoll_prod: "[| A \ B; 2 \ A |] ==> A+B \ A*B" +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) @@ -401,7 +401,7 @@ \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" +lemma nat_cons_lepoll: "nat \ A \ cons(u,A) \ A" apply (unfold lepoll_def) apply (erule exE) apply (rule_tac x = @@ -419,13 +419,13 @@ inj_converse_fun [THEN apply_funtype]) done -lemma nat_cons_eqpoll: "nat \ A ==> cons(u,A) \ A" +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" +lemma nat_succ_eqpoll: "nat \ A \ succ(A) \ A" apply (unfold succ_def) apply (erule subset_imp_lepoll [THEN nat_cons_eqpoll]) done @@ -435,19 +435,19 @@ apply (blast intro: Card_nat le_refl Card_is_Ord) done -lemma InfCard_is_Card: "InfCard(K) ==> Card(K)" +lemma InfCard_is_Card: "InfCard(K) \ Card(K)" apply (unfold InfCard_def) apply (erule conjunct1) done lemma InfCard_Un: - "[| InfCard(K); Card(L) |] ==> InfCard(K \ L)" + "\InfCard(K); Card(L)\ \ InfCard(K \ L)" apply (unfold 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)" +lemma InfCard_is_Limit: "InfCard(K) \ Limit(K)" apply (unfold InfCard_def) apply (erule conjE) apply (frule Card_is_Ord) @@ -467,7 +467,7 @@ (*A general fact about ordermap*) lemma ordermap_eqpoll_pred: - "[| well_ord(A,r); x \ A |] ==> ordermap(A,r)`x \ Order.pred(A,x,r)" + "\well_ord(A,r); x \ A\ \ ordermap(A,r)`x \ Order.pred(A,x,r)" apply (unfold eqpoll_def) apply (rule exI) apply (simp add: ordermap_eq_image well_ord_is_wf) @@ -491,7 +491,7 @@ subsubsection\Characterising initial segments of the well-ordering\ lemma csquareD: - "[| <, > \ csquare_rel(K); x x \ z & y \ z" + "\<, > \ csquare_rel(K); x \ x \ z & y \ z" apply (unfold csquare_rel_def) apply (erule rev_mp) apply (elim ltE) @@ -501,14 +501,14 @@ done lemma pred_csquare_subset: - "z Order.pred(K*K, , csquare_rel(K)) \ succ(z)*succ(z)" + "z Order.pred(K*K, , csquare_rel(K)) \ succ(z)*succ(z)" apply (unfold Order.pred_def) apply (safe del: SigmaI dest!: csquareD) apply (unfold lt_def, auto) done lemma csquare_ltI: - "[| x <, > \ csquare_rel(K)" + "\x \ <, > \ csquare_rel(K)" apply (unfold csquare_rel_def) apply (subgoal_tac "x z; y \ z; z <, > \ csquare_rel(K) | x=z & y=z" + "\x \ z; y \ z; z \ <, > \ csquare_rel(K) | x=z & y=z" apply (unfold csquare_rel_def) apply (subgoal_tac "xThe cardinality of initial segments\ lemma ordermap_z_lt: - "[| Limit(K); x y) |] ==> + "\Limit(K); x y)\ \ ordermap(K*K, csquare_rel(K)) ` < ordermap(K*K, csquare_rel(K)) ` " apply (subgoal_tac "z K \ K \ K" +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)" +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" +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) @@ -689,7 +689,7 @@ done (*Corollary 10.13 (1), for cardinal multiplication*) -lemma InfCard_cmult_eq: "[| InfCard(K); InfCard(L) |] ==> K \ L = K \ L" +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]) @@ -698,13 +698,13 @@ subset_Un_iff2 [THEN iffD1] le_imp_subset) done -lemma InfCard_cdouble_eq: "InfCard(K) ==> K \ K = K" +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" +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) @@ -713,7 +713,7 @@ apply (simp add: InfCard_cdouble_eq) done -lemma InfCard_cadd_eq: "[| InfCard(K); InfCard(L) |] ==> K \ L = K \ L" +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]) @@ -723,7 +723,7 @@ (*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. + might be InfCard(K) \ |list(K)| = K. *) subsection\For Every Cardinal Number There Exists A Greater One\ @@ -751,7 +751,7 @@ done (*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*) -lemma K_lt_jump_cardinal: "Ord(K) ==> K < jump_cardinal(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) @@ -763,9 +763,9 @@ (*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)" + "\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]) @@ -787,7 +787,7 @@ subsection\Basic Properties of Successor Cardinals\ -lemma csucc_basic: "Ord(K) ==> Card(csucc(K)) & K < csucc(K)" +lemma csucc_basic: "Ord(K) \ Card(csucc(K)) & K < csucc(K)" apply (unfold csucc_def) apply (rule LeastI) apply (blast intro: Card_jump_cardinal K_lt_jump_cardinal Ord_jump_cardinal)+ @@ -797,16 +797,16 @@ lemmas lt_csucc = csucc_basic [THEN conjunct2] -lemma Ord_0_lt_csucc: "Ord(K) ==> 0 < csucc(K)" +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" +lemma csucc_le: "\Card(L); K \ csucc(K) \ L" apply (unfold 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" +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) @@ -818,10 +818,10 @@ done lemma Card_lt_csucc_iff: - "[| Card(K'); Card(K) |] ==> K' < csucc(K) \ K' \ K" + "\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))" +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]) @@ -832,7 +832,7 @@ 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" + 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 @@ -842,7 +842,7 @@ thus False using cons by auto qed } - hence [simp]: "~ cons(a,A) \ A" using a FA by auto + 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|)" @@ -864,18 +864,18 @@ qed lemma Finite_imp_succ_cardinal_Diff: - "[| Finite(A); a \ A |] ==> succ(|A-{a}|) = |A|" + "\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|" +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" +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 @@ -883,13 +883,13 @@ qed lemma card_Un_Int: - "[|Finite(A); Finite(B)|] ==> |A| #+ |B| = |A \ B| #+ |A \ B|" + "\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|" + "\Finite(A); Finite(B); A \ B = 0\ \ |A \ B| = |A| #+ |B|" by (simp add: Finite_Un card_Un_Int) lemma card_partition: @@ -924,7 +924,7 @@ finally show ?thesis . qed -lemma Ord_subset_natD [rule_format]: "Ord(i) ==> i \ nat \ i \ nat | i=nat" +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 @@ -934,7 +934,7 @@ by (blast dest: nat_le_Limit le_imp_subset) qed -lemma Ord_nat_subset_into_Card: "[| Ord(i); i \ nat |] ==> Card(i)" +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 -r f2094906e491 -r e44d86131648 src/ZF/Cardinal_AC.thy --- a/src/ZF/Cardinal_AC.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Cardinal_AC.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,7 +19,7 @@ text\The theorem \<^term>\||A|| = |A|\\ lemmas cardinal_idem = cardinal_eqpoll [THEN cardinal_cong, simp] -lemma cardinal_eqE: "|X| = |Y| ==> X \ Y" +lemma cardinal_eqE: "|X| = |Y| \ X \ Y" apply (rule AC_well_ord [THEN exE]) apply (rule AC_well_ord [THEN exE]) apply (rule well_ord_cardinal_eqE, assumption+) @@ -29,11 +29,11 @@ by (blast intro: cardinal_cong cardinal_eqE) lemma cardinal_disjoint_Un: - "[| |A|=|B|; |C|=|D|; A \ C = 0; B \ D = 0 |] - ==> |A \ C| = |B \ D|" + "\|A|=|B|; |C|=|D|; A \ C = 0; B \ D = 0\ + \ |A \ C| = |B \ D|" by (simp add: cardinal_eqpoll_iff eqpoll_disjoint_Un) -lemma lepoll_imp_cardinal_le: "A \ B ==> |A| \ |B|" +lemma lepoll_imp_cardinal_le: "A \ B \ |A| \ |B|" apply (rule AC_well_ord [THEN exE]) apply (erule well_ord_lepoll_imp_cardinal_le, assumption) done @@ -59,7 +59,7 @@ apply (rule well_ord_cadd_cmult_distrib, assumption+) done -lemma InfCard_square_eq: "InfCard(|A|) ==> A*A \ A" +lemma InfCard_square_eq: "InfCard(|A|) \ A*A \ A" apply (rule AC_well_ord [THEN exE]) apply (erule well_ord_InfCard_square_eq, assumption) done @@ -79,7 +79,7 @@ finally show ?thesis . qed -lemma le_Card_iff: "Card(K) ==> |A| \ K \ A \ K" +lemma le_Card_iff: "Card(K) \ |A| \ K \ A \ K" apply (erule Card_cardinal_eq [THEN subst], rule iffI, erule Card_le_imp_lepoll) apply (erule lepoll_imp_cardinal_le) @@ -109,7 +109,7 @@ by (force intro: cardinal_lt_imp_lt lesspoll_cardinal_lt) qed -lemma cardinal_le_imp_lepoll: " i \ |A| ==> i \ A" +lemma cardinal_le_imp_lepoll: " i \ |A| \ i \ A" by (blast intro: lt_Ord Card_le_imp_lepoll Ord_cardinal_le le_trans) @@ -141,11 +141,11 @@ text\Kunen's Lemma 10.21\ lemma cardinal_UN_le: assumes K: "InfCard(K)" - shows "(!!i. i\K ==> |X(i)| \ K) ==> |\i\K. X(i)| \ K" + shows "(\i. i\K \ |X(i)| \ K) \ |\i\K. X(i)| \ K" proof (simp add: K InfCard_is_Card le_Card_iff) have [intro]: "Ord(K)" by (blast intro: InfCard_is_Card Card_is_Ord K) - assume "!!i. i\K ==> X(i) \ K" - hence "!!i. i\K ==> \f. f \ inj(X(i), K)" by (simp add: lepoll_def) + assume "\i. i\K \ X(i) \ K" + hence "\i. i\K \ \f. f \ inj(X(i), K)" by (simp add: lepoll_def) with AC_Pi obtain f where f: "f \ (\i\K. inj(X(i), K))" by force { fix z @@ -173,15 +173,15 @@ text\The same again, using \<^term>\csucc\\ lemma cardinal_UN_lt_csucc: - "[| InfCard(K); \i. i\K \ |X(i)| < csucc(K) |] - ==> |\i\K. X(i)| < csucc(K)" + "\InfCard(K); \i. i\K \ |X(i)| < csucc(K)\ + \ |\i\K. X(i)| < csucc(K)" by (simp add: Card_lt_csucc_iff cardinal_UN_le InfCard_is_Card Card_cardinal) text\The same again, for a union of ordinals. In use, j(i) is a bit like rank(i), the least ordinal j such that i:Vfrom(A,j).\ lemma cardinal_UN_Ord_lt_csucc: - "[| InfCard(K); \i. i\K \ j(i) < csucc(K) |] - ==> (\i\K. j(i)) < csucc(K)" + "\InfCard(K); \i. i\K \ j(i) < csucc(K)\ + \ (\i\K. j(i)) < csucc(K)" apply (rule cardinal_UN_lt_csucc [THEN Card_lt_imp_lt], assumption) apply (blast intro: Ord_cardinal_le [THEN lt_trans1] elim: ltE) apply (blast intro!: Ord_UN elim: ltE) diff -r f2094906e491 -r e44d86131648 src/ZF/Coind/Dynamic.thy --- a/src/ZF/Coind/Dynamic.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/Dynamic.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,28 +11,28 @@ domains "EvalRel" \ "ValEnv * Exp * Val" intros eval_constI: - " [| ve \ ValEnv; c \ Const |] ==> + " \ve \ ValEnv; c \ Const\ \ :EvalRel" eval_varI: - " [| ve \ ValEnv; x \ ExVar; x \ ve_dom(ve) |] ==> + " \ve \ ValEnv; x \ ExVar; x \ ve_dom(ve)\ \ :EvalRel" eval_fnI: - " [| ve \ ValEnv; x \ ExVar; e \ Exp |] ==> + " \ve \ ValEnv; x \ ExVar; e \ Exp\ \ :EvalRel " eval_fixI: - " [| ve \ ValEnv; x \ ExVar; e \ Exp; f \ ExVar; cl \ Val; - v_clos(x,e,ve_owr(ve,f,cl))=cl |] ==> + " \ve \ ValEnv; x \ ExVar; e \ Exp; f \ ExVar; cl \ Val; + v_clos(x,e,ve_owr(ve,f,cl))=cl\ \ :EvalRel " eval_appI1: - " [| ve \ ValEnv; e1 \ Exp; e2 \ Exp; c1 \ Const; c2 \ Const; + " \ve \ ValEnv; e1 \ Exp; e2 \ Exp; c1 \ Const; c2 \ Const; :EvalRel; - :EvalRel |] ==> + :EvalRel\ \ :EvalRel " eval_appI2: - " [| ve \ ValEnv; vem \ ValEnv; e1 \ Exp; e2 \ Exp; em \ Exp; xm \ ExVar; v \ Val; + " \ve \ ValEnv; vem \ ValEnv; e1 \ Exp; e2 \ Exp; em \ Exp; xm \ ExVar; v \ Val; :EvalRel; :EvalRel; - :EvalRel |] ==> + :EvalRel\ \ :EvalRel " type_intros c_appI ve_appI ve_empI ve_owrI Exp.intros Val_ValEnv.intros diff -r f2094906e491 -r e44d86131648 src/ZF/Coind/ECR.thy --- a/src/ZF/Coind/ECR.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/ECR.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,13 +13,13 @@ domains "HasTyRel" \ "Val * Ty" intros htr_constI [intro!]: - "[| c \ Const; t \ Ty; isof(c,t) |] ==> \ HasTyRel" + "\c \ Const; t \ Ty; isof(c,t)\ \ \ HasTyRel" htr_closI [intro]: - "[| x \ ExVar; e \ Exp; t \ Ty; ve \ ValEnv; te \ TyEnv; + "\x \ ExVar; e \ Exp; t \ Ty; ve \ ValEnv; te \ TyEnv; \ ElabRel; ve_dom(ve) = te_dom(te); - {.y \ ve_dom(ve)} \ Pow(HasTyRel) |] - ==> \ HasTyRel" + {.y \ ve_dom(ve)} \ Pow(HasTyRel)\ + \ \ HasTyRel" monos Pow_mono type_intros Val_ValEnv.intros @@ -27,18 +27,18 @@ definition hastyenv :: "[i,i] => o" where - "hastyenv(ve,te) == + "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; + "\x \ ExVar; e \ Exp; t \ Ty; ve \ ValEnv; te \ TyEnv; \ ElabRel; ve_dom(ve) = te_dom(te); {.y \ ve_dom(ve)} \ - Pow({} \ HasTyRel) |] - ==> \ HasTyRel" + Pow({} \ HasTyRel)\ + \ \ HasTyRel" apply (rule singletonI [THEN HasTyRel.coinduct], auto) done @@ -55,12 +55,12 @@ HasTyRel.dom_subset [THEN subsetD, THEN SigmaD1, THEN ValNEE] lemma hastyenv_owr: - "[| ve \ ValEnv; te \ TyEnv; hastyenv(ve,te); \ HasTyRel |] - ==> hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))" + "\ve \ ValEnv; te \ TyEnv; hastyenv(ve,te); \ 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)" + "\ve \ ValEnv; te \ TyEnv; isofenv(ve,te)\ \ hastyenv(ve,te)" apply (unfold isofenv_def hastyenv_def) apply (force intro: te_appI ve_domI) done @@ -71,20 +71,20 @@ (* ############################################################ *) lemma consistency_const: - "[| c \ Const; hastyenv(ve,te); \ ElabRel |] - ==> \ HasTyRel" + "\c \ Const; hastyenv(ve,te); \ ElabRel\ + \ \ HasTyRel" by blast lemma consistency_var: - "[| x \ ve_dom(ve); hastyenv(ve,te); \ ElabRel |] ==> + "\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); + "\ve \ ValEnv; x \ ExVar; e \ Exp; hastyenv(ve,te); \ ElabRel - |] ==> \ HasTyRel" +\ \ \ HasTyRel" by (unfold hastyenv_def, blast) declare ElabRel.dom_subset [THEN subsetD, dest] @@ -94,9 +94,9 @@ declare Val_ValEnv.intros [simp, intro!] lemma consistency_fix: - "[| ve \ ValEnv; x \ ExVar; e \ Exp; f \ ExVar; cl \ Val; + "\ve \ ValEnv; x \ ExVar; e \ Exp; f \ ExVar; cl \ Val; v_clos(x,e,ve_owr(ve,f,cl)) = cl; - hastyenv(ve,te); \ ElabRel |] ==> + hastyenv(ve,te); \ ElabRel\ \ \ HasTyRel" apply (unfold hastyenv_def) apply (erule elab_fixE, safe) @@ -111,7 +111,7 @@ lemma consistency_app1: - "[| ve \ ValEnv; e1 \ Exp; e2 \ Exp; c1 \ Const; c2 \ Const; + "\ve \ ValEnv; e1 \ Exp; e2 \ Exp; c1 \ Const; c2 \ Const; \ EvalRel; \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; @@ -119,12 +119,12 @@ \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; hastyenv(ve, te); - \ ElabRel |] - ==> \ HasTyRel" + \ 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; + "\ve \ ValEnv; vem \ ValEnv; e1 \ Exp; e2 \ Exp; em \ Exp; xm \ ExVar; v \ Val; \ EvalRel; \t te. hastyenv(ve,te) \ @@ -134,8 +134,8 @@ \ EvalRel; \t te. hastyenv(ve_owr(vem,xm,v2),te) \ \ ElabRel \ \ HasTyRel; - hastyenv(ve,te); \ ElabRel |] - ==> \ HasTyRel" + hastyenv(ve,te); \ ElabRel\ + \ \ 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+) @@ -151,7 +151,7 @@ lemma consistency [rule_format]: " \ EvalRel - ==> (\t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel)" + \ (\t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel)" apply (erule EvalRel.induct) apply (simp_all add: consistency_const consistency_var consistency_fn consistency_fix consistency_app1) @@ -159,9 +159,9 @@ done lemma basic_consistency: - "[| ve \ ValEnv; te \ TyEnv; isofenv(ve,te); - \ EvalRel; \ ElabRel |] - ==> isof(c,t)" + "\ve \ ValEnv; te \ TyEnv; isofenv(ve,te); + \ EvalRel; \ ElabRel\ + \ isof(c,t)" by (blast dest: consistency intro!: basic_consistency_lem) end diff -r f2094906e491 -r e44d86131648 src/ZF/Coind/Language.thy --- a/src/ZF/Coind/Language.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/Language.thy Tue Sep 27 16:51:35 2022 +0100 @@ -12,8 +12,8 @@ Const :: i and (* Abstract type of constants *) c_app :: "[i,i] => i" (* Abstract constructor for fun application*) where - constNEE: "c \ Const ==> c \ 0" and - c_appI: "[| c1 \ Const; c2 \ Const |] ==> c_app(c1,c2) \ Const" + constNEE: "c \ Const \ c \ 0" and + c_appI: "\c1 \ Const; c2 \ Const\ \ c_app(c1,c2) \ Const" consts diff -r f2094906e491 -r e44d86131648 src/ZF/Coind/Map.thy --- a/src/ZF/Coind/Map.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/Map.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,25 +9,25 @@ definition TMap :: "[i,i] => i" where - "TMap(A,B) == {m \ Pow(A*\(B)).\a \ A. m``{a} \ B}" + "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))" + "PMap(A,B) \ TMap(A,cons(0,B))" -(* Note: 0 \ B ==> TMap(A,B) = PMap(A,B) *) +(* Note: 0 \ B \ TMap(A,B) = PMap(A,B) *) definition map_emp :: i where - "map_emp == 0" + "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}" + "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}" + "map_app(m,a) \ m``{a}" lemma "{0,1} \ {1} \ TMap(I, {0,1})" by (unfold TMap_def, blast) @@ -51,13 +51,13 @@ 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)" +lemma qbeta: "a \ A \ Sigma(A,B)``{a} = B(a)" by fast -lemma qbeta_emp: "a\A ==> Sigma(A,B)``{a} = 0" +lemma qbeta_emp: "a\A \ Sigma(A,B)``{a} = 0" by fast -lemma image_Sigma1: "a \ A ==> Sigma(A,B)``{a}=0" +lemma image_Sigma1: "a \ A \ Sigma(A,B)``{a}=0" by fast @@ -68,7 +68,7 @@ (* Lemmas *) lemma MapQU_lemma: - "A \ univ(X) ==> Pow(A * \(quniv(X))) \ quniv(X)" + "A \ univ(X) \ Pow(A * \(quniv(X))) \ quniv(X)" apply (unfold quniv_def) apply (rule Pow_mono) apply (rule subset_trans [OF Sigma_mono product_univ]) @@ -80,7 +80,7 @@ (* Theorems *) lemma mapQU: - "[| m \ PMap(A,quniv(B)); !!x. x \ A ==> x \ univ(B) |] ==> m \ quniv(B)" + "\m \ PMap(A,quniv(B)); \x. x \ A \ x \ univ(B)\ \ m \ quniv(B)" apply (unfold PMap_def TMap_def) apply (blast intro!: MapQU_lemma [THEN subsetD]) done @@ -89,7 +89,7 @@ (* Monotonicity *) (* ############################################################ *) -lemma PMap_mono: "B \ C ==> PMap(A,B)<=PMap(A,C)" +lemma PMap_mono: "B \ C \ PMap(A,B)<=PMap(A,C)" by (unfold PMap_def TMap_def, blast) @@ -105,7 +105,7 @@ (** map_owr **) lemma pmap_owrI: - "[| m \ PMap(A,B); a \ A; b \ B |] ==> map_owr(m,a,b):PMap(A,B)" + "\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*) @@ -118,15 +118,15 @@ (** map_app **) lemma tmap_app_notempty: - "[| m \ TMap(A,B); a \ domain(m) |] ==> map_app(m,a) \0" + "\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" + "\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" + "\m \ PMap(A,B); a \ domain(m)\ \ map_app(m,a):B" apply (unfold PMap_def) apply (frule tmap_app_notempty, assumption) apply (drule tmap_appI, auto) @@ -135,11 +135,11 @@ (** domain **) lemma tmap_domainD: - "[| m \ TMap(A,B); a \ domain(m) |] ==> a \ A" + "\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" + "\m \ PMap(A,B); a \ domain(m)\ \ a \ A" by (unfold PMap_def TMap_def, blast) @@ -164,7 +164,7 @@ by (unfold map_emp_def, blast) lemma map_domain_owr: - "b \ 0 ==> domain(map_owr(f,a,b)) = {a} \ domain(f)" + "b \ 0 \ domain(map_owr(f,a,b)) = {a} \ domain(f)" apply (unfold map_owr_def) apply (auto simp add: domain_Sigma) done @@ -178,7 +178,7 @@ 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)" +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 -r f2094906e491 -r e44d86131648 src/ZF/Coind/Static.thy --- a/src/ZF/Coind/Static.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/Static.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,13 +10,13 @@ ***) axiomatization isof :: "[i,i] => o" - where isof_app: "[|isof(c1,t_fun(t1,t2)); isof(c2,t1)|] ==> isof(c_app(c1,c2),t2)" + where isof_app: "\isof(c1,t_fun(t1,t2)); isof(c2,t1)\ \ isof(c_app(c1,c2),t2)" (*Its extension to environments*) definition isofenv :: "[i,i] => o" where - "isofenv(ve,te) == + "isofenv(ve,te) \ ve_dom(ve) = te_dom(te) & (\x \ ve_dom(ve). \c \ Const. ve_app(ve,x) = v_const(c) & isof(c,te_app(te,x)))" @@ -30,23 +30,23 @@ domains "ElabRel" \ "TyEnv * Exp * Ty" intros constI [intro!]: - "[| te \ TyEnv; c \ Const; t \ Ty; isof(c,t) |] ==> + "\te \ TyEnv; c \ Const; t \ Ty; isof(c,t)\ \ \ ElabRel" varI [intro!]: - "[| te \ TyEnv; x \ ExVar; x \ te_dom(te) |] ==> + "\te \ TyEnv; x \ ExVar; x \ te_dom(te)\ \ \ ElabRel" fnI [intro!]: - "[| te \ TyEnv; x \ ExVar; e \ Exp; t1 \ Ty; t2 \ Ty; - \ ElabRel |] ==> + "\te \ TyEnv; x \ ExVar; e \ Exp; t1 \ Ty; t2 \ Ty; + \ ElabRel\ \ \ ElabRel" fixI [intro!]: - "[| te \ TyEnv; f \ ExVar; x \ ExVar; t1 \ Ty; t2 \ Ty; - \ ElabRel |] ==> + "\te \ TyEnv; f \ ExVar; x \ ExVar; t1 \ Ty; t2 \ Ty; + \ ElabRel\ \ \ ElabRel" appI [intro]: - "[| te \ TyEnv; e1 \ Exp; e2 \ Exp; t1 \ Ty; t2 \ Ty; + "\te \ TyEnv; e1 \ Exp; e2 \ Exp; t1 \ Ty; t2 \ Ty; \ ElabRel; - \ ElabRel |] ==> \ ElabRel" + \ ElabRel\ \ \ ElabRel" type_intros te_appI Exp.intros Ty.intros diff -r f2094906e491 -r e44d86131648 src/ZF/Coind/Types.thy --- a/src/ZF/Coind/Types.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/Types.thy Tue Sep 27 16:51:35 2022 +0100 @@ -43,7 +43,7 @@ by simp (*redundant??*) -lemma te_app_owr2: "x \ y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)" +lemma te_app_owr2: "x \ y \ te_app(te_owr(te,x,t),y) = te_app(te,y)" by auto lemma te_app_owr [simp]: @@ -51,7 +51,7 @@ by auto lemma te_appI: - "[| te \ TyEnv; x \ ExVar; x \ te_dom(te) |] ==> te_app(te,x) \ Ty" + "\te \ TyEnv; x \ ExVar; x \ te_dom(te)\ \ te_app(te,x) \ Ty" apply (erule_tac P = "x \ te_dom (te) " in rev_mp) apply (erule TyEnv.induct, auto) done diff -r f2094906e491 -r e44d86131648 src/ZF/Coind/Values.thy --- a/src/ZF/Coind/Values.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Coind/Values.thy Tue Sep 27 16:51:35 2022 +0100 @@ -34,24 +34,24 @@ definition ve_emp :: i where - "ve_emp == ve_mk(map_emp)" + "ve_emp \ ve_mk(map_emp)" (* Elimination rules *) lemma ValEnvE: - "[| ve \ ValEnv; !!m.[| ve=ve_mk(m); m \ PMap(ExVar,Val) |] ==> Q |] ==> Q" + "\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 - |] ==> + "\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) @@ -66,7 +66,7 @@ declare v_closNE [THEN notE, elim!] -lemma v_constNE [simp]: "c \ Const ==> v_const(c) \ 0" +lemma v_constNE [simp]: "c \ Const \ v_const(c) \ 0" apply (unfold QPair_def QInl_def QInr_def Val_ValEnv.con_defs) apply (drule constNEE, auto) done @@ -74,28 +74,28 @@ (* Proving that the empty set is not a value *) -lemma ValNEE: "v \ Val ==> v \ 0" +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}" + "\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))" + \ 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" +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" +lemma ve_domI: "\ve \ ValEnv; x \ ve_dom(ve)\ \ x \ ExVar" apply (erule ValEnvE, simp) apply (blast dest: pmap_domainD) done @@ -107,7 +107,7 @@ done lemma ve_owrI: - "[|ve \ ValEnv; x \ ExVar; v \ Val |] ==> ve_owr(ve,x,v):ValEnv" + "\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 diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/AC_in_L.thy --- a/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,17 +17,17 @@ domains "rlist(A,r)" \ "list(A) * list(A)" intros shorterI: - "[| length(l') < length(l); l' \ list(A); l \ list(A) |] - ==> \ rlist(A,r)" + "\length(l') < length(l); l' \ list(A); l \ list(A)\ + \ \ rlist(A,r)" sameI: - "[| \ rlist(A,r); a \ A |] - ==> \ rlist(A,r)" + "\ \ rlist(A,r); a \ A\ + \ \ rlist(A,r)" diffI: - "[| length(l') = length(l); \ r; - l' \ list(A); l \ list(A); a' \ A; a \ A |] - ==> \ rlist(A,r)" + "\length(l') = length(l); \ r; + l' \ list(A); l \ list(A); a' \ A; a \ A\ + \ \ rlist(A,r)" type_intros list.intros @@ -40,7 +40,7 @@ subsubsection\Linearity\ lemma rlist_Nil_Cons [intro]: - "[|a \ A; l \ list(A)|] ==> <[], Cons(a,l)> \ rlist(A, r)" + "\a \ A; l \ list(A)\ \ <[], Cons(a,l)> \ rlist(A, r)" by (simp add: shorterI) lemma linear_rlist: @@ -82,13 +82,13 @@ lemma not_rlist_Nil [simp]: " \ rlist(A,r)" by (blast intro: elim: rlist_NilE) -lemma rlist_imp_length_le: " \ rlist(A,r) ==> length(l') \ length(l)" +lemma rlist_imp_length_le: " \ rlist(A,r) \ length(l') \ length(l)" apply (erule rlist.induct) apply (simp_all add: leI) done lemma wf_on_rlist_n: - "[| n \ nat; wf[A](r) |] ==> wf[{l \ list(A). length(l) = n}](rlist(A,r))" + "\n \ nat; wf[A](r)\ \ wf[{l \ list(A). length(l) = n}](rlist(A,r))" apply (induct_tac n) apply (rule wf_onI2, simp) apply (rule wf_onI2, clarify) @@ -112,7 +112,7 @@ by (blast intro: length_type) -lemma wf_on_rlist: "wf[A](r) ==> wf[list(A)](rlist(A,r))" +lemma wf_on_rlist: "wf[A](r) \ wf[list(A)](rlist(A,r))" apply (subst list_eq_UN_length) apply (rule wf_on_Union) apply (rule wf_imp_wf_on [OF wf_Memrel [of nat]]) @@ -125,14 +125,14 @@ done -lemma wf_rlist: "wf(r) ==> wf(rlist(field(r),r))" +lemma wf_rlist: "wf(r) \ wf(rlist(field(r),r))" apply (simp add: wf_iff_wf_on_field) apply (rule wf_on_subset_A [OF _ field_rlist]) apply (blast intro: wf_on_rlist) done lemma well_ord_rlist: - "well_ord(A,r) ==> well_ord(list(A), rlist(A,r))" + "well_ord(A,r) \ well_ord(list(A), rlist(A,r))" apply (rule well_ordI) apply (simp add: well_ord_def wf_on_rlist) apply (simp add: well_ord_def tot_ord_def linear_rlist) @@ -167,20 +167,20 @@ "enum(f, Forall(p)) = f ` " lemma (in Nat_Times_Nat) fn_type [TC,simp]: - "[|x \ nat; y \ nat|] ==> fn` \ nat" + "\x \ nat; y \ nat\ \ fn` \ nat" by (blast intro: inj_is_fun [OF fn_inj] apply_funtype) lemma (in Nat_Times_Nat) fn_iff: - "[|x \ nat; y \ nat; u \ nat; v \ nat|] - ==> (fn` = fn`) \ (x=u & y=v)" + "\x \ nat; y \ nat; u \ nat; v \ nat\ + \ (fn` = fn`) \ (x=u & y=v)" by (blast dest: inj_apply_equality [OF fn_inj]) lemma (in Nat_Times_Nat) enum_type [TC,simp]: - "p \ formula ==> enum(fn,p) \ nat" + "p \ formula \ enum(fn,p) \ nat" by (induct_tac p, simp_all) lemma (in Nat_Times_Nat) enum_inject [rule_format]: - "p \ formula ==> \q\formula. enum(fn,p) = enum(fn,q) \ p=q" + "p \ formula \ \q\formula. enum(fn,p) = enum(fn,q) \ p=q" apply (induct_tac p, simp_all) apply (rule ballI) apply (erule formula.cases) @@ -232,7 +232,7 @@ definition env_form_r :: "[i,i,i]=>i" where \ \wellordering on (environment, formula) pairs\ - "env_form_r(f,r,A) == + "env_form_r(f,r,A) \ rmult(list(A), rlist(A, r), formula, measure(formula, enum(f)))" @@ -240,12 +240,12 @@ env_form_map :: "[i,i,i,i]=>i" where \ \map from (environment, formula) pairs to ordinals\ "env_form_map(f,r,A,z) - == ordermap(list(A) * formula, env_form_r(f,r,A)) ` z" + \ ordermap(list(A) * formula, env_form_r(f,r,A)) ` z" definition DPow_ord :: "[i,i,i,i,i]=>o" where \ \predicate that holds if \<^term>\k\ is a valid index for \<^term>\X\\ - "DPow_ord(f,r,A,X,k) == + "DPow_ord(f,r,A,X,k) \ \env \ list(A). \p \ formula. arity(p) \ succ(length(env)) & X = {x\A. sats(A, p, Cons(x,env))} & @@ -254,61 +254,61 @@ definition DPow_least :: "[i,i,i,i]=>i" where \ \function yielding the smallest index for \<^term>\X\\ - "DPow_least(f,r,A,X) == \ k. DPow_ord(f,r,A,X,k)" + "DPow_least(f,r,A,X) \ \ k. DPow_ord(f,r,A,X,k)" definition DPow_r :: "[i,i,i]=>i" where \ \a wellordering on \<^term>\DPow(A)\\ - "DPow_r(f,r,A) == measure(DPow(A), DPow_least(f,r,A))" + "DPow_r(f,r,A) \ measure(DPow(A), DPow_least(f,r,A))" lemma (in Nat_Times_Nat) well_ord_env_form_r: "well_ord(A,r) - ==> well_ord(list(A) * formula, env_form_r(fn,r,A))" + \ well_ord(list(A) * formula, env_form_r(fn,r,A))" by (simp add: env_form_r_def well_ord_rmult well_ord_rlist well_ord_formula) lemma (in Nat_Times_Nat) Ord_env_form_map: - "[|well_ord(A,r); z \ list(A) * formula|] - ==> Ord(env_form_map(fn,r,A,z))" + "\well_ord(A,r); z \ list(A) * formula\ + \ Ord(env_form_map(fn,r,A,z))" by (simp add: env_form_map_def Ord_ordermap well_ord_env_form_r) lemma DPow_imp_ex_DPow_ord: - "X \ DPow(A) ==> \k. DPow_ord(fn,r,A,X,k)" + "X \ DPow(A) \ \k. DPow_ord(fn,r,A,X,k)" apply (simp add: DPow_ord_def) apply (blast dest!: DPowD) done lemma (in Nat_Times_Nat) DPow_ord_imp_Ord: - "[|DPow_ord(fn,r,A,X,k); well_ord(A,r)|] ==> Ord(k)" + "\DPow_ord(fn,r,A,X,k); well_ord(A,r)\ \ Ord(k)" apply (simp add: DPow_ord_def, clarify) apply (simp add: Ord_env_form_map) done lemma (in Nat_Times_Nat) DPow_imp_DPow_least: - "[|X \ DPow(A); well_ord(A,r)|] - ==> DPow_ord(fn, r, A, X, DPow_least(fn,r,A,X))" + "\X \ DPow(A); well_ord(A,r)\ + \ DPow_ord(fn, r, A, X, DPow_least(fn,r,A,X))" apply (simp add: DPow_least_def) apply (blast dest: DPow_imp_ex_DPow_ord intro: DPow_ord_imp_Ord LeastI) done lemma (in Nat_Times_Nat) env_form_map_inject: - "[|env_form_map(fn,r,A,u) = env_form_map(fn,r,A,v); well_ord(A,r); - u \ list(A) * formula; v \ list(A) * formula|] - ==> u=v" + "\env_form_map(fn,r,A,u) = env_form_map(fn,r,A,v); well_ord(A,r); + u \ list(A) * formula; v \ list(A) * formula\ + \ u=v" apply (simp add: env_form_map_def) apply (rule inj_apply_equality [OF bij_is_inj, OF ordermap_bij, OF well_ord_env_form_r], assumption+) done lemma (in Nat_Times_Nat) DPow_ord_unique: - "[|DPow_ord(fn,r,A,X,k); DPow_ord(fn,r,A,Y,k); well_ord(A,r)|] - ==> X=Y" + "\DPow_ord(fn,r,A,X,k); DPow_ord(fn,r,A,Y,k); well_ord(A,r)\ + \ X=Y" apply (simp add: DPow_ord_def, clarify) apply (drule env_form_map_inject, auto) done lemma (in Nat_Times_Nat) well_ord_DPow_r: - "well_ord(A,r) ==> well_ord(DPow(A), DPow_r(fn,r,A))" + "well_ord(A,r) \ well_ord(DPow(A), DPow_r(fn,r,A))" apply (simp add: DPow_r_def) apply (rule well_ord_measure) apply (simp add: DPow_least_def) @@ -332,7 +332,7 @@ rlimit :: "[i,i=>i]=>i" where \ \Expresses the wellordering at limit ordinals. The conditional lets us remove the premise \<^term>\Limit(i)\ from some theorems.\ - "rlimit(i,r) == + "rlimit(i,r) \ if Limit(i) then {z: Lset(i) * Lset(i). \x' x. z = & @@ -344,10 +344,10 @@ Lset_new :: "i=>i" where \ \This constant denotes the set of elements introduced at level \<^term>\succ(i)\\ - "Lset_new(i) == {x \ Lset(succ(i)). lrank(x) = i}" + "Lset_new(i) \ {x \ Lset(succ(i)). lrank(x) = i}" lemma Limit_Lset_eq2: - "Limit(i) ==> Lset(i) = (\j\i. Lset_new(j))" + "Limit(i) \ Lset(i) = (\j\i. Lset_new(j))" apply (simp add: Limit_Lset_eq) apply (rule equalityI) apply safe @@ -362,14 +362,14 @@ done lemma wf_on_Lset: - "wf[Lset(succ(j))](r(succ(j))) ==> wf[Lset_new(j)](rlimit(i,r))" + "wf[Lset(succ(j))](r(succ(j))) \ wf[Lset_new(j)](rlimit(i,r))" apply (simp add: wf_on_def Lset_new_def) apply (erule wf_subset) apply (simp add: rlimit_def, force) done lemma wf_on_rlimit: - "(\j wf[Lset(i)](rlimit(i,r))" + "(\j wf[Lset(i)](rlimit(i,r))" apply (case_tac "Limit(i)") prefer 2 apply (simp add: rlimit_def wf_on_any_0) @@ -382,8 +382,8 @@ done lemma linear_rlimit: - "[|Limit(i); \j linear(Lset(i), rlimit(i,r))" + "\Limit(i); \j + \ linear(Lset(i), rlimit(i,r))" apply (frule Limit_is_Ord) apply (simp add: Limit_Lset_eq2 Lset_new_def) apply (simp add: linear_def rlimit_def Ball_def lt_Ord Lset_iff_lrank_lt) @@ -397,13 +397,13 @@ done lemma well_ord_rlimit: - "[|Limit(i); \j well_ord(Lset(i), rlimit(i,r))" + "\Limit(i); \j + \ well_ord(Lset(i), rlimit(i,r))" by (blast intro: well_ordI wf_on_rlimit well_ord_is_wf linear_rlimit well_ord_is_linear) lemma rlimit_cong: - "(!!j. j r'(j) = r(j)) ==> rlimit(i,r) = rlimit(i,r')" + "(\j. j r'(j) = r(j)) \ rlimit(i,r) = rlimit(i,r')" apply (simp add: rlimit_def, clarify) apply (rule refl iff_refl Collect_cong ex_cong conj_cong)+ apply (simp add: Limit_is_Ord Lset_lrank_lt) @@ -414,7 +414,7 @@ definition L_r :: "[i, i] => i" where - "L_r(f) == %i. + "L_r(f) \ %i. transrec3(i, 0, \x r. DPow_r(f, r, Lset(x)), \x r. rlimit(x, \y. r`y))" @@ -427,25 +427,25 @@ text\The limit case is non-trivial because of the distinction between object-level and meta-level abstraction.\ -lemma [simp]: "Limit(i) ==> L_r(f,i) = rlimit(i, L_r(f))" +lemma [simp]: "Limit(i) \ L_r(f,i) = rlimit(i, L_r(f))" by (simp cong: rlimit_cong add: transrec3_Limit L_r_def ltD) lemma (in Nat_Times_Nat) L_r_type: - "Ord(i) ==> L_r(fn,i) \ Lset(i) * Lset(i)" + "Ord(i) \ L_r(fn,i) \ Lset(i) * Lset(i)" apply (induct i rule: trans_induct3) apply (simp_all add: Lset_succ DPow_r_type well_ord_DPow_r rlimit_def Transset_subset_DPow [OF Transset_Lset], blast) done lemma (in Nat_Times_Nat) well_ord_L_r: - "Ord(i) ==> well_ord(Lset(i), L_r(fn,i))" + "Ord(i) \ well_ord(Lset(i), L_r(fn,i))" apply (induct i rule: trans_induct3) apply (simp_all add: well_ord0 Lset_succ L_r_type well_ord_DPow_r well_ord_rlimit ltD) done lemma well_ord_L_r: - "Ord(i) ==> \r. well_ord(Lset(i), r)" + "Ord(i) \ \r. well_ord(Lset(i), r)" apply (insert nat_times_nat_lepoll_nat) apply (unfold lepoll_def) apply (blast intro: Nat_Times_Nat.well_ord_L_r Nat_Times_Nat.intro) @@ -453,7 +453,7 @@ text\Every constructible set is well-ordered! Therefore the Wellordering Theorem and - the Axiom of Choice hold in \<^term>\L\!!\ + the Axiom of Choice hold in \<^term>\L\\\ theorem L_implies_AC: assumes x: "L(x)" shows "\r. well_ord(x,r)" using Transset_Lset x apply (simp add: Transset_def L_def) diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/DPow_absolute.thy --- a/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,10 +13,10 @@ subsubsection\The Operator \<^term>\is_formula_rec\\ text\The three arguments of \<^term>\p\ are always 2, 1, 0. It is buried - within 11 quantifiers!!\ + within 11 quantifiers\\ (* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" - "is_formula_rec(M,MH,p,z) == + "is_formula_rec(M,MH,p,z) \ \dp[M]. \i[M]. \f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) & 2 1 0 successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)" @@ -24,7 +24,7 @@ definition formula_rec_fm :: "[i, i, i]=>i" where - "formula_rec_fm(mh,p,z) == + "formula_rec_fm(mh,p,z) \ Exists(Exists(Exists( And(finite_ordinal_fm(2), And(depth_fm(p#+3,2), @@ -32,42 +32,42 @@ And(fun_apply_fm(0,p#+3,z#+3), is_transrec_fm(mh,1,0))))))))" lemma is_formula_rec_type [TC]: - "[| p \ formula; x \ nat; z \ nat |] - ==> formula_rec_fm(p,x,z) \ formula" + "\p \ formula; x \ nat; z \ nat\ + \ formula_rec_fm(p,x,z) \ formula" by (simp add: formula_rec_fm_def) lemma sats_formula_rec_fm: assumes MH_iff_sats: - "!!a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10. - [|a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A; a8\A; a9\A; a10\A|] - ==> MH(a2, a1, a0) \ + "\a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10. + \a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A; a8\A; a9\A; a10\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3, Cons(a4,Cons(a5,Cons(a6,Cons(a7, Cons(a8,Cons(a9,Cons(a10,env))))))))))))" shows - "[|x \ nat; z \ nat; env \ list(A)|] - ==> sats(A, formula_rec_fm(p,x,z), env) \ + "\x \ nat; z \ nat; env \ list(A)\ + \ sats(A, formula_rec_fm(p,x,z), env) \ is_formula_rec(##A, MH, nth(x,env), nth(z,env))" by (simp add: formula_rec_fm_def sats_is_transrec_fm is_formula_rec_def MH_iff_sats [THEN iff_sym]) lemma formula_rec_iff_sats: assumes MH_iff_sats: - "!!a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10. - [|a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A; a8\A; a9\A; a10\A|] - ==> MH(a2, a1, a0) \ + "\a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10. + \a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A; a8\A; a9\A; a10\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3, Cons(a4,Cons(a5,Cons(a6,Cons(a7, Cons(a8,Cons(a9,Cons(a10,env))))))))))))" shows - "[|nth(i,env) = x; nth(k,env) = z; - i \ nat; k \ nat; env \ list(A)|] - ==> is_formula_rec(##A, MH, x, z) \ sats(A, formula_rec_fm(p,i,k), env)" + "\nth(i,env) = x; nth(k,env) = z; + i \ nat; k \ nat; env \ list(A)\ + \ is_formula_rec(##A, MH, x, z) \ sats(A, formula_rec_fm(p,i,k), env)" by (simp add: sats_formula_rec_fm [OF MH_iff_sats]) theorem formula_rec_reflection: assumes MH_reflection: - "!!f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), + "\f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), \i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]" shows "REFLECTS[\x. is_formula_rec(L, MH(L,x), f(x), h(x)), \i x. is_formula_rec(##Lset(i), MH(##Lset(i),x), f(x), h(x))]" @@ -79,25 +79,25 @@ subsubsection\The Operator \<^term>\is_satisfies\\ -(* is_satisfies(M,A,p,z) == is_formula_rec (M, satisfies_MH(M,A), p, z) *) +(* is_satisfies(M,A,p,z) \ is_formula_rec (M, satisfies_MH(M,A), p, z) *) definition satisfies_fm :: "[i,i,i]=>i" where - "satisfies_fm(x) == formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))" + "satisfies_fm(x) \ formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))" lemma is_satisfies_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> satisfies_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ satisfies_fm(x,y,z) \ formula" by (simp add: satisfies_fm_def) lemma sats_satisfies_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, satisfies_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, satisfies_fm(x,y,z), env) \ is_satisfies(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: satisfies_fm_def is_satisfies_def sats_formula_rec_fm) lemma satisfies_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_satisfies(##A, x, y, z) \ sats(A, satisfies_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_satisfies(##A, x, y, z) \ sats(A, satisfies_fm(i,j,k), env)" by (simp) theorem satisfies_reflection: @@ -121,56 +121,56 @@ (the comprehension).\ definition is_DPow_sats :: "[i=>o,i,i,i,i] => o" where - "is_DPow_sats(M,A,env,p,x) == + "is_DPow_sats(M,A,env,p,x) \ \n1[M]. \e[M]. \sp[M]. is_satisfies(M,A,p,sp) \ is_Cons(M,x,env,e) \ fun_apply(M, sp, e, n1) \ number1(M, n1)" lemma (in M_satisfies) DPow_sats_abs: - "[| M(A); env \ list(A); p \ formula; M(x) |] - ==> is_DPow_sats(M,A,env,p,x) \ sats(A, p, Cons(x,env))" + "\M(A); env \ list(A); p \ formula; M(x)\ + \ is_DPow_sats(M,A,env,p,x) \ sats(A, p, Cons(x,env))" apply (subgoal_tac "M(env)") apply (simp add: is_DPow_sats_def satisfies_closed satisfies_abs) apply (blast dest: transM) done lemma (in M_satisfies) Collect_DPow_sats_abs: - "[| M(A); env \ list(A); p \ formula |] - ==> Collect(A, is_DPow_sats(M,A,env,p)) = + "\M(A); env \ list(A); p \ formula\ + \ Collect(A, is_DPow_sats(M,A,env,p)) = {x \ A. sats(A, p, Cons(x,env))}" by (simp add: DPow_sats_abs transM [of _ A]) subsubsection\The Operator \<^term>\is_DPow_sats\, Internalized\ -(* is_DPow_sats(M,A,env,p,x) == +(* is_DPow_sats(M,A,env,p,x) \ \n1[M]. \e[M]. \sp[M]. is_satisfies(M,A,p,sp) \ is_Cons(M,x,env,e) \ fun_apply(M, sp, e, n1) \ number1(M, n1) *) definition DPow_sats_fm :: "[i,i,i,i]=>i" where - "DPow_sats_fm(A,env,p,x) == + "DPow_sats_fm(A,env,p,x) \ Forall(Forall(Forall( Implies(satisfies_fm(A#+3,p#+3,0), Implies(Cons_fm(x#+3,env#+3,1), Implies(fun_apply_fm(0,1,2), number1_fm(2)))))))" lemma is_DPow_sats_type [TC]: - "[| A \ nat; x \ nat; y \ nat; z \ nat |] - ==> DPow_sats_fm(A,x,y,z) \ formula" + "\A \ nat; x \ nat; y \ nat; z \ nat\ + \ DPow_sats_fm(A,x,y,z) \ formula" by (simp add: DPow_sats_fm_def) lemma sats_DPow_sats_fm [simp]: - "[| u \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, DPow_sats_fm(u,x,y,z), env) \ + "\u \ nat; x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, DPow_sats_fm(u,x,y,z), env) \ is_DPow_sats(##A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" by (simp add: DPow_sats_fm_def is_DPow_sats_def) lemma DPow_sats_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)|] - ==> is_DPow_sats(##A,nu,nx,ny,nz) \ + "\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)\ + \ is_DPow_sats(##A,nu,nx,ny,nz) \ sats(A, DPow_sats_fm(u,x,y,z), env)" by simp @@ -187,23 +187,23 @@ locale M_DPow = M_satisfies + assumes sep: - "[| M(A); env \ list(A); p \ formula |] - ==> separation(M, \x. is_DPow_sats(M,A,env,p,x))" + "\M(A); env \ list(A); p \ formula\ + \ separation(M, \x. is_DPow_sats(M,A,env,p,x))" and rep: "M(A) - ==> strong_replacement (M, + \ strong_replacement (M, \ep z. \env[M]. \p[M]. mem_formula(M,p) & mem_list(M,A,env) & pair(M,env,p,ep) & is_Collect(M, A, \x. is_DPow_sats(M,A,env,p,x), z))" lemma (in M_DPow) sep': - "[| M(A); env \ list(A); p \ formula |] - ==> separation(M, \x. sats(A, p, Cons(x,env)))" + "\M(A); env \ list(A); p \ formula\ + \ separation(M, \x. sats(A, p, Cons(x,env)))" by (insert sep [of A env p], simp add: DPow_sats_abs) lemma (in M_DPow) rep': "M(A) - ==> strong_replacement (M, + \ strong_replacement (M, \ep z. \env\list(A). \p\formula. ep = & z = {x \ A . sats(A, p, Cons(x, env))})" by (insert rep [of A], simp add: Collect_DPow_sats_abs) @@ -213,7 +213,7 @@ "univalent (M, A, \xy z. \x\B. \y\C. xy = \x,y\ \ z = f(x,y))" by (simp add: univalent_def, blast) -lemma (in M_DPow) DPow'_closed: "M(A) ==> M(DPow'(A))" +lemma (in M_DPow) DPow'_closed: "M(A) \ M(DPow'(A))" apply (simp add: DPow'_eq) apply (fast intro: rep' sep' univalent_pair_eq) done @@ -221,14 +221,14 @@ text\Relativization of the Operator \<^term>\DPow'\\ definition is_DPow' :: "[i=>o,i,i] => o" where - "is_DPow'(M,A,Z) == + "is_DPow'(M,A,Z) \ \X[M]. X \ Z \ subset(M,X,A) & (\env[M]. \p[M]. mem_formula(M,p) & mem_list(M,A,env) & is_Collect(M, A, is_DPow_sats(M,A,env,p), X))" lemma (in M_DPow) DPow'_abs: - "[|M(A); M(Z)|] ==> is_DPow'(M,A,Z) \ Z = DPow'(A)" + "\M(A); M(Z)\ \ is_DPow'(M,A,Z) \ Z = DPow'(A)" apply (rule iffI) prefer 2 apply (simp add: is_DPow'_def DPow'_def Collect_DPow_sats_abs) apply (rule M_equalityI) @@ -242,8 +242,8 @@ subsubsection\The Instance of Separation\ lemma DPow_separation: - "[| L(A); env \ list(A); p \ formula |] - ==> separation(L, \x. is_DPow_sats(L,A,env,p,x))" + "\L(A); env \ list(A); p \ formula\ + \ separation(L, \x. is_DPow_sats(L,A,env,p,x))" apply (rule gen_separation_multi [OF DPow_sats_reflection, of "{A,env,p}"], auto intro: transL) apply (rule_tac env="[A,env,p]" in DPow_LsetI) @@ -271,7 +271,7 @@ lemma DPow_replacement: "L(A) - ==> strong_replacement (L, + \ strong_replacement (L, \ep z. \env[L]. \p[L]. mem_formula(L,p) & mem_list(L,A,env) & pair(L,env,p,ep) & is_Collect(L, A, \x. is_DPow_sats(L,A,env,p,x), z))" @@ -309,35 +309,35 @@ enclosed within a single quantifier.\ (* is_Collect :: "[i=>o,i,i=>o,i] => o" - "is_Collect(M,A,P,z) == \x[M]. x \ z \ x \ A & P(x)" *) + "is_Collect(M,A,P,z) \ \x[M]. x \ z \ x \ A & P(x)" *) definition Collect_fm :: "[i, i, i]=>i" where - "Collect_fm(A,is_P,z) == + "Collect_fm(A,is_P,z) \ Forall(Iff(Member(0,succ(z)), And(Member(0,succ(A)), is_P)))" lemma is_Collect_type [TC]: - "[| is_P \ formula; x \ nat; y \ nat |] - ==> Collect_fm(x,is_P,y) \ formula" + "\is_P \ formula; x \ nat; y \ nat\ + \ Collect_fm(x,is_P,y) \ formula" by (simp add: Collect_fm_def) lemma sats_Collect_fm: assumes is_P_iff_sats: - "!!a. a \ A ==> is_P(a) \ sats(A, p, Cons(a, env))" + "\a. a \ A \ is_P(a) \ sats(A, p, Cons(a, env))" shows - "[|x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, Collect_fm(x,p,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, Collect_fm(x,p,y), env) \ is_Collect(##A, nth(x,env), is_P, nth(y,env))" by (simp add: Collect_fm_def is_Collect_def is_P_iff_sats [THEN iff_sym]) lemma Collect_iff_sats: assumes is_P_iff_sats: - "!!a. a \ A ==> is_P(a) \ sats(A, p, Cons(a, env))" + "\a. a \ A \ is_P(a) \ sats(A, p, Cons(a, env))" shows - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_Collect(##A, x, is_P, y) \ sats(A, Collect_fm(i,p,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_Collect(##A, x, is_P, y) \ sats(A, Collect_fm(i,p,j), env)" by (simp add: sats_Collect_fm [OF is_P_iff_sats]) @@ -345,7 +345,7 @@ which is essential for handling free variable references.\ theorem Collect_reflection: assumes is_P_reflection: - "!!h f g. REFLECTS[\x. is_P(L, f(x), g(x)), + "\h f g. REFLECTS[\x. is_P(L, f(x), g(x)), \i x. is_P(##Lset(i), f(x), g(x))]" shows "REFLECTS[\x. is_Collect(L, f(x), is_P(L,x), g(x)), \i x. is_Collect(##Lset(i), f(x), is_P(##Lset(i), x), g(x))]" @@ -360,37 +360,37 @@ and not the usual 1, 0! It is enclosed within two quantifiers.\ (* is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" - "is_Replace(M,A,P,z) == \u[M]. u \ z \ (\x[M]. x\A & P(x,u))" *) + "is_Replace(M,A,P,z) \ \u[M]. u \ z \ (\x[M]. x\A & P(x,u))" *) definition Replace_fm :: "[i, i, i]=>i" where - "Replace_fm(A,is_P,z) == + "Replace_fm(A,is_P,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,A#+2), is_P))))" lemma is_Replace_type [TC]: - "[| is_P \ formula; x \ nat; y \ nat |] - ==> Replace_fm(x,is_P,y) \ formula" + "\is_P \ formula; x \ nat; y \ nat\ + \ Replace_fm(x,is_P,y) \ formula" by (simp add: Replace_fm_def) lemma sats_Replace_fm: assumes is_P_iff_sats: - "!!a b. [|a \ A; b \ A|] - ==> is_P(a,b) \ sats(A, p, Cons(a,Cons(b,env)))" + "\a b. \a \ A; b \ A\ + \ is_P(a,b) \ sats(A, p, Cons(a,Cons(b,env)))" shows - "[|x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, Replace_fm(x,p,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, Replace_fm(x,p,y), env) \ is_Replace(##A, nth(x,env), is_P, nth(y,env))" by (simp add: Replace_fm_def is_Replace_def is_P_iff_sats [THEN iff_sym]) lemma Replace_iff_sats: assumes is_P_iff_sats: - "!!a b. [|a \ A; b \ A|] - ==> is_P(a,b) \ sats(A, p, Cons(a,Cons(b,env)))" + "\a b. \a \ A; b \ A\ + \ is_P(a,b) \ sats(A, p, Cons(a,Cons(b,env)))" shows - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_Replace(##A, x, is_P, y) \ sats(A, Replace_fm(i,p,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_Replace(##A, x, is_P, y) \ sats(A, Replace_fm(i,p,j), env)" by (simp add: sats_Replace_fm [OF is_P_iff_sats]) @@ -398,7 +398,7 @@ which is essential for handling free variable references.\ theorem Replace_reflection: assumes is_P_reflection: - "!!h f g. REFLECTS[\x. is_P(L, f(x), g(x), h(x)), + "\h f g. REFLECTS[\x. is_P(L, f(x), g(x), h(x)), \i x. is_P(##Lset(i), f(x), g(x), h(x))]" shows "REFLECTS[\x. is_Replace(L, f(x), is_P(L,x), g(x)), \i x. is_Replace(##Lset(i), f(x), is_P(##Lset(i), x), g(x))]" @@ -410,7 +410,7 @@ subsubsection\The Operator \<^term>\is_DPow'\, Internalized\ -(* "is_DPow'(M,A,Z) == +(* "is_DPow'(M,A,Z) \ \X[M]. X \ Z \ subset(M,X,A) & (\env[M]. \p[M]. mem_formula(M,p) & mem_list(M,A,env) & @@ -418,7 +418,7 @@ definition DPow'_fm :: "[i,i]=>i" where - "DPow'_fm(A,Z) == + "DPow'_fm(A,Z) \ Forall( Iff(Member(0,succ(Z)), And(subset_fm(0,succ(A)), @@ -429,19 +429,19 @@ DPow_sats_fm(A#+4, 2, 1, 0), 2))))))))" lemma is_DPow'_type [TC]: - "[| x \ nat; y \ nat |] ==> DPow'_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ DPow'_fm(x,y) \ formula" by (simp add: DPow'_fm_def) lemma sats_DPow'_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, DPow'_fm(x,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, DPow'_fm(x,y), env) \ is_DPow'(##A, nth(x,env), nth(y,env))" by (simp add: DPow'_fm_def is_DPow'_def sats_subset_fm' sats_Collect_fm) lemma DPow'_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_DPow'(##A, x, y) \ sats(A, DPow'_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_DPow'(##A, x, y) \ sats(A, DPow'_fm(i,j), env)" by (simp) theorem DPow'_reflection: @@ -457,39 +457,39 @@ definition transrec_body :: "[i=>o,i,i,i,i] => o" where - "transrec_body(M,g,x) == + "transrec_body(M,g,x) \ \y z. \gy[M]. y \ x & fun_apply(M,g,y,gy) & is_DPow'(M,gy,z)" lemma (in M_DPow) transrec_body_abs: - "[|M(x); M(g); M(z)|] - ==> transrec_body(M,g,x,y,z) \ y \ x & z = DPow'(g`y)" + "\M(x); M(g); M(z)\ + \ transrec_body(M,g,x,y,z) \ y \ x & z = DPow'(g`y)" by (simp add: transrec_body_def DPow'_abs transM [of _ x]) locale M_Lset = M_DPow + assumes strong_rep: - "[|M(x); M(g)|] ==> strong_replacement(M, \y z. transrec_body(M,g,x,y,z))" + "\M(x); M(g)\ \ strong_replacement(M, \y z. transrec_body(M,g,x,y,z))" and transrec_rep: - "M(i) ==> transrec_replacement(M, \x f u. + "M(i) \ transrec_replacement(M, \x f u. \r[M]. is_Replace(M, x, transrec_body(M,f,x), r) & big_union(M, r, u), i)" lemma (in M_Lset) strong_rep': - "[|M(x); M(g)|] - ==> strong_replacement(M, \y z. y \ x & z = DPow'(g`y))" + "\M(x); M(g)\ + \ strong_replacement(M, \y z. y \ x & z = DPow'(g`y))" by (insert strong_rep [of x g], simp add: transrec_body_abs) lemma (in M_Lset) DPow_apply_closed: - "[|M(f); M(x); y\x|] ==> M(DPow'(f`y))" + "\M(f); M(x); y\x\ \ M(DPow'(f`y))" by (blast intro: DPow'_closed dest: transM) lemma (in M_Lset) RepFun_DPow_apply_closed: - "[|M(f); M(x)|] ==> M({DPow'(f`y). y\x})" + "\M(f); M(x)\ \ M({DPow'(f`y). y\x})" by (blast intro: DPow_apply_closed RepFun_closed2 strong_rep') lemma (in M_Lset) RepFun_DPow_abs: - "[|M(x); M(f); M(r) |] - ==> is_Replace(M, x, \y z. transrec_body(M,f,x,y,z), r) \ + "\M(x); M(f); M(r)\ + \ is_Replace(M, x, \y z. transrec_body(M,f,x,y,z), r) \ r = {DPow'(f`y). y\x}" apply (simp add: transrec_body_abs RepFun_def) apply (rule iff_trans) @@ -498,7 +498,7 @@ done lemma (in M_Lset) transrec_rep': - "M(i) ==> transrec_replacement(M, \x f u. u = (\y\x. DPow'(f ` y)), i)" + "M(i) \ transrec_replacement(M, \x f u. u = (\y\x. DPow'(f ` y)), i)" apply (insert transrec_rep [of i]) apply (simp add: RepFun_DPow_apply_closed RepFun_DPow_abs transrec_replacement_def) @@ -512,18 +512,18 @@ \ \We can use the term language below because \<^term>\is_Lset\ will not have to be internalized: it isn't used in any instance of separation.\ - "is_Lset(M,a,z) == is_transrec(M, %x f u. u = (\y\x. DPow'(f`y)), a, z)" + "is_Lset(M,a,z) \ is_transrec(M, %x f u. u = (\y\x. DPow'(f`y)), a, z)" lemma (in M_Lset) Lset_abs: - "[|Ord(i); M(i); M(z)|] - ==> is_Lset(M,i,z) \ z = Lset(i)" + "\Ord(i); M(i); M(z)\ + \ is_Lset(M,i,z) \ z = Lset(i)" apply (simp add: is_Lset_def Lset_eq_transrec_DPow') apply (rule transrec_abs) apply (simp_all add: transrec_rep' relation2_def RepFun_DPow_apply_closed) done lemma (in M_Lset) Lset_closed: - "[|Ord(i); M(i)|] ==> M(Lset(i))" + "\Ord(i); M(i)\ \ M(Lset(i))" apply (simp add: Lset_eq_transrec_DPow') apply (rule transrec_closed [OF transrec_rep']) apply (simp_all add: relation2_def RepFun_DPow_apply_closed) @@ -542,7 +542,7 @@ by (intro FOL_reflections function_reflections DPow'_reflection) lemma strong_rep: - "[|L(x); L(g)|] ==> strong_replacement(L, \y z. transrec_body(L,g,x,y,z))" + "\L(x); L(g)\ \ strong_replacement(L, \y z. transrec_body(L,g,x,y,z))" apply (unfold transrec_body_def) apply (rule strong_replacementI) apply (rule_tac u="{x,g,B}" @@ -577,8 +577,8 @@ lemma transrec_rep: - "[|L(j)|] - ==> transrec_replacement(L, \x f u. + "\L(j)\ + \ transrec_replacement(L, \x f u. \r[L]. is_Replace(L, x, transrec_body(L,f,x), r) & big_union(L, r, u), j)" apply (rule L.transrec_replacementI, assumption) @@ -614,7 +614,7 @@ definition constructible :: "[i=>o,i] => o" where - "constructible(M,x) == + "constructible(M,x) \ \i[M]. \Li[M]. ordinal(M,i) & is_Lset(M,i,Li) & x \ Li" theorem V_equals_L_in_L: diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Datatype_absolute.thy --- a/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,25 +11,25 @@ definition directed :: "i=>o" where - "directed(A) == A\0 & (\x\A. \y\A. x \ y \ A)" + "directed(A) \ A\0 & (\x\A. \y\A. x \ y \ A)" definition contin :: "(i=>i) => o" where - "contin(h) == (\A. directed(A) \ h(\A) = (\X\A. h(X)))" + "contin(h) \ (\A. directed(A) \ h(\A) = (\X\A. h(X)))" -lemma bnd_mono_iterates_subset: "[|bnd_mono(D, h); n \ nat|] ==> h^n (0) \ D" +lemma bnd_mono_iterates_subset: "\bnd_mono(D, h); n \ nat\ \ h^n (0) \ D" apply (induct_tac n) apply (simp_all add: bnd_mono_def, blast) done lemma bnd_mono_increasing [rule_format]: - "[|i \ nat; j \ nat; bnd_mono(D,h)|] ==> i \ j \ h^i(0) \ h^j(0)" + "\i \ nat; j \ nat; bnd_mono(D,h)\ \ i \ j \ h^i(0) \ h^j(0)" apply (rule_tac m=i and n=j in diff_induct, simp_all) apply (blast del: subsetI intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h]) done -lemma directed_iterates: "bnd_mono(D,h) ==> directed({h^n (0). n\nat})" +lemma directed_iterates: "bnd_mono(D,h) \ directed({h^n (0). n\nat})" apply (simp add: directed_def, clarify) apply (rename_tac i j) apply (rule_tac x="i \ j" in bexI) @@ -42,8 +42,8 @@ lemma contin_iterates_eq: - "[|bnd_mono(D, h); contin(h)|] - ==> h(\n\nat. h^n (0)) = (\n\nat. h^n (0))" + "\bnd_mono(D, h); contin(h)\ + \ h(\n\nat. h^n (0)) = (\n\nat. h^n (0))" apply (simp add: contin_def directed_iterates) apply (rule trans) apply (rule equalityI) @@ -56,14 +56,14 @@ done lemma lfp_subset_Union: - "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) \ (\n\nat. h^n(0))" + "\bnd_mono(D, h); contin(h)\ \ lfp(D,h) \ (\n\nat. h^n(0))" apply (rule lfp_lowerbound) apply (simp add: contin_iterates_eq) apply (simp add: contin_def bnd_mono_iterates_subset UN_subset_iff) done lemma Union_subset_lfp: - "bnd_mono(D,h) ==> (\n\nat. h^n(0)) \ lfp(D,h)" + "bnd_mono(D,h) \ (\n\nat. h^n(0)) \ lfp(D,h)" apply (simp add: UN_subset_iff) apply (rule ballI) apply (induct_tac n, simp_all) @@ -73,23 +73,23 @@ done lemma lfp_eq_Union: - "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) = (\n\nat. h^n(0))" + "\bnd_mono(D, h); contin(h)\ \ lfp(D,h) = (\n\nat. h^n(0))" by (blast del: subsetI intro: lfp_subset_Union Union_subset_lfp) subsubsection\Some Standard Datatype Constructions Preserve Continuity\ -lemma contin_imp_mono: "[|X\Y; contin(F)|] ==> F(X) \ F(Y)" +lemma contin_imp_mono: "\X\Y; contin(F)\ \ F(X) \ F(Y)" apply (simp add: contin_def) apply (drule_tac x="{X,Y}" in spec) apply (simp add: directed_def subset_Un_iff2 Un_commute) done -lemma sum_contin: "[|contin(F); contin(G)|] ==> contin(\X. F(X) + G(X))" +lemma sum_contin: "\contin(F); contin(G)\ \ contin(\X. F(X) + G(X))" by (simp add: contin_def, blast) -lemma prod_contin: "[|contin(F); contin(G)|] ==> contin(\X. F(X) * G(X))" +lemma prod_contin: "\contin(F); contin(G)\ \ contin(\X. F(X) * G(X))" apply (subgoal_tac "\B C. F(B) \ F(B \ C)") prefer 2 apply (simp add: Un_upper1 contin_imp_mono) apply (subgoal_tac "\B C. G(C) \ G(B \ C)") @@ -115,38 +115,38 @@ definition iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where - "iterates_MH(M,isF,v,n,g,z) == + "iterates_MH(M,isF,v,n,g,z) \ is_nat_case(M, v, \m u. \gm[M]. fun_apply(M,g,m,gm) & isF(gm,u), n, z)" definition is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o" where - "is_iterates(M,isF,v,n,Z) == + "is_iterates(M,isF,v,n,Z) \ \sn[M]. \msn[M]. successor(M,n,sn) & membership(M,sn,msn) & is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)" definition iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where - "iterates_replacement(M,isF,v) == + "iterates_replacement(M,isF,v) \ \n[M]. n\nat \ wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))" lemma (in M_basic) iterates_MH_abs: - "[| relation1(M,isF,F); M(n); M(g); M(z) |] - ==> iterates_MH(M,isF,v,n,g,z) \ z = nat_case(v, \m. F(g`m), n)" + "\relation1(M,isF,F); M(n); M(g); M(z)\ + \ iterates_MH(M,isF,v,n,g,z) \ z = nat_case(v, \m. F(g`m), n)" by (simp add: nat_case_abs [of _ "\m. F(g ` m)"] relation1_def iterates_MH_def) lemma (in M_trancl) iterates_imp_wfrec_replacement: - "[|relation1(M,isF,F); n \ nat; iterates_replacement(M,isF,v)|] - ==> wfrec_replacement(M, \n f z. z = nat_case(v, \m. F(f`m), n), + "\relation1(M,isF,F); n \ nat; iterates_replacement(M,isF,v)\ + \ wfrec_replacement(M, \n f z. z = nat_case(v, \m. F(f`m), n), Memrel(succ(n)))" by (simp add: iterates_replacement_def iterates_MH_abs) theorem (in M_trancl) iterates_abs: - "[| iterates_replacement(M,isF,v); relation1(M,isF,F); - n \ nat; M(v); M(z); \x[M]. M(F(x)) |] - ==> is_iterates(M,isF,v,n,z) \ z = iterates(F,n,v)" + "\iterates_replacement(M,isF,v); relation1(M,isF,F); + n \ nat; M(v); M(z); \x[M]. M(F(x))\ + \ is_iterates(M,isF,v,n,z) \ z = iterates(F,n,v)" apply (frule iterates_imp_wfrec_replacement, assumption+) apply (simp add: wf_Memrel trans_Memrel relation_Memrel is_iterates_def relation2_def iterates_MH_abs @@ -157,9 +157,9 @@ lemma (in M_trancl) iterates_closed [intro,simp]: - "[| iterates_replacement(M,isF,v); relation1(M,isF,F); - n \ nat; M(v); \x[M]. M(F(x)) |] - ==> M(iterates(F,n,v))" + "\iterates_replacement(M,isF,v); relation1(M,isF,F); + n \ nat; M(v); \x[M]. M(F(x))\ + \ M(iterates(F,n,v))" apply (frule iterates_imp_wfrec_replacement, assumption+) apply (simp add: wf_Memrel trans_Memrel relation_Memrel relation2_def iterates_MH_abs @@ -210,12 +210,12 @@ definition is_list_functor :: "[i=>o,i,i,i] => o" where - "is_list_functor(M,A,X,Z) == + "is_list_functor(M,A,X,Z) \ \n1[M]. \AX[M]. number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" lemma (in M_basic) list_functor_abs [simp]: - "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) \ (Z = {0} + A*X)" + "\M(A); M(X); M(Z)\ \ is_list_functor(M,A,X,Z) \ (Z = {0} + A*X)" by (simp add: is_list_functor_def singleton_0 nat_into_M) @@ -263,7 +263,7 @@ definition is_formula_functor :: "[i=>o,i,i] => o" where - "is_formula_functor(M,X,Z) == + "is_formula_functor(M,X,Z) \ \nat'[M]. \natnat[M]. \natnatsum[M]. \XX[M]. \X3[M]. omega(M,nat') & cartprod(M,nat',nat',natnat) & is_sum(M,natnat,natnat,natnatsum) & @@ -271,8 +271,8 @@ is_sum(M,natnatsum,X3,Z)" lemma (in M_trancl) formula_functor_abs [simp]: - "[| M(X); M(Z) |] - ==> is_formula_functor(M,X,Z) \ + "\M(X); M(Z)\ + \ is_formula_functor(M,X,Z) \ Z = ((nat*nat) + (nat*nat)) + (X*X + X)" by (simp add: is_formula_functor_def) @@ -281,7 +281,7 @@ definition list_N :: "[i,i] => i" where - "list_N(A,n) == (\X. {0} + A * X)^n (0)" + "list_N(A,n) \ (\X. {0} + A * X)^n (0)" lemma Nil_in_list_N [simp]: "[] \ list_N(A,succ(n))" by (simp add: list_N_def Nil_def) @@ -299,25 +299,25 @@ by (simp add: list_N_def) lemma list_N_imp_list: - "[| l \ list_N(A,n); n \ nat |] ==> l \ list(A)" + "\l \ list_N(A,n); n \ nat\ \ l \ list(A)" by (force simp add: list_eq_Union list_N_def) lemma list_N_imp_length_lt [rule_format]: - "n \ nat ==> \l \ list_N(A,n). length(l) < n" + "n \ nat \ \l \ list_N(A,n). length(l) < n" apply (induct_tac n) apply (auto simp add: list_N_0 list_N_succ Nil_def [symmetric] Cons_def [symmetric]) done lemma list_imp_list_N [rule_format]: - "l \ list(A) ==> \n\nat. length(l) < n \ l \ list_N(A, n)" + "l \ list(A) \ \n\nat. length(l) < n \ l \ list_N(A, n)" apply (induct_tac l) apply (force elim: natE)+ done lemma list_N_imp_eq_length: - "[|n \ nat; l \ list_N(A, n); l \ list_N(A, succ(n))|] - ==> n = length(l)" + "\n \ nat; l \ list_N(A, n); l \ list_N(A, succ(n))\ + \ n = length(l)" apply (rule le_anti_sym) prefer 2 apply (simp add: list_N_imp_length_lt) apply (frule list_N_imp_list, simp) @@ -328,7 +328,7 @@ text\Express \<^term>\list_rec\ without using \<^term>\rank\ or \<^term>\Vset\, neither of which is absolute.\ lemma (in M_trivial) list_rec_eq: - "l \ list(A) ==> + "l \ list(A) \ list_rec(a,g,l) = transrec (succ(length(l)), \x h. Lambda (list(A), @@ -342,19 +342,19 @@ definition is_list_N :: "[i=>o,i,i,i] => o" where - "is_list_N(M,A,n,Z) == + "is_list_N(M,A,n,Z) \ \zero[M]. empty(M,zero) & is_iterates(M, is_list_functor(M,A), zero, n, Z)" definition mem_list :: "[i=>o,i,i] => o" where - "mem_list(M,A,l) == + "mem_list(M,A,l) \ \n[M]. \listn[M]. finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \ listn" definition is_list :: "[i=>o,i,i] => o" where - "is_list(M,A,Z) == \l[M]. l \ Z \ mem_list(M,A,l)" + "is_list(M,A,Z) \ \l[M]. l \ Z \ mem_list(M,A,l)" subsubsection\Towards Absoluteness of \<^term>\formula_rec\\ @@ -365,13 +365,13 @@ "depth(Nand(p,q)) = succ(depth(p) \ depth(q))" "depth(Forall(p)) = succ(depth(p))" -lemma depth_type [TC]: "p \ formula ==> depth(p) \ nat" +lemma depth_type [TC]: "p \ formula \ depth(p) \ nat" by (induct_tac p, simp_all) definition formula_N :: "i => i" where - "formula_N(n) == (\X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)" + "formula_N(n) \ (\X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)" lemma Member_in_formula_N [simp]: "Member(x,y) \ formula_N(succ(n)) \ x \ nat & y \ nat" @@ -400,11 +400,11 @@ by (simp add: formula_N_def) lemma formula_N_imp_formula: - "[| p \ formula_N(n); n \ nat |] ==> p \ formula" + "\p \ formula_N(n); n \ nat\ \ p \ formula" by (force simp add: formula_eq_Union formula_N_def) lemma formula_N_imp_depth_lt [rule_format]: - "n \ nat ==> \p \ formula_N(n). depth(p) < n" + "n \ nat \ \p \ formula_N(n). depth(p) < n" apply (induct_tac n) apply (auto simp add: formula_N_0 formula_N_succ depth_type formula_N_imp_formula Un_least_lt_iff @@ -413,15 +413,15 @@ done lemma formula_imp_formula_N [rule_format]: - "p \ formula ==> \n\nat. depth(p) < n \ p \ formula_N(n)" + "p \ formula \ \n\nat. depth(p) < n \ p \ formula_N(n)" apply (induct_tac p) apply (simp_all add: succ_Un_distrib Un_least_lt_iff) apply (force elim: natE)+ done lemma formula_N_imp_eq_depth: - "[|n \ nat; p \ formula_N(n); p \ formula_N(succ(n))|] - ==> n = depth(p)" + "\n \ nat; p \ formula_N(n); p \ formula_N(succ(n))\ + \ n = depth(p)" apply (rule le_anti_sym) prefer 2 apply (simp add: formula_N_imp_depth_lt) apply (frule formula_N_imp_formula, simp) @@ -432,13 +432,13 @@ text\This result and the next are unused.\ lemma formula_N_mono [rule_format]: - "[| m \ nat; n \ nat |] ==> m\n \ formula_N(m) \ formula_N(n)" + "\m \ nat; n \ nat\ \ m\n \ formula_N(m) \ formula_N(n)" apply (rule_tac m = m and n = n in diff_induct) apply (simp_all add: formula_N_0 formula_N_succ, blast) done lemma formula_N_distrib: - "[| m \ nat; n \ nat |] ==> formula_N(m \ n) = formula_N(m) \ formula_N(n)" + "\m \ nat; n \ nat\ \ formula_N(m \ n) = formula_N(m) \ formula_N(n)" apply (rule_tac i = m and j = n in Ord_linear_le, auto) apply (simp_all add: subset_Un_iff [THEN iffD1] subset_Un_iff2 [THEN iffD1] le_imp_subset formula_N_mono) @@ -446,26 +446,26 @@ definition is_formula_N :: "[i=>o,i,i] => o" where - "is_formula_N(M,n,Z) == + "is_formula_N(M,n,Z) \ \zero[M]. empty(M,zero) & is_iterates(M, is_formula_functor(M), zero, n, Z)" definition mem_formula :: "[i=>o,i] => o" where - "mem_formula(M,p) == + "mem_formula(M,p) \ \n[M]. \formn[M]. finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \ formn" definition is_formula :: "[i=>o,i] => o" where - "is_formula(M,Z) == \p[M]. p \ Z \ mem_formula(M,p)" + "is_formula(M,Z) \ \p[M]. p \ Z \ mem_formula(M,p)" locale M_datatypes = M_trancl + assumes list_replacement1: - "M(A) ==> iterates_replacement(M, is_list_functor(M,A), 0)" + "M(A) \ iterates_replacement(M, is_list_functor(M,A), 0)" and list_replacement2: - "M(A) ==> strong_replacement(M, + "M(A) \ strong_replacement(M, \n y. n\nat & is_iterates(M, is_list_functor(M,A), 0, n, y))" and formula_replacement1: "iterates_replacement(M, is_formula_functor(M), 0)" @@ -473,13 +473,13 @@ "strong_replacement(M, \n y. n\nat & is_iterates(M, is_formula_functor(M), 0, n, y))" and nth_replacement: - "M(l) ==> iterates_replacement(M, %l t. is_tl(M,l,t), l)" + "M(l) \ iterates_replacement(M, %l t. is_tl(M,l,t), l)" subsubsection\Absoluteness of the List Construction\ lemma (in M_datatypes) list_replacement2': - "M(A) ==> strong_replacement(M, \n y. n\nat & y = (\X. {0} + A * X)^n (0))" + "M(A) \ strong_replacement(M, \n y. n\nat & y = (\X. {0} + A * X)^n (0))" apply (insert list_replacement2 [of A]) apply (rule strong_replacement_cong [THEN iffD1]) apply (rule conj_cong [OF iff_refl iterates_abs [of "is_list_functor(M,A)"]]) @@ -487,7 +487,7 @@ done lemma (in M_datatypes) list_closed [intro,simp]: - "M(A) ==> M(list(A))" + "M(A) \ M(list(A))" apply (insert list_replacement1) by (simp add: RepFun_closed2 list_eq_Union list_replacement2' relation1_def @@ -497,29 +497,29 @@ lemmas (in M_datatypes) list_into_M = transM [OF _ list_closed] lemma (in M_datatypes) list_N_abs [simp]: - "[|M(A); n\nat; M(Z)|] - ==> is_list_N(M,A,n,Z) \ Z = list_N(A,n)" + "\M(A); n\nat; M(Z)\ + \ is_list_N(M,A,n,Z) \ Z = list_N(A,n)" apply (insert list_replacement1) apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M iterates_abs [of "is_list_functor(M,A)" _ "\X. {0} + A*X"]) done lemma (in M_datatypes) list_N_closed [intro,simp]: - "[|M(A); n\nat|] ==> M(list_N(A,n))" + "\M(A); n\nat\ \ M(list_N(A,n))" apply (insert list_replacement1) apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M iterates_closed [of "is_list_functor(M,A)"]) done lemma (in M_datatypes) mem_list_abs [simp]: - "M(A) ==> mem_list(M,A,l) \ l \ list(A)" + "M(A) \ mem_list(M,A,l) \ l \ list(A)" apply (insert list_replacement1) apply (simp add: mem_list_def list_N_def relation1_def list_eq_Union iterates_closed [of "is_list_functor(M,A)"]) done lemma (in M_datatypes) list_abs [simp]: - "[|M(A); M(Z)|] ==> is_list(M,A,Z) \ Z = list(A)" + "\M(A); M(Z)\ \ is_list(M,A,Z) \ Z = list(A)" apply (simp add: is_list_def, safe) apply (rule M_equalityI, simp_all) done @@ -545,8 +545,8 @@ lemmas (in M_datatypes) formula_into_M = transM [OF _ formula_closed] lemma (in M_datatypes) formula_N_abs [simp]: - "[|n\nat; M(Z)|] - ==> is_formula_N(M,n,Z) \ Z = formula_N(n)" + "\n\nat; M(Z)\ + \ is_formula_N(M,n,Z) \ Z = formula_N(n)" apply (insert formula_replacement1) apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M iterates_abs [of "is_formula_functor(M)" _ @@ -554,7 +554,7 @@ done lemma (in M_datatypes) formula_N_closed [intro,simp]: - "n\nat ==> M(formula_N(n))" + "n\nat \ M(formula_N(n))" apply (insert formula_replacement1) apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M iterates_closed [of "is_formula_functor(M)"]) @@ -568,7 +568,7 @@ done lemma (in M_datatypes) formula_abs [simp]: - "[|M(Z)|] ==> is_formula(M,Z) \ Z = formula" + "\M(Z)\ \ is_formula(M,Z) \ Z = formula" apply (simp add: is_formula_def, safe) apply (rule M_equalityI, simp_all) done @@ -589,28 +589,28 @@ definition is_eclose_n :: "[i=>o,i,i,i] => o" where - "is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z)" + "is_eclose_n(M,A,n,Z) \ is_iterates(M, big_union(M), A, n, Z)" definition mem_eclose :: "[i=>o,i,i] => o" where - "mem_eclose(M,A,l) == + "mem_eclose(M,A,l) \ \n[M]. \eclosen[M]. finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \ eclosen" definition is_eclose :: "[i=>o,i,i] => o" where - "is_eclose(M,A,Z) == \u[M]. u \ Z \ mem_eclose(M,A,u)" + "is_eclose(M,A,Z) \ \u[M]. u \ Z \ mem_eclose(M,A,u)" locale M_eclose = M_datatypes + assumes eclose_replacement1: - "M(A) ==> iterates_replacement(M, big_union(M), A)" + "M(A) \ iterates_replacement(M, big_union(M), A)" and eclose_replacement2: - "M(A) ==> strong_replacement(M, + "M(A) \ strong_replacement(M, \n y. n\nat & is_iterates(M, big_union(M), A, n, y))" lemma (in M_eclose) eclose_replacement2': - "M(A) ==> strong_replacement(M, \n y. n\nat & y = Union^n (A))" + "M(A) \ strong_replacement(M, \n y. n\nat & y = Union^n (A))" apply (insert eclose_replacement2 [of A]) apply (rule strong_replacement_cong [THEN iffD1]) apply (rule conj_cong [OF iff_refl iterates_abs [of "big_union(M)"]]) @@ -618,28 +618,28 @@ done lemma (in M_eclose) eclose_closed [intro,simp]: - "M(A) ==> M(eclose(A))" + "M(A) \ M(eclose(A))" apply (insert eclose_replacement1) by (simp add: RepFun_closed2 eclose_eq_Union eclose_replacement2' relation1_def iterates_closed [of "big_union(M)"]) lemma (in M_eclose) is_eclose_n_abs [simp]: - "[|M(A); n\nat; M(Z)|] ==> is_eclose_n(M,A,n,Z) \ Z = Union^n (A)" + "\M(A); n\nat; M(Z)\ \ is_eclose_n(M,A,n,Z) \ Z = Union^n (A)" apply (insert eclose_replacement1) apply (simp add: is_eclose_n_def relation1_def nat_into_M iterates_abs [of "big_union(M)" _ "Union"]) done lemma (in M_eclose) mem_eclose_abs [simp]: - "M(A) ==> mem_eclose(M,A,l) \ l \ eclose(A)" + "M(A) \ mem_eclose(M,A,l) \ l \ eclose(A)" apply (insert eclose_replacement1) apply (simp add: mem_eclose_def relation1_def eclose_eq_Union iterates_closed [of "big_union(M)"]) done lemma (in M_eclose) eclose_abs [simp]: - "[|M(A); M(Z)|] ==> is_eclose(M,A,Z) \ Z = eclose(A)" + "\M(A); M(Z)\ \ is_eclose(M,A,Z) \ Z = eclose(A)" apply (simp add: is_eclose_def, safe) apply (rule M_equalityI, simp_all) done @@ -651,14 +651,14 @@ definition is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where - "is_transrec(M,MH,a,z) == + "is_transrec(M,MH,a,z) \ \sa[M]. \esa[M]. \mesa[M]. upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) & is_wfrec(M,MH,mesa,a,z)" definition transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where - "transrec_replacement(M,MH,a) == + "transrec_replacement(M,MH,a) \ \sa[M]. \esa[M]. \mesa[M]. upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) & wfrec_replacement(M,MH,mesa)" @@ -667,30 +667,30 @@ \trans_wfrec_abs\ rather than \trans_wfrec_abs\, which I haven't even proved yet.\ theorem (in M_eclose) transrec_abs: - "[|transrec_replacement(M,MH,i); relation2(M,MH,H); + "\transrec_replacement(M,MH,i); relation2(M,MH,H); Ord(i); M(i); M(z); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> is_transrec(M,MH,i,z) \ z = transrec(i,H)" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ is_transrec(M,MH,i,z) \ z = transrec(i,H)" by (simp add: trans_wfrec_abs transrec_replacement_def is_transrec_def transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel) theorem (in M_eclose) transrec_closed: - "[|transrec_replacement(M,MH,i); relation2(M,MH,H); + "\transrec_replacement(M,MH,i); relation2(M,MH,H); Ord(i); M(i); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> M(transrec(i,H))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ M(transrec(i,H))" by (simp add: trans_wfrec_closed transrec_replacement_def is_transrec_def transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel) text\Helps to prove instances of \<^term>\transrec_replacement\\ lemma (in M_eclose) transrec_replacementI: - "[|M(a); + "\M(a); strong_replacement (M, \x z. \y[M]. pair(M, x, y, z) & - is_wfrec(M,MH,Memrel(eclose({a})),x,y))|] - ==> transrec_replacement(M,MH,a)" + is_wfrec(M,MH,Memrel(eclose({a})),x,y))\ + \ transrec_replacement(M,MH,a)" by (simp add: transrec_replacement_def wfrec_replacement_def) @@ -699,14 +699,14 @@ definition is_length :: "[i=>o,i,i,i] => o" where - "is_length(M,A,l,n) == + "is_length(M,A,l,n) \ \sn[M]. \list_n[M]. \list_sn[M]. is_list_N(M,A,n,list_n) & l \ list_n & successor(M,n,sn) & is_list_N(M,A,sn,list_sn) & l \ list_sn" lemma (in M_datatypes) length_abs [simp]: - "[|M(A); l \ list(A); n \ nat|] ==> is_length(M,A,l,n) \ n = length(l)" + "\M(A); l \ list(A); n \ nat\ \ is_length(M,A,l,n) \ n = length(l)" apply (subgoal_tac "M(l) & M(n)") prefer 2 apply (blast dest: transM) apply (simp add: is_length_def) @@ -716,14 +716,14 @@ text\Proof is trivial since \<^term>\length\ returns natural numbers.\ lemma (in M_trivial) length_closed [intro,simp]: - "l \ list(A) ==> M(length(l))" + "l \ list(A) \ M(length(l))" by (simp add: nat_into_M) subsection \Absoluteness for the List Operator \<^term>\nth\\ lemma nth_eq_hd_iterates_tl [rule_format]: - "xs \ list(A) ==> \n \ nat. nth(n,xs) = hd' (tl'^n (xs))" + "xs \ list(A) \ \n \ nat. nth(n,xs) = hd' (tl'^n (xs))" apply (induct_tac xs) apply (simp add: iterates_tl_Nil hd'_Nil, clarify) apply (erule natE) @@ -732,14 +732,14 @@ done lemma (in M_basic) iterates_tl'_closed: - "[|n \ nat; M(x)|] ==> M(tl'^n (x))" + "\n \ nat; M(x)\ \ M(tl'^n (x))" apply (induct_tac n, simp) apply (simp add: tl'_Cons tl'_closed) done text\Immediate by type-checking\ lemma (in M_datatypes) nth_closed [intro,simp]: - "[|xs \ list(A); n \ nat; M(A)|] ==> M(nth(n,xs))" + "\xs \ list(A); n \ nat; M(A)\ \ M(nth(n,xs))" apply (case_tac "n < length(xs)") apply (blast intro: nth_type transM) apply (simp add: not_lt_iff_le nth_eq_0) @@ -747,12 +747,12 @@ definition is_nth :: "[i=>o,i,i,i] => o" where - "is_nth(M,n,l,Z) == + "is_nth(M,n,l,Z) \ \X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" lemma (in M_datatypes) nth_abs [simp]: - "[|M(A); n \ nat; l \ list(A); M(Z)|] - ==> is_nth(M,n,l,Z) \ Z = nth(n,l)" + "\M(A); n \ nat; l \ list(A); M(Z)\ + \ is_nth(M,n,l,Z) \ Z = nth(n,l)" apply (subgoal_tac "M(l)") prefer 2 apply (blast intro: transM) apply (simp add: is_nth_def nth_eq_hd_iterates_tl nat_into_M @@ -766,11 +766,11 @@ definition is_Member :: "[i=>o,i,i,i] => o" where \ \because \<^term>\Member(x,y) \ Inl(Inl(\x,y\))\\ - "is_Member(M,x,y,Z) == + "is_Member(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" lemma (in M_trivial) Member_abs [simp]: - "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) \ (Z = Member(x,y))" + "\M(x); M(y); M(Z)\ \ is_Member(M,x,y,Z) \ (Z = Member(x,y))" by (simp add: is_Member_def Member_def) lemma (in M_trivial) Member_in_M_iff [iff]: @@ -780,11 +780,11 @@ definition is_Equal :: "[i=>o,i,i,i] => o" where \ \because \<^term>\Equal(x,y) \ Inl(Inr(\x,y\))\\ - "is_Equal(M,x,y,Z) == + "is_Equal(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" lemma (in M_trivial) Equal_abs [simp]: - "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) \ (Z = Equal(x,y))" + "\M(x); M(y); M(Z)\ \ is_Equal(M,x,y,Z) \ (Z = Equal(x,y))" by (simp add: is_Equal_def Equal_def) lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) \ M(x) & M(y)" @@ -793,11 +793,11 @@ definition is_Nand :: "[i=>o,i,i,i] => o" where \ \because \<^term>\Nand(x,y) \ Inr(Inl(\x,y\))\\ - "is_Nand(M,x,y,Z) == + "is_Nand(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" lemma (in M_trivial) Nand_abs [simp]: - "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) \ (Z = Nand(x,y))" + "\M(x); M(y); M(Z)\ \ is_Nand(M,x,y,Z) \ (Z = Nand(x,y))" by (simp add: is_Nand_def Nand_def) lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) \ M(x) & M(y)" @@ -806,10 +806,10 @@ definition is_Forall :: "[i=>o,i,i] => o" where \ \because \<^term>\Forall(x) \ Inr(Inr(p))\\ - "is_Forall(M,p,Z) == \u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" + "is_Forall(M,p,Z) \ \u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" lemma (in M_trivial) Forall_abs [simp]: - "[|M(x); M(Z)|] ==> is_Forall(M,x,Z) \ (Z = Forall(x))" + "\M(x); M(Z)\ \ is_Forall(M,x,Z) \ (Z = Forall(x))" by (simp add: is_Forall_def Forall_def) lemma (in M_trivial) Forall_in_M_iff [iff]: "M(Forall(x)) \ M(x)" @@ -822,7 +822,7 @@ definition formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where \ \the instance of \<^term>\formula_case\ in \<^term>\formula_rec\\ - "formula_rec_case(a,b,c,d,h) == + "formula_rec_case(a,b,c,d,h) \ formula_case (a, b, \u v. c(u, v, h ` succ(depth(u)) ` u, h ` succ(depth(v)) ` v), @@ -832,7 +832,7 @@ Express \<^term>\formula_rec\ without using \<^term>\rank\ or \<^term>\Vset\, neither of which is absolute.\ lemma (in M_trivial) formula_rec_eq: - "p \ formula ==> + "p \ formula \ formula_rec(a,b,c,d,p) = transrec (succ(depth(p)), \x h. Lambda (formula, formula_rec_case(a,b,c,d,h))) ` p" @@ -855,14 +855,14 @@ definition is_depth :: "[i=>o,i,i] => o" where - "is_depth(M,p,n) == + "is_depth(M,p,n) \ \sn[M]. \formula_n[M]. \formula_sn[M]. is_formula_N(M,n,formula_n) & p \ formula_n & successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \ formula_sn" lemma (in M_datatypes) depth_abs [simp]: - "[|p \ formula; n \ nat|] ==> is_depth(M,p,n) \ n = depth(p)" + "\p \ formula; n \ nat\ \ is_depth(M,p,n) \ n = depth(p)" apply (subgoal_tac "M(p) & M(n)") prefer 2 apply (blast dest: transM) apply (simp add: is_depth_def) @@ -872,7 +872,7 @@ text\Proof is trivial since \<^term>\depth\ returns natural numbers.\ lemma (in M_trivial) depth_closed [intro,simp]: - "p \ formula ==> M(depth(p))" + "p \ formula \ M(depth(p))" by (simp add: nat_into_M) @@ -882,7 +882,7 @@ is_formula_case :: "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where \ \no constraint on non-formulas\ - "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) == + "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) \ (\x[M]. \y[M]. finite_ordinal(M,x) \ finite_ordinal(M,y) \ is_Member(M,x,y,p) \ is_a(x,y,z)) & (\x[M]. \y[M]. finite_ordinal(M,x) \ finite_ordinal(M,y) \ @@ -892,10 +892,10 @@ (\x[M]. mem_formula(M,x) \ is_Forall(M,x,p) \ is_d(x,z))" lemma (in M_datatypes) formula_case_abs [simp]: - "[| Relation2(M,nat,nat,is_a,a); Relation2(M,nat,nat,is_b,b); + "\Relation2(M,nat,nat,is_a,a); Relation2(M,nat,nat,is_b,b); Relation2(M,formula,formula,is_c,c); Relation1(M,formula,is_d,d); - p \ formula; M(z) |] - ==> is_formula_case(M,is_a,is_b,is_c,is_d,p,z) \ + p \ formula; M(z)\ + \ is_formula_case(M,is_a,is_b,is_c,is_d,p,z) \ z = formula_case(a,b,c,d,p)" apply (simp add: formula_into_M is_formula_case_def) apply (erule formula.cases) @@ -903,11 +903,11 @@ done lemma (in M_datatypes) formula_case_closed [intro,simp]: - "[|p \ formula; + "\p \ formula; \x[M]. \y[M]. x\nat \ y\nat \ M(a(x,y)); \x[M]. \y[M]. x\nat \ y\nat \ M(b(x,y)); \x[M]. \y[M]. x\formula \ y\formula \ M(c(x,y)); - \x[M]. x\formula \ M(d(x))|] ==> M(formula_case(a,b,c,d,p))" + \x[M]. x\formula \ M(d(x))\ \ M(formula_case(a,b,c,d,p))" by (erule formula.cases, simp_all) @@ -916,7 +916,7 @@ definition is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where \ \predicate to relativize the functional \<^term>\formula_rec\\ - "is_formula_rec(M,MH,p,z) == + "is_formula_rec(M,MH,p,z) \ \dp[M]. \i[M]. \f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)" @@ -924,14 +924,14 @@ text\Sufficient conditions to relativize the instance of \<^term>\formula_case\ in \<^term>\formula_rec\\ lemma (in M_datatypes) Relation1_formula_rec_case: - "[|Relation2(M, nat, nat, is_a, a); + "\Relation2(M, nat, nat, is_a, a); Relation2(M, nat, nat, is_b, b); Relation2 (M, formula, formula, is_c, \u v. c(u, v, h`succ(depth(u))`u, h`succ(depth(v))`v)); Relation1(M, formula, is_d, \u. d(u, h ` succ(depth(u)) ` u)); - M(h) |] - ==> Relation1(M, formula, + M(h)\ + \ Relation1(M, formula, is_formula_case (M, is_a, is_b, is_c, is_d), formula_rec_case(a, b, c, d, h))" apply (simp (no_asm) add: formula_rec_case_def Relation1_def) @@ -945,38 +945,38 @@ locale Formula_Rec = M_eclose + fixes a and is_a and b and is_b and c and is_c and d and is_d and MH defines - "MH(u::i,f,z) == + "MH(u::i,f,z) \ \fml[M]. is_formula(M,fml) \ is_lambda (M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)" - assumes a_closed: "[|x\nat; y\nat|] ==> M(a(x,y))" + assumes a_closed: "\x\nat; y\nat\ \ M(a(x,y))" and a_rel: "Relation2(M, nat, nat, is_a, a)" - and b_closed: "[|x\nat; y\nat|] ==> M(b(x,y))" + and b_closed: "\x\nat; y\nat\ \ M(b(x,y))" and b_rel: "Relation2(M, nat, nat, is_b, b)" - and c_closed: "[|x \ formula; y \ formula; M(gx); M(gy)|] - ==> M(c(x, y, gx, gy))" + and c_closed: "\x \ formula; y \ formula; M(gx); M(gy)\ + \ M(c(x, y, gx, gy))" and c_rel: - "M(f) ==> + "M(f) \ Relation2 (M, formula, formula, is_c(f), \u v. c(u, v, f ` succ(depth(u)) ` u, f ` succ(depth(v)) ` v))" - and d_closed: "[|x \ formula; M(gx)|] ==> M(d(x, gx))" + and d_closed: "\x \ formula; M(gx)\ \ M(d(x, gx))" and d_rel: - "M(f) ==> + "M(f) \ Relation1(M, formula, is_d(f), \u. d(u, f ` succ(depth(u)) ` u))" - and fr_replace: "n \ nat ==> transrec_replacement(M,MH,n)" + and fr_replace: "n \ nat \ transrec_replacement(M,MH,n)" and fr_lam_replace: - "M(g) ==> + "M(g) \ strong_replacement (M, \x y. x \ formula & y = \x, formula_rec_case(a,b,c,d,g,x)\)" lemma (in Formula_Rec) formula_rec_case_closed: - "[|M(g); p \ formula|] ==> M(formula_rec_case(a, b, c, d, g, p))" + "\M(g); p \ formula\ \ M(formula_rec_case(a, b, c, d, g, p))" by (simp add: formula_rec_case_def a_closed b_closed c_closed d_closed) lemma (in Formula_Rec) formula_rec_lam_closed: - "M(g) ==> M(Lambda (formula, formula_rec_case(a,b,c,d,g)))" + "M(g) \ M(Lambda (formula, formula_rec_case(a,b,c,d,g)))" by (simp add: lam_closed2 fr_lam_replace formula_rec_case_closed) lemma (in Formula_Rec) MH_rel2: @@ -990,20 +990,20 @@ lemma (in Formula_Rec) fr_transrec_closed: "n \ nat - ==> M(transrec + \ M(transrec (n, \x h. Lambda(formula, formula_rec_case(a, b, c, d, h))))" by (simp add: transrec_closed [OF fr_replace MH_rel2] nat_into_M formula_rec_lam_closed) text\The main two results: \<^term>\formula_rec\ is absolute for \<^term>\M\.\ theorem (in Formula_Rec) formula_rec_closed: - "p \ formula ==> M(formula_rec(a,b,c,d,p))" + "p \ formula \ M(formula_rec(a,b,c,d,p))" by (simp add: formula_rec_eq fr_transrec_closed transM [OF _ formula_closed]) theorem (in Formula_Rec) formula_rec_abs: - "[| p \ formula; M(z)|] - ==> is_formula_rec(M,MH,p,z) \ z = formula_rec(a,b,c,d,p)" + "\p \ formula; M(z)\ + \ is_formula_rec(M,MH,p,z) \ z = formula_rec(a,b,c,d,p)" by (simp add: is_formula_rec_def formula_rec_eq transM [OF _ formula_closed] transrec_abs [OF fr_replace MH_rel2] depth_type fr_transrec_closed formula_rec_lam_closed eq_commute) diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Formula.thy --- a/src/ZF/Constructible/Formula.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Formula.thy Tue Sep 27 16:51:35 2022 +0100 @@ -22,46 +22,46 @@ definition Neg :: "i=>i" where - "Neg(p) == Nand(p,p)" + "Neg(p) \ Nand(p,p)" definition And :: "[i,i]=>i" where - "And(p,q) == Neg(Nand(p,q))" + "And(p,q) \ Neg(Nand(p,q))" definition Or :: "[i,i]=>i" where - "Or(p,q) == Nand(Neg(p),Neg(q))" + "Or(p,q) \ Nand(Neg(p),Neg(q))" definition Implies :: "[i,i]=>i" where - "Implies(p,q) == Nand(p,Neg(q))" + "Implies(p,q) \ Nand(p,Neg(q))" definition Iff :: "[i,i]=>i" where - "Iff(p,q) == And(Implies(p,q), Implies(q,p))" + "Iff(p,q) \ And(Implies(p,q), Implies(q,p))" definition Exists :: "i=>i" where - "Exists(p) == Neg(Forall(Neg(p)))" + "Exists(p) \ Neg(Forall(Neg(p)))" -lemma Neg_type [TC]: "p \ formula ==> Neg(p) \ formula" +lemma Neg_type [TC]: "p \ formula \ Neg(p) \ formula" by (simp add: Neg_def) -lemma And_type [TC]: "[| p \ formula; q \ formula |] ==> And(p,q) \ formula" +lemma And_type [TC]: "\p \ formula; q \ formula\ \ And(p,q) \ formula" by (simp add: And_def) -lemma Or_type [TC]: "[| p \ formula; q \ formula |] ==> Or(p,q) \ formula" +lemma Or_type [TC]: "\p \ formula; q \ formula\ \ Or(p,q) \ formula" by (simp add: Or_def) lemma Implies_type [TC]: - "[| p \ formula; q \ formula |] ==> Implies(p,q) \ formula" + "\p \ formula; q \ formula\ \ Implies(p,q) \ formula" by (simp add: Implies_def) lemma Iff_type [TC]: - "[| p \ formula; q \ formula |] ==> Iff(p,q) \ formula" + "\p \ formula; q \ formula\ \ Iff(p,q) \ formula" by (simp add: Iff_def) -lemma Exists_type [TC]: "p \ formula ==> Exists(p) \ formula" +lemma Exists_type [TC]: "p \ formula \ Exists(p) \ formula" by (simp add: Exists_def) @@ -85,7 +85,7 @@ abbreviation sats :: "[i,i,i] => o" where - "sats(A,p,env) == satisfies(A,p)`env = 1" + "sats(A,p,env) \ satisfies(A,p)`env = 1" lemma sats_Member_iff [simp]: "env \ list(A) \ sats(A, Member(x,y), env) \ nth(x,env) \ nth(y,env)" @@ -97,12 +97,12 @@ lemma sats_Nand_iff [simp]: "env \ list(A) - ==> (sats(A, Nand(p,q), env)) \ ~ (sats(A,p,env) & sats(A,q,env))" + \ (sats(A, Nand(p,q), env)) \ \ (sats(A,p,env) & sats(A,q,env))" by (simp add: Bool.and_def Bool.not_def cond_def) lemma sats_Forall_iff [simp]: "env \ list(A) - ==> sats(A, Forall(p), env) \ (\x\A. sats(A, p, Cons(x,env)))" + \ sats(A, Forall(p), env) \ (\x\A. sats(A, p, Cons(x,env)))" by simp declare satisfies.simps [simp del] @@ -111,80 +111,80 @@ lemma sats_Neg_iff [simp]: "env \ list(A) - ==> sats(A, Neg(p), env) \ ~ sats(A,p,env)" + \ sats(A, Neg(p), env) \ \ sats(A,p,env)" by (simp add: Neg_def) lemma sats_And_iff [simp]: "env \ list(A) - ==> (sats(A, And(p,q), env)) \ sats(A,p,env) & sats(A,q,env)" + \ (sats(A, And(p,q), env)) \ sats(A,p,env) & sats(A,q,env)" by (simp add: And_def) lemma sats_Or_iff [simp]: "env \ list(A) - ==> (sats(A, Or(p,q), env)) \ sats(A,p,env) | sats(A,q,env)" + \ (sats(A, Or(p,q), env)) \ sats(A,p,env) | sats(A,q,env)" by (simp add: Or_def) lemma sats_Implies_iff [simp]: "env \ list(A) - ==> (sats(A, Implies(p,q), env)) \ (sats(A,p,env) \ sats(A,q,env))" + \ (sats(A, Implies(p,q), env)) \ (sats(A,p,env) \ sats(A,q,env))" by (simp add: Implies_def, blast) lemma sats_Iff_iff [simp]: "env \ list(A) - ==> (sats(A, Iff(p,q), env)) \ (sats(A,p,env) \ sats(A,q,env))" + \ (sats(A, Iff(p,q), env)) \ (sats(A,p,env) \ sats(A,q,env))" by (simp add: Iff_def, blast) lemma sats_Exists_iff [simp]: "env \ list(A) - ==> sats(A, Exists(p), env) \ (\x\A. sats(A, p, Cons(x,env)))" + \ sats(A, Exists(p), env) \ (\x\A. sats(A, p, Cons(x,env)))" by (simp add: Exists_def) subsubsection\Derived rules to help build up formulas\ lemma mem_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; env \ list(A)|] - ==> (x\y) \ sats(A, Member(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; env \ list(A)\ + \ (x\y) \ sats(A, Member(i,j), env)" by (simp add: satisfies.simps) lemma equal_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; env \ list(A)|] - ==> (x=y) \ sats(A, Equal(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; env \ list(A)\ + \ (x=y) \ sats(A, Equal(i,j), env)" by (simp add: satisfies.simps) lemma not_iff_sats: - "[| P \ sats(A,p,env); env \ list(A)|] - ==> (~P) \ sats(A, Neg(p), env)" + "\P \ sats(A,p,env); env \ list(A)\ + \ (\P) \ sats(A, Neg(p), env)" by simp lemma conj_iff_sats: - "[| P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)|] - ==> (P & Q) \ sats(A, And(p,q), env)" + "\P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)\ + \ (P & Q) \ sats(A, And(p,q), env)" by (simp) lemma disj_iff_sats: - "[| P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)|] - ==> (P | Q) \ sats(A, Or(p,q), env)" + "\P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)\ + \ (P | Q) \ sats(A, Or(p,q), env)" by (simp) lemma iff_iff_sats: - "[| P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)|] - ==> (P \ Q) \ sats(A, Iff(p,q), env)" + "\P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)\ + \ (P \ Q) \ sats(A, Iff(p,q), env)" by (simp) lemma imp_iff_sats: - "[| P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)|] - ==> (P \ Q) \ sats(A, Implies(p,q), env)" + "\P \ sats(A,p,env); Q \ sats(A,q,env); env \ list(A)\ + \ (P \ Q) \ sats(A, Implies(p,q), env)" by (simp) lemma ball_iff_sats: - "[| !!x. x\A ==> P(x) \ sats(A, p, Cons(x, env)); env \ list(A)|] - ==> (\x\A. P(x)) \ sats(A, Forall(p), env)" + "\\x. x\A \ P(x) \ sats(A, p, Cons(x, env)); env \ list(A)\ + \ (\x\A. P(x)) \ sats(A, Forall(p), env)" by (simp) lemma bex_iff_sats: - "[| !!x. x\A ==> P(x) \ sats(A, p, Cons(x, env)); env \ list(A)|] - ==> (\x\A. P(x)) \ sats(A, Exists(p), env)" + "\\x. x\A \ P(x) \ sats(A, p, Cons(x, env)); env \ list(A)\ + \ (\x\A. P(x)) \ sats(A, Exists(p), env)" by (simp) lemmas FOL_iff_sats = @@ -206,7 +206,7 @@ "arity(Forall(p)) = Arith.pred(arity(p))" -lemma arity_type [TC]: "p \ formula ==> arity(p) \ nat" +lemma arity_type [TC]: "p \ formula \ arity(p) \ nat" by (induct_tac p, simp_all) lemma arity_Neg [simp]: "arity(Neg(p)) = arity(p)" @@ -229,8 +229,8 @@ lemma arity_sats_iff [rule_format]: - "[| p \ formula; extra \ list(A) |] - ==> \env \ list(A). + "\p \ formula; extra \ list(A)\ + \ \env \ list(A). arity(p) \ length(env) \ sats(A, p, env @ extra) \ sats(A, p, env)" apply (induct_tac p) @@ -239,9 +239,9 @@ done lemma arity_sats1_iff: - "[| arity(p) \ succ(length(env)); p \ formula; x \ A; env \ list(A); - extra \ list(A) |] - ==> sats(A, p, Cons(x, env @ extra)) \ sats(A, p, Cons(x, env))" + "\arity(p) \ succ(length(env)); p \ formula; x \ A; env \ list(A); + extra \ list(A)\ + \ sats(A, p, Cons(x, env @ extra)) \ sats(A, p, Cons(x, env))" apply (insert arity_sats_iff [of p extra A "Cons(x,env)"]) apply simp done @@ -251,12 +251,12 @@ definition incr_var :: "[i,i]=>i" where - "incr_var(x,nq) == if x if x incr_var(x,nq) = x" +lemma incr_var_lt: "x incr_var(x,nq) = x" by (simp add: incr_var_def) -lemma incr_var_le: "nq\x ==> incr_var(x,nq) = succ(x)" +lemma incr_var_le: "nq\x \ incr_var(x,nq) = succ(x)" apply (simp add: incr_var_def) apply (blast dest: lt_trans1) done @@ -276,10 +276,10 @@ (\nq \ nat. Forall (incr_bv(p) ` succ(nq)))" -lemma [TC]: "x \ nat ==> incr_var(x,nq) \ nat" +lemma [TC]: "x \ nat \ incr_var(x,nq) \ nat" by (simp add: incr_var_def) -lemma incr_bv_type [TC]: "p \ formula ==> incr_bv(p) \ nat -> formula" +lemma incr_bv_type [TC]: "p \ formula \ incr_bv(p) \ nat -> formula" by (induct_tac p, simp_all) text\Obviously, \<^term>\DPow\ is closed under complements and finite @@ -287,8 +287,8 @@ parameters to be combined.\ lemma sats_incr_bv_iff [rule_format]: - "[| p \ formula; env \ list(A); x \ A |] - ==> \bvs \ list(A). + "\p \ formula; env \ list(A); x \ A\ + \ \bvs \ list(A). sats(A, incr_bv(p) ` length(bvs), bvs @ Cons(x,env)) \ sats(A, p, bvs@env)" apply (induct_tac p) @@ -299,8 +299,8 @@ (*the following two lemmas prevent huge case splits in arity_incr_bv_lemma*) lemma incr_var_lemma: - "[| x \ nat; y \ nat; nq \ x |] - ==> succ(x) \ incr_var(y,nq) = succ(x \ y)" + "\x \ nat; y \ nat; nq \ x\ + \ succ(x) \ incr_var(y,nq) = succ(x \ y)" apply (simp add: incr_var_def Ord_Un_if, auto) apply (blast intro: leI) apply (simp add: not_lt_iff_le) @@ -309,14 +309,14 @@ done lemma incr_And_lemma: - "y < x ==> y \ succ(x) = succ(x \ y)" + "y < x \ y \ succ(x) = succ(x \ y)" apply (simp add: Ord_Un_if lt_Ord lt_Ord2 succ_lt_iff) apply (blast dest: lt_asym) done lemma arity_incr_bv_lemma [rule_format]: "p \ formula - ==> \n \ nat. arity (incr_bv(p) ` n) = + \ \n \ nat. arity (incr_bv(p) ` n) = (if n < arity(p) then succ(arity(p)) else arity(p))" apply (induct_tac p) apply (simp_all add: imp_disj not_lt_iff_le Un_least_lt_iff lt_Un_iff le_Un_iff @@ -340,24 +340,24 @@ definition incr_bv1 :: "i => i" where - "incr_bv1(p) == incr_bv(p)`1" + "incr_bv1(p) \ incr_bv(p)`1" -lemma incr_bv1_type [TC]: "p \ formula ==> incr_bv1(p) \ formula" +lemma incr_bv1_type [TC]: "p \ formula \ incr_bv1(p) \ formula" by (simp add: incr_bv1_def) (*For renaming all but the bound variable at level 0*) lemma sats_incr_bv1_iff: - "[| p \ formula; env \ list(A); x \ A; y \ A |] - ==> sats(A, incr_bv1(p), Cons(x, Cons(y, env))) \ + "\p \ formula; env \ list(A); x \ A; y \ A\ + \ sats(A, incr_bv1(p), Cons(x, Cons(y, env))) \ sats(A, p, Cons(x,env))" apply (insert sats_incr_bv_iff [of p env A y "Cons(x,Nil)"]) apply (simp add: incr_bv1_def) done lemma formula_add_params1 [rule_format]: - "[| p \ formula; n \ nat; x \ A |] - ==> \bvs \ list(A). \env \ list(A). + "\p \ formula; n \ nat; x \ A\ + \ \bvs \ list(A). \env \ list(A). length(bvs) = n \ sats(A, iterates(incr_bv1, n, p), Cons(x, bvs@env)) \ sats(A, p, Cons(x,env))" @@ -369,15 +369,15 @@ lemma arity_incr_bv1_eq: "p \ formula - ==> arity(incr_bv1(p)) = + \ arity(incr_bv1(p)) = (if 1 < arity(p) then succ(arity(p)) else arity(p))" apply (insert arity_incr_bv_lemma [of p 1]) apply (simp add: incr_bv1_def) done lemma arity_iterates_incr_bv1_eq: - "[| p \ formula; n \ nat |] - ==> arity(incr_bv1^n(p)) = + "\p \ formula; n \ nat\ + \ arity(incr_bv1^n(p)) = (if 1 < arity(p) then n #+ arity(p) else arity(p))" apply (induct_tac n) apply (simp_all add: arity_incr_bv1_eq) @@ -392,26 +392,26 @@ text\The definable powerset operation: Kunen's definition VI 1.1, page 165.\ definition DPow :: "i => i" where - "DPow(A) == {X \ Pow(A). + "DPow(A) \ {X \ Pow(A). \env \ list(A). \p \ formula. arity(p) \ succ(length(env)) & X = {x\A. sats(A, p, Cons(x,env))}}" lemma DPowI: - "[|env \ list(A); p \ formula; arity(p) \ succ(length(env))|] - ==> {x\A. sats(A, p, Cons(x,env))} \ DPow(A)" + "\env \ list(A); p \ formula; arity(p) \ succ(length(env))\ + \ {x\A. sats(A, p, Cons(x,env))} \ DPow(A)" by (simp add: DPow_def, blast) text\With this rule we can specify \<^term>\p\ later.\ lemma DPowI2 [rule_format]: - "[|\x\A. P(x) \ sats(A, p, Cons(x,env)); - env \ list(A); p \ formula; arity(p) \ succ(length(env))|] - ==> {x\A. P(x)} \ DPow(A)" + "\\x\A. P(x) \ sats(A, p, Cons(x,env)); + env \ list(A); p \ formula; arity(p) \ succ(length(env))\ + \ {x\A. P(x)} \ DPow(A)" by (simp add: DPow_def, blast) lemma DPowD: "X \ DPow(A) - ==> X \ A & + \ X \ A & (\env \ list(A). \p \ formula. arity(p) \ succ(length(env)) & X = {x\A. sats(A, p, Cons(x,env))})" @@ -420,8 +420,8 @@ lemmas DPow_imp_subset = DPowD [THEN conjunct1] (*Kunen's Lemma VI 1.2*) -lemma "[| p \ formula; env \ list(A); arity(p) \ succ(length(env)) |] - ==> {x\A. sats(A, p, Cons(x,env))} \ DPow(A)" +lemma "\p \ formula; env \ list(A); arity(p) \ succ(length(env))\ + \ {x\A. sats(A, p, Cons(x,env))} \ DPow(A)" by (blast intro: DPowI) lemma DPow_subset_Pow: "DPow(A) \ Pow(A)" @@ -434,14 +434,14 @@ apply (auto simp add: Un_least_lt_iff) done -lemma Compl_in_DPow: "X \ DPow(A) ==> (A-X) \ DPow(A)" +lemma Compl_in_DPow: "X \ DPow(A) \ (A-X) \ DPow(A)" apply (simp add: DPow_def, clarify, auto) apply (rule bexI) apply (rule_tac x="Neg(p)" in bexI) apply auto done -lemma Int_in_DPow: "[| X \ DPow(A); Y \ DPow(A) |] ==> X \ Y \ DPow(A)" +lemma Int_in_DPow: "\X \ DPow(A); Y \ DPow(A)\ \ X \ Y \ DPow(A)" apply (simp add: DPow_def, auto) apply (rename_tac envp p envq q) apply (rule_tac x="envp@envq" in bexI) @@ -454,13 +454,13 @@ apply (simp add: arity_sats1_iff formula_add_params1, blast) done -lemma Un_in_DPow: "[| X \ DPow(A); Y \ DPow(A) |] ==> X \ Y \ DPow(A)" +lemma Un_in_DPow: "\X \ DPow(A); Y \ DPow(A)\ \ X \ Y \ DPow(A)" apply (subgoal_tac "X \ Y = A - ((A-X) \ (A-Y))") apply (simp add: Int_in_DPow Compl_in_DPow) apply (simp add: DPow_def, blast) done -lemma singleton_in_DPow: "a \ A ==> {a} \ DPow(A)" +lemma singleton_in_DPow: "a \ A \ {a} \ DPow(A)" apply (simp add: DPow_def) apply (rule_tac x="Cons(a,Nil)" in bexI) apply (rule_tac x="Equal(0,1)" in bexI) @@ -468,13 +468,13 @@ apply (force simp add: succ_Un_distrib [symmetric]) done -lemma cons_in_DPow: "[| a \ A; X \ DPow(A) |] ==> cons(a,X) \ DPow(A)" +lemma cons_in_DPow: "\a \ A; X \ DPow(A)\ \ cons(a,X) \ DPow(A)" apply (rule cons_eq [THEN subst]) apply (blast intro: singleton_in_DPow Un_in_DPow) done (*Part of Lemma 1.3*) -lemma Fin_into_DPow: "X \ Fin(A) ==> X \ DPow(A)" +lemma Fin_into_DPow: "X \ Fin(A) \ X \ DPow(A)" apply (erule Fin.induct) apply (rule empty_in_DPow) apply (blast intro: cons_in_DPow) @@ -491,10 +491,10 @@ oops *) -lemma Finite_Pow_subset_Pow: "Finite(A) ==> Pow(A) \ DPow(A)" +lemma Finite_Pow_subset_Pow: "Finite(A) \ Pow(A) \ DPow(A)" by (blast intro: Fin_into_DPow Finite_into_Fin Fin_subset) -lemma Finite_DPow_eq_Pow: "Finite(A) ==> DPow(A) = Pow(A)" +lemma Finite_DPow_eq_Pow: "Finite(A) \ DPow(A) = Pow(A)" apply (rule equalityI) apply (rule DPow_subset_Pow) apply (erule Finite_Pow_subset_Pow) @@ -515,18 +515,18 @@ definition subset_fm :: "[i,i]=>i" where - "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))" + "subset_fm(x,y) \ Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))" -lemma subset_type [TC]: "[| x \ nat; y \ nat |] ==> subset_fm(x,y) \ formula" +lemma subset_type [TC]: "\x \ nat; y \ nat\ \ subset_fm(x,y) \ formula" by (simp add: subset_fm_def) lemma arity_subset_fm [simp]: - "[| x \ nat; y \ nat |] ==> arity(subset_fm(x,y)) = succ(x) \ succ(y)" + "\x \ nat; y \ nat\ \ arity(subset_fm(x,y)) = succ(x) \ succ(y)" by (simp add: subset_fm_def succ_Un_distrib [symmetric]) lemma sats_subset_fm [simp]: - "[|x < length(env); y \ nat; env \ list(A); Transset(A)|] - ==> sats(A, subset_fm(x,y), env) \ nth(x,env) \ nth(y,env)" + "\x < length(env); y \ nat; env \ list(A); Transset(A)\ + \ sats(A, subset_fm(x,y), env) \ nth(x,env) \ nth(y,env)" apply (frule lt_length_in_nat, assumption) apply (simp add: subset_fm_def Transset_def) apply (blast intro: nth_type) @@ -536,18 +536,18 @@ definition transset_fm :: "i=>i" where - "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))" + "transset_fm(x) \ Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))" -lemma transset_type [TC]: "x \ nat ==> transset_fm(x) \ formula" +lemma transset_type [TC]: "x \ nat \ transset_fm(x) \ formula" by (simp add: transset_fm_def) lemma arity_transset_fm [simp]: - "x \ nat ==> arity(transset_fm(x)) = succ(x)" + "x \ nat \ arity(transset_fm(x)) = succ(x)" by (simp add: transset_fm_def succ_Un_distrib [symmetric]) lemma sats_transset_fm [simp]: - "[|x < length(env); env \ list(A); Transset(A)|] - ==> sats(A, transset_fm(x), env) \ Transset(nth(x,env))" + "\x < length(env); env \ list(A); Transset(A)\ + \ sats(A, transset_fm(x), env) \ Transset(nth(x,env))" apply (frule lt_nat_in_nat, erule length_type) apply (simp add: transset_fm_def Transset_def) apply (blast intro: nth_type) @@ -557,19 +557,19 @@ definition ordinal_fm :: "i=>i" where - "ordinal_fm(x) == + "ordinal_fm(x) \ And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))" -lemma ordinal_type [TC]: "x \ nat ==> ordinal_fm(x) \ formula" +lemma ordinal_type [TC]: "x \ nat \ ordinal_fm(x) \ formula" by (simp add: ordinal_fm_def) lemma arity_ordinal_fm [simp]: - "x \ nat ==> arity(ordinal_fm(x)) = succ(x)" + "x \ nat \ arity(ordinal_fm(x)) = succ(x)" by (simp add: ordinal_fm_def succ_Un_distrib [symmetric]) lemma sats_ordinal_fm: - "[|x < length(env); env \ list(A); Transset(A)|] - ==> sats(A, ordinal_fm(x), env) \ Ord(nth(x,env))" + "\x < length(env); env \ list(A); Transset(A)\ + \ sats(A, ordinal_fm(x), env) \ Ord(nth(x,env))" apply (frule lt_nat_in_nat, erule length_type) apply (simp add: ordinal_fm_def Ord_def Transset_def) apply (blast intro: nth_type) @@ -577,7 +577,7 @@ text\The subset consisting of the ordinals is definable. Essential lemma for \Ord_in_Lset\. This result is the objective of the present subsection.\ -theorem Ords_in_DPow: "Transset(A) ==> {x \ A. Ord(x)} \ DPow(A)" +theorem Ords_in_DPow: "Transset(A) \ {x \ A. Ord(x)} \ DPow(A)" apply (simp add: DPow_def Collect_subset) apply (rule_tac x=Nil in bexI) apply (rule_tac x="ordinal_fm(0)" in bexI) @@ -589,40 +589,40 @@ definition Lset :: "i=>i" where - "Lset(i) == transrec(i, %x f. \y\x. DPow(f`y))" + "Lset(i) \ transrec(i, %x f. \y\x. DPow(f`y))" definition L :: "i=>o" where \ \Kunen's definition VI 1.5, page 167\ - "L(x) == \i. Ord(i) & x \ Lset(i)" + "L(x) \ \i. Ord(i) & x \ Lset(i)" text\NOT SUITABLE FOR REWRITING -- RECURSIVE!\ lemma Lset: "Lset(i) = (\j\i. DPow(Lset(j)))" by (subst Lset_def [THEN def_transrec], simp) -lemma LsetI: "[|y\x; A \ DPow(Lset(y))|] ==> A \ Lset(x)" +lemma LsetI: "\y\x; A \ DPow(Lset(y))\ \ A \ Lset(x)" by (subst Lset, blast) -lemma LsetD: "A \ Lset(x) ==> \y\x. A \ DPow(Lset(y))" +lemma LsetD: "A \ Lset(x) \ \y\x. A \ DPow(Lset(y))" apply (insert Lset [of x]) apply (blast intro: elim: equalityE) done subsubsection\Transitivity\ -lemma elem_subset_in_DPow: "[|X \ A; X \ A|] ==> X \ DPow(A)" +lemma elem_subset_in_DPow: "\X \ A; X \ A\ \ X \ DPow(A)" apply (simp add: Transset_def DPow_def) apply (rule_tac x="[X]" in bexI) apply (rule_tac x="Member(0,1)" in bexI) apply (auto simp add: Un_least_lt_iff) done -lemma Transset_subset_DPow: "Transset(A) ==> A \ DPow(A)" +lemma Transset_subset_DPow: "Transset(A) \ A \ DPow(A)" apply clarify apply (simp add: Transset_def) apply (blast intro: elem_subset_in_DPow) done -lemma Transset_DPow: "Transset(A) ==> Transset(DPow(A))" +lemma Transset_DPow: "Transset(A) \ Transset(DPow(A))" apply (simp add: Transset_def) apply (blast intro: elem_subset_in_DPow dest: DPowD) done @@ -634,7 +634,7 @@ apply (blast intro!: Transset_Union_family Transset_Un Transset_DPow) done -lemma mem_Lset_imp_subset_Lset: "a \ Lset(i) ==> a \ Lset(i)" +lemma mem_Lset_imp_subset_Lset: "a \ Lset(i) \ a \ Lset(i)" apply (insert Transset_Lset) apply (simp add: Transset_def) done @@ -666,7 +666,7 @@ text\Useful with Reflection to bump up the ordinal\ -lemma subset_Lset_ltD: "[|A \ Lset(i); i < j|] ==> A \ Lset(j)" +lemma subset_Lset_ltD: "\A \ Lset(i); i < j\ \ A \ Lset(j)" by (blast dest: ltD [THEN Lset_mono_mem]) subsubsection\0, successor and limit equations for Lset\ @@ -707,16 +707,16 @@ subsubsection\Lset applied to Limit ordinals\ lemma Limit_Lset_eq: - "Limit(i) ==> Lset(i) = (\y\i. Lset(y))" + "Limit(i) \ Lset(i) = (\y\i. Lset(y))" by (simp add: Lset_Union [symmetric] Limit_Union_eq) -lemma lt_LsetI: "[| a \ Lset(j); j a \ Lset(i)" +lemma lt_LsetI: "\a \ Lset(j); j \ a \ Lset(i)" by (blast dest: Lset_mono [OF le_imp_subset [OF leI]]) lemma Limit_LsetE: - "[| a \ Lset(i); ~R ==> Limit(i); - !!x. [| x Lset(x) |] ==> R - |] ==> R" + "\a \ Lset(i); \R \ Limit(i); + \x. \x Lset(x)\ \ R +\ \ R" apply (rule classical) apply (rule Limit_Lset_eq [THEN equalityD1, THEN subsetD, THEN UN_E]) prefer 2 apply assumption @@ -726,7 +726,7 @@ subsubsection\Basic closure properties\ -lemma zero_in_Lset: "y \ x ==> 0 \ Lset(x)" +lemma zero_in_Lset: "y \ x \ 0 \ Lset(x)" by (subst Lset, blast intro: empty_in_DPow) lemma notin_Lset: "x \ Lset(x)" @@ -738,7 +738,7 @@ subsection\Constructible Ordinals: Kunen's VI 1.9 (b)\ -lemma Ords_of_Lset_eq: "Ord(i) ==> {x\Lset(i). Ord(x)} = i" +lemma Ords_of_Lset_eq: "Ord(i) \ {x\Lset(i). Ord(x)} = i" apply (erule trans_induct3) apply (simp_all add: Lset_succ Limit_Lset_eq Limit_Union_eq) txt\The successor case remains.\ @@ -758,22 +758,22 @@ done -lemma Ord_subset_Lset: "Ord(i) ==> i \ Lset(i)" +lemma Ord_subset_Lset: "Ord(i) \ i \ Lset(i)" by (subst Ords_of_Lset_eq [symmetric], assumption, fast) -lemma Ord_in_Lset: "Ord(i) ==> i \ Lset(succ(i))" +lemma Ord_in_Lset: "Ord(i) \ i \ Lset(succ(i))" apply (simp add: Lset_succ) apply (subst Ords_of_Lset_eq [symmetric], assumption, rule Ords_in_DPow [OF Transset_Lset]) done -lemma Ord_in_L: "Ord(i) ==> L(i)" +lemma Ord_in_L: "Ord(i) \ L(i)" by (simp add: L_def, blast intro: Ord_in_Lset) subsubsection\Unions\ lemma Union_in_Lset: - "X \ Lset(i) ==> \(X) \ Lset(succ(i))" + "X \ Lset(i) \ \(X) \ Lset(succ(i))" apply (insert Transset_Lset) apply (rule LsetI [OF succI1]) apply (simp add: Transset_def DPow_def) @@ -785,20 +785,20 @@ apply (simp add: succ_Un_distrib [symmetric], blast) done -theorem Union_in_L: "L(X) ==> L(\(X))" +theorem Union_in_L: "L(X) \ L(\(X))" by (simp add: L_def, blast dest: Union_in_Lset) subsubsection\Finite sets and ordered pairs\ -lemma singleton_in_Lset: "a \ Lset(i) ==> {a} \ Lset(succ(i))" +lemma singleton_in_Lset: "a \ Lset(i) \ {a} \ Lset(succ(i))" by (simp add: Lset_succ singleton_in_DPow) lemma doubleton_in_Lset: - "[| a \ Lset(i); b \ Lset(i) |] ==> {a,b} \ Lset(succ(i))" + "\a \ Lset(i); b \ Lset(i)\ \ {a,b} \ Lset(succ(i))" by (simp add: Lset_succ empty_in_DPow cons_in_DPow) lemma Pair_in_Lset: - "[| a \ Lset(i); b \ Lset(i); Ord(i) |] ==> \ Lset(succ(succ(i)))" + "\a \ Lset(i); b \ Lset(i); Ord(i)\ \ \ Lset(succ(succ(i)))" apply (unfold Pair_def) apply (blast intro: doubleton_in_Lset) done @@ -808,21 +808,21 @@ text\Hard work is finding a single \<^term>\j \ i\ such that \<^term>\{a,b} \ Lset(j)\\ lemma doubleton_in_LLimit: - "[| a \ Lset(i); b \ Lset(i); Limit(i) |] ==> {a,b} \ Lset(i)" + "\a \ Lset(i); b \ Lset(i); Limit(i)\ \ {a,b} \ Lset(i)" apply (erule Limit_LsetE, assumption) apply (erule Limit_LsetE, assumption) apply (blast intro: lt_LsetI [OF doubleton_in_Lset] Lset_UnI1 Lset_UnI2 Limit_has_succ Un_least_lt) done -theorem doubleton_in_L: "[| L(a); L(b) |] ==> L({a, b})" +theorem doubleton_in_L: "\L(a); L(b)\ \ L({a, b})" apply (simp add: L_def, clarify) apply (drule Ord2_imp_greater_Limit, assumption) apply (blast intro: lt_LsetI doubleton_in_LLimit Limit_is_Ord) done lemma Pair_in_LLimit: - "[| a \ Lset(i); b \ Lset(i); Limit(i) |] ==> \ Lset(i)" + "\a \ Lset(i); b \ Lset(i); Limit(i)\ \ \ Lset(i)" txt\Infer that a, b occur at ordinals x,xa < i.\ apply (erule Limit_LsetE, assumption) apply (erule Limit_LsetE, assumption) @@ -836,18 +836,18 @@ text\The rank function for the constructible universe\ definition lrank :: "i=>i" where \ \Kunen's definition VI 1.7\ - "lrank(x) == \ i. x \ Lset(succ(i))" + "lrank(x) \ \ i. x \ Lset(succ(i))" -lemma L_I: "[|x \ Lset(i); Ord(i)|] ==> L(x)" +lemma L_I: "\x \ Lset(i); Ord(i)\ \ L(x)" by (simp add: L_def, blast) -lemma L_D: "L(x) ==> \i. Ord(i) & x \ Lset(i)" +lemma L_D: "L(x) \ \i. Ord(i) & x \ Lset(i)" by (simp add: L_def) lemma Ord_lrank [simp]: "Ord(lrank(a))" by (simp add: lrank_def) -lemma Lset_lrank_lt [rule_format]: "Ord(i) ==> x \ Lset(i) \ lrank(x) < i" +lemma Lset_lrank_lt [rule_format]: "Ord(i) \ x \ Lset(i) \ lrank(x) < i" apply (erule trans_induct3) apply simp apply (simp only: lrank_def) @@ -859,7 +859,7 @@ text\Kunen's VI 1.8. The proof is much harder than the text would suggest. For a start, it needs the previous lemma, which is proved by induction.\ -lemma Lset_iff_lrank_lt: "Ord(i) ==> x \ Lset(i) \ L(x) & lrank(x) < i" +lemma Lset_iff_lrank_lt: "Ord(i) \ x \ Lset(i) \ L(x) & lrank(x) < i" apply (simp add: L_def, auto) apply (blast intro: Lset_lrank_lt) apply (unfold lrank_def) @@ -872,7 +872,7 @@ by (simp add: Lset_iff_lrank_lt) text\Kunen's VI 1.9 (a)\ -lemma lrank_of_Ord: "Ord(i) ==> lrank(i) = i" +lemma lrank_of_Ord: "Ord(i) \ lrank(i) = i" apply (unfold lrank_def) apply (rule Least_equality) apply (erule Ord_in_Lset) @@ -893,7 +893,7 @@ apply auto done -lemma lrank_Lset: "Ord(i) ==> lrank(Lset(i)) = i" +lemma lrank_Lset: "Ord(i) \ lrank(Lset(i)) = i" apply (unfold lrank_def) apply (rule Least_equality) apply (rule Lset_in_Lset_succ) @@ -905,7 +905,7 @@ done text\Kunen's VI 1.11\ -lemma Lset_subset_Vset: "Ord(i) ==> Lset(i) \ Vset(i)" +lemma Lset_subset_Vset: "Ord(i) \ Lset(i) \ Vset(i)" apply (erule trans_induct) apply (subst Lset) apply (subst Vset) @@ -915,7 +915,7 @@ done text\Kunen's VI 1.12\ -lemma Lset_subset_Vset': "i \ nat ==> Lset(i) = Vset(i)" +lemma Lset_subset_Vset': "i \ nat \ Lset(i) = Vset(i)" apply (erule nat_induct) apply (simp add: Vfrom_0) apply (simp add: Lset_succ Vset_succ Finite_Vset Finite_DPow_eq_Pow) @@ -923,24 +923,24 @@ text\Every set of constructible sets is included in some \<^term>\Lset\\ lemma subset_Lset: - "(\x\A. L(x)) ==> \i. Ord(i) & A \ Lset(i)" + "(\x\A. L(x)) \ \i. Ord(i) & A \ Lset(i)" by (rule_tac x = "\x\A. succ(lrank(x))" in exI, force) lemma subset_LsetE: - "[|\x\A. L(x); - !!i. [|Ord(i); A \ Lset(i)|] ==> P|] - ==> P" + "\\x\A. L(x); + \i. \Ord(i); A \ Lset(i)\ \ P\ + \ P" by (blast dest: subset_Lset) subsubsection\For L to satisfy the Powerset axiom\ lemma LPow_env_typing: - "[| y \ Lset(i); Ord(i); y \ X |] - ==> \z \ Pow(X). y \ Lset(succ(lrank(z)))" + "\y \ Lset(i); Ord(i); y \ X\ + \ \z \ Pow(X). y \ Lset(succ(lrank(z)))" by (auto intro: L_I iff: Lset_succ_lrank_iff) lemma LPow_in_Lset: - "[|X \ Lset(i); Ord(i)|] ==> \j. Ord(j) & {y \ Pow(X). L(y)} \ Lset(j)" + "\X \ Lset(i); Ord(i)\ \ \j. Ord(j) & {y \ Pow(X). L(y)} \ Lset(j)" apply (rule_tac x="succ(\y \ Pow(X). succ(lrank(y)))" in exI) apply simp apply (rule LsetI [OF succI1]) @@ -958,34 +958,34 @@ apply (auto intro: L_I iff: Lset_succ_lrank_iff) done -theorem LPow_in_L: "L(X) ==> L({y \ Pow(X). L(y)})" +theorem LPow_in_L: "L(X) \ L({y \ Pow(X). L(y)})" by (blast intro: L_I dest: L_D LPow_in_Lset) subsection\Eliminating \<^term>\arity\ from the Definition of \<^term>\Lset\\ -lemma nth_zero_eq_0: "n \ nat ==> nth(n,[0]) = 0" +lemma nth_zero_eq_0: "n \ nat \ nth(n,[0]) = 0" by (induct_tac n, auto) lemma sats_app_0_iff [rule_format]: - "[| p \ formula; 0 \ A |] - ==> \env \ list(A). sats(A,p, env@[0]) \ sats(A,p,env)" + "\p \ formula; 0 \ A\ + \ \env \ list(A). sats(A,p, env@[0]) \ sats(A,p,env)" apply (induct_tac p) apply (simp_all del: app_Cons add: app_Cons [symmetric] add: nth_zero_eq_0 nth_append not_lt_iff_le nth_eq_0) done lemma sats_app_zeroes_iff: - "[| p \ formula; 0 \ A; env \ list(A); n \ nat |] - ==> sats(A,p,env @ repeat(0,n)) \ sats(A,p,env)" + "\p \ formula; 0 \ A; env \ list(A); n \ nat\ + \ sats(A,p,env @ repeat(0,n)) \ sats(A,p,env)" apply (induct_tac n, simp) apply (simp del: repeat.simps add: repeat_succ_app sats_app_0_iff app_assoc [symmetric]) done lemma exists_bigger_env: - "[| p \ formula; 0 \ A; env \ list(A) |] - ==> \env' \ list(A). arity(p) \ succ(length(env')) & + "\p \ formula; 0 \ A; env \ list(A)\ + \ \env' \ list(A). arity(p) \ succ(length(env')) & (\a\A. sats(A,p,Cons(a,env')) \ sats(A,p,Cons(a,env)))" apply (rule_tac x="env @ repeat(0,arity(p))" in bexI) apply (simp del: app_Cons add: app_Cons [symmetric] @@ -996,7 +996,7 @@ text\A simpler version of \<^term>\DPow\: no arity check!\ definition DPow' :: "i => i" where - "DPow'(A) == {X \ Pow(A). + "DPow'(A) \ {X \ Pow(A). \env \ list(A). \p \ formula. X = {x\A. sats(A, p, Cons(x,env))}}" @@ -1006,12 +1006,12 @@ lemma DPow'_0: "DPow'(0) = {0}" by (auto simp add: DPow'_def) -lemma DPow'_subset_DPow: "0 \ A ==> DPow'(A) \ DPow(A)" +lemma DPow'_subset_DPow: "0 \ A \ DPow'(A) \ DPow(A)" apply (auto simp add: DPow'_def DPow_def) apply (frule exists_bigger_env, assumption+, force) done -lemma DPow_eq_DPow': "Transset(A) ==> DPow(A) = DPow'(A)" +lemma DPow_eq_DPow': "Transset(A) \ DPow(A) = DPow'(A)" apply (drule Transset_0_disj) apply (erule disjE) apply (simp add: DPow'_0 Finite_DPow_eq_Pow) @@ -1032,9 +1032,9 @@ text\With this rule we can specify \<^term>\p\ later and don't worry about arities at all!\ lemma DPow_LsetI [rule_format]: - "[|\x\Lset(i). P(x) \ sats(Lset(i), p, Cons(x,env)); - env \ list(Lset(i)); p \ formula|] - ==> {x\Lset(i). P(x)} \ DPow(Lset(i))" + "\\x\Lset(i). P(x) \ sats(Lset(i), p, Cons(x,env)); + env \ list(Lset(i)); p \ formula\ + \ {x\Lset(i). P(x)} \ DPow(Lset(i))" by (simp add: DPow_eq_DPow' [OF Transset_Lset] DPow'_def, blast) end diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Internalize.thy --- a/src/ZF/Constructible/Internalize.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Internalize.thy Tue Sep 27 16:51:35 2022 +0100 @@ -8,24 +8,24 @@ subsubsection\The Formula \<^term>\is_Inl\, Internalized\ -(* is_Inl(M,a,z) == \zero[M]. empty(M,zero) & pair(M,zero,a,z) *) +(* is_Inl(M,a,z) \ \zero[M]. empty(M,zero) & pair(M,zero,a,z) *) definition Inl_fm :: "[i,i]=>i" where - "Inl_fm(a,z) == Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))" + "Inl_fm(a,z) \ Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))" lemma Inl_type [TC]: - "[| x \ nat; z \ nat |] ==> Inl_fm(x,z) \ formula" + "\x \ nat; z \ nat\ \ Inl_fm(x,z) \ formula" by (simp add: Inl_fm_def) lemma sats_Inl_fm [simp]: - "[| x \ nat; z \ nat; env \ list(A)|] - ==> sats(A, Inl_fm(x,z), env) \ is_Inl(##A, nth(x,env), nth(z,env))" + "\x \ nat; z \ nat; env \ list(A)\ + \ sats(A, Inl_fm(x,z), env) \ is_Inl(##A, nth(x,env), nth(z,env))" by (simp add: Inl_fm_def is_Inl_def) lemma Inl_iff_sats: - "[| nth(i,env) = x; nth(k,env) = z; - i \ nat; k \ nat; env \ list(A)|] - ==> is_Inl(##A, x, z) \ sats(A, Inl_fm(i,k), env)" + "\nth(i,env) = x; nth(k,env) = z; + i \ nat; k \ nat; env \ list(A)\ + \ is_Inl(##A, x, z) \ sats(A, Inl_fm(i,k), env)" by simp theorem Inl_reflection: @@ -38,24 +38,24 @@ subsubsection\The Formula \<^term>\is_Inr\, Internalized\ -(* is_Inr(M,a,z) == \n1[M]. number1(M,n1) & pair(M,n1,a,z) *) +(* is_Inr(M,a,z) \ \n1[M]. number1(M,n1) & pair(M,n1,a,z) *) definition Inr_fm :: "[i,i]=>i" where - "Inr_fm(a,z) == Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))" + "Inr_fm(a,z) \ Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))" lemma Inr_type [TC]: - "[| x \ nat; z \ nat |] ==> Inr_fm(x,z) \ formula" + "\x \ nat; z \ nat\ \ Inr_fm(x,z) \ formula" by (simp add: Inr_fm_def) lemma sats_Inr_fm [simp]: - "[| x \ nat; z \ nat; env \ list(A)|] - ==> sats(A, Inr_fm(x,z), env) \ is_Inr(##A, nth(x,env), nth(z,env))" + "\x \ nat; z \ nat; env \ list(A)\ + \ sats(A, Inr_fm(x,z), env) \ is_Inr(##A, nth(x,env), nth(z,env))" by (simp add: Inr_fm_def is_Inr_def) lemma Inr_iff_sats: - "[| nth(i,env) = x; nth(k,env) = z; - i \ nat; k \ nat; env \ list(A)|] - ==> is_Inr(##A, x, z) \ sats(A, Inr_fm(i,k), env)" + "\nth(i,env) = x; nth(k,env) = z; + i \ nat; k \ nat; env \ list(A)\ + \ is_Inr(##A, x, z) \ sats(A, Inr_fm(i,k), env)" by simp theorem Inr_reflection: @@ -68,23 +68,23 @@ subsubsection\The Formula \<^term>\is_Nil\, Internalized\ -(* is_Nil(M,xs) == \zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *) +(* is_Nil(M,xs) \ \zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *) definition Nil_fm :: "i=>i" where - "Nil_fm(x) == Exists(And(empty_fm(0), Inl_fm(0,succ(x))))" + "Nil_fm(x) \ Exists(And(empty_fm(0), Inl_fm(0,succ(x))))" -lemma Nil_type [TC]: "x \ nat ==> Nil_fm(x) \ formula" +lemma Nil_type [TC]: "x \ nat \ Nil_fm(x) \ formula" by (simp add: Nil_fm_def) lemma sats_Nil_fm [simp]: - "[| x \ nat; env \ list(A)|] - ==> sats(A, Nil_fm(x), env) \ is_Nil(##A, nth(x,env))" + "\x \ nat; env \ list(A)\ + \ sats(A, Nil_fm(x), env) \ is_Nil(##A, nth(x,env))" by (simp add: Nil_fm_def is_Nil_def) lemma Nil_iff_sats: - "[| nth(i,env) = x; i \ nat; env \ list(A)|] - ==> is_Nil(##A, x) \ sats(A, Nil_fm(i), env)" + "\nth(i,env) = x; i \ nat; env \ list(A)\ + \ is_Nil(##A, x) \ sats(A, Nil_fm(i), env)" by simp theorem Nil_reflection: @@ -98,26 +98,26 @@ subsubsection\The Formula \<^term>\is_Cons\, Internalized\ -(* "is_Cons(M,a,l,Z) == \p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *) +(* "is_Cons(M,a,l,Z) \ \p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *) definition Cons_fm :: "[i,i,i]=>i" where - "Cons_fm(a,l,Z) == + "Cons_fm(a,l,Z) \ Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))" lemma Cons_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> Cons_fm(x,y,z) \ formula" + "\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) \ + "\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)" + "\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: @@ -129,24 +129,24 @@ subsubsection\The Formula \<^term>\is_quasilist\, Internalized\ -(* is_quasilist(M,xs) == is_Nil(M,z) | (\x[M]. \l[M]. is_Cons(M,x,l,z))" *) +(* is_quasilist(M,xs) \ is_Nil(M,z) | (\x[M]. \l[M]. is_Cons(M,x,l,z))" *) definition quasilist_fm :: "i=>i" where - "quasilist_fm(x) == + "quasilist_fm(x) \ Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))" -lemma quasilist_type [TC]: "x \ nat ==> quasilist_fm(x) \ formula" +lemma quasilist_type [TC]: "x \ nat \ quasilist_fm(x) \ formula" by (simp add: quasilist_fm_def) lemma sats_quasilist_fm [simp]: - "[| x \ nat; env \ list(A)|] - ==> sats(A, quasilist_fm(x), env) \ is_quasilist(##A, nth(x,env))" + "\x \ nat; env \ list(A)\ + \ sats(A, quasilist_fm(x), env) \ is_quasilist(##A, nth(x,env))" by (simp add: quasilist_fm_def is_quasilist_def) lemma quasilist_iff_sats: - "[| nth(i,env) = x; i \ nat; env \ list(A)|] - ==> is_quasilist(##A, x) \ sats(A, quasilist_fm(i), env)" + "\nth(i,env) = x; i \ nat; env \ list(A)\ + \ is_quasilist(##A, x) \ sats(A, quasilist_fm(i), env)" by simp theorem quasilist_reflection: @@ -162,30 +162,30 @@ subsubsection\The Formula \<^term>\is_hd\, Internalized\ -(* "is_hd(M,xs,H) == +(* "is_hd(M,xs,H) \ (is_Nil(M,xs) \ empty(M,H)) & - (\x[M]. \l[M]. ~ is_Cons(M,x,l,xs) | H=x) & + (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | H=x) & (is_quasilist(M,xs) | empty(M,H))" *) definition hd_fm :: "[i,i]=>i" where - "hd_fm(xs,H) == + "hd_fm(xs,H) \ And(Implies(Nil_fm(xs), empty_fm(H)), And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(H#+2,1)))), Or(quasilist_fm(xs), empty_fm(H))))" lemma hd_type [TC]: - "[| x \ nat; y \ nat |] ==> hd_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ hd_fm(x,y) \ formula" by (simp add: hd_fm_def) lemma sats_hd_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, hd_fm(x,y), env) \ is_hd(##A, nth(x,env), nth(y,env))" + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, hd_fm(x,y), env) \ is_hd(##A, nth(x,env), nth(y,env))" by (simp add: hd_fm_def is_hd_def) lemma hd_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_hd(##A, x, y) \ sats(A, hd_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_hd(##A, x, y) \ sats(A, hd_fm(i,j), env)" by simp theorem hd_reflection: @@ -199,30 +199,30 @@ subsubsection\The Formula \<^term>\is_tl\, Internalized\ -(* "is_tl(M,xs,T) == +(* "is_tl(M,xs,T) \ (is_Nil(M,xs) \ T=xs) & - (\x[M]. \l[M]. ~ is_Cons(M,x,l,xs) | T=l) & + (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | T=l) & (is_quasilist(M,xs) | empty(M,T))" *) definition tl_fm :: "[i,i]=>i" where - "tl_fm(xs,T) == + "tl_fm(xs,T) \ And(Implies(Nil_fm(xs), Equal(T,xs)), And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))), Or(quasilist_fm(xs), empty_fm(T))))" lemma tl_type [TC]: - "[| x \ nat; y \ nat |] ==> tl_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ tl_fm(x,y) \ formula" by (simp add: tl_fm_def) lemma sats_tl_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, tl_fm(x,y), env) \ is_tl(##A, nth(x,env), nth(y,env))" + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, tl_fm(x,y), env) \ is_tl(##A, nth(x,env), nth(y,env))" by (simp add: tl_fm_def is_tl_def) lemma tl_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_tl(##A, x, y) \ sats(A, tl_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_tl(##A, x, y) \ sats(A, tl_fm(i,j), env)" by simp theorem tl_reflection: @@ -237,34 +237,34 @@ subsubsection\The Operator \<^term>\is_bool_of_o\\ (* is_bool_of_o :: "[i=>o, o, i] => o" - "is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))" *) + "is_bool_of_o(M,P,z) \ (P & number1(M,z)) | (\P & empty(M,z))" *) text\The formula \<^term>\p\ has no free variables.\ definition bool_of_o_fm :: "[i, i]=>i" where - "bool_of_o_fm(p,z) == + "bool_of_o_fm(p,z) \ Or(And(p,number1_fm(z)), And(Neg(p),empty_fm(z)))" lemma is_bool_of_o_type [TC]: - "[| p \ formula; z \ nat |] ==> bool_of_o_fm(p,z) \ formula" + "\p \ formula; z \ nat\ \ bool_of_o_fm(p,z) \ formula" by (simp add: bool_of_o_fm_def) lemma sats_bool_of_o_fm: assumes p_iff_sats: "P \ sats(A, p, env)" shows - "[|z \ nat; env \ list(A)|] - ==> sats(A, bool_of_o_fm(p,z), env) \ + "\z \ nat; env \ list(A)\ + \ sats(A, bool_of_o_fm(p,z), env) \ is_bool_of_o(##A, P, nth(z,env))" by (simp add: bool_of_o_fm_def is_bool_of_o_def p_iff_sats [THEN iff_sym]) lemma is_bool_of_o_iff_sats: - "[| P \ sats(A, p, env); nth(k,env) = z; k \ nat; env \ list(A)|] - ==> is_bool_of_o(##A, P, z) \ sats(A, bool_of_o_fm(p,k), env)" + "\P \ sats(A, p, env); nth(k,env) = z; k \ nat; env \ list(A)\ + \ is_bool_of_o(##A, P, z) \ sats(A, bool_of_o_fm(p,k), env)" by (simp add: sats_bool_of_o_fm) theorem bool_of_o_reflection: - "REFLECTS [P(L), \i. P(##Lset(i))] ==> + "REFLECTS [P(L), \i. P(##Lset(i))] \ REFLECTS[\x. is_bool_of_o(L, P(L,x), f(x)), \i x. is_bool_of_o(##Lset(i), P(##Lset(i),x), f(x))]" apply (simp (no_asm) only: is_bool_of_o_def) @@ -280,12 +280,12 @@ \<^term>\p\ will be enclosed by three quantifiers.\ (* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" - "is_lambda(M, A, is_b, z) == + "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))" *) definition lambda_fm :: "[i, i, i]=>i" where - "lambda_fm(p,A,z) == + "lambda_fm(p,A,z) \ Forall(Iff(Member(0,succ(z)), Exists(Exists(And(Member(1,A#+3), And(pair_fm(1,0,2), p))))))" @@ -294,24 +294,24 @@ the corresponding quantified variables with de Bruijn indices 1, 0.\ lemma is_lambda_type [TC]: - "[| p \ formula; x \ nat; y \ nat |] - ==> lambda_fm(p,x,y) \ formula" + "\p \ formula; x \ nat; y \ nat\ + \ lambda_fm(p,x,y) \ formula" by (simp add: lambda_fm_def) lemma sats_lambda_fm: assumes is_b_iff_sats: - "!!a0 a1 a2. - [|a0\A; a1\A; a2\A|] - ==> is_b(a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,env))))" + "\a0 a1 a2. + \a0\A; a1\A; a2\A\ + \ is_b(a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,env))))" shows - "[|x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, lambda_fm(p,x,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, lambda_fm(p,x,y), env) \ is_lambda(##A, nth(x,env), is_b, nth(y,env))" by (simp add: lambda_fm_def is_lambda_def is_b_iff_sats [THEN iff_sym]) theorem is_lambda_reflection: assumes is_b_reflection: - "!!f g h. REFLECTS[\x. is_b(L, f(x), g(x), h(x)), + "\f g h. REFLECTS[\x. is_b(L, f(x), g(x), h(x)), \i x. is_b(##Lset(i), f(x), g(x), h(x))]" shows "REFLECTS[\x. is_lambda(L, A(x), is_b(L,x), f(x)), \i x. is_lambda(##Lset(i), A(x), is_b(##Lset(i),x), f(x))]" @@ -321,28 +321,28 @@ subsubsection\The Operator \<^term>\is_Member\, Internalized\ -(* "is_Member(M,x,y,Z) == +(* "is_Member(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *) definition Member_fm :: "[i,i,i]=>i" where - "Member_fm(x,y,Z) == + "Member_fm(x,y,Z) \ Exists(Exists(And(pair_fm(x#+2,y#+2,1), And(Inl_fm(1,0), Inl_fm(0,Z#+2)))))" lemma is_Member_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> Member_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ Member_fm(x,y,z) \ formula" by (simp add: Member_fm_def) lemma sats_Member_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, Member_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, Member_fm(x,y,z), env) \ is_Member(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: Member_fm_def is_Member_def) lemma Member_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_Member(##A, x, y, z) \ sats(A, Member_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_Member(##A, x, y, z) \ sats(A, Member_fm(i,j,k), env)" by (simp) theorem Member_reflection: @@ -354,28 +354,28 @@ subsubsection\The Operator \<^term>\is_Equal\, Internalized\ -(* "is_Equal(M,x,y,Z) == +(* "is_Equal(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *) definition Equal_fm :: "[i,i,i]=>i" where - "Equal_fm(x,y,Z) == + "Equal_fm(x,y,Z) \ Exists(Exists(And(pair_fm(x#+2,y#+2,1), And(Inr_fm(1,0), Inl_fm(0,Z#+2)))))" lemma is_Equal_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> Equal_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ Equal_fm(x,y,z) \ formula" by (simp add: Equal_fm_def) lemma sats_Equal_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, Equal_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, Equal_fm(x,y,z), env) \ is_Equal(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: Equal_fm_def is_Equal_def) lemma Equal_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_Equal(##A, x, y, z) \ sats(A, Equal_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_Equal(##A, x, y, z) \ sats(A, Equal_fm(i,j,k), env)" by (simp) theorem Equal_reflection: @@ -387,28 +387,28 @@ subsubsection\The Operator \<^term>\is_Nand\, Internalized\ -(* "is_Nand(M,x,y,Z) == +(* "is_Nand(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *) definition Nand_fm :: "[i,i,i]=>i" where - "Nand_fm(x,y,Z) == + "Nand_fm(x,y,Z) \ Exists(Exists(And(pair_fm(x#+2,y#+2,1), And(Inl_fm(1,0), Inr_fm(0,Z#+2)))))" lemma is_Nand_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> Nand_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ Nand_fm(x,y,z) \ formula" by (simp add: Nand_fm_def) lemma sats_Nand_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, Nand_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, Nand_fm(x,y,z), env) \ is_Nand(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: Nand_fm_def is_Nand_def) lemma Nand_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_Nand(##A, x, y, z) \ sats(A, Nand_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_Nand(##A, x, y, z) \ sats(A, Nand_fm(i,j,k), env)" by (simp) theorem Nand_reflection: @@ -420,26 +420,26 @@ subsubsection\The Operator \<^term>\is_Forall\, Internalized\ -(* "is_Forall(M,p,Z) == \u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *) +(* "is_Forall(M,p,Z) \ \u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *) definition Forall_fm :: "[i,i]=>i" where - "Forall_fm(x,Z) == + "Forall_fm(x,Z) \ Exists(And(Inr_fm(succ(x),0), Inr_fm(0,succ(Z))))" lemma is_Forall_type [TC]: - "[| x \ nat; y \ nat |] ==> Forall_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ Forall_fm(x,y) \ formula" by (simp add: Forall_fm_def) lemma sats_Forall_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, Forall_fm(x,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, Forall_fm(x,y), env) \ is_Forall(##A, nth(x,env), nth(y,env))" by (simp add: Forall_fm_def is_Forall_def) lemma Forall_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_Forall(##A, x, y) \ sats(A, Forall_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_Forall(##A, x, y) \ sats(A, Forall_fm(i,j), env)" by (simp) theorem Forall_reflection: @@ -452,28 +452,28 @@ subsubsection\The Operator \<^term>\is_and\, Internalized\ -(* is_and(M,a,b,z) == (number1(M,a) & z=b) | - (~number1(M,a) & empty(M,z)) *) +(* is_and(M,a,b,z) \ (number1(M,a) & z=b) | + (\number1(M,a) & empty(M,z)) *) definition and_fm :: "[i,i,i]=>i" where - "and_fm(a,b,z) == + "and_fm(a,b,z) \ Or(And(number1_fm(a), Equal(z,b)), And(Neg(number1_fm(a)),empty_fm(z)))" lemma is_and_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> and_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ and_fm(x,y,z) \ formula" by (simp add: and_fm_def) lemma sats_and_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, and_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, and_fm(x,y,z), env) \ is_and(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: and_fm_def is_and_def) lemma is_and_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_and(##A, x, y, z) \ sats(A, and_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_and(##A, x, y, z) \ sats(A, and_fm(i,j,k), env)" by simp theorem is_and_reflection: @@ -486,29 +486,29 @@ subsubsection\The Operator \<^term>\is_or\, Internalized\ -(* is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) | - (~number1(M,a) & z=b) *) +(* is_or(M,a,b,z) \ (number1(M,a) & number1(M,z)) | + (\number1(M,a) & z=b) *) definition or_fm :: "[i,i,i]=>i" where - "or_fm(a,b,z) == + "or_fm(a,b,z) \ Or(And(number1_fm(a), number1_fm(z)), And(Neg(number1_fm(a)), Equal(z,b)))" lemma is_or_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> or_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ or_fm(x,y,z) \ formula" by (simp add: or_fm_def) lemma sats_or_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, or_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, or_fm(x,y,z), env) \ is_or(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: or_fm_def is_or_def) lemma is_or_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_or(##A, x, y, z) \ sats(A, or_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_or(##A, x, y, z) \ sats(A, or_fm(i,j,k), env)" by simp theorem is_or_reflection: @@ -522,27 +522,27 @@ subsubsection\The Operator \<^term>\is_not\, Internalized\ -(* is_not(M,a,z) == (number1(M,a) & empty(M,z)) | - (~number1(M,a) & number1(M,z)) *) +(* is_not(M,a,z) \ (number1(M,a) & empty(M,z)) | + (\number1(M,a) & number1(M,z)) *) definition not_fm :: "[i,i]=>i" where - "not_fm(a,z) == + "not_fm(a,z) \ Or(And(number1_fm(a), empty_fm(z)), And(Neg(number1_fm(a)), number1_fm(z)))" lemma is_not_type [TC]: - "[| x \ nat; z \ nat |] ==> not_fm(x,z) \ formula" + "\x \ nat; z \ nat\ \ not_fm(x,z) \ formula" by (simp add: not_fm_def) lemma sats_is_not_fm [simp]: - "[| x \ nat; z \ nat; env \ list(A)|] - ==> sats(A, not_fm(x,z), env) \ is_not(##A, nth(x,env), nth(z,env))" + "\x \ nat; z \ nat; env \ list(A)\ + \ sats(A, not_fm(x,z), env) \ is_not(##A, nth(x,env), nth(z,env))" by (simp add: not_fm_def is_not_def) lemma is_not_iff_sats: - "[| nth(i,env) = x; nth(k,env) = z; - i \ nat; k \ nat; env \ list(A)|] - ==> is_not(##A, x, z) \ sats(A, not_fm(i,k), env)" + "\nth(i,env) = x; nth(k,env) = z; + i \ nat; k \ nat; env \ list(A)\ + \ is_not(##A, x, z) \ sats(A, not_fm(i,k), env)" by simp theorem is_not_reflection: @@ -579,7 +579,7 @@ (* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" - "M_is_recfun(M,MH,r,a,f) == + "M_is_recfun(M,MH,r,a,f) \ \z[M]. z \ f \ 2 1 0 new def (\x[M]. \f_r_sx[M]. \y[M]. @@ -593,7 +593,7 @@ text\The three arguments of \<^term>\p\ are always 2, 1, 0 and z\ definition is_recfun_fm :: "[i, i, i, i]=>i" where - "is_recfun_fm(p,r,a,f) == + "is_recfun_fm(p,r,a,f) \ Forall(Iff(Member(0,succ(f)), Exists(Exists(Exists( And(p, @@ -605,31 +605,31 @@ And(restriction_fm(f#+7,0,4), Member(2,r#+7)))))))))))))))" lemma is_recfun_type [TC]: - "[| p \ formula; x \ nat; y \ nat; z \ nat |] - ==> is_recfun_fm(p,x,y,z) \ formula" + "\p \ formula; x \ nat; y \ nat; z \ nat\ + \ is_recfun_fm(p,x,y,z) \ formula" by (simp add: is_recfun_fm_def) lemma sats_is_recfun_fm: assumes MH_iff_sats: - "!!a0 a1 a2 a3. - [|a0\A; a1\A; a2\A; a3\A|] - ==> MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))" + "\a0 a1 a2 a3. + \a0\A; a1\A; a2\A; a3\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))" shows - "[|x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, is_recfun_fm(p,x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, is_recfun_fm(p,x,y,z), env) \ M_is_recfun(##A, MH, nth(x,env), nth(y,env), nth(z,env))" by (simp add: is_recfun_fm_def M_is_recfun_iff MH_iff_sats [THEN iff_sym]) lemma is_recfun_iff_sats: assumes MH_iff_sats: - "!!a0 a1 a2 a3. - [|a0\A; a1\A; a2\A; a3\A|] - ==> MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))" + "\a0 a1 a2 a3. + \a0\A; a1\A; a2\A; a3\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,env)))))" shows - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> M_is_recfun(##A, MH, x, y, z) \ sats(A, is_recfun_fm(p,i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ M_is_recfun(##A, MH, x, y, z) \ sats(A, is_recfun_fm(p,i,j,k), env)" by (simp add: sats_is_recfun_fm [OF MH_iff_sats]) text\The additional variable in the premise, namely \<^term>\f'\, is essential. @@ -637,7 +637,7 @@ The same thing occurs in \is_wfrec_reflection\.\ theorem is_recfun_reflection: assumes MH_reflection: - "!!f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), + "\f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), \i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]" shows "REFLECTS[\x. M_is_recfun(L, MH(L,x), f(x), g(x), h(x)), \i x. M_is_recfun(##Lset(i), MH(##Lset(i),x), f(x), g(x), h(x))]" @@ -652,11 +652,11 @@ \<^term>\p\ is enclosed by 5 quantifiers.\ (* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o" - "is_wfrec(M,MH,r,a,z) == + "is_wfrec(M,MH,r,a,z) \ \f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" *) definition is_wfrec_fm :: "[i, i, i, i]=>i" where - "is_wfrec_fm(p,r,a,z) == + "is_wfrec_fm(p,r,a,z) \ Exists(And(is_recfun_fm(p, succ(r), succ(a), 0), Exists(Exists(Exists(Exists( And(Equal(2,a#+5), And(Equal(1,4), And(Equal(0,z#+5), p)))))))))" @@ -668,18 +668,18 @@ environments in both calls to MH have the same length.\ lemma is_wfrec_type [TC]: - "[| p \ formula; x \ nat; y \ nat; z \ nat |] - ==> is_wfrec_fm(p,x,y,z) \ formula" + "\p \ formula; x \ nat; y \ nat; z \ nat\ + \ is_wfrec_fm(p,x,y,z) \ formula" by (simp add: is_wfrec_fm_def) lemma sats_is_wfrec_fm: assumes MH_iff_sats: - "!!a0 a1 a2 a3 a4. - [|a0\A; a1\A; a2\A; a3\A; a4\A|] - ==> MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" + "\a0 a1 a2 a3 a4. + \a0\A; a1\A; a2\A; a3\A; a4\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" shows - "[|x \ nat; y < length(env); z < length(env); env \ list(A)|] - ==> sats(A, is_wfrec_fm(p,x,y,z), env) \ + "\x \ nat; y < length(env); z < length(env); env \ list(A)\ + \ sats(A, is_wfrec_fm(p,x,y,z), env) \ is_wfrec(##A, MH, nth(x,env), nth(y,env), nth(z,env))" apply (frule_tac x=z in lt_length_in_nat, assumption) apply (frule lt_length_in_nat, assumption) @@ -689,18 +689,18 @@ lemma is_wfrec_iff_sats: assumes MH_iff_sats: - "!!a0 a1 a2 a3 a4. - [|a0\A; a1\A; a2\A; a3\A; a4\A|] - ==> MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" + "\a0 a1 a2 a3 a4. + \a0\A; a1\A; a2\A; a3\A; a4\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3,Cons(a4,env))))))" shows - "[|nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j < length(env); k < length(env); env \ list(A)|] - ==> is_wfrec(##A, MH, x, y, z) \ sats(A, is_wfrec_fm(p,i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j < length(env); k < length(env); env \ list(A)\ + \ is_wfrec(##A, MH, x, y, z) \ sats(A, is_wfrec_fm(p,i,j,k), env)" by (simp add: sats_is_wfrec_fm [OF MH_iff_sats]) theorem is_wfrec_reflection: assumes MH_reflection: - "!!f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), + "\f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), \i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]" shows "REFLECTS[\x. is_wfrec(L, MH(L,x), f(x), g(x), h(x)), \i x. is_wfrec(##Lset(i), MH(##Lset(i),x), f(x), g(x), h(x))]" @@ -715,28 +715,28 @@ definition cartprod_fm :: "[i,i,i]=>i" where -(* "cartprod(M,A,B,z) == +(* "cartprod(M,A,B,z) \ \u[M]. u \ z \ (\x[M]. x\A & (\y[M]. y\B & pair(M,x,y,u)))" *) - "cartprod_fm(A,B,z) == + "cartprod_fm(A,B,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), Exists(And(Member(0,succ(succ(succ(B)))), pair_fm(1,0,2)))))))" lemma cartprod_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> cartprod_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ cartprod_fm(x,y,z) \ formula" by (simp add: cartprod_fm_def) lemma sats_cartprod_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, cartprod_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, cartprod_fm(x,y,z), env) \ cartprod(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: cartprod_fm_def cartprod_def) lemma cartprod_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> cartprod(##A, x, y, z) \ sats(A, cartprod_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ cartprod(##A, x, y, z) \ sats(A, cartprod_fm(i,j,k), env)" by (simp) theorem cartprod_reflection: @@ -749,14 +749,14 @@ subsubsection\Binary Sums, Internalized\ -(* "is_sum(M,A,B,Z) == +(* "is_sum(M,A,B,Z) \ \A0[M]. \n1[M]. \s1[M]. \B1[M]. 3 2 1 0 number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) & cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" *) definition sum_fm :: "[i,i,i]=>i" where - "sum_fm(A,B,Z) == + "sum_fm(A,B,Z) \ Exists(Exists(Exists(Exists( And(number1_fm(2), And(cartprod_fm(2,A#+4,3), @@ -764,19 +764,19 @@ And(cartprod_fm(1,B#+4,0), union_fm(3,0,Z#+4)))))))))" lemma sum_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> sum_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ sum_fm(x,y,z) \ formula" by (simp add: sum_fm_def) lemma sats_sum_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, sum_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, sum_fm(x,y,z), env) \ is_sum(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: sum_fm_def is_sum_def) lemma sum_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_sum(##A, x, y, z) \ sats(A, sum_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_sum(##A, x, y, z) \ sats(A, sum_fm(i,j,k), env)" by simp theorem sum_reflection: @@ -789,24 +789,24 @@ subsubsection\The Operator \<^term>\quasinat\\ -(* "is_quasinat(M,z) == empty(M,z) | (\m[M]. successor(M,m,z))" *) +(* "is_quasinat(M,z) \ empty(M,z) | (\m[M]. successor(M,m,z))" *) definition quasinat_fm :: "i=>i" where - "quasinat_fm(z) == Or(empty_fm(z), Exists(succ_fm(0,succ(z))))" + "quasinat_fm(z) \ Or(empty_fm(z), Exists(succ_fm(0,succ(z))))" lemma quasinat_type [TC]: - "x \ nat ==> quasinat_fm(x) \ formula" + "x \ nat \ quasinat_fm(x) \ formula" by (simp add: quasinat_fm_def) lemma sats_quasinat_fm [simp]: - "[| x \ nat; env \ list(A)|] - ==> sats(A, quasinat_fm(x), env) \ is_quasinat(##A, nth(x,env))" + "\x \ nat; env \ list(A)\ + \ sats(A, quasinat_fm(x), env) \ is_quasinat(##A, nth(x,env))" by (simp add: quasinat_fm_def is_quasinat_def) lemma quasinat_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; env \ list(A)|] - ==> is_quasinat(##A, x) \ sats(A, quasinat_fm(i), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; env \ list(A)\ + \ is_quasinat(##A, x) \ sats(A, quasinat_fm(i), env)" by simp theorem quasinat_reflection: @@ -823,43 +823,43 @@ stand for \<^term>\m\ and \<^term>\b\, respectively.\ (* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" - "is_nat_case(M, a, is_b, k, z) == + "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))" *) text\The formula \<^term>\is_b\ has free variables 1 and 0.\ definition is_nat_case_fm :: "[i, i, i, i]=>i" where - "is_nat_case_fm(a,is_b,k,z) == + "is_nat_case_fm(a,is_b,k,z) \ And(Implies(empty_fm(k), Equal(z,a)), And(Forall(Implies(succ_fm(0,succ(k)), Forall(Implies(Equal(0,succ(succ(z))), is_b)))), Or(quasinat_fm(k), empty_fm(z))))" lemma is_nat_case_type [TC]: - "[| is_b \ formula; - x \ nat; y \ nat; z \ nat |] - ==> is_nat_case_fm(x,is_b,y,z) \ formula" + "\is_b \ formula; + x \ nat; y \ nat; z \ nat\ + \ is_nat_case_fm(x,is_b,y,z) \ formula" by (simp add: is_nat_case_fm_def) lemma sats_is_nat_case_fm: assumes is_b_iff_sats: - "!!a. a \ A ==> is_b(a,nth(z, env)) \ + "\a. a \ A \ is_b(a,nth(z, env)) \ sats(A, p, Cons(nth(z,env), Cons(a, env)))" shows - "[|x \ nat; y \ nat; z < length(env); env \ list(A)|] - ==> sats(A, is_nat_case_fm(x,p,y,z), env) \ + "\x \ nat; y \ nat; z < length(env); env \ list(A)\ + \ sats(A, is_nat_case_fm(x,p,y,z), env) \ is_nat_case(##A, nth(x,env), is_b, nth(y,env), nth(z,env))" apply (frule lt_length_in_nat, assumption) apply (simp add: is_nat_case_fm_def is_nat_case_def is_b_iff_sats [THEN iff_sym]) done lemma is_nat_case_iff_sats: - "[| (!!a. a \ A ==> is_b(a,z) \ + "\(\a. a \ A \ is_b(a,z) \ sats(A, p, Cons(z, Cons(a,env)))); nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k < length(env); env \ list(A)|] - ==> is_nat_case(##A, x, is_b, y, z) \ sats(A, is_nat_case_fm(i,p,j,k), env)" + i \ nat; j \ nat; k < length(env); env \ list(A)\ + \ is_nat_case(##A, x, is_b, y, z) \ sats(A, is_nat_case_fm(i,p,j,k), env)" by (simp add: sats_is_nat_case_fm [of A is_b]) @@ -868,7 +868,7 @@ argument, we cannot prove reflection for \<^term>\iterates_MH\.\ theorem is_nat_case_reflection: assumes is_b_reflection: - "!!h f g. REFLECTS[\x. is_b(L, h(x), f(x), g(x)), + "\h f g. REFLECTS[\x. is_b(L, h(x), f(x), g(x)), \i x. is_b(##Lset(i), h(x), f(x), g(x))]" shows "REFLECTS[\x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)), \i x. is_nat_case(##Lset(i), f(x), is_b(##Lset(i), x), g(x), h(x))]" @@ -881,31 +881,31 @@ subsection\The Operator \<^term>\iterates_MH\, Needed for Iteration\ (* iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" - "iterates_MH(M,isF,v,n,g,z) == + "iterates_MH(M,isF,v,n,g,z) \ is_nat_case(M, v, \m u. \gm[M]. fun_apply(M,g,m,gm) & isF(gm,u), n, z)" *) definition iterates_MH_fm :: "[i, i, i, i, i]=>i" where - "iterates_MH_fm(isF,v,n,g,z) == + "iterates_MH_fm(isF,v,n,g,z) \ is_nat_case_fm(v, Exists(And(fun_apply_fm(succ(succ(succ(g))),2,0), Forall(Implies(Equal(0,2), isF)))), n, z)" lemma iterates_MH_type [TC]: - "[| p \ formula; - v \ nat; x \ nat; y \ nat; z \ nat |] - ==> iterates_MH_fm(p,v,x,y,z) \ formula" + "\p \ formula; + v \ nat; x \ nat; y \ nat; z \ nat\ + \ iterates_MH_fm(p,v,x,y,z) \ formula" by (simp add: iterates_MH_fm_def) lemma sats_iterates_MH_fm: assumes is_F_iff_sats: - "!!a b c d. [| a \ A; b \ A; c \ A; d \ A|] - ==> is_F(a,b) \ + "\a b c d. \a \ A; b \ A; c \ A; d \ A\ + \ is_F(a,b) \ sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d,env)))))" shows - "[|v \ nat; x \ nat; y \ nat; z < length(env); env \ list(A)|] - ==> sats(A, iterates_MH_fm(p,v,x,y,z), env) \ + "\v \ nat; x \ nat; y \ nat; z < length(env); env \ list(A)\ + \ sats(A, iterates_MH_fm(p,v,x,y,z), env) \ iterates_MH(##A, is_F, nth(v,env), nth(x,env), nth(y,env), nth(z,env))" apply (frule lt_length_in_nat, assumption) apply (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm @@ -916,13 +916,13 @@ lemma iterates_MH_iff_sats: assumes is_F_iff_sats: - "!!a b c d. [| a \ A; b \ A; c \ A; d \ A|] - ==> is_F(a,b) \ + "\a b c d. \a \ A; b \ A; c \ A; d \ A\ + \ is_F(a,b) \ sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d,env)))))" shows - "[| nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i' \ nat; i \ nat; j \ nat; k < length(env); env \ list(A)|] - ==> iterates_MH(##A, is_F, v, x, y, z) \ + "\nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i' \ nat; i \ nat; j \ nat; k < length(env); env \ list(A)\ + \ iterates_MH(##A, is_F, v, x, y, z) \ sats(A, iterates_MH_fm(p,i',i,j,k), env)" by (simp add: sats_iterates_MH_fm [OF is_F_iff_sats]) @@ -931,7 +931,7 @@ argument, we cannot prove reflection for \<^term>\list_N\.\ theorem iterates_MH_reflection: assumes p_reflection: - "!!f g h. REFLECTS[\x. p(L, h(x), f(x), g(x)), + "\f g h. REFLECTS[\x. p(L, h(x), f(x), g(x)), \i x. p(##Lset(i), h(x), f(x), g(x))]" shows "REFLECTS[\x. iterates_MH(L, p(L,x), e(x), f(x), g(x), h(x)), \i x. iterates_MH(##Lset(i), p(##Lset(i),x), e(x), f(x), g(x), h(x))]" @@ -946,13 +946,13 @@ text\The three arguments of \<^term>\p\ are always 2, 1, 0; \<^term>\p\ is enclosed by 9 (??) quantifiers.\ -(* "is_iterates(M,isF,v,n,Z) == +(* "is_iterates(M,isF,v,n,Z) \ \sn[M]. \msn[M]. successor(M,n,sn) & membership(M,sn,msn) & 1 0 is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"*) definition is_iterates_fm :: "[i, i, i, i]=>i" where - "is_iterates_fm(p,v,n,Z) == + "is_iterates_fm(p,v,n,Z) \ Exists(Exists( And(succ_fm(n#+2,1), And(Memrel_fm(1,0), @@ -964,21 +964,21 @@ lemma is_iterates_type [TC]: - "[| p \ formula; x \ nat; y \ nat; z \ nat |] - ==> is_iterates_fm(p,x,y,z) \ formula" + "\p \ formula; x \ nat; y \ nat; z \ nat\ + \ is_iterates_fm(p,x,y,z) \ formula" by (simp add: is_iterates_fm_def) lemma sats_is_iterates_fm: assumes is_F_iff_sats: - "!!a b c d e f g h i j k. - [| a \ A; b \ A; c \ A; d \ A; e \ A; f \ A; - g \ A; h \ A; i \ A; j \ A; k \ A|] - ==> is_F(a,b) \ + "\a b c d e f g h i j k. + \a \ A; b \ A; c \ A; d \ A; e \ A; f \ A; + g \ A; h \ A; i \ A; j \ A; k \ A\ + \ is_F(a,b) \ sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d, Cons(e, Cons(f, Cons(g, Cons(h, Cons(i, Cons(j, Cons(k, env))))))))))))" shows - "[|x \ nat; y < length(env); z < length(env); env \ list(A)|] - ==> sats(A, is_iterates_fm(p,x,y,z), env) \ + "\x \ nat; y < length(env); z < length(env); env \ list(A)\ + \ sats(A, is_iterates_fm(p,x,y,z), env) \ is_iterates(##A, is_F, nth(x,env), nth(y,env), nth(z,env))" apply (frule_tac x=z in lt_length_in_nat, assumption) apply (frule lt_length_in_nat, assumption) @@ -989,16 +989,16 @@ lemma is_iterates_iff_sats: assumes is_F_iff_sats: - "!!a b c d e f g h i j k. - [| a \ A; b \ A; c \ A; d \ A; e \ A; f \ A; - g \ A; h \ A; i \ A; j \ A; k \ A|] - ==> is_F(a,b) \ + "\a b c d e f g h i j k. + \a \ A; b \ A; c \ A; d \ A; e \ A; f \ A; + g \ A; h \ A; i \ A; j \ A; k \ A\ + \ is_F(a,b) \ sats(A, p, Cons(b, Cons(a, Cons(c, Cons(d, Cons(e, Cons(f, Cons(g, Cons(h, Cons(i, Cons(j, Cons(k, env))))))))))))" shows - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j < length(env); k < length(env); env \ list(A)|] - ==> is_iterates(##A, is_F, x, y, z) \ + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j < length(env); k < length(env); env \ list(A)\ + \ is_iterates(##A, is_F, x, y, z) \ sats(A, is_iterates_fm(p,i,j,k), env)" by (simp add: sats_is_iterates_fm [OF is_F_iff_sats]) @@ -1007,7 +1007,7 @@ argument, we cannot prove reflection for \<^term>\list_N\.\ theorem is_iterates_reflection: assumes p_reflection: - "!!f g h. REFLECTS[\x. p(L, h(x), f(x), g(x)), + "\f g h. REFLECTS[\x. p(L, h(x), f(x), g(x)), \i x. p(##Lset(i), h(x), f(x), g(x))]" shows "REFLECTS[\x. is_iterates(L, p(L,x), f(x), g(x), h(x)), \i x. is_iterates(##Lset(i), p(##Lset(i),x), f(x), g(x), h(x))]" @@ -1019,19 +1019,19 @@ subsubsection\The Formula \<^term>\is_eclose_n\, Internalized\ -(* is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z) *) +(* is_eclose_n(M,A,n,Z) \ is_iterates(M, big_union(M), A, n, Z) *) definition eclose_n_fm :: "[i,i,i]=>i" where - "eclose_n_fm(A,n,Z) == is_iterates_fm(big_union_fm(1,0), A, n, Z)" + "eclose_n_fm(A,n,Z) \ is_iterates_fm(big_union_fm(1,0), A, n, Z)" lemma eclose_n_fm_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> eclose_n_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ eclose_n_fm(x,y,z) \ formula" by (simp add: eclose_n_fm_def) lemma sats_eclose_n_fm [simp]: - "[| x \ nat; y < length(env); z < length(env); env \ list(A)|] - ==> sats(A, eclose_n_fm(x,y,z), env) \ + "\x \ nat; y < length(env); z < length(env); env \ list(A)\ + \ sats(A, eclose_n_fm(x,y,z), env) \ is_eclose_n(##A, nth(x,env), nth(y,env), nth(z,env))" apply (frule_tac x=z in lt_length_in_nat, assumption) apply (frule_tac x=y in lt_length_in_nat, assumption) @@ -1040,9 +1040,9 @@ done lemma eclose_n_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j < length(env); k < length(env); env \ list(A)|] - ==> is_eclose_n(##A, x, y, z) \ sats(A, eclose_n_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j < length(env); k < length(env); env \ list(A)\ + \ is_eclose_n(##A, x, y, z) \ sats(A, eclose_n_fm(i,j,k), env)" by (simp) theorem eclose_n_reflection: @@ -1055,29 +1055,29 @@ subsubsection\Membership in \<^term>\eclose(A)\\ -(* mem_eclose(M,A,l) == +(* mem_eclose(M,A,l) \ \n[M]. \eclosen[M]. finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \ eclosen *) definition mem_eclose_fm :: "[i,i]=>i" where - "mem_eclose_fm(x,y) == + "mem_eclose_fm(x,y) \ Exists(Exists( And(finite_ordinal_fm(1), And(eclose_n_fm(x#+2,1,0), Member(y#+2,0)))))" lemma mem_eclose_type [TC]: - "[| x \ nat; y \ nat |] ==> mem_eclose_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ mem_eclose_fm(x,y) \ formula" by (simp add: mem_eclose_fm_def) lemma sats_mem_eclose_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, mem_eclose_fm(x,y), env) \ mem_eclose(##A, nth(x,env), nth(y,env))" + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, mem_eclose_fm(x,y), env) \ mem_eclose(##A, nth(x,env), nth(y,env))" by (simp add: mem_eclose_fm_def mem_eclose_def) lemma mem_eclose_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> mem_eclose(##A, x, y) \ sats(A, mem_eclose_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ mem_eclose(##A, x, y) \ sats(A, mem_eclose_fm(i,j), env)" by simp theorem mem_eclose_reflection: @@ -1090,25 +1090,25 @@ subsubsection\The Predicate ``Is \<^term>\eclose(A)\''\ -(* is_eclose(M,A,Z) == \l[M]. l \ Z \ mem_eclose(M,A,l) *) +(* is_eclose(M,A,Z) \ \l[M]. l \ Z \ mem_eclose(M,A,l) *) definition is_eclose_fm :: "[i,i]=>i" where - "is_eclose_fm(A,Z) == + "is_eclose_fm(A,Z) \ Forall(Iff(Member(0,succ(Z)), mem_eclose_fm(succ(A),0)))" lemma is_eclose_type [TC]: - "[| x \ nat; y \ nat |] ==> is_eclose_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ is_eclose_fm(x,y) \ formula" by (simp add: is_eclose_fm_def) lemma sats_is_eclose_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, is_eclose_fm(x,y), env) \ is_eclose(##A, nth(x,env), nth(y,env))" + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, is_eclose_fm(x,y), env) \ is_eclose(##A, nth(x,env), nth(y,env))" by (simp add: is_eclose_fm_def is_eclose_def) lemma is_eclose_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_eclose(##A, x, y) \ sats(A, is_eclose_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_eclose(##A, x, y) \ sats(A, is_eclose_fm(i,j), env)" by simp theorem is_eclose_reflection: @@ -1123,28 +1123,28 @@ definition list_functor_fm :: "[i,i,i]=>i" where -(* "is_list_functor(M,A,X,Z) == +(* "is_list_functor(M,A,X,Z) \ \n1[M]. \AX[M]. number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *) - "list_functor_fm(A,X,Z) == + "list_functor_fm(A,X,Z) \ Exists(Exists( And(number1_fm(1), And(cartprod_fm(A#+2,X#+2,0), sum_fm(1,0,Z#+2)))))" lemma list_functor_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> list_functor_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ list_functor_fm(x,y,z) \ formula" by (simp add: list_functor_fm_def) lemma sats_list_functor_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, list_functor_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, list_functor_fm(x,y,z), env) \ is_list_functor(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: list_functor_fm_def is_list_functor_def) lemma list_functor_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> is_list_functor(##A, x, y, z) \ sats(A, list_functor_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ is_list_functor(##A, x, y, z) \ sats(A, list_functor_fm(i,j,k), env)" by simp theorem list_functor_reflection: @@ -1158,24 +1158,24 @@ subsubsection\The Formula \<^term>\is_list_N\, Internalized\ -(* "is_list_N(M,A,n,Z) == +(* "is_list_N(M,A,n,Z) \ \zero[M]. empty(M,zero) & is_iterates(M, is_list_functor(M,A), zero, n, Z)" *) definition list_N_fm :: "[i,i,i]=>i" where - "list_N_fm(A,n,Z) == + "list_N_fm(A,n,Z) \ Exists( And(empty_fm(0), is_iterates_fm(list_functor_fm(A#+9#+3,1,0), 0, n#+1, Z#+1)))" lemma list_N_fm_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> list_N_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ list_N_fm(x,y,z) \ formula" by (simp add: list_N_fm_def) lemma sats_list_N_fm [simp]: - "[| x \ nat; y < length(env); z < length(env); env \ list(A)|] - ==> sats(A, list_N_fm(x,y,z), env) \ + "\x \ nat; y < length(env); z < length(env); env \ list(A)\ + \ sats(A, list_N_fm(x,y,z), env) \ is_list_N(##A, nth(x,env), nth(y,env), nth(z,env))" apply (frule_tac x=z in lt_length_in_nat, assumption) apply (frule_tac x=y in lt_length_in_nat, assumption) @@ -1183,9 +1183,9 @@ done lemma list_N_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j < length(env); k < length(env); env \ list(A)|] - ==> is_list_N(##A, x, y, z) \ sats(A, list_N_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j < length(env); k < length(env); env \ list(A)\ + \ is_list_N(##A, x, y, z) \ sats(A, list_N_fm(i,j,k), env)" by (simp) theorem list_N_reflection: @@ -1200,29 +1200,29 @@ subsubsection\The Predicate ``Is A List''\ -(* mem_list(M,A,l) == +(* mem_list(M,A,l) \ \n[M]. \listn[M]. finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \ listn *) definition mem_list_fm :: "[i,i]=>i" where - "mem_list_fm(x,y) == + "mem_list_fm(x,y) \ Exists(Exists( And(finite_ordinal_fm(1), And(list_N_fm(x#+2,1,0), Member(y#+2,0)))))" lemma mem_list_type [TC]: - "[| x \ nat; y \ nat |] ==> mem_list_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ mem_list_fm(x,y) \ formula" by (simp add: mem_list_fm_def) lemma sats_mem_list_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, mem_list_fm(x,y), env) \ mem_list(##A, nth(x,env), nth(y,env))" + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, mem_list_fm(x,y), env) \ mem_list(##A, nth(x,env), nth(y,env))" by (simp add: mem_list_fm_def mem_list_def) lemma mem_list_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> mem_list(##A, x, y) \ sats(A, mem_list_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ mem_list(##A, x, y) \ sats(A, mem_list_fm(i,j), env)" by simp theorem mem_list_reflection: @@ -1235,25 +1235,25 @@ subsubsection\The Predicate ``Is \<^term>\list(A)\''\ -(* is_list(M,A,Z) == \l[M]. l \ Z \ mem_list(M,A,l) *) +(* is_list(M,A,Z) \ \l[M]. l \ Z \ mem_list(M,A,l) *) definition is_list_fm :: "[i,i]=>i" where - "is_list_fm(A,Z) == + "is_list_fm(A,Z) \ Forall(Iff(Member(0,succ(Z)), mem_list_fm(succ(A),0)))" lemma is_list_type [TC]: - "[| x \ nat; y \ nat |] ==> is_list_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ is_list_fm(x,y) \ formula" by (simp add: is_list_fm_def) lemma sats_is_list_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, is_list_fm(x,y), env) \ is_list(##A, nth(x,env), nth(y,env))" + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, is_list_fm(x,y), env) \ is_list(##A, nth(x,env), nth(y,env))" by (simp add: is_list_fm_def is_list_def) lemma is_list_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_list(##A, x, y) \ sats(A, is_list_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_list(##A, x, y) \ sats(A, is_list_fm(i,j), env)" by simp theorem is_list_reflection: @@ -1267,14 +1267,14 @@ subsubsection\The Formula Functor, Internalized\ definition formula_functor_fm :: "[i,i]=>i" where -(* "is_formula_functor(M,X,Z) == +(* "is_formula_functor(M,X,Z) \ \nat'[M]. \natnat[M]. \natnatsum[M]. \XX[M]. \X3[M]. 4 3 2 1 0 omega(M,nat') & cartprod(M,nat',nat',natnat) & is_sum(M,natnat,natnat,natnatsum) & cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & is_sum(M,natnatsum,X3,Z)" *) - "formula_functor_fm(X,Z) == + "formula_functor_fm(X,Z) \ Exists(Exists(Exists(Exists(Exists( And(omega_fm(4), And(cartprod_fm(4,4,3), @@ -1283,19 +1283,19 @@ And(sum_fm(1,X#+5,0), sum_fm(2,0,Z#+5)))))))))))" lemma formula_functor_type [TC]: - "[| x \ nat; y \ nat |] ==> formula_functor_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ formula_functor_fm(x,y) \ formula" by (simp add: formula_functor_fm_def) lemma sats_formula_functor_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, formula_functor_fm(x,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, formula_functor_fm(x,y), env) \ is_formula_functor(##A, nth(x,env), nth(y,env))" by (simp add: formula_functor_fm_def is_formula_functor_def) lemma formula_functor_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> is_formula_functor(##A, x, y) \ sats(A, formula_functor_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ is_formula_functor(##A, x, y) \ sats(A, formula_functor_fm(i,j), env)" by simp theorem formula_functor_reflection: @@ -1309,23 +1309,23 @@ subsubsection\The Formula \<^term>\is_formula_N\, Internalized\ -(* "is_formula_N(M,n,Z) == +(* "is_formula_N(M,n,Z) \ \zero[M]. empty(M,zero) & is_iterates(M, is_formula_functor(M), zero, n, Z)" *) definition formula_N_fm :: "[i,i]=>i" where - "formula_N_fm(n,Z) == + "formula_N_fm(n,Z) \ Exists( And(empty_fm(0), is_iterates_fm(formula_functor_fm(1,0), 0, n#+1, Z#+1)))" lemma formula_N_fm_type [TC]: - "[| x \ nat; y \ nat |] ==> formula_N_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ formula_N_fm(x,y) \ formula" by (simp add: formula_N_fm_def) lemma sats_formula_N_fm [simp]: - "[| x < length(env); y < length(env); env \ list(A)|] - ==> sats(A, formula_N_fm(x,y), env) \ + "\x < length(env); y < length(env); env \ list(A)\ + \ sats(A, formula_N_fm(x,y), env) \ is_formula_N(##A, nth(x,env), nth(y,env))" apply (frule_tac x=y in lt_length_in_nat, assumption) apply (frule lt_length_in_nat, assumption) @@ -1333,9 +1333,9 @@ done lemma formula_N_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i < length(env); j < length(env); env \ list(A)|] - ==> is_formula_N(##A, x, y) \ sats(A, formula_N_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i < length(env); j < length(env); env \ list(A)\ + \ is_formula_N(##A, x, y) \ sats(A, formula_N_fm(i,j), env)" by (simp) theorem formula_N_reflection: @@ -1350,28 +1350,28 @@ subsubsection\The Predicate ``Is A Formula''\ -(* mem_formula(M,p) == +(* mem_formula(M,p) \ \n[M]. \formn[M]. finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \ formn *) definition mem_formula_fm :: "i=>i" where - "mem_formula_fm(x) == + "mem_formula_fm(x) \ Exists(Exists( And(finite_ordinal_fm(1), And(formula_N_fm(1,0), Member(x#+2,0)))))" lemma mem_formula_type [TC]: - "x \ nat ==> mem_formula_fm(x) \ formula" + "x \ nat \ mem_formula_fm(x) \ formula" by (simp add: mem_formula_fm_def) lemma sats_mem_formula_fm [simp]: - "[| x \ nat; env \ list(A)|] - ==> sats(A, mem_formula_fm(x), env) \ mem_formula(##A, nth(x,env))" + "\x \ nat; env \ list(A)\ + \ sats(A, mem_formula_fm(x), env) \ mem_formula(##A, nth(x,env))" by (simp add: mem_formula_fm_def mem_formula_def) lemma mem_formula_iff_sats: - "[| nth(i,env) = x; i \ nat; env \ list(A)|] - ==> mem_formula(##A, x) \ sats(A, mem_formula_fm(i), env)" + "\nth(i,env) = x; i \ nat; env \ list(A)\ + \ mem_formula(##A, x) \ sats(A, mem_formula_fm(i), env)" by simp theorem mem_formula_reflection: @@ -1385,23 +1385,23 @@ subsubsection\The Predicate ``Is \<^term>\formula\''\ -(* is_formula(M,Z) == \p[M]. p \ Z \ mem_formula(M,p) *) +(* is_formula(M,Z) \ \p[M]. p \ Z \ mem_formula(M,p) *) definition is_formula_fm :: "i=>i" where - "is_formula_fm(Z) == Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))" + "is_formula_fm(Z) \ Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))" lemma is_formula_type [TC]: - "x \ nat ==> is_formula_fm(x) \ formula" + "x \ nat \ is_formula_fm(x) \ formula" by (simp add: is_formula_fm_def) lemma sats_is_formula_fm [simp]: - "[| x \ nat; env \ list(A)|] - ==> sats(A, is_formula_fm(x), env) \ is_formula(##A, nth(x,env))" + "\x \ nat; env \ list(A)\ + \ sats(A, is_formula_fm(x), env) \ is_formula(##A, nth(x,env))" by (simp add: is_formula_fm_def is_formula_def) lemma is_formula_iff_sats: - "[| nth(i,env) = x; i \ nat; env \ list(A)|] - ==> is_formula(##A, x) \ sats(A, is_formula_fm(i), env)" + "\nth(i,env) = x; i \ nat; env \ list(A)\ + \ is_formula(##A, x) \ sats(A, is_formula_fm(i), env)" by simp theorem is_formula_reflection: @@ -1420,14 +1420,14 @@ the corresponding quantified variables with de Bruijn indices 2, 1, 0.\ (* is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" - "is_transrec(M,MH,a,z) == + "is_transrec(M,MH,a,z) \ \sa[M]. \esa[M]. \mesa[M]. 2 1 0 upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) & is_wfrec(M,MH,mesa,a,z)" *) definition is_transrec_fm :: "[i, i, i]=>i" where - "is_transrec_fm(p,a,z) == + "is_transrec_fm(p,a,z) \ Exists(Exists(Exists( And(upair_fm(a#+3,a#+3,2), And(is_eclose_fm(2,1), @@ -1435,20 +1435,20 @@ lemma is_transrec_type [TC]: - "[| p \ formula; x \ nat; z \ nat |] - ==> is_transrec_fm(p,x,z) \ formula" + "\p \ formula; x \ nat; z \ nat\ + \ is_transrec_fm(p,x,z) \ formula" by (simp add: is_transrec_fm_def) lemma sats_is_transrec_fm: assumes MH_iff_sats: - "!!a0 a1 a2 a3 a4 a5 a6 a7. - [|a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A|] - ==> MH(a2, a1, a0) \ + "\a0 a1 a2 a3 a4 a5 a6 a7. + \a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3, Cons(a4,Cons(a5,Cons(a6,Cons(a7,env)))))))))" shows - "[|x < length(env); z < length(env); env \ list(A)|] - ==> sats(A, is_transrec_fm(p,x,z), env) \ + "\x < length(env); z < length(env); env \ list(A)\ + \ sats(A, is_transrec_fm(p,x,z), env) \ is_transrec(##A, MH, nth(x,env), nth(z,env))" apply (frule_tac x=z in lt_length_in_nat, assumption) apply (frule_tac x=x in lt_length_in_nat, assumption) @@ -1458,20 +1458,20 @@ lemma is_transrec_iff_sats: assumes MH_iff_sats: - "!!a0 a1 a2 a3 a4 a5 a6 a7. - [|a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A|] - ==> MH(a2, a1, a0) \ + "\a0 a1 a2 a3 a4 a5 a6 a7. + \a0\A; a1\A; a2\A; a3\A; a4\A; a5\A; a6\A; a7\A\ + \ MH(a2, a1, a0) \ sats(A, p, Cons(a0,Cons(a1,Cons(a2,Cons(a3, Cons(a4,Cons(a5,Cons(a6,Cons(a7,env)))))))))" shows - "[|nth(i,env) = x; nth(k,env) = z; - i < length(env); k < length(env); env \ list(A)|] - ==> is_transrec(##A, MH, x, z) \ sats(A, is_transrec_fm(p,i,k), env)" + "\nth(i,env) = x; nth(k,env) = z; + i < length(env); k < length(env); env \ list(A)\ + \ is_transrec(##A, MH, x, z) \ sats(A, is_transrec_fm(p,i,k), env)" by (simp add: sats_is_transrec_fm [OF MH_iff_sats]) theorem is_transrec_reflection: assumes MH_reflection: - "!!f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), + "\f' f g h. REFLECTS[\x. MH(L, f'(x), f(x), g(x), h(x)), \i x. MH(##Lset(i), f'(x), f(x), g(x), h(x))]" shows "REFLECTS[\x. is_transrec(L, MH(L,x), f(x), h(x)), \i x. is_transrec(##Lset(i), MH(##Lset(i),x), f(x), h(x))]" diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/L_axioms.thy Tue Sep 27 16:51:35 2022 +0100 @@ -8,7 +8,7 @@ text \The class L satisfies the premises of locale \M_trivial\\ -lemma transL: "[| y\x; L(x) |] ==> L(y)" +lemma transL: "\y\x; L(x)\ \ L(y)" apply (insert Transset_Lset) apply (simp add: Transset_def L_def, blast) done @@ -52,8 +52,8 @@ 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)" + "\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 @@ -64,8 +64,8 @@ done lemma LReplace_in_L: - "[|L(X); univalent(L,X,Q)|] - ==> \Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \ Y" + "\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) @@ -79,8 +79,8 @@ done lemma strong_replacementI [rule_format]: - "[| \B[L]. separation(L, %u. \x[L]. x\B & P(x,u)) |] - ==> strong_replacement(L,P)" + "\\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) @@ -127,22 +127,22 @@ definition L_F0 :: "[i=>o,i] => i" where - "L_F0(P,y) == \ b. (\z. L(z) \ P()) \ (\z\Lset(b). P())" + "L_F0(P,y) \ \ b. (\z. L(z) \ P()) \ (\z\Lset(b). P())" definition L_FF :: "[i=>o,i] => i" where - "L_FF(P) == \a. \y\Lset(a). L_F0(P,y)" + "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" + "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) & + "REFLECTS[P,Q] \ (\Cl. Closed_Unbounded(Cl) & (\a. Cl(a) \ (\x \ Lset(a). P(x) \ Q(a,x))))" @@ -154,15 +154,15 @@ done theorem Not_reflection: - "REFLECTS[P,Q] ==> REFLECTS[\x. ~P(x), \a x. ~Q(a,x)]" + "REFLECTS[P,Q] \ REFLECTS[\x. \P(x), \a x. \Q(a,x)]" apply (unfold 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)]" + "\REFLECTS[P,Q]; REFLECTS[P',Q']\ + \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) @@ -170,8 +170,8 @@ done theorem Or_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] - ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" + "\REFLECTS[P,Q]; REFLECTS[P',Q']\ + \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) @@ -179,8 +179,8 @@ done theorem Imp_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] - ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" + "\REFLECTS[P,Q]; REFLECTS[P',Q']\ + \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) @@ -188,8 +188,8 @@ done theorem Iff_reflection: - "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] - ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" + "\REFLECTS[P,Q]; REFLECTS[P',Q']\ + \ REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) @@ -204,7 +204,7 @@ 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)]" + \ 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) apply (elim meta_exE) apply (rule meta_exI) @@ -213,7 +213,7 @@ 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)]" + \ 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) apply (elim meta_exE) apply (rule meta_exI) @@ -222,14 +222,14 @@ 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)]" + \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" apply (unfold 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)]" + \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" apply (unfold rall_def) apply (intro Imp_reflection All_reflection, assumption) done @@ -238,7 +238,7 @@ 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)]" + \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z[##Lset(a)]. Q(a,x,z)]" apply (unfold setclass_def rex_def) apply (erule Rex_reflection [unfolded rex_def Bex_def]) done @@ -246,7 +246,7 @@ 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)]" + \ REFLECTS[\x. \z[L]. P(x,z), \a x. \z[##Lset(a)]. Q(a,x,z)]" apply (unfold setclass_def rall_def) apply (erule Rall_reflection [unfolded rall_def Ball_def]) done @@ -257,17 +257,17 @@ Rex_reflection Rall_reflection Rex_reflection' Rall_reflection' lemma ReflectsD: - "[|REFLECTS[P,Q]; Ord(i)|] - ==> \j. ix \ Lset(j). P(x) \ Q(j,x))" + "\REFLECTS[P,Q]; Ord(i)\ + \ \j. ix \ Lset(j). P(x) \ Q(j,x))" apply (unfold 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" + "\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" @@ -279,46 +279,46 @@ subsubsection\Some numbers to help write de Bruijn indices\ abbreviation - digit3 :: i (\3\) where "3 == succ(2)" + digit3 :: i (\3\) where "3 \ succ(2)" abbreviation - digit4 :: i (\4\) where "4 == succ(3)" + digit4 :: i (\4\) where "4 \ succ(3)" abbreviation - digit5 :: i (\5\) where "5 == succ(4)" + digit5 :: i (\5\) where "5 \ succ(4)" abbreviation - digit6 :: i (\6\) where "6 == succ(5)" + digit6 :: i (\6\) where "6 \ succ(5)" abbreviation - digit7 :: i (\7\) where "7 == succ(6)" + digit7 :: i (\7\) where "7 \ succ(6)" abbreviation - digit8 :: i (\8\) where "8 == succ(7)" + digit8 :: i (\8\) where "8 \ succ(7)" abbreviation - digit9 :: i (\9\) where "9 == succ(8)" + 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))))" + "empty_fm(x) \ Forall(Neg(Member(0,succ(x))))" lemma empty_type [TC]: - "x \ nat ==> empty_fm(x) \ formula" + "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))" + "\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)" + "\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: @@ -330,8 +330,8 @@ 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" + "\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) @@ -343,32 +343,32 @@ definition upair_fm :: "[i,i,i]=>i" where - "upair_fm(x,y,z) == + "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" + "\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) \ + "\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)" + "\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) \ + "\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) @@ -386,25 +386,25 @@ definition pair_fm :: "[i,i,i]=>i" where - "pair_fm(x,y,z) == + "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" + "\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) \ + "\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)" + "\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: @@ -419,24 +419,24 @@ definition union_fm :: "[i,i,i]=>i" where - "union_fm(x,y,z) == + "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" + "\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) \ + "\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)" + "\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: @@ -451,25 +451,25 @@ definition cons_fm :: "[i,i,i]=>i" where - "cons_fm(x,y,z) == + "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" + "\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) \ + "\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)" + "\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: @@ -484,22 +484,22 @@ definition succ_fm :: "[i,i]=>i" where - "succ_fm(x,y) == cons_fm(x,x,y)" + "succ_fm(x,y) \ cons_fm(x,x,y)" lemma succ_type [TC]: - "[| x \ nat; y \ nat |] ==> succ_fm(x,y) \ formula" + "\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) \ + "\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)" + "\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: @@ -512,24 +512,24 @@ subsubsection\The Number 1, Internalized\ -(* "number1(M,a) == (\x[M]. empty(M,x) & successor(M,x,a))" *) +(* "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))))" + "number1_fm(a) \ Exists(And(empty_fm(0), succ_fm(0,succ(a))))" lemma number1_type [TC]: - "x \ nat ==> number1_fm(x) \ formula" + "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))" + "\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)" + "\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: @@ -542,27 +542,27 @@ subsubsection\Big Union, Internalized\ -(* "big_union(M,A,z) == \x[M]. x \ z \ (\y[M]. y\A & x \ y)" *) +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -582,8 +582,8 @@ \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))" + "\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: @@ -594,8 +594,8 @@ done lemma sats_transset_fm': - "[|x \ nat; env \ list(A)|] - ==> sats(A, transset_fm(x), env) \ transitive_set(##A, nth(x,env))" + "\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: @@ -606,13 +606,13 @@ done lemma sats_ordinal_fm': - "[|x \ nat; env \ list(A)|] - ==> sats(A, ordinal_fm(x), env) \ ordinal(##A,nth(x,env))" + "\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)" + "\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: @@ -626,7 +626,7 @@ definition Memrel_fm :: "[i,i]=>i" where - "Memrel_fm(A,r) == + "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)))), @@ -634,19 +634,19 @@ pair_fm(1,0,2))))))))" lemma Memrel_type [TC]: - "[| x \ nat; y \ nat |] ==> Memrel_fm(x,y) \ formula" + "\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) \ + "\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)" + "\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: @@ -660,7 +660,7 @@ definition pred_set_fm :: "[i,i,i,i]=>i" where - "pred_set_fm(A,x,r,B) == + "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))), @@ -668,20 +668,20 @@ lemma pred_set_type [TC]: - "[| A \ nat; x \ nat; r \ nat; B \ nat |] - ==> pred_set_fm(A,x,r,B) \ formula" + "\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) \ + "\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)" + "\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: @@ -695,29 +695,29 @@ subsubsection\Domain of a Relation, Internalized\ -(* "is_domain(M,r,z) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -730,29 +730,29 @@ subsubsection\Range of a Relation, Internalized\ -(* "is_range(M,r,z) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -765,30 +765,30 @@ subsubsection\Field of a Relation, Internalized\ -(* "is_field(M,r,z) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -802,30 +802,30 @@ subsubsection\Image under a Relation, Internalized\ -(* "image(M,r,A,z) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -838,30 +838,30 @@ subsubsection\Pre-Image under a Relation, Internalized\ -(* "pre_image(M,r,A,z) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -874,30 +874,30 @@ subsubsection\Function Application, Internalized\ -(* "fun_apply(M,f,x,y) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -911,26 +911,26 @@ subsubsection\The Concept of Relation, Internalized\ -(* "is_relation(M,r) == +(* "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) == + "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" + "\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))" + "\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)" + "\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: @@ -943,12 +943,12 @@ subsubsection\The Concept of Function, Internalized\ -(* "is_function(M,r) == +(* "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) == + "function_fm(r) \ Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,3,1), Implies(pair_fm(4,2,0), @@ -956,18 +956,18 @@ Implies(Member(0,r#+5), Equal(3,2))))))))))" lemma function_type [TC]: - "[| x \ nat |] ==> function_fm(x) \ formula" + "\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))" + "\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)" + "\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: @@ -980,13 +980,13 @@ subsubsection\Typed Functions, Internalized\ -(* "typed_function(M,A,B,r) == +(* "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) == + "typed_function_fm(A,B,r) \ And(function_fm(r), And(relation_fm(r), And(domain_fm(r,A), @@ -994,19 +994,19 @@ 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" + "\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) \ + "\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)" + "\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 = @@ -1039,14 +1039,14 @@ subsubsection\Composition of Relations, Internalized\ -(* "composition(M,r,s,t) == +(* "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) == + "composition_fm(r,s,t) \ Forall(Iff(Member(0,succ(t)), Exists(Exists(Exists(Exists(Exists( And(pair_fm(4,2,5), @@ -1055,19 +1055,19 @@ 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" + "\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) \ + "\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)" + "\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: @@ -1080,13 +1080,13 @@ subsubsection\Injections, Internalized\ -(* "injection(M,A,B,f) == +(* "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) == + "injection_fm(A,B,f) \ And(typed_function_fm(A,B,f), Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,2,1), @@ -1096,19 +1096,19 @@ lemma injection_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> injection_fm(x,y,z) \ formula" + "\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) \ + "\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)" + "\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: @@ -1122,31 +1122,31 @@ subsubsection\Surjections, Internalized\ (* surjection :: "[i=>o,i,i,i] => o" - "surjection(M,A,B,f) == + "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -1161,25 +1161,25 @@ 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)" *) + "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))" + "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" + "\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) \ + "\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)" + "\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: @@ -1193,30 +1193,30 @@ subsubsection\Restriction of a Relation, Internalized\ -(* "restriction(M,r,A,z) == +(* "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) == + "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" + "\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) \ + "\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)" + "\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: @@ -1229,7 +1229,7 @@ subsubsection\Order-Isomorphisms, Internalized\ (* order_isomorphism :: "[i=>o,i,i,i,i,i] => o" - "order_isomorphism(M,A,r,B,s,f) == + "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]. @@ -1239,7 +1239,7 @@ definition order_isomorphism_fm :: "[i,i,i,i,i]=>i" where - "order_isomorphism_fm(A,r,B,s,f) == + "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))), @@ -1251,22 +1251,22 @@ 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" + "\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) \ + "\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(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) \ + 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 @@ -1281,13 +1281,13 @@ text\A limit ordinal is a non-empty, successor-closed ordinal\ -(* "limit_ordinal(M,a) == - ordinal(M,a) & ~ empty(M,a) & +(* "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) == + "limit_ordinal_fm(x) \ And(ordinal_fm(x), And(Neg(empty_fm(x)), Forall(Implies(Member(0,succ(x)), @@ -1295,18 +1295,18 @@ succ_fm(1,0)))))))" lemma limit_ordinal_type [TC]: - "x \ nat ==> limit_ordinal_fm(x) \ formula" + "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))" + "\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)" + "\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: @@ -1319,30 +1319,30 @@ 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))" *) +(* "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) == + "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" + "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))" + "\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)" + "\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: @@ -1355,27 +1355,27 @@ subsubsection\Omega: The Set of Natural Numbers\ -(* omega(M,a) == limit_ordinal(M,a) & (\x[M]. x\a \ ~ limit_ordinal(M,x)) *) +(* omega(M,a) \ limit_ordinal(M,a) & (\x[M]. x\a \ \ limit_ordinal(M,x)) *) definition omega_fm :: "i=>i" where - "omega_fm(x) == + "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" + "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))" + "\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)" + "\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: diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/MetaExists.thy --- a/src/ZF/Constructible/MetaExists.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/MetaExists.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,9 +11,9 @@ definition ex :: "(('a::{}) \ prop) \ prop" (binder \\\ 0) where - "ex(P) == (\Q. (\x. PROP P(x) \ PROP Q) \ PROP Q)" + "ex(P) \ (\Q. (\x. PROP P(x) \ PROP Q) \ PROP Q)" -lemma meta_exI: "PROP P(x) ==> (\x. PROP P(x))" +lemma meta_exI: "PROP P(x) \ (\x. PROP P(x))" proof (unfold ex_def) assume P: "PROP P(x)" fix Q @@ -21,7 +21,7 @@ from P show "PROP Q" by (rule PQ) qed -lemma meta_exE: "[|\x. PROP P(x); \x. PROP P(x) ==> PROP R |] ==> PROP R" +lemma meta_exE: "\\x. PROP P(x); \x. PROP P(x) \ PROP R\ \ PROP R" proof (unfold ex_def) assume QPQ: "\Q. (\x. PROP P(x) \ PROP Q) \ PROP Q" assume PR: "\x. PROP P(x) \ PROP R" diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Normal.thy --- a/src/ZF/Constructible/Normal.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Normal.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,34 +19,34 @@ definition Closed :: "(i=>o) => o" where - "Closed(P) == \I. I \ 0 \ (\i\I. Ord(i) \ P(i)) \ P(\(I))" + "Closed(P) \ \I. I \ 0 \ (\i\I. Ord(i) \ P(i)) \ P(\(I))" definition Unbounded :: "(i=>o) => o" where - "Unbounded(P) == \i. Ord(i) \ (\j. i P(j))" + "Unbounded(P) \ \i. Ord(i) \ (\j. i P(j))" definition Closed_Unbounded :: "(i=>o) => o" where - "Closed_Unbounded(P) == Closed(P) \ Unbounded(P)" + "Closed_Unbounded(P) \ Closed(P) \ Unbounded(P)" subsubsection\Simple facts about c.u. classes\ lemma ClosedI: - "[| !!I. [| I \ 0; \i\I. Ord(i) \ P(i) |] ==> P(\(I)) |] - ==> Closed(P)" + "\\I. \I \ 0; \i\I. Ord(i) \ P(i)\ \ P(\(I))\ + \ Closed(P)" by (simp add: Closed_def) lemma ClosedD: - "[| Closed(P); I \ 0; !!i. i\I ==> Ord(i); !!i. i\I ==> P(i) |] - ==> P(\(I))" + "\Closed(P); I \ 0; \i. i\I \ Ord(i); \i. i\I \ P(i)\ + \ P(\(I))" by (simp add: Closed_def) lemma UnboundedD: - "[| Unbounded(P); Ord(i) |] ==> \j. i P(j)" + "\Unbounded(P); Ord(i)\ \ \j. i P(j)" by (simp add: Unbounded_def) -lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) ==> Unbounded(C)" +lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) \ Unbounded(C)" by (simp add: Closed_Unbounded_def) @@ -82,11 +82,11 @@ fixes P and A fixes next_greater \ \the next ordinal satisfying class \<^term>\A\\ fixes sup_greater \ \sup of those ordinals over all \<^term>\A\\ - assumes closed: "a\A ==> Closed(P(a))" - and unbounded: "a\A ==> Unbounded(P(a))" + assumes closed: "a\A \ Closed(P(a))" + and unbounded: "a\A \ Unbounded(P(a))" and A_non0: "A\0" - defines "next_greater(a,x) == \ y. x P(a,y)" - and "sup_greater(x) == \a\A. next_greater(a,x)" + defines "next_greater(a,x) \ \ y. x P(a,y)" + and "sup_greater(x) \ \a\A. next_greater(a,x)" text\Trivial that the intersection is closed.\ @@ -106,7 +106,7 @@ text\\<^term>\next_greater\ works as expected: it returns a larger value and one that belongs to class \<^term>\P(a)\.\ lemma (in cub_family) next_greater_lemma: - "[| Ord(x); a\A |] ==> P(a, next_greater(a,x)) \ x < next_greater(a,x)" + "\Ord(x); a\A\ \ P(a, next_greater(a,x)) \ x < next_greater(a,x)" apply (simp add: next_greater_def) apply (rule exE [OF UnboundedD [OF unbounded]]) apply assumption+ @@ -114,29 +114,29 @@ done lemma (in cub_family) next_greater_in_P: - "[| Ord(x); a\A |] ==> P(a, next_greater(a,x))" + "\Ord(x); a\A\ \ P(a, next_greater(a,x))" by (blast dest: next_greater_lemma) lemma (in cub_family) next_greater_gt: - "[| Ord(x); a\A |] ==> x < next_greater(a,x)" + "\Ord(x); a\A\ \ x < next_greater(a,x)" by (blast dest: next_greater_lemma) lemma (in cub_family) sup_greater_gt: - "Ord(x) ==> x < sup_greater(x)" + "Ord(x) \ x < sup_greater(x)" apply (simp add: sup_greater_def) apply (insert A_non0) apply (blast intro: UN_upper_lt next_greater_gt Ord_next_greater) done lemma (in cub_family) next_greater_le_sup_greater: - "a\A ==> next_greater(a,x) \ sup_greater(x)" + "a\A \ next_greater(a,x) \ sup_greater(x)" apply (simp add: sup_greater_def) apply (blast intro: UN_upper_le Ord_next_greater) done lemma (in cub_family) omega_sup_greater_eq_UN: - "[| Ord(x); a\A |] - ==> sup_greater^\ (x) = + "\Ord(x); a\A\ + \ sup_greater^\ (x) = (\n\nat. next_greater(a, sup_greater^n (x)))" apply (simp add: iterates_omega_def) apply (rule le_anti_sym) @@ -153,7 +153,7 @@ done lemma (in cub_family) P_omega_sup_greater: - "[| Ord(x); a\A |] ==> P(a, sup_greater^\ (x))" + "\Ord(x); a\A\ \ P(a, sup_greater^\ (x))" apply (simp add: omega_sup_greater_eq_UN) apply (rule ClosedD [OF closed]) apply (blast intro: ltD, auto) @@ -162,7 +162,7 @@ done lemma (in cub_family) omega_sup_greater_gt: - "Ord(x) ==> x < sup_greater^\ (x)" + "Ord(x) \ x < sup_greater^\ (x)" apply (simp add: iterates_omega_def) apply (rule UN_upper_lt [of 1], simp_all) apply (blast intro: sup_greater_gt) @@ -180,8 +180,8 @@ theorem Closed_Unbounded_INT: - "(!!a. a\A ==> Closed_Unbounded(P(a))) - ==> Closed_Unbounded(\x. \a\A. P(a, x))" + "(\a. a\A \ Closed_Unbounded(P(a))) + \ Closed_Unbounded(\x. \a\A. P(a, x))" apply (case_tac "A=0", simp) apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro]) apply (simp_all add: Closed_Unbounded_def) @@ -192,8 +192,8 @@ by auto theorem Closed_Unbounded_Int: - "[| Closed_Unbounded(P); Closed_Unbounded(Q) |] - ==> Closed_Unbounded(\x. P(x) \ Q(x))" + "\Closed_Unbounded(P); Closed_Unbounded(Q)\ + \ Closed_Unbounded(\x. P(x) \ Q(x))" apply (simp only: Int_iff_INT2) apply (rule Closed_Unbounded_INT, auto) done @@ -203,43 +203,43 @@ definition mono_le_subset :: "(i=>i) => o" where - "mono_le_subset(M) == \i j. i\j \ M(i) \ M(j)" + "mono_le_subset(M) \ \i j. i\j \ M(i) \ M(j)" definition mono_Ord :: "(i=>i) => o" where - "mono_Ord(F) == \i j. i F(i) < F(j)" + "mono_Ord(F) \ \i j. i F(i) < F(j)" definition cont_Ord :: "(i=>i) => o" where - "cont_Ord(F) == \l. Limit(l) \ F(l) = (\i \l. Limit(l) \ F(l) = (\ii) => o" where - "Normal(F) == mono_Ord(F) \ cont_Ord(F)" + "Normal(F) \ mono_Ord(F) \ cont_Ord(F)" subsubsection\Immediate properties of the definitions\ lemma NormalI: - "[|!!i j. i F(i) < F(j); !!l. Limit(l) ==> F(l) = (\i Normal(F)" + "\\i j. i F(i) < F(j); \l. Limit(l) \ F(l) = (\i + \ Normal(F)" by (simp add: Normal_def mono_Ord_def cont_Ord_def) -lemma mono_Ord_imp_Ord: "[| Ord(i); mono_Ord(F) |] ==> Ord(F(i))" +lemma mono_Ord_imp_Ord: "\Ord(i); mono_Ord(F)\ \ Ord(F(i))" apply (auto simp add: mono_Ord_def) apply (blast intro: lt_Ord) done -lemma mono_Ord_imp_mono: "[| i F(i) < F(j)" +lemma mono_Ord_imp_mono: "\i \ F(i) < F(j)" by (simp add: mono_Ord_def) -lemma Normal_imp_Ord [simp]: "[| Normal(F); Ord(i) |] ==> Ord(F(i))" +lemma Normal_imp_Ord [simp]: "\Normal(F); Ord(i)\ \ Ord(F(i))" by (simp add: Normal_def mono_Ord_imp_Ord) -lemma Normal_imp_cont: "[| Normal(F); Limit(l) |] ==> F(l) = (\iNormal(F); Limit(l)\ \ F(l) = (\i F(i) < F(j)" +lemma Normal_imp_mono: "\i \ F(i) < F(j)" by (simp add: Normal_def mono_Ord_def) lemma Normal_increasing: @@ -271,7 +271,7 @@ text\The proof is from Drake, pages 113--114.\ -lemma mono_Ord_imp_le_subset: "mono_Ord(F) ==> mono_le_subset(F)" +lemma mono_Ord_imp_le_subset: "mono_Ord(F) \ mono_le_subset(F)" apply (simp add: mono_le_subset_def, clarify) apply (subgoal_tac "F(i)\F(j)", blast dest: le_imp_subset) apply (simp add: le_iff) @@ -280,8 +280,8 @@ text\The following equation is taken for granted in any set theory text.\ lemma cont_Ord_Union: - "[| cont_Ord(F); mono_le_subset(F); X=0 \ F(0)=0; \x\X. Ord(x) |] - ==> F(\(X)) = (\y\X. F(y))" + "\cont_Ord(F); mono_le_subset(F); X=0 \ F(0)=0; \x\X. Ord(x)\ + \ F(\(X)) = (\y\X. F(y))" apply (frule Ord_set_cases) apply (erule disjE, force) apply (thin_tac "X=0 \ Q" for Q, auto) @@ -306,12 +306,12 @@ done lemma Normal_Union: - "[| X\0; \x\X. Ord(x); Normal(F) |] ==> F(\(X)) = (\y\X. F(y))" + "\X\0; \x\X. Ord(x); Normal(F)\ \ F(\(X)) = (\y\X. F(y))" apply (simp add: Normal_def) apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) done -lemma Normal_imp_fp_Closed: "Normal(F) ==> Closed(\i. F(i) = i)" +lemma Normal_imp_fp_Closed: "Normal(F) \ Closed(\i. F(i) = i)" apply (simp add: Closed_def ball_conj_distrib, clarify) apply (frule Ord_set_cases) apply (auto simp add: Normal_Union) @@ -319,19 +319,19 @@ lemma iterates_Normal_increasing: - "[| n\nat; x < F(x); Normal(F) |] - ==> F^n (x) < F^(succ(n)) (x)" + "\n\nat; x < F(x); Normal(F)\ + \ F^n (x) < F^(succ(n)) (x)" apply (induct n rule: nat_induct) apply (simp_all add: Normal_imp_mono) done lemma Ord_iterates_Normal: - "[| n\nat; Normal(F); Ord(x) |] ==> Ord(F^n (x))" + "\n\nat; Normal(F); Ord(x)\ \ Ord(F^n (x))" by (simp) text\THIS RESULT IS UNUSED\ lemma iterates_omega_Limit: - "[| Normal(F); x < F(x) |] ==> Limit(F^\ (x))" + "\Normal(F); x < F(x)\ \ Limit(F^\ (x))" apply (frule lt_Ord) apply (simp add: iterates_omega_def) apply (rule increasing_LimitI) @@ -347,7 +347,7 @@ done lemma iterates_omega_fixedpoint: - "[| Normal(F); Ord(a) |] ==> F(F^\ (a)) = F^\ (a)" + "\Normal(F); Ord(a)\ \ F(F^\ (a)) = F^\ (a)" apply (frule Normal_increasing, assumption) apply (erule leE) apply (simp_all add: iterates_omega_triv [OF sym]) (*for subgoal 2*) @@ -363,12 +363,12 @@ done lemma iterates_omega_increasing: - "[| Normal(F); Ord(a) |] ==> a \ F^\ (a)" + "\Normal(F); Ord(a)\ \ a \ F^\ (a)" apply (unfold iterates_omega_def) apply (rule UN_upper_le [of 0], simp_all) done -lemma Normal_imp_fp_Unbounded: "Normal(F) ==> Unbounded(\i. F(i) = i)" +lemma Normal_imp_fp_Unbounded: "Normal(F) \ Unbounded(\i. F(i) = i)" apply (unfold Unbounded_def, clarify) apply (rule_tac x="F^\ (succ(i))" in exI) apply (simp add: iterates_omega_fixedpoint) @@ -377,7 +377,7 @@ theorem Normal_imp_fp_Closed_Unbounded: - "Normal(F) ==> Closed_Unbounded(\i. F(i) = i)" + "Normal(F) \ Closed_Unbounded(\i. F(i) = i)" by (simp add: Closed_Unbounded_def Normal_imp_fp_Closed Normal_imp_fp_Unbounded) @@ -391,17 +391,17 @@ \ definition normalize :: "[i=>i, i] => i" where - "normalize(F,a) == transrec2(a, F(0), \x r. F(succ(x)) \ succ(r))" + "normalize(F,a) \ transrec2(a, F(0), \x r. F(succ(x)) \ succ(r))" lemma Ord_normalize [simp, intro]: - "[| Ord(a); !!x. Ord(x) ==> Ord(F(x)) |] ==> Ord(normalize(F, a))" + "\Ord(a); \x. Ord(x) \ Ord(F(x))\ \ Ord(normalize(F, a))" apply (induct a rule: trans_induct3) apply (simp_all add: ltD def_transrec2 [OF normalize_def]) done lemma normalize_increasing: - assumes ab: "a < b" and F: "!!x. Ord(x) ==> Ord(F(x))" + assumes ab: "a < b" and F: "\x. Ord(x) \ Ord(F(x))" shows "normalize(F,a) < normalize(F,b)" proof - { fix x @@ -428,14 +428,14 @@ qed theorem Normal_normalize: - "(!!x. Ord(x) ==> Ord(F(x))) ==> Normal(normalize(F))" + "(\x. Ord(x) \ Ord(F(x))) \ Normal(normalize(F))" apply (rule NormalI) apply (blast intro!: normalize_increasing) apply (simp add: def_transrec2 [OF normalize_def]) done theorem le_normalize: - assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "!!x. Ord(x) ==> Ord(F(x))" + assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "\x. Ord(x) \ Ord(F(x))" shows "F(a) \ normalize(F,a)" using a proof (induct a rule: trans_induct3) @@ -457,10 +457,10 @@ definition Aleph :: "i => i" (\\_\ [90] 90) where - "Aleph(a) == transrec2(a, nat, \x r. csucc(r))" + "Aleph(a) \ transrec2(a, nat, \x r. csucc(r))" lemma Card_Aleph [simp, intro]: - "Ord(a) ==> Card(Aleph(a))" + "Ord(a) \ Card(Aleph(a))" apply (erule trans_induct3) apply (simp_all add: Card_csucc Card_nat Card_is_Ord def_transrec2 [OF Aleph_def]) diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Rank.thy --- a/src/ZF/Constructible/Rank.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Rank.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,24 +11,24 @@ 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]. + "\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]. + "\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]. + "\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, + "\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))" @@ -37,8 +37,8 @@ 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" + "\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) apply (elim conjE CollectE) apply (erule wellfounded_on_induct, assumption+) @@ -51,12 +51,12 @@ 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" + "\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) *) +(*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) @@ -64,8 +64,8 @@ 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) |] ==> \ r" + "\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)\ \ \ r" apply (frule wellordered_is_trans_on, assumption) apply (rule notI) apply (drule_tac x2=y and x=x and r2=r in @@ -77,9 +77,9 @@ text\Simple consequence of Lemma 6.1\ lemma (in M_ordertype) wellordered_iso_pred_eq: - "[| wellordered(M,A,r); + "\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" + 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) @@ -95,11 +95,11 @@ 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)); + "\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" + 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) @@ -127,8 +127,8 @@ lemma ord_iso_converse1: - "[| f: ord_iso(A,r,B,s); : s; a:A; b:B |] - ==> \ r" + "\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]) @@ -138,28 +138,28 @@ 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) & + "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) == + "omap(M,A,r,f) \ \z[M]. z \ f \ (\a\A. \x[M]. \g[M]. z = & 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)" + "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 \ + "\omap(M,A,r,f); M(A); M(f)\ + \ z \ f \ (\a\A. \x[M]. \g[M]. z = & Ord(x) & g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))" apply (simp add: omap_def) @@ -170,18 +170,18 @@ 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" + "\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)" + "\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 \ + "\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) @@ -192,15 +192,15 @@ 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)" + "\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)" + "\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) @@ -220,32 +220,32 @@ 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)" + "\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" + "\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" + "\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)" + "\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) @@ -258,8 +258,8 @@ 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))" + "\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) @@ -279,8 +279,8 @@ 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))" + "\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) @@ -304,9 +304,9 @@ 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)\)" + "\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]) @@ -315,8 +315,8 @@ 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" + "\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) @@ -334,21 +334,21 @@ 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))" + "\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))" + "\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)" + "\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) @@ -363,7 +363,7 @@ done lemma (in M_ordertype) otype_exists: - "[| wellordered(M,A,r); M(A); M(r) |] ==> \i[M]. otype(M,A,r,i)" + "\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) @@ -371,8 +371,8 @@ 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)))" + "\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) @@ -383,7 +383,7 @@ lemma (in M_ordertype) relativized_imp_well_ord: - "[| wellordered(M,A,r); M(A); M(r) |] ==> well_ord(A,r)" + "\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 @@ -392,14 +392,14 @@ 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)" + "\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)" + "\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) @@ -410,14 +410,14 @@ subsubsection\Ordinal Addition\ -(*FIXME: update to use new techniques!!*) +(*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) == + "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, @@ -426,10 +426,10 @@ 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) | + "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) & @@ -438,7 +438,7 @@ definition (*NEEDS RELATIVIZATION*) omult_eqns :: "[i,i,i,i] => o" where - "omult_eqns(i,x,g,z) == + "omult_eqns(i,x,g,z) \ Ord(x) & (x=0 \ z=0) & (\j. x = succ(j) \ z = g`j ++ i) & @@ -446,14 +446,14 @@ definition is_omult_fun :: "[i=>o,i,i,i] => o" where - "is_omult_fun(M,i,j,f) == + "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) == + "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" @@ -461,14 +461,14 @@ locale M_ord_arith = M_ordertype + assumes oadd_strong_replacement: - "[| M(i); M(j) |] ==> + "\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) |] ==> + "\M(i); M(j)\ \ strong_replacement(M, \x z. \y[M]. z = & (\g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & @@ -478,8 +478,8 @@ 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) \ + "\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 @@ -492,7 +492,7 @@ lemma (in M_ord_arith) oadd_strong_replacement': - "[| M(i); M(j) |] ==> + "\M(i); M(j)\ \ strong_replacement(M, \x z. \y[M]. z = & (\g[M]. is_recfun(Memrel(succ(j)),x,%x g. i \ g``x,g) & @@ -504,14 +504,14 @@ 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)" + "\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)" + "\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) @@ -522,8 +522,8 @@ 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})" + "\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]) @@ -531,8 +531,8 @@ 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" + "\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. k is_oadd(M,i,j,k) \ k = i++j" + "\M(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" + "\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)" + "\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) @@ -564,7 +564,7 @@ subsubsection\Ordinal Multiplication\ lemma omult_eqns_unique: - "[| omult_eqns(i,x,g,z); omult_eqns(i,x,g,z') |] ==> z=z'" + "\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 @@ -579,26 +579,26 @@ 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" + "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)" + "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)" + "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)" +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))" + "\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) @@ -608,15 +608,15 @@ 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)" + "\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)" + "\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) @@ -630,24 +630,24 @@ done lemma (in M_ord_arith) is_omult_fun_apply_0: - "[| 0 < j; is_omult_fun(M,i,j,f) |] ==> f`0 = 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" + "\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)" + "\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" + "\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) @@ -660,7 +660,7 @@ done lemma (in M_ord_arith) omult_abs: - "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_omult(M,i,j,k) \ k = i**j" + "\M(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 @@ -675,22 +675,22 @@ locale M_wfrank = M_trancl + assumes wfrank_separation: - "M(r) ==> + "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)))" + \ (\f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))" and wfrank_strong_replacement: - "M(r) ==> + "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) ==> + "M(r) \ separation (M, \x. \rplus[M]. tran_closure(M,r,rplus) \ - ~ (\f[M]. \rangef[M]. + \ (\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)))" @@ -700,15 +700,15 @@ agree with the "real" ones.\ lemma (in M_wfrank) wfrank_separation': - "M(r) ==> + "M(r) \ separation - (M, \x. ~ (\f[M]. is_recfun(r^+, x, %x f. range(f), f)))" + (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) ==> + "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))" @@ -717,9 +717,9 @@ done lemma (in M_wfrank) Ord_wfrank_separation': - "M(r) ==> + "M(r) \ separation (M, \x. - ~ (\f[M]. is_recfun(r^+, x, \x. range, f) \ Ord(range(f))))" + \ (\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 @@ -728,14 +728,14 @@ well-founded relations within the class M.\ definition wellfoundedrank :: "[i=>o,i,i] => i" where - "wellfoundedrank(M,r,A) == + "wellfoundedrank(M,r,A) \ {p. x\A, \y[M]. \f[M]. p = & 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)" + "\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) @@ -745,7 +745,7 @@ done lemma (in M_wfrank) M_wellfoundedrank: - "[| wellfounded(M,r); M(r); M(A) |] ==> M(wellfoundedrank(M,r,A))" + "\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) @@ -758,8 +758,8 @@ 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))" + "\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 @@ -790,8 +790,8 @@ done lemma (in M_wfrank) Ord_range_wellfoundedrank: - "[| wellfounded(M,r); r \ A*A; M(r); M(A) |] - ==> Ord (range(wellfoundedrank(M,r,A)))" + "\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) @@ -819,8 +819,8 @@ done lemma (in M_wfrank) function_wellfoundedrank: - "[| wellfounded(M,r); M(r); M(A)|] - ==> function(wellfoundedrank(M,r,A))" + "\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) @@ -829,8 +829,8 @@ done lemma (in M_wfrank) domain_wellfoundedrank: - "[| wellfounded(M,r); M(r); M(A)|] - ==> domain(wellfoundedrank(M,r,A)) = A" + "\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) @@ -847,8 +847,8 @@ done lemma (in M_wfrank) wellfoundedrank_type: - "[| wellfounded(M,r); M(r); M(A)|] - ==> wellfoundedrank(M,r,A) \ A -> range(wellfoundedrank(M,r,A))" + "\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) @@ -857,15 +857,15 @@ 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)" + "\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)" + "\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) @@ -882,9 +882,9 @@ lemma (in M_wfrank) wellfoundedrank_lt: - "[| \ r; - wellfounded(M,r); r \ A*A; M(r); M(A)|] - ==> wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,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 @@ -908,8 +908,8 @@ 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))" + "\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) @@ -919,12 +919,12 @@ done lemma (in M_wfrank) wellfounded_imp_wf: - "[|wellfounded(M,r); relation(r); M(r)|] ==> wf(r)" + "\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)" + "\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) @@ -932,11 +932,11 @@ theorem (in M_wfrank) wf_abs: - "[|relation(r); M(r)|] ==> wellfounded(M,r) \ wf(r)" + "\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)" + "\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 -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Rank_Separation.thy --- a/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 16:51:35 2022 +0100 @@ -24,8 +24,8 @@ by (intro FOL_reflections function_reflections) lemma well_ord_iso_separation: - "[| L(A); L(f); L(r) |] - ==> separation (L, \x. x\A \ (\y[L]. (\p[L]. + "\L(A); L(f); L(r)\ + \ separation (L, \x. x\A \ (\y[L]. (\p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \ r)))" apply (rule gen_separation_multi [OF well_ord_iso_Reflects, of "{A,f,r}"], auto) @@ -47,8 +47,8 @@ lemma obase_separation: \ \part of the order type formalization\ - "[| L(A); L(r) |] - ==> separation(L, \a. \x[L]. \g[L]. \mx[L]. \par[L]. + "\L(A); L(r)\ + \ separation(L, \a. \x[L]. \g[L]. \mx[L]. \par[L]. ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))" apply (rule gen_separation_multi [OF obase_reflects, of "{A,r}"], auto) @@ -60,19 +60,19 @@ subsubsection\Separation for a Theorem about \<^term>\obase\\ lemma obase_equals_reflects: - "REFLECTS[\x. x\A \ ~(\y[L]. \g[L]. + "REFLECTS[\x. x\A \ \(\y[L]. \g[L]. ordinal(L,y) & (\my[L]. \pxr[L]. membership(L,y,my) & pred_set(L,A,x,r,pxr) & order_isomorphism(L,pxr,r,y,my,g))), - \i x. x\A \ ~(\y \ Lset(i). \g \ Lset(i). + \i x. x\A \ \(\y \ Lset(i). \g \ Lset(i). ordinal(##Lset(i),y) & (\my \ Lset(i). \pxr \ Lset(i). membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) & order_isomorphism(##Lset(i),pxr,r,y,my,g)))]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma obase_equals_separation: - "[| L(A); L(r) |] - ==> separation (L, \x. x\A \ ~(\y[L]. \g[L]. + "\L(A); L(r)\ + \ separation (L, \x. x\A \ \(\y[L]. \g[L]. ordinal(L,y) & (\my[L]. \pxr[L]. membership(L,y,my) & pred_set(L,A,x,r,pxr) & order_isomorphism(L,pxr,r,y,my,g))))" @@ -96,8 +96,8 @@ by (intro FOL_reflections function_reflections fun_plus_reflections) lemma omap_replacement: - "[| L(A); L(r) |] - ==> strong_replacement(L, + "\L(A); L(r)\ + \ strong_replacement(L, \a z. \x[L]. \g[L]. \mx[L]. \par[L]. ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) & pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))" @@ -133,17 +133,17 @@ lemma wfrank_Reflects: "REFLECTS[\x. \rplus[L]. tran_closure(L,r,rplus) \ - ~ (\f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)), + \ (\f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)), \i x. \rplus \ Lset(i). tran_closure(##Lset(i),r,rplus) \ - ~ (\f \ Lset(i). + \ (\f \ Lset(i). M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f))]" by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection) lemma wfrank_separation: - "L(r) ==> + "L(r) \ separation (L, \x. \rplus[L]. tran_closure(L,r,rplus) \ - ~ (\f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))" + \ (\f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))" apply (rule gen_separation [OF wfrank_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ @@ -167,7 +167,7 @@ is_recfun_reflection tran_closure_reflection) lemma wfrank_strong_replacement: - "L(r) ==> + "L(r) \ strong_replacement(L, \x z. \rplus[L]. tran_closure(L,r,rplus) \ (\y[L]. \f[L]. pair(L,x,y,z) & @@ -186,12 +186,12 @@ lemma Ord_wfrank_Reflects: "REFLECTS[\x. \rplus[L]. tran_closure(L,r,rplus) \ - ~ (\f[L]. \rangef[L]. + \ (\f[L]. \rangef[L]. is_range(L,f,rangef) \ M_is_recfun(L, \x f y. is_range(L,f,y), rplus, x, f) \ ordinal(L,rangef)), \i x. \rplus \ Lset(i). tran_closure(##Lset(i),r,rplus) \ - ~ (\f \ Lset(i). \rangef \ Lset(i). + \ (\f \ Lset(i). \rangef \ Lset(i). is_range(##Lset(i),f,rangef) \ M_is_recfun(##Lset(i), \x f y. is_range(##Lset(i),f,y), rplus, x, f) \ @@ -200,10 +200,10 @@ tran_closure_reflection ordinal_reflection) lemma Ord_wfrank_separation: - "L(r) ==> + "L(r) \ separation (L, \x. \rplus[L]. tran_closure(L,r,rplus) \ - ~ (\f[L]. \rangef[L]. + \ (\f[L]. \rangef[L]. is_range(L,f,rangef) \ M_is_recfun(L, \x f y. is_range(L,f,y), rplus, x, f) \ ordinal(L,rangef)))" diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 16:51:35 2022 +0100 @@ -8,7 +8,7 @@ text\This theory proves all instances needed for locales \M_trancl\ and \M_datatypes\\ -lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> ji = succ(j); Ord(i)\ \ jFirst, The Defining Formula\ -(* "rtran_closure_mem(M,A,r,p) == +(* "rtran_closure_mem(M,A,r,p) \ \nnat[M]. \n[M]. \n'[M]. omega(M,nnat) & n\nnat & successor(M,n,n') & (\f[M]. typed_function(M,n',A,f) & @@ -30,7 +30,7 @@ fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \ r)))"*) definition rtran_closure_mem_fm :: "[i,i,i]=>i" where - "rtran_closure_mem_fm(A,r,p) == + "rtran_closure_mem_fm(A,r,p) \ Exists(Exists(Exists( And(omega_fm(2), And(Member(1,2), @@ -49,19 +49,19 @@ lemma rtran_closure_mem_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> rtran_closure_mem_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ rtran_closure_mem_fm(x,y,z) \ formula" by (simp add: rtran_closure_mem_fm_def) lemma sats_rtran_closure_mem_fm [simp]: - "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, rtran_closure_mem_fm(x,y,z), env) \ + "\x \ nat; y \ nat; z \ nat; env \ list(A)\ + \ sats(A, rtran_closure_mem_fm(x,y,z), env) \ rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def) lemma rtran_closure_mem_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i \ nat; j \ nat; k \ nat; env \ list(A)|] - ==> rtran_closure_mem(##A, x, y, z) \ sats(A, rtran_closure_mem_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i \ nat; j \ nat; k \ nat; env \ list(A)\ + \ rtran_closure_mem(##A, x, y, z) \ sats(A, rtran_closure_mem_fm(i,j,k), env)" by (simp) lemma rtran_closure_mem_reflection: @@ -73,7 +73,7 @@ text\Separation for \<^term>\rtrancl(r)\.\ lemma rtrancl_separation: - "[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))" + "\L(r); L(A)\ \ separation (L, rtran_closure_mem(L,A,r))" apply (rule gen_separation_multi [OF rtran_closure_mem_reflection, of "{r,A}"], auto) apply (rule_tac env="[r,A]" in DPow_LsetI) @@ -83,30 +83,30 @@ subsubsection\Reflexive/Transitive Closure, Internalized\ -(* "rtran_closure(M,r,s) == +(* "rtran_closure(M,r,s) \ \A[M]. is_field(M,r,A) \ (\p[M]. p \ s \ rtran_closure_mem(M,A,r,p))" *) definition rtran_closure_fm :: "[i,i]=>i" where - "rtran_closure_fm(r,s) == + "rtran_closure_fm(r,s) \ Forall(Implies(field_fm(succ(r),0), Forall(Iff(Member(0,succ(succ(s))), rtran_closure_mem_fm(1,succ(succ(r)),0)))))" lemma rtran_closure_type [TC]: - "[| x \ nat; y \ nat |] ==> rtran_closure_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ rtran_closure_fm(x,y) \ formula" by (simp add: rtran_closure_fm_def) lemma sats_rtran_closure_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, rtran_closure_fm(x,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, rtran_closure_fm(x,y), env) \ rtran_closure(##A, nth(x,env), nth(y,env))" by (simp add: rtran_closure_fm_def rtran_closure_def) lemma rtran_closure_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> rtran_closure(##A, x, y) \ sats(A, rtran_closure_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ rtran_closure(##A, x, y) \ sats(A, rtran_closure_fm(i,j), env)" by simp theorem rtran_closure_reflection: @@ -119,27 +119,27 @@ subsubsection\Transitive Closure of a Relation, Internalized\ -(* "tran_closure(M,r,t) == +(* "tran_closure(M,r,t) \ \s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *) definition tran_closure_fm :: "[i,i]=>i" where - "tran_closure_fm(r,s) == + "tran_closure_fm(r,s) \ Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))" lemma tran_closure_type [TC]: - "[| x \ nat; y \ nat |] ==> tran_closure_fm(x,y) \ formula" + "\x \ nat; y \ nat\ \ tran_closure_fm(x,y) \ formula" by (simp add: tran_closure_fm_def) lemma sats_tran_closure_fm [simp]: - "[| x \ nat; y \ nat; env \ list(A)|] - ==> sats(A, tran_closure_fm(x,y), env) \ + "\x \ nat; y \ nat; env \ list(A)\ + \ sats(A, tran_closure_fm(x,y), env) \ tran_closure(##A, nth(x,env), nth(y,env))" by (simp add: tran_closure_fm_def tran_closure_def) lemma tran_closure_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; - i \ nat; j \ nat; env \ list(A)|] - ==> tran_closure(##A, x, y) \ sats(A, tran_closure_fm(i,j), env)" + "\nth(i,env) = x; nth(j,env) = y; + i \ nat; j \ nat; env \ list(A)\ + \ tran_closure(##A, x, y) \ sats(A, tran_closure_fm(i,j), env)" by simp theorem tran_closure_reflection: @@ -163,7 +163,7 @@ tran_closure_reflection) lemma wellfounded_trancl_separation: - "[| L(r); L(Z) |] ==> + "\L(r); L(Z)\ \ separation (L, \x. \w[L]. \wx[L]. \rp[L]. w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp)" @@ -204,7 +204,7 @@ lemma list_replacement1: - "L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)" + "L(A) \ iterates_replacement(L, is_list_functor(L,A), 0)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,A,n,0,Memrel(succ(n))}" @@ -226,7 +226,7 @@ is_iterates_reflection list_functor_reflection) lemma list_replacement2: - "L(A) ==> strong_replacement(L, + "L(A) \ strong_replacement(L, \n y. n\nat & is_iterates(L, is_list_functor(L,A), 0, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{A,B,0,nat}" @@ -292,30 +292,30 @@ subsubsection\The Formula \<^term>\is_nth\, Internalized\ -(* "is_nth(M,n,l,Z) == +(* "is_nth(M,n,l,Z) \ \X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *) definition nth_fm :: "[i,i,i]=>i" where - "nth_fm(n,l,Z) == + "nth_fm(n,l,Z) \ Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0), hd_fm(0,succ(Z))))" lemma nth_fm_type [TC]: - "[| x \ nat; y \ nat; z \ nat |] ==> nth_fm(x,y,z) \ formula" + "\x \ nat; y \ nat; z \ nat\ \ nth_fm(x,y,z) \ formula" by (simp add: nth_fm_def) lemma sats_nth_fm [simp]: - "[| x < length(env); y \ nat; z \ nat; env \ list(A)|] - ==> sats(A, nth_fm(x,y,z), env) \ + "\x < length(env); y \ nat; z \ nat; env \ list(A)\ + \ sats(A, nth_fm(x,y,z), env) \ is_nth(##A, nth(x,env), nth(y,env), nth(z,env))" apply (frule lt_length_in_nat, assumption) apply (simp add: nth_fm_def is_nth_def sats_is_iterates_fm) done lemma nth_iff_sats: - "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; - i < length(env); j \ nat; k \ nat; env \ list(A)|] - ==> is_nth(##A, x, y, z) \ sats(A, nth_fm(i,j,k), env)" + "\nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; + i < length(env); j \ nat; k \ nat; env \ list(A)\ + \ is_nth(##A, x, y, z) \ sats(A, nth_fm(i,j,k), env)" by (simp) theorem nth_reflection: @@ -343,7 +343,7 @@ iterates_MH_reflection tl_reflection) lemma nth_replacement: - "L(w) ==> iterates_replacement(L, is_tl(L), w)" + "L(w) \ iterates_replacement(L, is_tl(L), w)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,w,Memrel(succ(n))}" @@ -390,7 +390,7 @@ iterates_MH_reflection) lemma eclose_replacement1: - "L(A) ==> iterates_replacement(L, big_union(L), A)" + "L(A) \ iterates_replacement(L, big_union(L), A)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,A,n,Memrel(succ(n))}" @@ -410,7 +410,7 @@ by (intro FOL_reflections function_reflections is_iterates_reflection) lemma eclose_replacement2: - "L(A) ==> strong_replacement(L, + "L(A) \ strong_replacement(L, \n y. n\nat & is_iterates(L, big_union(L), A, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{A,B,nat}" diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Reflection.thy Tue Sep 27 16:51:35 2022 +0100 @@ -6,11 +6,11 @@ theory Reflection imports Normal begin -lemma all_iff_not_ex_not: "(\x. P(x)) \ (~ (\x. ~ P(x)))" -by blast +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 +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}, @@ -33,105 +33,106 @@ 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) |] - ==> \ Mset(a)" - defines "M(x) == \a. Ord(a) & x \ Mset(a)" - and "Reflects(Cl,P,Q) == Closed_Unbounded(Cl) & + and Pair_in_Mset : "\x \ Mset(a); y \ Mset(a); Limit(a)\ + \ \ 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()) \ + defines "F0(P,y) \ \ b. (\z. M(z) & P()) \ (\z\Mset(b). P())" - and "FF(P) == \a. \y\Mset(a). F0(P,y)" - and "ClEx(P,a) == Limit(a) & normalize(FF(P),a) = a" + and "FF(P) \ \a. \y\Mset(a). F0(P,y)" + and "ClEx(P,a) \ Limit(a) & normalize(FF(P),a) = a" -lemma (in reflection) Mset_mono: "i\j ==> Mset(i) \ Mset(j)" -apply (insert Mset_mono_le) -apply (simp add: mono_le_subset_def leI) -done +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 (in reflection) ClEx_eq: - "ClEx(P) == \a. Limit(a) & normalize(FF(P),a) = a" +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 (in reflection) Triv_reflection [intro]: - "Reflects(Ord, P, \a x. P(x))" -by (simp add: Reflects_def) +theorem Triv_reflection [intro]: + "Reflects(Ord, P, \a x. P(x))" + by (simp add: Reflects_def) -theorem (in reflection) Not_reflection [intro]: - "Reflects(Cl,P,Q) ==> Reflects(Cl, \x. ~P(x), \a x. ~Q(a,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 (in reflection) And_reflection [intro]: - "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) & C'(a), \x. P(x) & P'(x), +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) + by (simp add: Reflects_def Closed_Unbounded_Int, blast) -theorem (in reflection) Or_reflection [intro]: - "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) & C'(a), \x. P(x) | P'(x), +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 (in reflection) Imp_reflection [intro]: - "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) & C'(a), +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 (in reflection) Iff_reflection [intro]: - "[| Reflects(Cl,P,Q); Reflects(C',P',Q') |] - ==> Reflects(\a. Cl(a) & C'(a), +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 (in reflection) F0_works: - "[| y\Mset(a); Ord(a); M(z); P() |] ==> \z\Mset(F0(P,y)). P()" +lemma F0_works: + "\y\Mset(a); Ord(a); M(z); P()\ \ \z\Mset(F0(P,y)). P()" apply (unfold F0_def M_def, clarify) apply (rule LeastI2) apply (blast intro: Mset_mono [THEN subsetD]) apply (blast intro: lt_Ord2, blast) done -lemma (in reflection) Ord_F0 [intro,simp]: "Ord(F0(P,y))" +lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))" by (simp add: F0_def) -lemma (in reflection) Ord_FF [intro,simp]: "Ord(FF(P,y))" +lemma Ord_FF [intro,simp]: "Ord(FF(P,y))" by (simp add: FF_def) -lemma (in reflection) cont_Ord_FF: "cont_Ord(FF(P))" +lemma cont_Ord_FF: "cont_Ord(FF(P))" apply (insert Mset_cont) apply (simp add: cont_Ord_def FF_def, blast) done text\Recall that \<^term>\F0\ depends upon \<^term>\y\Mset(a)\, while \<^term>\FF\ depends only upon \<^term>\a\.\ -lemma (in reflection) FF_works: - "[| M(z); y\Mset(a); P(); Ord(a) |] ==> \z\Mset(FF(P,a)). P()" +lemma FF_works: + "\M(z); y\Mset(a); P(); Ord(a)\ \ \z\Mset(FF(P,a)). P()" 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 (in reflection) FFN_works: - "[| M(z); y\Mset(a); P(); Ord(a) |] - ==> \z\Mset(normalize(FF(P),a)). P()" +lemma FFN_works: + "\M(z); y\Mset(a); P(); Ord(a)\ + \ \z\Mset(normalize(FF(P),a)). P()" 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\ @@ -139,11 +140,13 @@ 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)" + assumes Cl_reflects: "\Cl(a); Ord(a)\ \ \x\Mset(a). P(x) \ Q(a,x)" + +begin -lemma (in ex_reflection) ClEx_downward: - "[| M(z); y\Mset(a); P(); Cl(a); ClEx(P,a) |] - ==> \z\Mset(a). Q(a,)" +lemma ClEx_downward: + "\M(z); y\Mset(a); P(); Cl(a); ClEx(P,a)\ + \ \z\Mset(a). Q(a,)" apply (simp add: ClEx_def, clarify) apply (frule Limit_is_Ord) apply (frule FFN_works [of concl: P], assumption+) @@ -151,35 +154,40 @@ apply (auto simp add: Limit_is_Ord Pair_in_Mset) done -lemma (in ex_reflection) ClEx_upward: - "[| z\Mset(a); y\Mset(a); Q(a,); Cl(a); ClEx(P,a) |] - ==> \z. M(z) & P()" +lemma ClEx_upward: + "\z\Mset(a); y\Mset(a); Q(a,); Cl(a); ClEx(P,a)\ + \ \z. M(z) & P()" 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 (in ex_reflection) ZF_ClEx_iff: - "[| y\Mset(a); Cl(a); ClEx(P,a) |] - ==> (\z. M(z) & P()) \ (\z\Mset(a). Q(a,))" +lemma ZF_ClEx_iff: + "\y\Mset(a); Cl(a); ClEx(P,a)\ + \ (\z. M(z) & P()) \ (\z\Mset(a). Q(a,))" by (blast intro: dest: ClEx_downward ClEx_upward) text\...and it is closed and unbounded\ -lemma (in ex_reflection) ZF_Closed_Unbounded_ClEx: +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 (in reflection) 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()) \ (\z\Mset(a). Q(a,))" +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()) \ (\z\Mset(a). Q(a,))" apply (unfold 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, @@ -194,9 +202,9 @@ apply (simp add: ex_reflection_axioms.intro, assumption+) *) -lemma (in reflection) Closed_Unbounded_ClEx: - "(!!a. [| Cl(a); Ord(a) |] ==> \x\Mset(a). P(x) \ Q(a,x)) - ==> Closed_Unbounded(ClEx(P))" +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) apply (rule ex_reflection.ZF_Closed_Unbounded_ClEx [of Mset _ _ Cl]) apply (rule ex_reflection.intro, rule reflection_axioms) @@ -205,9 +213,9 @@ subsection\Packaging the Quantifier Reflection Rules\ -lemma (in reflection) Ex_reflection_0: +lemma Ex_reflection_0: "Reflects(Cl,P0,Q0) - ==> Reflects(\a. Cl(a) & ClEx(P0,a), + \ Reflects(\a. Cl(a) & ClEx(P0,a), \x. \z. M(z) & P0(), \a x. \z\Mset(a). Q0(a,))" apply (simp add: Reflects_def) @@ -217,9 +225,9 @@ apply (rule_tac Cl=Cl in ClEx_iff, assumption+, blast) done -lemma (in reflection) All_reflection_0: +lemma All_reflection_0: "Reflects(Cl,P0,Q0) - ==> Reflects(\a. Cl(a) & ClEx(\x.~P0(x), a), + \ Reflects(\a. Cl(a) & ClEx(\x.\P0(x), a), \x. \z. M(z) \ P0(), \a x. \z\Mset(a). Q0(a,))" apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) @@ -227,17 +235,17 @@ apply (erule Ex_reflection_0) done -theorem (in reflection) Ex_reflection [intro]: +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), + \ 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 (in reflection) All_reflection [intro]: +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), + \ 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))" @@ -245,16 +253,16 @@ text\And again, this time using class-bounded quantifiers\ -theorem (in reflection) Rex_reflection [intro]: +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), + \ 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 (in reflection) Rall_reflection [intro]: +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), + \ 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) @@ -268,7 +276,7 @@ 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 (in reflection) +schematic_goal "Reflects(?Cl, \x. \y. M(y) & x \ y, \a x. \y\Mset(a). x \ y)" @@ -277,7 +285,7 @@ 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 (in reflection) +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)" @@ -285,31 +293,31 @@ text\Example 2\ -schematic_goal (in reflection) +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 (in reflection) +lemma "Reflects (\a. (Ord(a) & - ClEx(\x. ~ (snd(x) \ fst(fst(x)) \ snd(x) \ snd(fst(x))), 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 (in reflection) +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 (in reflection) +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)" @@ -329,21 +337,21 @@ text\Example 3. Warning: the following examples make sense only if \<^term>\P\ is quantifier-free, since it is not being relativized.\ -schematic_goal (in reflection) +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 (in reflection) +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 (in reflection) +schematic_goal "Reflects(?Cl, \x. \y. M(y) & y = Replace(x,P), \a x. \y\Mset(a). y = Replace(x,P))" @@ -351,7 +359,7 @@ text\Example 4: Axiom of Choice. Possibly wrong, since \\\ needs to be relativized.\ -schematic_goal (in reflection) +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)))" @@ -359,3 +367,5 @@ end +end + diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Relative.thy Tue Sep 27 16:51:35 2022 +0100 @@ -12,115 +12,115 @@ definition empty :: "[i=>o,i] => o" where - "empty(M,z) == \x[M]. x \ z" + "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" + "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)" + "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) & + "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" + "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)" + "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)" + "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)" + "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)" + "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)" + "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)" + "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)" + "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))" + "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" + "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" + "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)" + "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) == + "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) == + "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) == + "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)" + "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)" + "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) == + "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) == + "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) == + "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) == + "image(M,r,A,z) \ \y[M]. y \ z \ (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w)))" definition @@ -129,46 +129,46 @@ \<^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) == + "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) == + "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) == + "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) == + "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) == + "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) == + "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) == + "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) == + "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) & @@ -176,104 +176,104 @@ definition injection :: "[i=>o,i,i,i] => o" where - "injection(M,A,B,f) == + "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) == + "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)" + "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) == + "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)" + "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))" + "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) & + "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)" + "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))" + "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))" + "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))" + "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) == + "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)" + "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) == + "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)" + "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) == + "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) == + "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) == + "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) == + "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)" @@ -291,7 +291,7 @@ definition extensionality :: "(i=>o) => o" where - "extensionality(M) == + "extensionality(M) \ \x[M]. \y[M]. (\z[M]. z \ x \ z \ y) \ x=y" definition @@ -300,42 +300,42 @@ 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) == + "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)" + "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)" + "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)" + "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) == + "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) == + "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) == + "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))" + "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$\ @@ -344,28 +344,28 @@ (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)" +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))" + "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))" + "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))" + "(\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))" + "\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]: @@ -373,13 +373,13 @@ 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))" + "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)) \ + "\\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) @@ -391,21 +391,21 @@ text\The separation axiom requires some lemmas\ lemma Collect_in_Vfrom: - "[| X \ Vfrom(A,j); Transset(A) |] ==> Collect(X,P) \ Vfrom(A, succ(j))" + "\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)" + "\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)" + "\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)" @@ -431,7 +431,7 @@ text\Powerset axiom\ lemma Pow_in_univ: - "[| X \ univ(A); Transset(A) |] ==> Pow(X) \ univ(A)" + "\X \ univ(A); Transset(A)\ \ Pow(X) \ univ(A)" apply (simp add: univ_def Pow_in_VLimit) done @@ -486,17 +486,17 @@ 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))" + "\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))))" + "\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)" + "\separation(M,P); M(z)\ \ \y[M]. \x[M]. x \ y \ x \ z & P(x)" by (simp add: separation_def) @@ -504,7 +504,7 @@ definition order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where - "order_isomorphism(M,A,r,B,s,f) == + "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]. @@ -513,12 +513,12 @@ definition pred_set :: "[i=>o,i,i,i,i] => o" where - "pred_set(M,A,x,r,B) == + "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) == + "membership(M,A,r) \ \p[M]. p \ r \ (\x[M]. x\A & (\y[M]. y\A & x\y & pair(M,x,y,p)))" @@ -527,7 +527,7 @@ text\The class M is assumed to be transitive and inhabited\ locale M_trans = fixes M - assumes transM: "[| y\x; M(x) |] ==> M(y)" + assumes transM: "\y\x; M(x)\ \ M(y)" and M_inhabited: "\x . M(x)" lemma (in M_trans) nonempty [simp]: "M(0)" @@ -562,15 +562,15 @@ 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))" + "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))" + "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))) \ + "M(A) \ (\x[M]. (x\A \ P(x))) \ (\x\A. P(x)) & (\x. P(x) \ M(x) \ x\A)" by (blast intro: transM) @@ -579,26 +579,26 @@ 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" + "\\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" + "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" + "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}" + "M(z) \ upair(M,a,b,z) \ z={a,b}" apply (simp add: upair_def) apply (blast intro: transM) done @@ -628,7 +628,7 @@ by blast lemma (in M_trans) pair_abs [simp]: - "M(z) ==> pair(M,a,b,z) \ z=" + "M(z) \ pair(M,a,b,z) \ z=" apply (simp add: pair_def Pair_def) apply (blast intro: transM) done @@ -646,11 +646,11 @@ by blast lemma (in M_trans) pair_components_in_M: - "[| \ A; M(A) |] ==> M(x) & M(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" + "\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) @@ -660,43 +660,43 @@ 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" + "\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" + "\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" + "\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)" + "\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))" + "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)" + "\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))" + "\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)" + "\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)" + "\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!]: @@ -716,7 +716,7 @@ subsubsection\Absoluteness for Separation and Replacement\ lemma (in M_trans) separation_closed [intro,simp]: - "[| separation(M,P); M(A) |] ==> M(Collect(A,P))" + "\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) @@ -728,7 +728,7 @@ 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)" + "\M(A); M(z)\ \ is_Collect(M,A,P,z) \ z = Collect(A,P)" unfolding is_Collect_def by (blast dest: transM) @@ -736,24 +736,24 @@ 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) \ + "\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))" + "\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))" + "\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") @@ -764,9 +764,9 @@ 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)" + "\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) @@ -783,8 +783,8 @@ 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))" + "\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 @@ -794,8 +794,8 @@ 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)))" + "\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) @@ -805,26 +805,26 @@ definition is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where - "is_lambda(M, A, is_b, z) == + "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))" + "\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))" + "\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)" + "\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) @@ -836,14 +836,14 @@ 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) \ + "\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" + "\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) @@ -862,7 +862,7 @@ real powerset.\ lemma (in M_trans) powerset_imp_subset_Pow: - "[| powerset(M,x,y); M(y) |] ==> y \ Pow(x)" + "\powerset(M,x,y); M(y)\ \ y \ Pow(x)" apply (simp add: powerset_def) apply (blast dest: transM) done @@ -892,23 +892,23 @@ subsubsection\Absoluteness for the Natural Numbers\ lemma (in M_trivial) nat_into_M [intro]: - "n \ nat ==> M(n)" + "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))" + "\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)" + "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)" + "\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) @@ -922,9 +922,9 @@ 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')" + "\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) @@ -932,43 +932,43 @@ text\These results constitute Theorem IV 5.1 of Kunen (page 126).\ lemma (in M_trans) lt_closed: - "[| j M(j)" + "\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)" + "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)" + "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)" + "M(a) \ limit_ordinal(M,a) \ Limit(a)" apply (unfold 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))" + "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" + "\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" + "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" + "\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) @@ -977,21 +977,21 @@ done lemma (in M_trivial) omega_abs [simp]: - "M(a) ==> omega(M,a) \ a = nat" + "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" + "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)" + "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))" + "M(a) \ number3(M,a) \ a = succ(succ(1))" by (simp add: number3_def) text\Kunen continued to 20...\ @@ -1011,48 +1011,48 @@ definition natnumber :: "[i=>o,i,i] => o" - "natnumber(M,n,x) == natnumber_aux(M,n)`x = 1" + "natnumber(M,n,x) \ natnumber_aux(M,n)`x = 1" lemma (in M_trivial) [simp]: - "natnumber(M,0,x) == x=0" + "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)" + "M(A) \ separation(M, \x. \y[M]. y\A \ x\y)" and Diff_separation: - "M(B) ==> separation(M, \x. x \ B)" + "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)))" + "\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)))" + "\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, + "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)))" + "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]. + "\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))" + "\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) ==> + "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, + "\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) & @@ -1060,9 +1060,9 @@ 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}}}" + "\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) @@ -1073,8 +1073,8 @@ done lemma (in M_basic) cartprod_iff: - "[| M(A); M(B); M(C) |] - ==> cartprod(M,A,B,C) \ + "\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 = })" apply (simp add: Pair_def cartprod_def, safe) @@ -1091,7 +1091,7 @@ done lemma (in M_basic) cartprod_closed_lemma: - "[| M(A); M(B) |] ==> \C[M]. cartprod(M,A,B,C)" + "\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) @@ -1104,15 +1104,15 @@ 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)" + "\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)" + "\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)" + "\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]: @@ -1120,7 +1120,7 @@ by (simp add: Inl_def) lemma (in M_trivial) Inl_abs [simp]: - "M(Z) ==> is_Inl(M,a,Z) \ (Z = Inl(a))" + "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]: @@ -1128,14 +1128,14 @@ by (simp add: Inr_def) lemma (in M_trivial) Inr_abs [simp]: - "M(Z) ==> is_Inr(M,a,Z) \ (Z = Inr(a))" + "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) ==> + "M(r) \ converse(r) = {z \ \(\(r)) * \(\(r)). \p\r. \x[M]. \y[M]. p = \x,y\ & z = \y,x\}" @@ -1146,13 +1146,13 @@ done lemma (in M_basic) converse_closed [intro,simp]: - "M(r) ==> M(converse(r))" + "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)" + "\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 @@ -1165,106 +1165,106 @@ subsubsection \image, preimage, domain, range\ lemma (in M_basic) image_closed [intro,simp]: - "[| M(A); M(r) |] ==> M(r``A)" + "\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" + "\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)" + "\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)" + "\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))" + "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)" + "\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))" + "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)" + "\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))" + "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)" + "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)" + "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)" + "\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" + "\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" + "\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)" + "\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)" + "\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)" + "\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 = + "\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) @@ -1275,13 +1275,13 @@ done lemma (in M_basic) comp_closed [intro,simp]: - "[| M(r); M(s) |] ==> M(r O s)" + "\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" + "\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 @@ -1295,54 +1295,54 @@ 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)" + "\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)" + "\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\}" + "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))" + "\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)" + "\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))" + "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)" + "\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)" + "\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))" + "\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]) @@ -1352,37 +1352,37 @@ by blast lemma Diff_Collect_eq: - "A - Collect(A,P) = Collect(A, %x. ~ P(x))" + "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)) = + "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))" + "\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))" + "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))" + "\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))" + "\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) @@ -1394,7 +1394,7 @@ 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" + "\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 @@ -1403,7 +1403,7 @@ done lemma (in M_basic) succ_fun_eq2: - "[|M(B); M(n->B)|] ==> + "\M(B); M(n->B)\ \ succ(n) -> B = \{z. p \ (n->B)*B, \f[M]. \b[M]. p = & z = {cons(, f)}}" apply (simp add: succ_fun_eq) @@ -1411,7 +1411,7 @@ done lemma (in M_basic) funspace_succ: - "[|M(n); M(B); M(n->B) |] ==> M(succ(n) -> B)" + "\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 @@ -1420,7 +1420,7 @@ 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)" + "\n\nat; M(B)\ \ M(n->B)" apply (induct_tac n, simp) apply (simp add: funspace_succ nat_into_M) done @@ -1430,38 +1430,38 @@ 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))" + "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))" + "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))" + "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)" + "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)" + "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)" + "\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" + "\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" + "\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) @@ -1470,15 +1470,15 @@ by (simp add: bool_of_o_def) lemma (in M_trivial) and_closed [intro,simp]: - "[| M(p); M(q) |] ==> M(p and q)" + "\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)" + "\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))" + "M(p) \ M(not(p))" by (simp add: Bool.not_def cond_def) @@ -1487,46 +1487,46 @@ 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)" + "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)" + "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)" +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))" + "\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))" + "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))" + "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) == + "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_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))" @@ -1534,28 +1534,28 @@ definition hd' :: "i => i" where \ \A version of \<^term>\hd\ that's always defined.\ - "hd'(xs) == if quasilist(xs) then hd(xs) else 0" + "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" + "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_hd(M,xs,H) \ (is_Nil(M,xs) \ empty(M,H)) & - (\x[M]. \l[M]. ~ is_Cons(M,x,l,xs) | H=x) & + (\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_tl(M,xs,T) \ (is_Nil(M,xs) \ T=xs) & - (\x[M]. \l[M]. ~ is_Cons(M,x,l,xs) | T=l) & + (\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'\\ @@ -1566,7 +1566,7 @@ lemma [iff]: "quasilist(Cons(x,l))" by (simp add: quasilist_def) -lemma list_imp_quasilist: "l \ list(A) ==> quasilist(l)" +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\\ @@ -1577,27 +1577,27 @@ 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" +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)" + "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))" + "\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)" + "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)" + "\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) @@ -1614,11 +1614,11 @@ 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" + "\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)" + "\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) @@ -1630,11 +1630,11 @@ 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" + "\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)" + "\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) @@ -1657,12 +1657,12 @@ lemma tl'_Cons: "tl'(Cons(a,l)) = l" by (simp add: tl'_def) -lemma iterates_tl_Nil: "n \ nat ==> tl'^n ([]) = []" +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))" +lemma (in M_basic) tl'_closed: "M(x) \ M(tl'(x))" apply (simp add: tl'_def) apply (force simp add: quasilist_def) done diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Satisfies_absolute.thy --- a/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,14 +11,14 @@ subsubsection\The Formula \<^term>\is_depth\, Internalized\ -(* "is_depth(M,p,n) == +(* "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) == + "depth_fm(p,n) \ Exists(Exists(Exists( And(formula_N_fm(n#+3,1), And(Neg(Member(p#+3,1)), @@ -26,21 +26,21 @@ And(formula_N_fm(2,0), Member(p#+3,0))))))))" lemma depth_fm_type [TC]: - "[| x \ nat; y \ nat |] ==> depth_fm(x,y) \ formula" + "\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) \ + "\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)" + "\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: @@ -59,7 +59,7 @@ (* 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) == + "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 \ @@ -68,7 +68,7 @@ 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) == + "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), @@ -87,31 +87,31 @@ 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" + "\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))))" + "\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))))" + "\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))))" + "\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)))" + "\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) \ + "\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 @@ -121,25 +121,25 @@ 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))))" + "\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))))" + "\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))))" + "\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)))" + "\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) \ + "\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]) @@ -150,16 +150,16 @@ 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)), + "\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)), + "\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)), + "\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)), + "\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))]" @@ -177,14 +177,14 @@ 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) == + "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" + "\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) @@ -196,12 +196,12 @@ \<^term>\c\, \<^term>\d\, etc., of the locale \Formula_Rec\.\ definition satisfies_a :: "[i,i,i]=>i" where - "satisfies_a(A) == + "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) == + "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, @@ -211,14 +211,14 @@ definition satisfies_b :: "[i,i,i]=>i" where - "satisfies_b(A) == + "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) == + "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, @@ -227,11 +227,11 @@ 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)" + "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) == + "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)) & @@ -242,11 +242,11 @@ 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)" + \ \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) == + "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) & @@ -261,7 +261,7 @@ 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 == + "satisfies_MH \ \M A u f z. \fml[M]. is_formula(M,fml) \ is_lambda (M, fml, @@ -272,7 +272,7 @@ definition is_satisfies :: "[i=>o,i,i,i]=>o" where - "is_satisfies(M,A) == is_formula_rec (M, satisfies_MH(M,A))" + "is_satisfies(M,A) \ is_formula_rec (M, satisfies_MH(M,A))" text\This lemma relates the fragments defined above to the original primitive @@ -291,32 +291,32 @@ locale M_satisfies = M_eclose + assumes Member_replacement: - "[|M(A); x \ nat; y \ nat|] - ==> strong_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(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(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(A); M(rp)\ + \ strong_replacement (M, \env z. \bo[M]. env \ list(A) & is_bool_of_o (M, @@ -328,11 +328,11 @@ and formula_rec_replacement: \ \For the \<^term>\transrec\\ - "[|n \ nat; M(A)|] ==> transrec_replacement(M, satisfies_MH(M,A), n)" + "\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)|] ==> + "\M(g); M(A)\ \ strong_replacement (M, \x y. mem_formula(M,x) & (\c[M]. is_formula_case(M, satisfies_is_a(M,A), @@ -343,60 +343,60 @@ lemma (in M_satisfies) Member_replacement': - "[|M(A); x \ nat; y \ nat|] - ==> strong_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(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(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(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))" + "\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))" + "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))" + "\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))" + "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))" + "\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') @@ -404,7 +404,7 @@ done lemma (in M_satisfies) c_rel: - "[|M(A); M(f)|] ==> + "\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, @@ -415,7 +415,7 @@ done lemma (in M_satisfies) d_closed: - "[|M(A); x \ formula; M(rx)|] ==> M(satisfies_d(A,x,rx))" + "\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') @@ -423,7 +423,7 @@ done lemma (in M_satisfies) d_rel: - "[|M(A); M(f)|] ==> + "\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 @@ -433,11 +433,11 @@ lemma (in M_satisfies) fr_replace: - "[|n \ nat; M(A)|] ==> transrec_replacement(M,satisfies_MH(M,A),n)" + "\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(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), @@ -446,7 +446,7 @@ by (blast intro: a_closed b_closed c_closed d_closed) lemma (in M_satisfies) fr_lam_replace: - "[|M(g); M(A)|] ==> + "\M(g); M(A)\ \ strong_replacement (M, \x y. x \ formula & y = \x, formula_rec_case(satisfies_a(A), @@ -464,7 +464,7 @@ Function \<^term>\satisfies\\ lemma (in M_satisfies) Formula_Rec_axioms_M: - "M(A) ==> + "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), @@ -478,7 +478,7 @@ theorem (in M_satisfies) Formula_Rec_M: - "M(A) ==> + "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), @@ -494,13 +494,13 @@ lemma (in M_satisfies) satisfies_closed: - "[|M(A); p \ formula|] ==> M(satisfies(A,p))" + "\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)" + "\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) @@ -509,14 +509,14 @@ subsubsection\The Operator \<^term>\is_depth_apply\, Internalized\ -(* is_depth_apply(M,h,p,z) == +(* 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) == + "depth_apply_fm(h,p,z) \ Exists(Exists(Exists( And(finite_ordinal_fm(2), And(depth_fm(p#+3,2), @@ -524,19 +524,19 @@ 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" + "\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) \ + "\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)" + "\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: @@ -550,7 +550,7 @@ subsubsection\The Operator \<^term>\satisfies_is_a\, Internalized\ -(* satisfies_is_a(M,A) == +(* 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, @@ -560,7 +560,7 @@ definition satisfies_is_a_fm :: "[i,i,i,i]=>i" where - "satisfies_is_a_fm(A,x,y,z) == + "satisfies_is_a_fm(A,x,y,z) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( @@ -571,13 +571,13 @@ 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" + "\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) \ + "\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) @@ -586,9 +586,9 @@ 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) \ + "\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 @@ -603,7 +603,7 @@ subsubsection\The Operator \<^term>\satisfies_is_b\, Internalized\ -(* satisfies_is_b(M,A) == +(* 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, @@ -612,7 +612,7 @@ definition satisfies_is_b_fm :: "[i,i,i,i]=>i" where - "satisfies_is_b_fm(A,x,y,z) == + "satisfies_is_b_fm(A,x,y,z) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( @@ -620,13 +620,13 @@ 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" + "\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) \ + "\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) @@ -635,9 +635,9 @@ 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) \ + "\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 @@ -652,7 +652,7 @@ subsubsection\The Operator \<^term>\satisfies_is_c\, Internalized\ -(* satisfies_is_c(M,A,h) == +(* 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)) & @@ -662,7 +662,7 @@ definition satisfies_is_c_fm :: "[i,i,i,i,i]=>i" where - "satisfies_is_c_fm(A,h,p,q,zz) == + "satisfies_is_c_fm(A,h,p,q,zz) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( @@ -673,22 +673,22 @@ 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" + "\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) \ + "\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(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) \ + 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 @@ -703,7 +703,7 @@ subsubsection\The Operator \<^term>\satisfies_is_d\, Internalized\ -(* satisfies_is_d(M,A,h) == +(* 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) & @@ -716,7 +716,7 @@ definition satisfies_is_d_fm :: "[i,i,i,i]=>i" where - "satisfies_is_d_fm(A,h,p,zz) == + "satisfies_is_d_fm(A,h,p,zz) \ Forall( Implies(is_list_fm(succ(A),0), lambda_fm( @@ -730,21 +730,21 @@ 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" + "\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) \ + "\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) \ + "\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 @@ -760,7 +760,7 @@ subsubsection\The Operator \<^term>\satisfies_MH\, Internalized\ -(* satisfies_MH == +(* satisfies_MH \ \M A u f zz. \fml[M]. is_formula(M,fml) \ is_lambda (M, fml, @@ -771,7 +771,7 @@ definition satisfies_MH_fm :: "[i,i,i,i]=>i" where - "satisfies_MH_fm(A,u,f,zz) == + "satisfies_MH_fm(A,u,f,zz) \ Forall( Implies(is_formula_fm(0), lambda_fm( @@ -783,21 +783,21 @@ 0, succ(zz))))" lemma satisfies_MH_type [TC]: - "[| A \ nat; u \ nat; x \ nat; z \ nat |] - ==> satisfies_MH_fm(A,u,x,z) \ formula" + "\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) \ + "\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) \ + "\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 @@ -833,8 +833,8 @@ lemma Member_replacement: - "[|L(A); x \ nat; y \ nat|] - ==> strong_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) & @@ -863,8 +863,8 @@ lemma Equal_replacement: - "[|L(A); x \ nat; y \ nat|] - ==> strong_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) & @@ -895,8 +895,8 @@ done lemma Nand_replacement: - "[|L(A); L(rp); L(rq)|] - ==> strong_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) & @@ -930,8 +930,8 @@ done lemma Forall_replacement: - "[|L(A); L(rp)|] - ==> strong_replacement + "\L(A); L(rp)\ + \ strong_replacement (L, \env z. \bo[L]. env \ list(A) & is_bool_of_o (L, @@ -960,7 +960,7 @@ lemma formula_rec_replacement: \ \For the \<^term>\transrec\\ - "[|n \ nat; L(A)|] ==> transrec_replacement(L, satisfies_MH(L,A), n)" + "\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}))}" @@ -996,7 +996,7 @@ lemma formula_rec_lambda_replacement: \ \For the \<^term>\transrec\\ - "[|L(g); L(A)|] ==> + "\L(g); L(A)\ \ strong_replacement (L, \x y. mem_formula(L,x) & (\c[L]. is_formula_case(L, satisfies_is_a(L,A), diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Separation.thy --- a/src/ZF/Constructible/Separation.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Separation.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,7 +9,7 @@ text\This theory proves all instances needed for locale \M_basic\\ text\Helps us solve for de Bruijn indices!\ -lemma nth_ConsI: "[|nth(n,l) = x; n \ nat|] ==> nth(succ(n), Cons(a,l)) = x" +lemma nth_ConsI: "\nth(n,l) = x; n \ nat\ \ nth(succ(n), Cons(a,l)) = x" by simp lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI @@ -17,20 +17,20 @@ fun_plus_iff_sats lemma Collect_conj_in_DPow: - "[| {x\A. P(x)} \ DPow(A); {x\A. Q(x)} \ DPow(A) |] - ==> {x\A. P(x) & Q(x)} \ DPow(A)" + "\{x\A. P(x)} \ DPow(A); {x\A. Q(x)} \ DPow(A)\ + \ {x\A. P(x) & Q(x)} \ DPow(A)" by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) lemma Collect_conj_in_DPow_Lset: - "[|z \ Lset(j); {x \ Lset(j). P(x)} \ DPow(Lset(j))|] - ==> {x \ Lset(j). x \ z & P(x)} \ DPow(Lset(j))" + "\z \ Lset(j); {x \ Lset(j). P(x)} \ DPow(Lset(j))\ + \ {x \ Lset(j). x \ z & P(x)} \ DPow(Lset(j))" apply (frule mem_Lset_imp_subset_Lset) apply (simp add: Collect_conj_in_DPow Collect_mem_eq subset_Int_iff2 elem_subset_in_DPow) done lemma separation_CollectI: - "(\z. L(z) ==> L({x \ z . P(x)})) ==> separation(L, \x. P(x))" + "(\z. L(z) \ L({x \ z . P(x)})) \ separation(L, \x. P(x))" apply (unfold separation_def, clarify) apply (rule_tac x="{x\z. P(x)}" in rexI) apply simp_all @@ -38,9 +38,9 @@ text\Reduces the original comprehension to the reflected one\ lemma reflection_imp_L_separation: - "[| \x\Lset(j). P(x) \ Q(x); + "\\x\Lset(j). P(x) \ Q(x); {x \ Lset(j) . Q(x)} \ DPow(Lset(j)); - Ord(j); z \ Lset(j)|] ==> L({x \ z . P(x)})" + Ord(j); z \ Lset(j)\ \ L({x \ z . P(x)})" apply (rule_tac i = "succ(j)" in L_I) prefer 2 apply simp apply (subgoal_tac "{x \ z. P(x)} = {x \ Lset(j). x \ z & (Q(x))}") @@ -54,7 +54,7 @@ lemma gen_separation: assumes reflection: "REFLECTS [P,Q]" and Lu: "L(u)" - and collI: "!!j. u \ Lset(j) + and collI: "\j. u \ Lset(j) \ Collect(Lset(j), Q(j)) \ DPow(Lset(j))" shows "separation(L,P)" apply (rule separation_CollectI) @@ -73,7 +73,7 @@ lemma gen_separation_multi: assumes reflection: "REFLECTS [P,Q]" and Lu: "L(u)" - and collI: "!!j. u \ Lset(j) + and collI: "\j. u \ Lset(j) \ Collect(Lset(j), Q(j)) \ DPow(Lset(j))" shows "separation(L,P)" apply (rule gen_separation [OF reflection Lu]) @@ -90,7 +90,7 @@ by (intro FOL_reflections) lemma Inter_separation: - "L(A) ==> separation(L, \x. \y[L]. y\A \ x\y)" + "L(A) \ separation(L, \x. \y[L]. y\A \ x\y)" apply (rule gen_separation [OF Inter_Reflects], simp) apply (rule DPow_LsetI) txt\I leave this one example of a manual proof. The tedium of manually @@ -109,7 +109,7 @@ by (intro FOL_reflections) lemma Diff_separation: - "L(B) ==> separation(L, \x. x \ B)" + "L(B) \ separation(L, \x. x \ B)" apply (rule gen_separation [OF Diff_Reflects], simp) apply (rule_tac env="[B]" in DPow_LsetI) apply (rule sep_rules | simp)+ @@ -124,8 +124,8 @@ by (intro FOL_reflections function_reflections) lemma cartprod_separation: - "[| L(A); L(B) |] - ==> separation(L, \z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)))" + "\L(A); L(B)\ + \ separation(L, \z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)))" apply (rule gen_separation_multi [OF cartprod_Reflects, of "{A,B}"], auto) apply (rule_tac env="[A,B]" in DPow_LsetI) apply (rule sep_rules | simp)+ @@ -139,8 +139,8 @@ by (intro FOL_reflections function_reflections) lemma image_separation: - "[| L(A); L(r) |] - ==> separation(L, \y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)))" + "\L(A); L(r)\ + \ separation(L, \y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)))" apply (rule gen_separation_multi [OF image_Reflects, of "{A,r}"], auto) apply (rule_tac env="[A,r]" in DPow_LsetI) apply (rule sep_rules | simp)+ @@ -156,7 +156,7 @@ by (intro FOL_reflections function_reflections) lemma converse_separation: - "L(r) ==> separation(L, + "L(r) \ separation(L, \z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)))" apply (rule gen_separation [OF converse_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) @@ -172,7 +172,7 @@ by (intro FOL_reflections function_reflections) lemma restrict_separation: - "L(A) ==> separation(L, \z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)))" + "L(A) \ separation(L, \z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)))" apply (rule gen_separation [OF restrict_Reflects], simp) apply (rule_tac env="[A]" in DPow_LsetI) apply (rule sep_rules | simp)+ @@ -191,8 +191,8 @@ by (intro FOL_reflections function_reflections) lemma comp_separation: - "[| L(r); L(s) |] - ==> separation(L, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. + "\L(r); L(s)\ + \ separation(L, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & xy\s & yz\r)" apply (rule gen_separation_multi [OF comp_Reflects, of "{r,s}"], auto) @@ -213,7 +213,7 @@ by (intro FOL_reflections function_reflections) lemma pred_separation: - "[| L(r); L(x) |] ==> separation(L, \y. \p[L]. p\r & pair(L,y,x,p))" + "\L(r); L(x)\ \ separation(L, \y. \p[L]. p\r & pair(L,y,x,p))" apply (rule gen_separation_multi [OF pred_Reflects, of "{r,x}"], auto) apply (rule_tac env="[r,x]" in DPow_LsetI) apply (rule sep_rules | simp)+ @@ -248,7 +248,7 @@ by (intro FOL_reflections function_reflections) lemma funspace_succ_replacement: - "L(n) ==> + "L(n) \ strong_replacement(L, \p z. \f[L]. \b[L]. \nb[L]. \cnbf[L]. pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & upair(L,cnbf,cnbf,z))" @@ -275,8 +275,8 @@ lemma is_recfun_separation: \ \for well-founded recursion\ - "[| L(r); L(f); L(g); L(a); L(b) |] - ==> separation(L, + "\L(r); L(f); L(g); L(a); L(b)\ + \ separation(L, \x. \xa[L]. \xb[L]. pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/WF_absolute.thy --- a/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,14 +10,14 @@ definition rtrancl_alt :: "[i,i]=>i" where - "rtrancl_alt(A,r) == + "rtrancl_alt(A,r) \ {p \ A*A. \n\nat. \f \ succ(n) -> A. (\x y. p = & f`0 = x & f`n = y) & (\i\n. \ r)}" lemma alt_rtrancl_lemma1 [rule_format]: "n \ nat - ==> \f \ succ(n) -> field(r). + \ \f \ succ(n) -> field(r). (\i\n. \f`i, f ` succ(i)\ \ r) \ \f`0, f`n\ \ r^*" apply (induct_tac n) apply (simp_all add: apply_funtype rtrancl_refl, clarify) @@ -61,7 +61,7 @@ definition rtran_closure_mem :: "[i=>o,i,i,i] => o" where \ \The property of belonging to \rtran_closure(r)\\ - "rtran_closure_mem(M,A,r,p) == + "rtran_closure_mem(M,A,r,p) \ \nnat[M]. \n[M]. \n'[M]. omega(M,nnat) & n\nnat & successor(M,n,n') & (\f[M]. typed_function(M,n',A,f) & @@ -74,28 +74,28 @@ definition rtran_closure :: "[i=>o,i,i] => o" where - "rtran_closure(M,r,s) == + "rtran_closure(M,r,s) \ \A[M]. is_field(M,r,A) \ (\p[M]. p \ s \ rtran_closure_mem(M,A,r,p))" definition tran_closure :: "[i=>o,i,i] => o" where - "tran_closure(M,r,t) == + "tran_closure(M,r,t) \ \s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" locale M_trancl = M_basic + assumes rtrancl_separation: - "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))" + "\M(r); M(A)\ \ separation (M, rtran_closure_mem(M,A,r))" and wellfounded_trancl_separation: - "[| M(r); M(Z) |] ==> + "\M(r); M(Z)\ \ separation (M, \x. \w[M]. \wx[M]. \rp[M]. w \ Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \ rp)" and M_nat [iff] : "M(nat)" lemma (in M_trancl) rtran_closure_mem_iff: - "[|M(A); M(r); M(p)|] - ==> rtran_closure_mem(M,A,r,p) \ + "\M(A); M(r); M(p)\ + \ rtran_closure_mem(M,A,r,p) \ (\n[M]. n\nat & (\f[M]. f \ succ(n) -> A & (\x[M]. \y[M]. p = & f`0 = x & f`n = y) & @@ -104,21 +104,21 @@ done lemma (in M_trancl) rtran_closure_rtrancl: - "M(r) ==> rtran_closure(M,r,rtrancl(r))" + "M(r) \ rtran_closure(M,r,rtrancl(r))" apply (simp add: rtran_closure_def rtran_closure_mem_iff rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def) apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) done lemma (in M_trancl) rtrancl_closed [intro,simp]: - "M(r) ==> M(rtrancl(r))" + "M(r) \ M(rtrancl(r))" apply (insert rtrancl_separation [of r "field(r)"]) apply (simp add: rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def rtran_closure_mem_iff) done lemma (in M_trancl) rtrancl_abs [simp]: - "[| M(r); M(z) |] ==> rtran_closure(M,r,z) \ z = rtrancl(r)" + "\M(r); M(z)\ \ rtran_closure(M,r,z) \ z = rtrancl(r)" apply (rule iffI) txt\Proving the right-to-left implication\ prefer 2 apply (blast intro: rtran_closure_rtrancl) @@ -129,20 +129,20 @@ done lemma (in M_trancl) trancl_closed [intro,simp]: - "M(r) ==> M(trancl(r))" + "M(r) \ M(trancl(r))" by (simp add: trancl_def) lemma (in M_trancl) trancl_abs [simp]: - "[| M(r); M(z) |] ==> tran_closure(M,r,z) \ z = trancl(r)" + "\M(r); M(z)\ \ tran_closure(M,r,z) \ z = trancl(r)" by (simp add: tran_closure_def trancl_def) lemma (in M_trancl) wellfounded_trancl_separation': - "[| M(r); M(Z) |] ==> separation (M, \x. \w[M]. w \ Z & \ r^+)" + "\M(r); M(Z)\ \ separation (M, \x. \w[M]. w \ Z & \ r^+)" by (insert wellfounded_trancl_separation [of r Z], simp) text\Alternative proof of \wf_on_trancl\; inspiration for the relativized version. Original version is on theory WF.\ -lemma "[| wf[A](r); r-``A \ A |] ==> wf[A](r^+)" +lemma "\wf[A](r); r-``A \ A\ \ wf[A](r^+)" apply (simp add: wf_on_def wf_def) apply (safe) apply (drule_tac x = "{x\A. \w. \w,x\ \ r^+ & w \ Z}" in spec) @@ -150,8 +150,8 @@ done lemma (in M_trancl) wellfounded_on_trancl: - "[| wellfounded_on(M,A,r); r-``A \ A; M(r); M(A) |] - ==> wellfounded_on(M,A,r^+)" + "\wellfounded_on(M,A,r); r-``A \ A; M(r); M(A)\ + \ wellfounded_on(M,A,r^+)" apply (simp add: wellfounded_on_def) apply (safe intro!: equalityI) apply (rename_tac Z x) @@ -168,7 +168,7 @@ done lemma (in M_trancl) wellfounded_trancl: - "[|wellfounded(M,r); M(r)|] ==> wellfounded(M,r^+)" + "\wellfounded(M,r); M(r)\ \ wellfounded(M,r^+)" apply (simp add: wellfounded_iff_wellfounded_on_field) apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl) apply blast @@ -181,13 +181,13 @@ (*first use is_recfun, then M_is_recfun*) lemma (in M_trancl) wfrec_relativize: - "[|wf(r); M(a); M(r); + "\wf(r); M(a); M(r); strong_replacement(M, \x z. \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r^+, x, \x f. H(x, restrict(f, r -`` {x})), g) & y = H(x, restrict(g, r -`` {x}))); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> wfrec(r,a,H) = z \ + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ wfrec(r,a,H) = z \ (\f[M]. is_recfun(r^+, a, \x f. H(x, restrict(f, r -`` {x})), f) & z = H(a,restrict(f,r-``{a})))" apply (frule wf_trancl) @@ -204,10 +204,10 @@ The premise \<^term>\relation(r)\ is necessary before we can replace \<^term>\r^+\ by \<^term>\r\.\ theorem (in M_trancl) trans_wfrec_relativize: - "[|wf(r); trans(r); relation(r); M(r); M(a); + "\wf(r); trans(r); relation(r); M(r); M(a); wfrec_replacement(M,MH,r); relation2(M,MH,H); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> wfrec(r,a,H) = z \ (\f[M]. is_recfun(r,a,H,f) & z = H(a,f))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ wfrec(r,a,H) = z \ (\f[M]. is_recfun(r,a,H,f) & z = H(a,f))" apply (frule wfrec_replacement', assumption+) apply (simp cong: is_recfun_cong add: wfrec_relativize trancl_eq_r @@ -215,18 +215,18 @@ done theorem (in M_trancl) trans_wfrec_abs: - "[|wf(r); trans(r); relation(r); M(r); M(a); M(z); + "\wf(r); trans(r); relation(r); M(r); M(a); M(z); wfrec_replacement(M,MH,r); relation2(M,MH,H); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> is_wfrec(M,MH,r,a,z) \ z=wfrec(r,a,H)" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ is_wfrec(M,MH,r,a,z) \ z=wfrec(r,a,H)" by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) lemma (in M_trancl) trans_eq_pair_wfrec_iff: - "[|wf(r); trans(r); relation(r); M(r); M(y); + "\wf(r); trans(r); relation(r); M(r); M(y); wfrec_replacement(M,MH,r); relation2(M,MH,H); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> y = \ + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ y = \ (\f[M]. is_recfun(r,x,H,f) & y = )" apply safe apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) @@ -240,10 +240,10 @@ text\Lemma with the awkward premise mentioning \wfrec\.\ lemma (in M_trancl) wfrec_closed_lemma [rule_format]: - "[|wf(r); M(r); + "\wf(r); M(r); strong_replacement(M, \x y. y = \x, wfrec(r, x, H)\); - \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> M(a) \ M(wfrec(r,a,H))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ M(a) \ M(wfrec(r,a,H))" apply (rule_tac a=a in wf_induct, assumption+) apply (subst wfrec, assumption, clarify) apply (drule_tac x1=x and x="\x\r -`` {x}. wfrec(r, x, H)" @@ -264,10 +264,10 @@ text\Useful version for transitive relations\ theorem (in M_trancl) trans_wfrec_closed: - "[|wf(r); trans(r); relation(r); M(r); M(a); + "\wf(r); trans(r); relation(r); M(r); M(a); wfrec_replacement(M,MH,r); relation2(M,MH,H); - \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> M(wfrec(r,a,H))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ M(wfrec(r,a,H))" apply (frule wfrec_replacement', assumption+) apply (frule wfrec_replacement_iff [THEN iffD1]) apply (rule wfrec_closed_lemma, assumption+) @@ -276,13 +276,13 @@ subsection\Absoluteness without assuming transitivity\ lemma (in M_trancl) eq_pair_wfrec_iff: - "[|wf(r); M(r); M(y); + "\wf(r); M(r); M(y); strong_replacement(M, \x z. \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r^+, x, \x f. H(x, restrict(f, r -`` {x})), g) & y = H(x, restrict(g, r -`` {x}))); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> y = \ + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ y = \ (\f[M]. is_recfun(r^+, x, \x f. H(x, restrict(f, r -`` {x})), f) & y = )" apply safe @@ -294,11 +294,11 @@ text\Full version not assuming transitivity, but maybe not very useful.\ theorem (in M_trancl) wfrec_closed: - "[|wf(r); M(r); M(a); + "\wf(r); M(r); M(a); wfrec_replacement(M,MH,r^+); relation2(M,MH, \x f. H(x, restrict(f, r -`` {x}))); - \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> M(wfrec(r,a,H))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ M(wfrec(r,a,H))" apply (frule wfrec_replacement' [of MH "r^+" "\x f. H(x, restrict(f, r -`` {x}))"]) prefer 4 diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/WFrec.thy Tue Sep 27 16:51:35 2022 +0100 @@ -12,7 +12,7 @@ (*Many of these might be useful in WF.thy*) lemma apply_recfun2: - "[| is_recfun(r,a,H,f); :f |] ==> i = H(x, restrict(f,r-``{x}))" + "\is_recfun(r,a,H,f); :f\ \ i = H(x, restrict(f,r-``{x}))" apply (frule apply_recfun) apply (blast dest: is_recfun_type fun_is_rel) apply (simp add: function_apply_equality [OF _ is_recfun_imp_function]) @@ -32,24 +32,24 @@ apply (fast intro: lam_type, simp) done -lemma is_recfun_imp_in_r: "[|is_recfun(r,a,H,f); \x,i\ \ f|] ==> \x, a\ \ r" +lemma is_recfun_imp_in_r: "\is_recfun(r,a,H,f); \x,i\ \ f\ \ \x, a\ \ r" by (blast dest: is_recfun_type fun_is_rel) lemma trans_Int_eq: - "[| trans(r); \ r |] ==> r -`` {x} \ r -`` {y} = r -`` {y}" + "\trans(r); \ r\ \ r -`` {x} \ r -`` {y} = r -`` {y}" by (blast intro: transD) lemma is_recfun_restrict_idem: - "is_recfun(r,a,H,f) ==> restrict(f, r -`` {a}) = f" + "is_recfun(r,a,H,f) \ restrict(f, r -`` {a}) = f" apply (drule is_recfun_type) apply (auto simp add: Pi_iff subset_Sigma_imp_relation restrict_idem) done lemma is_recfun_cong_lemma: - "[| is_recfun(r,a,H,f); r = r'; a = a'; f = f'; - !!x g. [| \ r'; relation(g); domain(g) \ r' -``{x} |] - ==> H(x,g) = H'(x,g) |] - ==> is_recfun(r',a',H',f')" + "\is_recfun(r,a,H,f); r = r'; a = a'; f = f'; + \x g. \ \ r'; relation(g); domain(g) \ r' -``{x}\ + \ H(x,g) = H'(x,g)\ + \ is_recfun(r',a',H',f')" apply (simp add: is_recfun_def) apply (erule trans) apply (rule lam_cong) @@ -59,10 +59,10 @@ text\For \is_recfun\ we need only pay attention to functions whose domains are initial segments of \<^term>\r\.\ lemma is_recfun_cong: - "[| r = r'; a = a'; f = f'; - !!x g. [| \ r'; relation(g); domain(g) \ r' -``{x} |] - ==> H(x,g) = H'(x,g) |] - ==> is_recfun(r,a,H,f) \ is_recfun(r',a',H',f')" + "\r = r'; a = a'; f = f'; + \x g. \ \ r'; relation(g); domain(g) \ r' -``{x}\ + \ H(x,g) = H'(x,g)\ + \ is_recfun(r,a,H,f) \ is_recfun(r',a',H',f')" apply (rule iffI) txt\Messy: fast and blast don't work for some reason\ apply (erule is_recfun_cong_lemma, auto) @@ -73,9 +73,9 @@ subsection\Reworking of the Recursion Theory Within \<^term>\M\\ lemma (in M_basic) is_recfun_separation': - "[| f \ r -`` {a} \ range(f); g \ r -`` {b} \ range(g); - M(r); M(f); M(g); M(a); M(b) |] - ==> separation(M, \x. \ (\x, a\ \ r \ \x, b\ \ r \ f ` x = g ` x))" + "\f \ r -`` {a} \ range(f); g \ r -`` {b} \ range(g); + M(r); M(f); M(g); M(a); M(b)\ + \ separation(M, \x. \ (\x, a\ \ r \ \x, b\ \ r \ f ` x = g ` x))" apply (insert is_recfun_separation [of r f g a b]) apply (simp add: vimage_singleton_iff) done @@ -88,10 +88,10 @@ but without them we'd have to undertake more work to set up the induction formula.\ lemma (in M_basic) is_recfun_equal [rule_format]: - "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); + "\is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); - M(f); M(g); M(r); M(x); M(a); M(b) |] - ==> \ r \ \ r \ f`x=g`x" + M(f); M(g); M(r); M(x); M(a); M(b)\ + \ \ r \ \ 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) @@ -111,10 +111,10 @@ done lemma (in M_basic) is_recfun_cut: - "[|is_recfun(r,a,H,f); is_recfun(r,b,H,g); + "\is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); - M(f); M(g); M(r); \ r |] - ==> restrict(f, r-``{b}) = g" + M(f); M(g); M(r); \ r\ + \ restrict(f, r-``{b}) = g" apply (frule_tac f=f in is_recfun_type) apply (rule fun_extension) apply (blast intro: transD restrict_type2) @@ -123,8 +123,8 @@ done lemma (in M_basic) is_recfun_functional: - "[|is_recfun(r,a,H,f); is_recfun(r,a,H,g); - wellfounded(M,r); trans(r); M(f); M(g); M(r) |] ==> f=g" + "\is_recfun(r,a,H,f); is_recfun(r,a,H,g); + wellfounded(M,r); trans(r); M(f); M(g); M(r)\ \ f=g" apply (rule fun_extension) apply (erule is_recfun_type)+ apply (blast intro!: is_recfun_equal dest: transM) @@ -132,8 +132,8 @@ text\Tells us that \is_recfun\ can (in principle) be relativized.\ lemma (in M_basic) is_recfun_relativize: - "[| M(r); M(f); \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> is_recfun(r,a,H,f) \ + "\M(r); M(f); \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ is_recfun(r,a,H,f) \ (\z[M]. z \ f \ (\x[M]. \ r & z = ))" apply (simp add: is_recfun_def lam_def) @@ -153,10 +153,10 @@ done lemma (in M_basic) is_recfun_restrict: - "[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \y,x\ \ r; + "\wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \y,x\ \ r; M(r); M(f); - \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> is_recfun(r, y, H, restrict(f, r -`` {y}))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ is_recfun(r, y, H, restrict(f, r -`` {y}))" apply (frule pair_components_in_M, assumption, clarify) apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff trans_Int_eq) @@ -170,14 +170,14 @@ done lemma (in M_basic) restrict_Y_lemma: - "[| wellfounded(M,r); trans(r); M(r); + "\wellfounded(M,r); trans(r); M(r); \x[M]. \g[M]. function(g) \ M(H(x,g)); M(Y); \b[M]. b \ Y \ (\x[M]. \ r & (\y[M]. b = \x,y\ & (\g[M]. is_recfun(r,x,H,g) \ y = H(x,g)))); - \x,a1\ \ r; is_recfun(r,x,H,f); M(f) |] - ==> restrict(Y, r -`` {x}) = f" + \x,a1\ \ r; is_recfun(r,x,H,f); M(f)\ + \ restrict(Y, r -`` {x}) = f" apply (subgoal_tac "\y \ r-``{x}. \z. :Y \ :f") apply (simp (no_asm_simp) add: restrict_def) apply (thin_tac "rall(M,P)" for P)+ \ \essential for efficiency\ @@ -199,8 +199,8 @@ text\For typical applications of Replacement for recursive definitions\ lemma (in M_basic) univalent_is_recfun: - "[|wellfounded(M,r); trans(r); M(r)|] - ==> univalent (M, A, \x p. + "\wellfounded(M,r); trans(r); M(r)\ + \ univalent (M, A, \x p. \y[M]. p = \x,y\ & (\f[M]. is_recfun(r,x,H,f) & y = H(x,f)))" apply (simp add: univalent_def) apply (blast dest: is_recfun_functional) @@ -210,12 +210,12 @@ text\Proof of the inductive step for \exists_is_recfun\, since we must prove two versions.\ lemma (in M_basic) exists_is_recfun_indstep: - "[|\y. \y, a1\ \ r \ (\f[M]. is_recfun(r, y, H, f)); + "\\y. \y, a1\ \ r \ (\f[M]. is_recfun(r, y, H, f)); wellfounded(M,r); trans(r); M(r); M(a1); strong_replacement(M, \x z. \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); - \x[M]. \g[M]. function(g) \ M(H(x,g))|] - ==> \f[M]. is_recfun(r,a1,H,f)" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ \f[M]. is_recfun(r,a1,H,f)" apply (drule_tac A="r-``{a1}" in strong_replacementD) apply blast txt\Discharge the "univalent" obligation of Replacement\ @@ -243,23 +243,23 @@ text\Relativized version, when we have the (currently weaker) premise \<^term>\wellfounded(M,r)\\ lemma (in M_basic) wellfounded_exists_is_recfun: - "[|wellfounded(M,r); trans(r); - separation(M, \x. ~ (\f[M]. is_recfun(r, x, H, f))); + "\wellfounded(M,r); trans(r); + separation(M, \x. \ (\f[M]. is_recfun(r, x, H, f))); strong_replacement(M, \x z. \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); M(r); M(a); - \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> \f[M]. is_recfun(r,a,H,f)" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ \f[M]. is_recfun(r,a,H,f)" apply (rule wellfounded_induct, assumption+, clarify) apply (rule exists_is_recfun_indstep, assumption+) done lemma (in M_basic) wf_exists_is_recfun [rule_format]: - "[|wf(r); trans(r); M(r); + "\wf(r); trans(r); M(r); strong_replacement(M, \x z. \y[M]. \g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); - \x[M]. \g[M]. function(g) \ M(H(x,g)) |] - ==> M(a) \ (\f[M]. is_recfun(r,a,H,f))" + \x[M]. \g[M]. function(g) \ M(H(x,g))\ + \ M(a) \ (\f[M]. is_recfun(r,a,H,f))" apply (rule wf_induct, assumption+) apply (frule wf_imp_relativized) apply (intro impI) @@ -272,7 +272,7 @@ definition M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where - "M_is_recfun(M,MH,r,a,f) == + "M_is_recfun(M,MH,r,a,f) \ \z[M]. z \ f \ (\x[M]. \y[M]. \xa[M]. \sx[M]. \r_sx[M]. \f_r_sx[M]. pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) & @@ -281,50 +281,50 @@ definition is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where - "is_wfrec(M,MH,r,a,z) == + "is_wfrec(M,MH,r,a,z) \ \f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" definition wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where - "wfrec_replacement(M,MH,r) == + "wfrec_replacement(M,MH,r) \ strong_replacement(M, \x z. \y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))" lemma (in M_basic) is_recfun_abs: - "[| \x[M]. \g[M]. function(g) \ M(H(x,g)); M(r); M(a); M(f); - relation2(M,MH,H) |] - ==> M_is_recfun(M,MH,r,a,f) \ is_recfun(r,a,H,f)" + "\\x[M]. \g[M]. function(g) \ M(H(x,g)); M(r); M(a); M(f); + relation2(M,MH,H)\ + \ M_is_recfun(M,MH,r,a,f) \ is_recfun(r,a,H,f)" apply (simp add: M_is_recfun_def relation2_def is_recfun_relativize) apply (rule rall_cong) apply (blast dest: transM) done lemma M_is_recfun_cong [cong]: - "[| r = r'; a = a'; f = f'; - !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) \ MH'(x,g,y) |] - ==> M_is_recfun(M,MH,r,a,f) \ M_is_recfun(M,MH',r',a',f')" + "\r = r'; a = a'; f = f'; + \x g y. \M(x); M(g); M(y)\ \ MH(x,g,y) \ MH'(x,g,y)\ + \ M_is_recfun(M,MH,r,a,f) \ M_is_recfun(M,MH',r',a',f')" by (simp add: M_is_recfun_def) lemma (in M_basic) is_wfrec_abs: - "[| \x[M]. \g[M]. function(g) \ M(H(x,g)); - relation2(M,MH,H); M(r); M(a); M(z) |] - ==> is_wfrec(M,MH,r,a,z) \ + "\\x[M]. \g[M]. function(g) \ M(H(x,g)); + relation2(M,MH,H); M(r); M(a); M(z)\ + \ is_wfrec(M,MH,r,a,z) \ (\g[M]. is_recfun(r,a,H,g) & z = H(a,g))" by (simp add: is_wfrec_def relation2_def is_recfun_abs) text\Relating \<^term>\wfrec_replacement\ to native constructs\ lemma (in M_basic) wfrec_replacement': - "[|wfrec_replacement(M,MH,r); + "\wfrec_replacement(M,MH,r); \x[M]. \g[M]. function(g) \ M(H(x,g)); - relation2(M,MH,H); M(r)|] - ==> strong_replacement(M, \x z. \y[M]. + relation2(M,MH,H); M(r)\ + \ strong_replacement(M, \x z. \y[M]. pair(M,x,y,z) & (\g[M]. is_recfun(r,x,H,g) & y = H(x,g)))" by (simp add: wfrec_replacement_def is_wfrec_abs) lemma wfrec_replacement_cong [cong]: - "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) \ MH'(x,y,z); - r=r' |] - ==> wfrec_replacement(M, %x y. MH(x,y), r) \ + "\\x y z. \M(x); M(y); M(z)\ \ MH(x,y,z) \ MH'(x,y,z); + r=r'\ + \ wfrec_replacement(M, %x y. MH(x,y), r) \ wfrec_replacement(M, %x y. MH'(x,y), r')" by (simp add: is_wfrec_def wfrec_replacement_def) diff -r f2094906e491 -r e44d86131648 src/ZF/Constructible/Wellorderings.thy --- a/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,69 +17,69 @@ definition irreflexive :: "[i=>o,i,i]=>o" where - "irreflexive(M,A,r) == \x[M]. x\A \ \ r" + "irreflexive(M,A,r) \ \x[M]. x\A \ \ r" definition transitive_rel :: "[i=>o,i,i]=>o" where - "transitive_rel(M,A,r) == + "transitive_rel(M,A,r) \ \x[M]. x\A \ (\y[M]. y\A \ (\z[M]. z\A \ \r \ \r \ \r))" definition linear_rel :: "[i=>o,i,i]=>o" where - "linear_rel(M,A,r) == + "linear_rel(M,A,r) \ \x[M]. x\A \ (\y[M]. y\A \ \r | x=y | \r)" definition wellfounded :: "[i=>o,i]=>o" where \ \EVERY non-empty set has an \r\-minimal element\ - "wellfounded(M,r) == - \x[M]. x\0 \ (\y[M]. y\x & ~(\z[M]. z\x & \ r))" + "wellfounded(M,r) \ + \x[M]. x\0 \ (\y[M]. y\x & \(\z[M]. z\x & \ r))" definition wellfounded_on :: "[i=>o,i,i]=>o" where \ \every non-empty SUBSET OF \A\ has an \r\-minimal element\ - "wellfounded_on(M,A,r) == - \x[M]. x\0 \ x\A \ (\y[M]. y\x & ~(\z[M]. z\x & \ r))" + "wellfounded_on(M,A,r) \ + \x[M]. x\0 \ x\A \ (\y[M]. y\x & \(\z[M]. z\x & \ r))" definition wellordered :: "[i=>o,i,i]=>o" where \ \linear and wellfounded on \A\\ - "wellordered(M,A,r) == + "wellordered(M,A,r) \ transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)" subsubsection \Trivial absoluteness proofs\ lemma (in M_basic) irreflexive_abs [simp]: - "M(A) ==> irreflexive(M,A,r) \ irrefl(A,r)" + "M(A) \ irreflexive(M,A,r) \ irrefl(A,r)" by (simp add: irreflexive_def irrefl_def) lemma (in M_basic) transitive_rel_abs [simp]: - "M(A) ==> transitive_rel(M,A,r) \ trans[A](r)" + "M(A) \ transitive_rel(M,A,r) \ trans[A](r)" by (simp add: transitive_rel_def trans_on_def) lemma (in M_basic) linear_rel_abs [simp]: - "M(A) ==> linear_rel(M,A,r) \ linear(A,r)" + "M(A) \ linear_rel(M,A,r) \ linear(A,r)" by (simp add: linear_rel_def linear_def) lemma (in M_basic) wellordered_is_trans_on: - "[| wellordered(M,A,r); M(A) |] ==> trans[A](r)" + "\wellordered(M,A,r); M(A)\ \ trans[A](r)" by (auto simp add: wellordered_def) lemma (in M_basic) wellordered_is_linear: - "[| wellordered(M,A,r); M(A) |] ==> linear(A,r)" + "\wellordered(M,A,r); M(A)\ \ linear(A,r)" by (auto simp add: wellordered_def) lemma (in M_basic) wellordered_is_wellfounded_on: - "[| wellordered(M,A,r); M(A) |] ==> wellfounded_on(M,A,r)" + "\wellordered(M,A,r); M(A)\ \ wellfounded_on(M,A,r)" by (auto simp add: wellordered_def) lemma (in M_basic) wellfounded_imp_wellfounded_on: - "[| wellfounded(M,r); M(A) |] ==> wellfounded_on(M,A,r)" + "\wellfounded(M,r); M(A)\ \ wellfounded_on(M,A,r)" by (auto simp add: wellfounded_def wellfounded_on_def) lemma (in M_basic) wellfounded_on_subset_A: - "[| wellfounded_on(M,A,r); B<=A |] ==> wellfounded_on(M,B,r)" + "\wellfounded_on(M,A,r); B<=A\ \ wellfounded_on(M,B,r)" by (simp add: wellfounded_on_def, blast) @@ -93,35 +93,35 @@ done lemma (in M_basic) wellfounded_on_imp_wellfounded: - "[|wellfounded_on(M,A,r); r \ A*A|] ==> wellfounded(M,r)" + "\wellfounded_on(M,A,r); r \ A*A\ \ wellfounded(M,r)" by (simp add: wellfounded_on_iff_wellfounded subset_Int_iff) lemma (in M_basic) wellfounded_on_field_imp_wellfounded: - "wellfounded_on(M, field(r), r) ==> wellfounded(M,r)" + "wellfounded_on(M, field(r), r) \ wellfounded(M,r)" by (simp add: wellfounded_def wellfounded_on_iff_wellfounded, fast) lemma (in M_basic) wellfounded_iff_wellfounded_on_field: - "M(r) ==> wellfounded(M,r) \ wellfounded_on(M, field(r), r)" + "M(r) \ wellfounded(M,r) \ wellfounded_on(M, field(r), r)" by (blast intro: wellfounded_imp_wellfounded_on wellfounded_on_field_imp_wellfounded) (*Consider the least z in domain(r) such that P(z) does not hold...*) lemma (in M_basic) wellfounded_induct: - "[| wellfounded(M,r); M(a); M(r); separation(M, \x. ~P(x)); - \x. M(x) & (\y. \ r \ P(y)) \ P(x) |] - ==> P(a)" + "\wellfounded(M,r); M(a); M(r); separation(M, \x. \P(x)); + \x. M(x) & (\y. \ r \ P(y)) \ P(x)\ + \ P(a)" apply (simp (no_asm_use) add: wellfounded_def) -apply (drule_tac x="{z \ domain(r). ~P(z)}" in rspec) +apply (drule_tac x="{z \ domain(r). \P(z)}" in rspec) apply (blast dest: transM)+ done lemma (in M_basic) wellfounded_on_induct: - "[| a\A; wellfounded_on(M,A,r); M(A); - separation(M, \x. x\A \ ~P(x)); - \x\A. M(x) & (\y\A. \ r \ P(y)) \ P(x) |] - ==> P(a)" + "\a\A; wellfounded_on(M,A,r); M(A); + separation(M, \x. x\A \ \P(x)); + \x\A. M(x) & (\y\A. \ r \ P(y)) \ P(x)\ + \ P(a)" apply (simp (no_asm_use) add: wellfounded_on_def) -apply (drule_tac x="{z\A. z\A \ ~P(z)}" in rspec) +apply (drule_tac x="{z\A. z\A \ \P(z)}" in rspec) apply (blast intro: transM)+ done @@ -129,11 +129,11 @@ subsubsection \Kunen's lemma IV 3.14, page 123\ lemma (in M_basic) linear_imp_relativized: - "linear(A,r) ==> linear_rel(M,A,r)" + "linear(A,r) \ linear_rel(M,A,r)" by (simp add: linear_def linear_rel_def) lemma (in M_basic) trans_on_imp_relativized: - "trans[A](r) ==> transitive_rel(M,A,r)" + "trans[A](r) \ transitive_rel(M,A,r)" by (unfold transitive_rel_def trans_on_def, blast) lemma (in M_basic) wf_on_imp_relativized: @@ -143,13 +143,13 @@ done lemma (in M_basic) wf_imp_relativized: - "wf(r) ==> wellfounded(M,r)" + "wf(r) \ wellfounded(M,r)" apply (simp add: wellfounded_def wf_def, clarify) apply (drule_tac x=x in spec, blast) done lemma (in M_basic) well_ord_imp_relativized: - "well_ord(A,r) ==> wellordered(M,A,r)" + "well_ord(A,r) \ wellordered(M,A,r)" by (simp add: wellordered_def well_ord_def tot_ord_def part_ord_def linear_imp_relativized trans_on_imp_relativized wf_on_imp_relativized) @@ -160,12 +160,12 @@ subsection\Relativized versions of order-isomorphisms and order types\ lemma (in M_basic) order_isomorphism_abs [simp]: - "[| M(A); M(B); M(f) |] - ==> order_isomorphism(M,A,r,B,s,f) \ f \ ord_iso(A,r,B,s)" + "\M(A); M(B); M(f)\ + \ order_isomorphism(M,A,r,B,s,f) \ f \ ord_iso(A,r,B,s)" by (simp add: order_isomorphism_def ord_iso_def) lemma (in M_trans) pred_set_abs [simp]: - "[| M(r); M(B) |] ==> pred_set(M,A,x,r,B) \ B = Order.pred(A,x,r)" + "\M(r); M(B)\ \ pred_set(M,A,x,r,B) \ B = Order.pred(A,x,r)" apply (simp add: pred_set_def Order.pred_def) apply (blast dest: transM) done @@ -175,7 +175,7 @@ using pred_separation [of r x] by (simp add: Order.pred_def) lemma (in M_basic) membership_abs [simp]: - "[| M(r); M(A) |] ==> membership(M,A,r) \ r = Memrel(A)" + "\M(r); M(A)\ \ membership(M,A,r) \ r = Memrel(A)" apply (simp add: membership_def Memrel_def, safe) apply (rule equalityI) apply clarify @@ -220,14 +220,14 @@ done lemma (in M_basic) wellfounded_on_asym: - "[| wellfounded_on(M,A,r); \r; a\A; x\A; M(A) |] ==> \r" + "\wellfounded_on(M,A,r); \r; a\A; x\A; M(A)\ \ \r" apply (simp add: wellfounded_on_def) apply (drule_tac x="{x,a}" in rspec) apply (blast dest: transM)+ done lemma (in M_basic) wellordered_asym: - "[| wellordered(M,A,r); \r; a\A; x\A; M(A) |] ==> \r" + "\wellordered(M,A,r); \r; a\A; x\A; M(A)\ \ \r" by (simp add: wellordered_def, blast dest: wellfounded_on_asym) end diff -r f2094906e491 -r e44d86131648 src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Epsilon.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,19 +9,19 @@ definition eclose :: "i=>i" where - "eclose(A) == \n\nat. nat_rec(n, A, %m r. \(r))" + "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)" + "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))" + "rank(a) \ transrec(a, %x f. \y\x. succ(f`y))" definition transrec2 :: "[i, i, [i,i]=>i] =>i" where - "transrec2(k, a, b) == + "transrec2(k, a, b) \ transrec(k, %i r. if(i=0, a, if(\j. i=succ(j), @@ -30,11 +30,11 @@ definition recursor :: "[i, [i,i]=>i, i]=>i" where - "recursor(a,b,k) == transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))" + "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)" + "rec(k,a,b) \ recursor(a,b,k)" subsection\Basic Closure Properties\ @@ -56,19 +56,19 @@ apply (erule UnionI, assumption) done -(* @{term"x \ eclose(A) ==> x \ eclose(A)"} *) +(* @{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)"} *) +(* @{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) + \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] @@ -76,7 +76,7 @@ (*Epsilon induction*) lemma eps_induct: - "[| !!x. \y\x. P(y) ==> P(x) |] ==> P(a)" + "\\x. \y\x. P(y) \ P(x)\ \ P(a)" by (rule arg_in_eclose_sing [THEN eclose_induct], blast) @@ -85,7 +85,7 @@ (** 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" + "\Transset(X); A<=X; n \ nat\ \ nat_rec(n, A, %m r. \(r)) \ X" apply (unfold Transset_def) apply (erule nat_induct) apply (simp add: nat_rec_0) @@ -93,17 +93,17 @@ done lemma eclose_least: - "[| Transset(X); A<=X |] ==> eclose(A) \ X" + "\Transset(X); A<=X\ \ eclose(A) \ X" apply (unfold eclose_def) apply (rule eclose_least_lemma [THEN UN_least], assumption+) done -(*COMPLETELY DIFFERENT induction principle from eclose_induct!!*) +(*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)" + "\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 apply (unfold Transset_def) @@ -111,49 +111,49 @@ apply (blast intro: arg_subset_eclose [THEN subsetD]) done -lemma Transset_eclose_eq_arg: "Transset(X) ==> eclose(X) = X" +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" +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" +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)" +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})" + "\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" +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" +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"} *) +(* @{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}) |] ==> + "\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) @@ -162,7 +162,7 @@ done lemma wfrec_eclose_eq2: - "k \ i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)" + "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 @@ -175,20 +175,20 @@ (*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))" + "\\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)" + "\\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)" +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 @@ -198,7 +198,7 @@ apply (frule eclose_subset, blast) done -lemma eclose_sing_Ord_eq: "Ord(i) ==> eclose({i}) = succ(i)" +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) @@ -207,7 +207,7 @@ 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)" + 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) @@ -229,25 +229,25 @@ apply (erule bspec, assumption) done -lemma rank_of_Ord: "Ord(i) ==> rank(i) = i" +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)" +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)" +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)" +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 @@ -305,14 +305,14 @@ (*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)" + "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)" +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) @@ -321,12 +321,12 @@ subsection\Corollaries of Leastness\ -lemma mem_eclose_subset: "A \ B ==> eclose(A)<=eclose(B)" +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)" +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) @@ -352,14 +352,14 @@ done lemma transrec2_Limit: - "Limit(i) ==> transrec2(i,a,b) = (\j transrec2(i,a,b) = (\j f(0) = a & + "(\x. f(x)\transrec2(x,a,b)) + \ f(0) = a & f(succ(i)) = b(i, f(i)) & (Limit(K) \ f(K) = (\j nat; + "\n \ nat; a \ C(0); - !!m z. [| m \ nat; z \ C(m) |] ==> b(m,z): C(succ(m)) |] - ==> rec(n,a,b) \ C(n)" + \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 -r f2094906e491 -r e44d86131648 src/ZF/EquivClass.thy --- a/src/ZF/EquivClass.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/EquivClass.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,49 +9,49 @@ definition quotient :: "[i,i]=>i" (infixl \'/'/\ 90) (*set of equiv classes*) where - "A//r == {r``{x} . x \ A}" + "A//r \ {r``{x} . x \ A}" definition congruent :: "[i,i=>i]=>o" where - "congruent(r,b) == \y z. :r \ b(y)=b(z)" + "congruent(r,b) \ \y z. :r \ b(y)=b(z)" definition congruent2 :: "[i,i,[i,i]=>i]=>o" where - "congruent2(r1,r2,b) == \y1 z1 y2 z2. + "congruent2(r1,r2,b) \ \y1 z1 y2 z2. :r1 \ :r2 \ b(y1,y2) = b(z1,z2)" abbreviation RESPECTS ::"[i=>i, i] => o" (infixr \respects\ 80) where - "f respects r == congruent(r,f)" + "f respects r \ congruent(r,f)" abbreviation RESPECTS2 ::"[i=>i=>i, i] => o" (infixr \respects2 \ 80) where - "f respects2 r == congruent2(r,r,f)" + "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 **) +(** first half: equiv(A,r) \ converse(r) O r = r **) lemma sym_trans_comp_subset: - "[| sym(r); trans(r) |] ==> converse(r) O r \ r" + "\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" + "\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" + "equiv(A,r) \ converse(r) O r = r" apply (unfold 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)" + "\converse(r) O r = r; domain(r) = A\ \ equiv(A,r)" apply (unfold equiv_def refl_def sym_def trans_def) apply (erule equalityE) apply (subgoal_tac "\x y. \ r \ \ r", blast+) @@ -61,63 +61,63 @@ (*Lemma for the next result*) lemma equiv_class_subset: - "[| sym(r); trans(r); : r |] ==> r``{a} \ r``{b}" + "\sym(r); trans(r); : r\ \ r``{a} \ r``{b}" by (unfold trans_def sym_def, blast) lemma equiv_class_eq: - "[| equiv(A,r); : r |] ==> r``{a} = r``{b}" + "\equiv(A,r); : r\ \ r``{a} = r``{b}" apply (unfold 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}" + "\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 |] ==> : r" + "\equiv(A,r); r``{b} \ r``{a}; b \ A\ \ : r" by (unfold equiv_def refl_def, blast) -lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b \ A |] ==> : r" +lemma eq_equiv_class: "\r``{a} = r``{b}; equiv(A,r); b \ A\ \ : 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}) |] ==> : r" + "\equiv(A,r); x: (r``{a} \ r``{b})\ \ : r" by (unfold equiv_def trans_def sym_def, blast) -lemma equiv_type: "equiv(A,r) ==> r \ A*A" +lemma equiv_type: "equiv(A,r) \ r \ A*A" by (unfold equiv_def, blast) lemma equiv_class_eq_iff: - "equiv(A,r) ==> : r \ r``{x} = r``{y} & x \ A & y \ A" + "equiv(A,r) \ : 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} \ : r" + "\equiv(A,r); x \ A; y \ A\ \ r``{x} = r``{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" +lemma quotientI [TC]: "x \ A \ r``{x}: A//r" apply (unfold quotient_def) apply (erule RepFunI) done lemma quotientE: - "[| X \ A//r; !!x. [| X = r``{x}; x \ A |] ==> P |] ==> P" + "\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" + "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)" + "\equiv(A,r); X \ A//r; Y \ A//r\ \ X=Y | (X \ Y \ 0)" apply (unfold quotient_def) apply (safe intro!: equiv_class_eq, assumption) apply (unfold equiv_def trans_def sym_def, blast) @@ -130,7 +130,7 @@ (*Conversion rule*) lemma UN_equiv_class: - "[| equiv(A,r); b respects r; a \ A |] ==> (\x\r``{a}. b(x)) = b(a)" + "\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) @@ -139,20 +139,20 @@ (*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" + "\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 + 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; + "\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) |] ==> :r |] - ==> X=Y" + \x y. \x \ A; y \ A; b(x)=b(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]) @@ -162,11 +162,11 @@ subsection\Defining Binary Operations upon Equivalence Classes\ lemma congruent2_implies_congruent: - "[| equiv(A,r1); congruent2(r1,r2,b); a \ A |] ==> congruent(r2,b(a))" + "\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 |] ==> + "\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) @@ -176,17 +176,17 @@ 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)" + "\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; + "\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" + \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) @@ -196,18 +196,18 @@ (*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; \ r1 |] ==> b(y,w) = b(z,w); - !! y z w. [| w \ A1; \ r2 |] ==> b(w,y) = b(w,z) - |] ==> congruent2(r1,r2,b)" + "\equiv(A1,r1); equiv(A2,r2); + \y z w. \w \ A2; \ r1\ \ b(y,w) = b(z,w); + \y z w. \w \ A1; \ 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; : r |] ==> b(w,y) = b(w,z)" + and commute: "\y z. \y \ A; z \ A\ \ b(y,z) = b(z,y)" + and congt: "\y z w. \w \ A; : 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]) @@ -219,10 +219,10 @@ (*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))" + "\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) diff -r f2094906e491 -r e44d86131648 src/ZF/Finite.thy --- a/src/ZF/Finite.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Finite.thy Tue Sep 27 16:51:35 2022 +0100 @@ -2,7 +2,7 @@ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge -prove: b \ Fin(A) ==> inj(b,b) \ surj(b,b) +prove: b \ Fin(A) \ inj(b,b) \ surj(b,b) *) section\Finite Powerset Operator and Finite Function Space\ @@ -25,7 +25,7 @@ domains "Fin(A)" \ "Pow(A)" intros emptyI: "0 \ Fin(A)" - consI: "[| a \ A; b \ Fin(A) |] ==> cons(a,b) \ 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] @@ -33,31 +33,31 @@ domains "FiniteFun(A,B)" \ "Fin(A*B)" intros emptyI: "0 \ A -||> B" - consI: "[| a \ A; b \ B; h \ A -||> B; a \ domain(h) |] - ==> cons(,h) \ A -||> B" + consI: "\a \ A; b \ B; h \ A -||> B; a \ domain(h)\ + \ cons(,h) \ A -||> B" type_intros Fin.intros subsection \Finite Powerset Operator\ -lemma Fin_mono: "A<=B ==> Fin(A) \ Fin(B)" +lemma Fin_mono: "A<=B \ Fin(A) \ Fin(B)" apply (unfold Fin.defs) apply (rule lfp_mono) apply (rule Fin.bnd_mono)+ apply blast done -(* @{term"A \ Fin(B) ==> A \ B"} *) +(* @{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); + "\b \ Fin(A); P(0); - !!x y. [| x \ A; y \ Fin(A); x\y; P(y) |] ==> P(cons(x,y)) - |] ==> P(b)" + \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!*) @@ -72,53 +72,53 @@ 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)" +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)" +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)" +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)" +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)" +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)" +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)" + "\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); + "\b \ Fin(A); P(b); - !!x y. [| x \ A; y \ Fin(A); x \ y; P(y) |] ==> P(y-{x}) - |] ==> P(0)" + \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)" +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]) @@ -129,29 +129,29 @@ subsection\Finite Function Space\ lemma FiniteFun_mono: - "[| A<=C; B<=D |] ==> A -||> B \ C -||> D" + "\A<=C; B<=D\ \ A -||> B \ C -||> D" apply (unfold 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" +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" +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)" +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" + "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) @@ -160,12 +160,12 @@ apply (fast intro!: FiniteFun.intros) done -lemma FiniteFun_subset: "[| c<=b; b \ A-||>B |] ==> c \ A-||>B" +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" +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") @@ -178,7 +178,7 @@ FiniteFun_mono [THEN [2] rev_subsetD]) done -lemma lam_FiniteFun: "A \ Fin(X) ==> (\x\A. b(x)) \ A -||> {b(x). x \ A}" +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: @@ -201,7 +201,7 @@ definition contents :: "i=>i" where - "contents(X) == THE x. X = {x}" + "contents(X) \ THE x. X = {x}" lemma contents_eq [simp]: "contents ({x}) = x" by (simp add: contents_def) diff -r f2094906e491 -r e44d86131648 src/ZF/Fixedpt.thy --- a/src/ZF/Fixedpt.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Fixedpt.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,15 +10,15 @@ definition (*monotone operator from Pow(D) to itself*) bnd_mono :: "[i,i=>i]=>o" where - "bnd_mono(D,h) == h(D)<=D & (\W X. W<=X \ X<=D \ h(W) \ h(X))" + "bnd_mono(D,h) \ h(D)<=D & (\W X. W<=X \ X<=D \ h(W) \ h(X))" definition lfp :: "[i,i=>i]=>i" where - "lfp(D,h) == \({X: Pow(D). h(X) \ X})" + "lfp(D,h) \ \({X: Pow(D). h(X) \ X})" definition gfp :: "[i,i=>i]=>i" where - "gfp(D,h) == \({X: Pow(D). X \ h(X)})" + "gfp(D,h) \ \({X: Pow(D). X \ h(X)})" text\The theorem is proved in the lattice of subsets of \<^term>\D\, namely \<^term>\Pow(D)\, with Inter as the greatest lower bound.\ @@ -26,33 +26,33 @@ subsection\Monotone Operators\ lemma bnd_monoI: - "[| h(D)<=D; - !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) \ h(X) - |] ==> bnd_mono(D,h)" + "\h(D)<=D; + \W X. \W<=D; X<=D; W<=X\ \ h(W) \ h(X) +\ \ bnd_mono(D,h)" by (unfold bnd_mono_def, clarify, blast) -lemma bnd_monoD1: "bnd_mono(D,h) ==> h(D) \ D" +lemma bnd_monoD1: "bnd_mono(D,h) \ h(D) \ D" apply (unfold bnd_mono_def) apply (erule conjunct1) done -lemma bnd_monoD2: "[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) \ h(X)" +lemma bnd_monoD2: "\bnd_mono(D,h); W<=X; X<=D\ \ h(W) \ h(X)" by (unfold bnd_mono_def, blast) lemma bnd_mono_subset: - "[| bnd_mono(D,h); X<=D |] ==> h(X) \ D" + "\bnd_mono(D,h); X<=D\ \ h(X) \ D" by (unfold bnd_mono_def, clarify, blast) lemma bnd_mono_Un: - "[| bnd_mono(D,h); A \ D; B \ D |] ==> h(A) \ h(B) \ h(A \ B)" + "\bnd_mono(D,h); A \ D; B \ D\ \ h(A) \ h(B) \ h(A \ B)" apply (unfold bnd_mono_def) apply (rule Un_least, blast+) done (*unused*) lemma bnd_mono_UN: - "[| bnd_mono(D,h); \i\I. A(i) \ D |] - ==> (\i\I. h(A(i))) \ h((\i\I. A(i)))" + "\bnd_mono(D,h); \i\I. A(i) \ D\ + \ (\i\I. h(A(i))) \ h((\i\I. A(i)))" apply (unfold bnd_mono_def) apply (rule UN_least) apply (elim conjE) @@ -63,7 +63,7 @@ (*Useful??*) lemma bnd_mono_Int: - "[| bnd_mono(D,h); A \ D; B \ D |] ==> h(A \ B) \ h(A) \ h(B)" + "\bnd_mono(D,h); A \ D; B \ D\ \ h(A \ B) \ h(A) \ h(B)" apply (rule Int_greatest) apply (erule bnd_monoD2, rule Int_lower1, assumption) apply (erule bnd_monoD2, rule Int_lower2, assumption) @@ -73,7 +73,7 @@ (*lfp is contained in each pre-fixedpoint*) lemma lfp_lowerbound: - "[| h(A) \ A; A<=D |] ==> lfp(D,h) \ A" + "\h(A) \ A; A<=D\ \ lfp(D,h) \ A" by (unfold lfp_def, blast) (*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*) @@ -81,29 +81,29 @@ by (unfold lfp_def Inter_def, blast) (*Used in datatype package*) -lemma def_lfp_subset: "A == lfp(D,h) ==> A \ D" +lemma def_lfp_subset: "A \ lfp(D,h) \ A \ D" apply simp apply (rule lfp_subset) done lemma lfp_greatest: - "[| h(D) \ D; !!X. [| h(X) \ X; X<=D |] ==> A<=X |] ==> A \ lfp(D,h)" + "\h(D) \ D; \X. \h(X) \ X; X<=D\ \ A<=X\ \ A \ lfp(D,h)" by (unfold lfp_def, blast) lemma lfp_lemma1: - "[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) \ A" + "\bnd_mono(D,h); h(A)<=A; A<=D\ \ h(lfp(D,h)) \ A" apply (erule bnd_monoD2 [THEN subset_trans]) apply (rule lfp_lowerbound, assumption+) done -lemma lfp_lemma2: "bnd_mono(D,h) ==> h(lfp(D,h)) \ lfp(D,h)" +lemma lfp_lemma2: "bnd_mono(D,h) \ h(lfp(D,h)) \ lfp(D,h)" apply (rule bnd_monoD1 [THEN lfp_greatest]) apply (rule_tac [2] lfp_lemma1) apply (assumption+) done lemma lfp_lemma3: - "bnd_mono(D,h) ==> lfp(D,h) \ h(lfp(D,h))" + "bnd_mono(D,h) \ lfp(D,h) \ h(lfp(D,h))" apply (rule lfp_lowerbound) apply (rule bnd_monoD2, assumption) apply (rule lfp_lemma2, assumption) @@ -111,7 +111,7 @@ apply (rule lfp_subset)+ done -lemma lfp_unfold: "bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))" +lemma lfp_unfold: "bnd_mono(D,h) \ lfp(D,h) = h(lfp(D,h))" apply (rule equalityI) apply (erule lfp_lemma3) apply (erule lfp_lemma2) @@ -119,7 +119,7 @@ (*Definition form, to control unfolding*) lemma def_lfp_unfold: - "[| A==lfp(D,h); bnd_mono(D,h) |] ==> A = h(A)" + "\A\lfp(D,h); bnd_mono(D,h)\ \ A = h(A)" apply simp apply (erule lfp_unfold) done @@ -127,17 +127,17 @@ subsection\General Induction Rule for Least Fixedpoints\ lemma Collect_is_pre_fixedpt: - "[| bnd_mono(D,h); !!x. x \ h(Collect(lfp(D,h),P)) ==> P(x) |] - ==> h(Collect(lfp(D,h),P)) \ Collect(lfp(D,h),P)" + "\bnd_mono(D,h); \x. x \ h(Collect(lfp(D,h),P)) \ P(x)\ + \ h(Collect(lfp(D,h),P)) \ Collect(lfp(D,h),P)" by (blast intro: lfp_lemma2 [THEN subsetD] bnd_monoD2 [THEN subsetD] lfp_subset [THEN subsetD]) (*This rule yields an induction hypothesis in which the components of a data structure may be assumed to be elements of lfp(D,h)*) lemma induct: - "[| bnd_mono(D,h); a \ lfp(D,h); - !!x. x \ h(Collect(lfp(D,h),P)) ==> P(x) - |] ==> P(a)" + "\bnd_mono(D,h); a \ lfp(D,h); + \x. x \ h(Collect(lfp(D,h),P)) \ P(x) +\ \ P(a)" apply (rule Collect_is_pre_fixedpt [THEN lfp_lowerbound, THEN subsetD, THEN CollectD2]) apply (rule_tac [3] lfp_subset [THEN Collect_subset [THEN subset_trans]], @@ -146,15 +146,15 @@ (*Definition form, to control unfolding*) lemma def_induct: - "[| A == lfp(D,h); bnd_mono(D,h); a:A; - !!x. x \ h(Collect(A,P)) ==> P(x) - |] ==> P(a)" + "\A \ lfp(D,h); bnd_mono(D,h); a:A; + \x. x \ h(Collect(A,P)) \ P(x) +\ \ P(a)" by (rule induct, blast+) (*This version is useful when "A" is not a subset of D - second premise could simply be h(D \ A) \ D or !!X. X<=D ==> h(X)<=D *) + second premise could simply be h(D \ A) \ D or \X. X<=D \ h(X)<=D *) lemma lfp_Int_lowerbound: - "[| h(D \ A) \ A; bnd_mono(D,h) |] ==> lfp(D,h) \ A" + "\h(D \ A) \ A; bnd_mono(D,h)\ \ lfp(D,h) \ A" apply (rule lfp_lowerbound [THEN subset_trans]) apply (erule bnd_mono_subset [THEN Int_greatest], blast+) done @@ -164,7 +164,7 @@ lemma lfp_mono: assumes hmono: "bnd_mono(D,h)" and imono: "bnd_mono(E,i)" - and subhi: "!!X. X<=D ==> h(X) \ i(X)" + and subhi: "\X. X<=D \ h(X) \ i(X)" shows "lfp(D,h) \ lfp(E,i)" apply (rule bnd_monoD1 [THEN lfp_greatest]) apply (rule imono) @@ -176,13 +176,13 @@ (*This (unused) version illustrates that monotonicity is not really needed, but both lfp's must be over the SAME set D; Inter is anti-monotonic!*) lemma lfp_mono2: - "[| i(D) \ D; !!X. X<=D ==> h(X) \ i(X) |] ==> lfp(D,h) \ lfp(D,i)" + "\i(D) \ D; \X. X<=D \ h(X) \ i(X)\ \ lfp(D,h) \ lfp(D,i)" apply (rule lfp_greatest, assumption) apply (rule lfp_lowerbound, blast, assumption) done lemma lfp_cong: - "[|D=D'; !!X. X \ D' ==> h(X) = h'(X)|] ==> lfp(D,h) = lfp(D',h')" + "\D=D'; \X. X \ D' \ h(X) = h'(X)\ \ lfp(D,h) = lfp(D',h')" apply (simp add: lfp_def) apply (rule_tac t=Inter in subst_context) apply (rule Collect_cong, simp_all) @@ -192,7 +192,7 @@ subsection\Proof of Knaster-Tarski Theorem using \<^term>\gfp\\ (*gfp contains each post-fixedpoint that is contained in D*) -lemma gfp_upperbound: "[| A \ h(A); A<=D |] ==> A \ gfp(D,h)" +lemma gfp_upperbound: "\A \ h(A); A<=D\ \ A \ gfp(D,h)" apply (unfold gfp_def) apply (rule PowI [THEN CollectI, THEN Union_upper]) apply (assumption+) @@ -202,41 +202,41 @@ by (unfold gfp_def, blast) (*Used in datatype package*) -lemma def_gfp_subset: "A==gfp(D,h) ==> A \ D" +lemma def_gfp_subset: "A\gfp(D,h) \ A \ D" apply simp apply (rule gfp_subset) done lemma gfp_least: - "[| bnd_mono(D,h); !!X. [| X \ h(X); X<=D |] ==> X<=A |] ==> + "\bnd_mono(D,h); \X. \X \ h(X); X<=D\ \ X<=A\ \ gfp(D,h) \ A" apply (unfold gfp_def) apply (blast dest: bnd_monoD1) done lemma gfp_lemma1: - "[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A \ h(gfp(D,h))" + "\bnd_mono(D,h); A<=h(A); A<=D\ \ A \ h(gfp(D,h))" apply (rule subset_trans, assumption) apply (erule bnd_monoD2) apply (rule_tac [2] gfp_subset) apply (simp add: gfp_upperbound) done -lemma gfp_lemma2: "bnd_mono(D,h) ==> gfp(D,h) \ h(gfp(D,h))" +lemma gfp_lemma2: "bnd_mono(D,h) \ gfp(D,h) \ h(gfp(D,h))" apply (rule gfp_least) apply (rule_tac [2] gfp_lemma1) apply (assumption+) done lemma gfp_lemma3: - "bnd_mono(D,h) ==> h(gfp(D,h)) \ gfp(D,h)" + "bnd_mono(D,h) \ h(gfp(D,h)) \ gfp(D,h)" apply (rule gfp_upperbound) apply (rule bnd_monoD2, assumption) apply (rule gfp_lemma2, assumption) apply (erule bnd_mono_subset, rule gfp_subset)+ done -lemma gfp_unfold: "bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))" +lemma gfp_unfold: "bnd_mono(D,h) \ gfp(D,h) = h(gfp(D,h))" apply (rule equalityI) apply (erule gfp_lemma2) apply (erule gfp_lemma3) @@ -244,7 +244,7 @@ (*Definition form, to control unfolding*) lemma def_gfp_unfold: - "[| A==gfp(D,h); bnd_mono(D,h) |] ==> A = h(A)" + "\A\gfp(D,h); bnd_mono(D,h)\ \ A = h(A)" apply simp apply (erule gfp_unfold) done @@ -253,11 +253,11 @@ subsection\Coinduction Rules for Greatest Fixed Points\ (*weak version*) -lemma weak_coinduct: "[| a: X; X \ h(X); X \ D |] ==> a \ gfp(D,h)" +lemma weak_coinduct: "\a: X; X \ h(X); X \ D\ \ a \ gfp(D,h)" by (blast intro: gfp_upperbound [THEN subsetD]) lemma coinduct_lemma: - "[| X \ h(X \ gfp(D,h)); X \ D; bnd_mono(D,h) |] ==> + "\X \ h(X \ gfp(D,h)); X \ D; bnd_mono(D,h)\ \ X \ gfp(D,h) \ h(X \ gfp(D,h))" apply (erule Un_least) apply (rule gfp_lemma2 [THEN subset_trans], assumption) @@ -268,8 +268,8 @@ (*strong version*) lemma coinduct: - "[| bnd_mono(D,h); a: X; X \ h(X \ gfp(D,h)); X \ D |] - ==> a \ gfp(D,h)" + "\bnd_mono(D,h); a: X; X \ h(X \ gfp(D,h)); X \ D\ + \ a \ gfp(D,h)" apply (rule weak_coinduct) apply (erule_tac [2] coinduct_lemma) apply (simp_all add: gfp_subset Un_subset_iff) @@ -277,7 +277,7 @@ (*Definition form, to control unfolding*) lemma def_coinduct: - "[| A == gfp(D,h); bnd_mono(D,h); a: X; X \ h(X \ A); X \ D |] ==> + "\A \ gfp(D,h); bnd_mono(D,h); a: X; X \ h(X \ A); X \ D\ \ a \ A" apply simp apply (rule coinduct, assumption+) @@ -285,16 +285,16 @@ (*The version used in the induction/coinduction package*) lemma def_Collect_coinduct: - "[| A == gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w))); - a: X; X \ D; !!z. z: X ==> P(X \ A, z) |] ==> + "\A \ gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w))); + a: X; X \ D; \z. z: X \ P(X \ A, z)\ \ a \ A" apply (rule def_coinduct, assumption+, blast+) done (*Monotonicity of gfp!*) lemma gfp_mono: - "[| bnd_mono(D,h); D \ E; - !!X. X<=D ==> h(X) \ i(X) |] ==> gfp(D,h) \ gfp(E,i)" + "\bnd_mono(D,h); D \ E; + \X. X<=D \ h(X) \ i(X)\ \ gfp(D,h) \ gfp(E,i)" apply (rule gfp_upperbound) apply (rule gfp_lemma2 [THEN subset_trans], assumption) apply (blast del: subsetI intro: gfp_subset) diff -r f2094906e491 -r e44d86131648 src/ZF/IMP/Com.thy --- a/src/ZF/IMP/Com.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/IMP/Com.thy Tue Sep 27 16:51:35 2022 +0100 @@ -24,16 +24,16 @@ abbreviation evala_syntax :: "[i, i] => o" (infixl \-a->\ 50) - where "p -a-> n == \ evala" + where "p -a-> n \ \ evala" inductive domains "evala" \ "(aexp \ (loc -> nat)) \ nat" intros - N: "[| n \ nat; sigma \ loc->nat |] ==> -a-> n" - X: "[| x \ loc; sigma \ loc->nat |] ==> -a-> sigma`x" - Op1: "[| -a-> n; f \ nat -> nat |] ==> -a-> f`n" - Op2: "[| -a-> n0; -a-> n1; f \ (nat\nat) -> nat |] - ==> -a-> f`" + N: "\n \ nat; sigma \ loc->nat\ \ -a-> n" + X: "\x \ loc; sigma \ loc->nat\ \ -a-> sigma`x" + Op1: "\ -a-> n; f \ nat -> nat\ \ -a-> f`n" + Op2: "\ -a-> n0; -a-> n1; f \ (nat\nat) -> nat\ + \ -a-> f`" type_intros aexp.intros apply_funtype @@ -54,20 +54,20 @@ abbreviation evalb_syntax :: "[i,i] => o" (infixl \-b->\ 50) - where "p -b-> b == \ evalb" + where "p -b-> b \ \ evalb" inductive domains "evalb" \ "(bexp \ (loc -> nat)) \ bool" intros - true: "[| sigma \ loc -> nat |] ==> -b-> 1" - false: "[| sigma \ loc -> nat |] ==> -b-> 0" - ROp: "[| -a-> n0; -a-> n1; f \ (nat*nat)->bool |] - ==> -b-> f` " - noti: "[| -b-> w |] ==> -b-> not(w)" - andi: "[| -b-> w0; -b-> w1 |] - ==> -b-> (w0 and w1)" - ori: "[| -b-> w0; -b-> w1 |] - ==> -b-> (w0 or w1)" + true: "\sigma \ loc -> nat\ \ -b-> 1" + false: "\sigma \ loc -> nat\ \ -b-> 0" + ROp: "\ -a-> n0; -a-> n1; f \ (nat*nat)->bool\ + \ -b-> f` " + noti: "\ -b-> w\ \ -b-> not(w)" + andi: "\ -b-> w0; -b-> w1\ + \ -b-> (w0 and w1)" + ori: "\ -b-> w0; -b-> w1\ + \ -b-> (w0 or w1)" type_intros bexp.intros apply_funtype and_type or_type bool_1I bool_0I not_type type_elims evala.dom_subset [THEN subsetD, elim_format] @@ -88,33 +88,33 @@ abbreviation evalc_syntax :: "[i, i] => o" (infixl \-c->\ 50) - where "p -c-> s == \ evalc" + where "p -c-> s \ \ evalc" inductive domains "evalc" \ "(com \ (loc -> nat)) \ (loc -> nat)" intros - skip: "[| sigma \ loc -> nat |] ==> <\,sigma> -c-> sigma" + skip: "\sigma \ loc -> nat\ \ <\,sigma> -c-> sigma" - assign: "[| m \ nat; x \ loc; -a-> m |] - ==> a,sigma> -c-> sigma(x:=m)" + assign: "\m \ nat; x \ loc; -a-> m\ + \ a,sigma> -c-> sigma(x:=m)" - semi: "[| -c-> sigma2; -c-> sigma1 |] - ==> c1, sigma> -c-> sigma1" + semi: "\ -c-> sigma2; -c-> sigma1\ + \ c1, sigma> -c-> sigma1" - if1: "[| b \ bexp; c1 \ com; sigma \ loc->nat; - -b-> 1; -c-> sigma1 |] - ==> <\ b \ c0 \ c1, sigma> -c-> sigma1" + if1: "\b \ bexp; c1 \ com; sigma \ loc->nat; + -b-> 1; -c-> sigma1\ + \ <\ b \ c0 \ c1, sigma> -c-> sigma1" - if0: "[| b \ bexp; c0 \ com; sigma \ loc->nat; - -b-> 0; -c-> sigma1 |] - ==> <\ b \ c0 \ c1, sigma> -c-> sigma1" + if0: "\b \ bexp; c0 \ com; sigma \ loc->nat; + -b-> 0; -c-> sigma1\ + \ <\ b \ c0 \ c1, sigma> -c-> sigma1" - while0: "[| c \ com; -b-> 0 |] - ==> <\ b \ c,sigma> -c-> sigma" + while0: "\c \ com; -b-> 0\ + \ <\ b \ c,sigma> -c-> sigma" - while1: "[| c \ com; -b-> 1; -c-> sigma2; - <\ b \ c, sigma2> -c-> sigma1 |] - ==> <\ b \ c, sigma> -c-> sigma1" + while1: "\c \ com; -b-> 1; -c-> sigma2; + <\ b \ c, sigma2> -c-> sigma1\ + \ <\ b \ c, sigma> -c-> sigma1" type_intros com.intros update_type type_elims evala.dom_subset [THEN subsetD, elim_format] diff -r f2094906e491 -r e44d86131648 src/ZF/IMP/Denotation.thy --- a/src/ZF/IMP/Denotation.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/IMP/Denotation.thy Tue Sep 27 16:51:35 2022 +0100 @@ -15,7 +15,7 @@ definition Gamma :: "[i,i,i] => i" (\\\) where - "\(b,cden) == + "\(b,cden) \ (\phi. {io \ (phi O cden). B(b,fst(io))=1} \ {io \ id(loc->nat). B(b,fst(io))=0})" @@ -45,28 +45,28 @@ subsection \Misc lemmas\ -lemma A_type [TC]: "[|a \ aexp; sigma \ loc->nat|] ==> A(a,sigma) \ nat" +lemma A_type [TC]: "\a \ aexp; sigma \ loc->nat\ \ A(a,sigma) \ nat" by (erule aexp.induct) simp_all -lemma B_type [TC]: "[|b \ bexp; sigma \ loc->nat|] ==> B(b,sigma) \ bool" +lemma B_type [TC]: "\b \ bexp; sigma \ loc->nat\ \ B(b,sigma) \ bool" by (erule bexp.induct, simp_all) -lemma C_subset: "c \ com ==> C(c) \ (loc->nat) \ (loc->nat)" +lemma C_subset: "c \ com \ C(c) \ (loc->nat) \ (loc->nat)" apply (erule com.induct) apply simp_all apply (blast dest: lfp_subset [THEN subsetD])+ done lemma C_type_D [dest]: - "[| \ C(c); c \ com |] ==> x \ loc->nat & y \ loc->nat" + "\ \ C(c); c \ com\ \ x \ loc->nat & y \ loc->nat" by (blast dest: C_subset [THEN subsetD]) -lemma C_type_fst [dest]: "[| x \ C(c); c \ com |] ==> fst(x) \ loc->nat" +lemma C_type_fst [dest]: "\x \ C(c); c \ com\ \ fst(x) \ loc->nat" by (auto dest!: C_subset [THEN subsetD]) lemma Gamma_bnd_mono: "cden \ (loc->nat) \ (loc->nat) - ==> bnd_mono ((loc->nat) \ (loc->nat), \(b,cden))" + \ bnd_mono ((loc->nat) \ (loc->nat), \(b,cden))" by (unfold bnd_mono_def Gamma_def) blast end diff -r f2094906e491 -r e44d86131648 src/ZF/IMP/Equiv.thy --- a/src/ZF/IMP/Equiv.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/IMP/Equiv.thy Tue Sep 27 16:51:35 2022 +0100 @@ -7,8 +7,8 @@ theory Equiv imports Denotation Com begin lemma aexp_iff [rule_format]: - "[| a \ aexp; sigma: loc -> nat |] - ==> \n. -a-> n \ A(a,sigma) = n" + "\a \ aexp; sigma: loc -> nat\ + \ \n. -a-> n \ A(a,sigma) = n" apply (erule aexp.induct) apply (force intro!: evala.intros)+ done @@ -26,8 +26,8 @@ lemma bexp_iff [rule_format]: - "[| b \ bexp; sigma: loc -> nat |] - ==> \w. -b-> w \ B(b,sigma) = w" + "\b \ bexp; sigma: loc -> nat\ + \ \w. -b-> w \ B(b,sigma) = w" apply (erule bexp.induct) apply (auto intro!: evalb.intros) done @@ -35,7 +35,7 @@ declare bexp_iff [THEN iffD1, simp] bexp_iff [THEN iffD2, intro!] -lemma com1: " -c-> sigma' ==> \ C(c)" +lemma com1: " -c-> sigma' \ \ C(c)" apply (erule evalc.induct) apply (simp_all (no_asm_simp)) txt \\assign\\ @@ -54,7 +54,7 @@ declare evalc.intros [intro] -lemma com2 [rule_format]: "c \ com ==> \x \ C(c). -c-> snd(x)" +lemma com2 [rule_format]: "c \ com \ \x \ C(c). -c-> snd(x)" apply (erule com.induct) txt \\skip\\ apply force @@ -76,7 +76,7 @@ subsection \Main theorem\ theorem com_equivalence: - "c \ com ==> C(c) = {io \ (loc->nat) \ (loc->nat). -c-> snd(io)}" + "c \ com \ C(c) = {io \ (loc->nat) \ (loc->nat). -c-> snd(io)}" by (force intro: C_subset [THEN subsetD] elim: com2 dest: com1) end diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Acc.thy --- a/src/ZF/Induct/Acc.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Acc.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,7 +17,7 @@ inductive domains "acc(r)" \ "field(r)" intros - vimage: "[| r-``{a}: Pow(acc(r)); a \ field(r) |] ==> a \ acc(r)" + vimage: "\r-``{a}: Pow(acc(r)); a \ field(r)\ \ a \ acc(r)" monos Pow_mono text \ @@ -28,16 +28,16 @@ The intended introduction rule: \ -lemma accI: "[| !!b. :r ==> b \ acc(r); a \ field(r) |] ==> a \ acc(r)" +lemma accI: "\\b. :r \ b \ acc(r); a \ field(r)\ \ a \ acc(r)" by (blast intro: acc.intros) -lemma acc_downward: "[| b \ acc(r); : r |] ==> a \ acc(r)" +lemma acc_downward: "\b \ acc(r); : r\ \ a \ acc(r)" by (erule acc.cases) blast lemma acc_induct [consumes 1, case_names vimage, induct set: acc]: - "[| a \ acc(r); - !!x. [| x \ acc(r); \y. :r \ P(y) |] ==> P(x) - |] ==> P(a)" + "\a \ acc(r); + \x. \x \ acc(r); \y. :r \ P(y)\ \ P(x) +\ \ P(a)" by (erule acc.induct) (blast intro: acc.intros) lemma wf_on_acc: "wf[acc(r)](r)" @@ -49,7 +49,7 @@ lemma acc_wfI: "field(r) \ acc(r) \ wf(r)" by (erule wf_on_acc [THEN wf_on_subset_A, THEN wf_on_field_imp_wf]) -lemma acc_wfD: "wf(r) ==> field(r) \ acc(r)" +lemma acc_wfD: "wf(r) \ field(r) \ acc(r)" apply (rule subsetI) apply (erule wf_induct2, assumption) apply (blast intro: accI)+ diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Binary_Trees.thy --- a/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,7 +17,7 @@ declare bt.intros [simp] -lemma Br_neq_left: "l \ bt(A) ==> Br(x, l, r) \ l" +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'" @@ -32,7 +32,7 @@ definitions. \ -lemma bt_mono: "A \ B ==> bt(A) \ bt(B)" +lemma bt_mono: "A \ B \ bt(A) \ bt(B)" apply (unfold bt.defs) apply (rule lfp_mono) apply (rule bt.bnd_mono)+ @@ -46,18 +46,18 @@ 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)" +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); + "\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) |] ==> + \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)" +\ \ bt_rec(c, h, t) \ C(t)" \ \Type checking for recursor -- example only; not really needed.\ apply (induct_tac t) apply simp_all @@ -71,7 +71,7 @@ "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" +lemma n_nodes_type [simp]: "t \ bt(A) \ n_nodes(t) \ nat" by (induct set: bt) auto consts n_nodes_aux :: "i => i" @@ -81,7 +81,7 @@ (\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" + "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) @@ -89,9 +89,9 @@ definition n_nodes_tail :: "i => i" where - "n_nodes_tail(t) == n_nodes_aux(t) ` 0" + "n_nodes_tail(t) \ n_nodes_aux(t) ` 0" -lemma "t \ bt(A) ==> n_nodes_tail(t) = n_nodes(t)" +lemma "t \ bt(A) \ n_nodes_tail(t) = n_nodes(t)" by (simp add: n_nodes_tail_def n_nodes_aux_eq) @@ -103,7 +103,7 @@ "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" +lemma n_leaves_type [simp]: "t \ bt(A) \ n_leaves(t) \ nat" by (induct set: bt) auto @@ -115,24 +115,24 @@ "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)" +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)" +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))" +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" +lemma bt_reflect_bt_reflect_ident: "t \ bt(A) \ bt_reflect(bt_reflect(t)) = t" by (induct set: bt) simp_all end diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Brouwer.thy --- a/src/ZF/Induct/Brouwer.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Brouwer.thy Tue Sep 27 16:51:35 2022 +0100 @@ -25,8 +25,8 @@ assumes b: "b \ brouwer" and cases: "P(Zero)" - "!!b. [| b \ brouwer; P(b) |] ==> P(Suc(b))" - "!!h. [| h \ nat -> brouwer; \i \ nat. P(h`i) |] ==> P(Lim(h))" + "\b. \b \ brouwer; P(b)\ \ P(Suc(b))" + "\h. \h \ nat -> brouwer; \i \ nat. P(h`i)\ \ P(Lim(h))" shows "P(b)" \ \A nicer induction rule than the standard one.\ using b @@ -57,7 +57,7 @@ lemma Well_induct2 [consumes 1, case_names step]: assumes w: "w \ Well(A, B)" - and step: "!!a f. [| a \ A; f \ B(a) -> Well(A,B); \y \ B(a). P(f`y) |] ==> P(Sup(a,f))" + and step: "\a f. \a \ A; f \ B(a) -> Well(A,B); \y \ B(a). P(f`y)\ \ P(Sup(a,f))" shows "P(w)" \ \A nicer induction rule than the standard one.\ using w diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Comb.thy --- a/src/ZF/Induct/Comb.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Comb.thy Tue Sep 27 16:51:35 2022 +0100 @@ -41,10 +41,10 @@ inductive domains "contract" \ "comb \ comb" intros - K: "[| p \ comb; q \ comb |] ==> K\p\q \\<^sup>1 p" - S: "[| p \ comb; q \ comb; r \ comb |] ==> S\p\q\r \\<^sup>1 (p\r)\(q\r)" - Ap1: "[| p\\<^sup>1q; r \ comb |] ==> p\r \\<^sup>1 q\r" - Ap2: "[| p\\<^sup>1q; r \ comb |] ==> r\p \\<^sup>1 r\q" + K: "\p \ comb; q \ comb\ \ K\p\q \\<^sup>1 p" + S: "\p \ comb; q \ comb; r \ comb\ \ S\p\q\r \\<^sup>1 (p\r)\(q\r)" + Ap1: "\p\\<^sup>1q; r \ comb\ \ p\r \\<^sup>1 q\r" + Ap2: "\p\\<^sup>1q; r \ comb\ \ r\p \\<^sup>1 r\q" type_intros comb.intros text \ @@ -55,18 +55,18 @@ consts parcontract :: i abbreviation parcontract_syntax :: "[i,i] => o" (infixl \\\<^sup>1\ 50) - where "p \\<^sup>1 q == \ parcontract" + where "p \\<^sup>1 q \ \ parcontract" abbreviation parcontract_multi :: "[i,i] => o" (infixl \\\ 50) - where "p \ q == \ parcontract^+" + where "p \ q \ \ parcontract^+" inductive domains "parcontract" \ "comb \ comb" intros - refl: "[| p \ comb |] ==> p \\<^sup>1 p" - K: "[| p \ comb; q \ comb |] ==> K\p\q \\<^sup>1 p" - S: "[| p \ comb; q \ comb; r \ comb |] ==> S\p\q\r \\<^sup>1 (p\r)\(q\r)" - Ap: "[| p\\<^sup>1q; r\\<^sup>1s |] ==> p\r \\<^sup>1 q\s" + refl: "\p \ comb\ \ p \\<^sup>1 p" + K: "\p \ comb; q \ comb\ \ K\p\q \\<^sup>1 p" + S: "\p \ comb; q \ comb; r \ comb\ \ S\p\q\r \\<^sup>1 (p\r)\(q\r)" + Ap: "\p\\<^sup>1q; r\\<^sup>1s\ \ p\r \\<^sup>1 q\s" type_intros comb.intros text \ @@ -84,7 +84,7 @@ subsection \Transitive closure preserves the Church-Rosser property\ lemma diamond_strip_lemmaD [rule_format]: - "[| diamond(r); :r^+ |] ==> + "\diamond(r); :r^+\ \ \y'. :r \ (\z. : r^+ & : r)" apply (unfold diamond_def) apply (erule trancl_induct) @@ -94,7 +94,7 @@ apply (blast intro: r_into_trancl trans_trancl [THEN transD]) done -lemma diamond_trancl: "diamond(r) ==> diamond(r^+)" +lemma diamond_trancl: "diamond(r) \ diamond(r^+)" apply (simp (no_asm_simp) add: diamond_def) apply (rule impI [THEN allI, THEN allI]) apply (erule trancl_induct) @@ -134,7 +134,7 @@ contract.Ap1 [THEN rtrancl_into_rtrancl2] contract.Ap2 [THEN rtrancl_into_rtrancl2] -lemma "p \ comb ==> I\p \ p" +lemma "p \ comb \ I\p \ p" \ \Example only: not used\ unfolding I_def by (blast intro: reduction_rls) @@ -150,13 +150,13 @@ and S_contractE [elim!]: "S \\<^sup>1 r" and Ap_contractE [elim!]: "p\q \\<^sup>1 r" -lemma I_contract_E: "I \\<^sup>1 r ==> P" +lemma I_contract_E: "I \\<^sup>1 r \ P" by (auto simp add: I_def) -lemma K1_contractD: "K\p \\<^sup>1 r ==> (\q. r = K\q & p \\<^sup>1 q)" +lemma K1_contractD: "K\p \\<^sup>1 r \ (\q. r = K\q & p \\<^sup>1 q)" by auto -lemma Ap_reduce1: "[| p \ q; r \ comb |] ==> p\r \ q\r" +lemma Ap_reduce1: "\p \ q; r \ comb\ \ p\r \ q\r" apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) apply (erule rtrancl_induct) @@ -165,7 +165,7 @@ apply (blast intro: contract_combD2 reduction_rls) done -lemma Ap_reduce2: "[| p \ q; r \ comb |] ==> r\p \ r\q" +lemma Ap_reduce2: "\p \ q; r \ comb\ \ r\p \ r\q" apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) apply (erule rtrancl_induct) @@ -215,15 +215,15 @@ subsection \Basic properties of parallel contraction\ lemma K1_parcontractD [dest!]: - "K\p \\<^sup>1 r ==> (\p'. r = K\p' & p \\<^sup>1 p')" + "K\p \\<^sup>1 r \ (\p'. r = K\p' & p \\<^sup>1 p')" by auto lemma S1_parcontractD [dest!]: - "S\p \\<^sup>1 r ==> (\p'. r = S\p' & p \\<^sup>1 p')" + "S\p \\<^sup>1 r \ (\p'. r = S\p' & p \\<^sup>1 p')" by auto lemma S2_parcontractD [dest!]: - "S\p\q \\<^sup>1 r ==> (\p' q'. r = S\p'\q' & p \\<^sup>1 p' & q \\<^sup>1 q')" + "S\p\q \\<^sup>1 r \ (\p' q'. r = S\p'\q' & p \\<^sup>1 p' & q \\<^sup>1 q')" by auto lemma diamond_parcontract: "diamond(parcontract)" @@ -238,10 +238,10 @@ \medskip Equivalence of \<^prop>\p \ q\ and \<^prop>\p \ q\. \ -lemma contract_imp_parcontract: "p\\<^sup>1q ==> p\\<^sup>1q" +lemma contract_imp_parcontract: "p\\<^sup>1q \ p\\<^sup>1q" by (induct set: contract) auto -lemma reduce_imp_parreduce: "p\q ==> p\q" +lemma reduce_imp_parreduce: "p\q \ p\q" apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) apply (erule rtrancl_induct) @@ -250,7 +250,7 @@ trans_trancl [THEN transD]) done -lemma parcontract_imp_reduce: "p\\<^sup>1q ==> p\q" +lemma parcontract_imp_reduce: "p\\<^sup>1q \ p\q" apply (induct set: parcontract) apply (blast intro: reduction_rls) apply (blast intro: reduction_rls) @@ -259,7 +259,7 @@ Ap_reduce1 Ap_reduce2 parcontract_combD1 parcontract_combD2) done -lemma parreduce_imp_reduce: "p\q ==> p\q" +lemma parreduce_imp_reduce: "p\q \ p\q" apply (frule trancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_parcontract_eq [THEN equalityD1, THEN subsetD]) apply (erule trancl_induct, erule parcontract_imp_reduce) diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Datatypes.thy --- a/src/ZF/Induct/Datatypes.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Datatypes.thy Tue Sep 27 16:51:35 2022 +0100 @@ -31,7 +31,7 @@ type definitions. \ -lemma data_mono: "[| A \ C; B \ D |] ==> data(A, B) \ data(C, D)" +lemma data_mono: "\A \ C; B \ D\ \ data(A, B) \ data(C, D)" apply (unfold data.defs) apply (rule lfp_mono) apply (rule data.bnd_mono)+ @@ -46,7 +46,7 @@ done lemma data_subset_univ: - "[| A \ univ(C); B \ univ(C) |] ==> data(A, B) \ univ(C)" + "\A \ univ(C); B \ univ(C)\ \ data(A, B) \ univ(C)" by (rule subset_trans [OF data_mono data_univ]) diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/FoldSet.thy --- a/src/ZF/Induct/FoldSet.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/FoldSet.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,18 +13,18 @@ inductive domains "fold_set(A, B, f,e)" \ "Fin(A)*B" intros - emptyI: "e\B ==> <0, e>\fold_set(A, B, f,e)" - consI: "[| x\A; x \C; \ fold_set(A, B,f,e); f(x,y):B |] - ==> \fold_set(A, B, f, e)" + emptyI: "e\B \ <0, e>\fold_set(A, B, f,e)" + consI: "\x\A; x \C; \ fold_set(A, B,f,e); f(x,y):B\ + \ \fold_set(A, B, f, e)" type_intros Fin.intros definition fold :: "[i, [i,i]=>i, i, i] => i" (\fold[_]'(_,_,_')\) where - "fold[B](f,e, A) == THE x. \fold_set(A, B, f,e)" + "fold[B](f,e, A) \ THE x. \fold_set(A, B, f,e)" definition setsum :: "[i=>i, i] => i" where - "setsum(g, C) == if Finite(C) then + "setsum(g, C) \ if Finite(C) then fold[int](%x y. g(x) $+ y, #0, C) else #0" (** foldSet **) @@ -34,38 +34,38 @@ (* add-hoc lemmas *) -lemma cons_lemma1: "[| x\C; x\B |] ==> cons(x,B)=cons(x,C) \ B = C" +lemma cons_lemma1: "\x\C; x\B\ \ cons(x,B)=cons(x,C) \ B = C" by (auto elim: equalityE) -lemma cons_lemma2: "[| cons(x, B)=cons(y, C); x\y; x\B; y\C |] - ==> B - {y} = C-{x} & x\C & y\B" +lemma cons_lemma2: "\cons(x, B)=cons(y, C); x\y; x\B; y\C\ + \ B - {y} = C-{x} & x\C & y\B" apply (auto elim: equalityE) done (* fold_set monotonicity *) lemma fold_set_mono_lemma: " \ fold_set(A, B, f, e) - ==> \D. A<=D \ \ fold_set(D, B, f, e)" + \ \D. A<=D \ \ fold_set(D, B, f, e)" apply (erule fold_set.induct) apply (auto intro: fold_set.intros) done -lemma fold_set_mono: " C<=A ==> fold_set(C, B, f, e) \ fold_set(A, B, f, e)" +lemma fold_set_mono: " C<=A \ fold_set(C, B, f, e) \ fold_set(A, B, f, e)" apply clarify apply (frule fold_set.dom_subset [THEN subsetD], clarify) apply (auto dest: fold_set_mono_lemma) done lemma fold_set_lemma: - "\fold_set(A, B, f, e) ==> \fold_set(C, B, f, e) & C<=A" + "\fold_set(A, B, f, e) \ \fold_set(C, B, f, e) & C<=A" apply (erule fold_set.induct) apply (auto intro!: fold_set.intros intro: fold_set_mono [THEN subsetD]) done (* Proving that fold_set is deterministic *) lemma Diff1_fold_set: - "[| \ fold_set(A, B, f,e); x\C; x\A; f(x, y):B |] - ==> \ fold_set(A, B, f, e)" + "\ \ fold_set(A, B, f,e); x\C; x\A; f(x, y):B\ + \ \ fold_set(A, B, f, e)" apply (frule fold_set.dom_subset [THEN subsetD]) apply (erule cons_Diff [THEN subst], rule fold_set.intros, auto) done @@ -73,25 +73,25 @@ locale fold_typing = fixes A and B and e and f - assumes ftype [intro,simp]: "[|x \ A; y \ B|] ==> f(x,y) \ B" + assumes ftype [intro,simp]: "\x \ A; y \ B\ \ f(x,y) \ B" and etype [intro,simp]: "e \ B" - and fcomm: "[|x \ A; y \ A; z \ B|] ==> f(x, f(y, z))=f(y, f(x, z))" + and fcomm: "\x \ A; y \ A; z \ B\ \ f(x, f(y, z))=f(y, f(x, z))" lemma (in fold_typing) Fin_imp_fold_set: - "C\Fin(A) ==> (\x. \ fold_set(A, B, f,e))" + "C\Fin(A) \ (\x. \ fold_set(A, B, f,e))" apply (erule Fin_induct) apply (auto dest: fold_set.dom_subset [THEN subsetD] intro: fold_set.intros etype ftype) done lemma Diff_sing_imp: - "[|C - {b} = D - {a}; a \ b; b \ C|] ==> C = cons(b,D) - {a}" + "\C - {b} = D - {a}; a \ b; b \ C\ \ C = cons(b,D) - {a}" by (blast elim: equalityE) lemma (in fold_typing) fold_set_determ_lemma [rule_format]: "n\nat - ==> \C. |C| + \ \C. |C| (\x. \ fold_set(A, B, f,e)\ (\y. \ fold_set(A, B, f,e) \ y=x))" apply (erule nat_induct) @@ -131,8 +131,8 @@ done lemma (in fold_typing) fold_set_determ: - "[| \fold_set(A, B, f, e); - \fold_set(A, B, f, e)|] ==> y=x" + "\\fold_set(A, B, f, e); + \fold_set(A, B, f, e)\ \ y=x" apply (frule fold_set.dom_subset [THEN subsetD], clarify) apply (drule Fin_into_Finite) apply (unfold Finite_def, clarify) @@ -143,7 +143,7 @@ (** The fold function **) lemma (in fold_typing) fold_equality: - " \ fold_set(A,B,f,e) ==> fold[B](f,e,C) = y" + " \ fold_set(A,B,f,e) \ fold[B](f,e,C) = y" apply (unfold fold_def) apply (frule fold_set.dom_subset [THEN subsetD], clarify) apply (rule the_equality) @@ -154,15 +154,15 @@ apply (auto dest: fold_set_lemma intro: ftype etype fcomm) done -lemma fold_0 [simp]: "e \ B ==> fold[B](f,e,0) = e" +lemma fold_0 [simp]: "e \ B \ fold[B](f,e,0) = e" apply (unfold fold_def) apply (blast elim!: empty_fold_setE intro: fold_set.intros) done text\This result is the right-to-left direction of the subsequent result\ lemma (in fold_typing) fold_set_imp_cons: - "[| \ fold_set(C, B, f, e); C \ Fin(A); c \ A; c\C |] - ==> \ fold_set(cons(c, C), B, f, e)" + "\ \ fold_set(C, B, f, e); C \ Fin(A); c \ A; c\C\ + \ \ fold_set(cons(c, C), B, f, e)" apply (frule FinD [THEN fold_set_mono, THEN subsetD]) apply assumption apply (frule fold_set.dom_subset [of A, THEN subsetD]) @@ -170,8 +170,8 @@ done lemma (in fold_typing) fold_cons_lemma [rule_format]: -"[| C \ Fin(A); c \ A; c\C |] - ==> \ fold_set(cons(c, C), B, f, e) \ +"\C \ Fin(A); c \ A; c\C\ + \ \ fold_set(cons(c, C), B, f, e) \ (\y. \ fold_set(C, B, f, e) & v = f(c, y))" apply auto prefer 2 apply (blast intro: fold_set_imp_cons) @@ -191,8 +191,8 @@ done lemma (in fold_typing) fold_cons: - "[| C\Fin(A); c\A; c\C|] - ==> fold[B](f, e, cons(c, C)) = f(c, fold[B](f, e, C))" + "\C\Fin(A); c\A; c\C\ + \ fold[B](f, e, cons(c, C)) = f(c, fold[B](f, e, C))" apply (unfold fold_def) apply (simp add: fold_cons_lemma) apply (rule the_equality, auto) @@ -203,14 +203,14 @@ done lemma (in fold_typing) fold_type [simp,TC]: - "C\Fin(A) ==> fold[B](f,e,C):B" + "C\Fin(A) \ fold[B](f,e,C):B" apply (erule Fin_induct) apply (simp_all add: fold_cons ftype etype) done lemma (in fold_typing) fold_commute [rule_format]: - "[| C\Fin(A); c\A |] - ==> (\y\B. f(c, fold[B](f, y, C)) = fold[B](f, f(c, y), C))" + "\C\Fin(A); c\A\ + \ (\y\B. f(c, fold[B](f, y, C)) = fold[B](f, f(c, y), C))" apply (erule Fin_induct) apply (simp_all add: fold_typing.fold_cons [of A B _ f] fold_typing.fold_type [of A B _ f] @@ -218,8 +218,8 @@ done lemma (in fold_typing) fold_nest_Un_Int: - "[| C\Fin(A); D\Fin(A) |] - ==> fold[B](f, fold[B](f, e, D), C) = + "\C\Fin(A); D\Fin(A)\ + \ fold[B](f, fold[B](f, e, D), C) = fold[B](f, fold[B](f, e, (C \ D)), C \ D)" apply (erule Fin_induct, auto) apply (simp add: Un_cons Int_cons_left fold_type fold_commute @@ -228,11 +228,11 @@ done lemma (in fold_typing) fold_nest_Un_disjoint: - "[| C\Fin(A); D\Fin(A); C \ D = 0 |] - ==> fold[B](f,e,C \ D) = fold[B](f, fold[B](f,e,D), C)" + "\C\Fin(A); D\Fin(A); C \ D = 0\ + \ fold[B](f,e,C \ D) = fold[B](f, fold[B](f,e,D), C)" by (simp add: fold_nest_Un_Int) -lemma Finite_cons_lemma: "Finite(C) ==> C\Fin(cons(c, C))" +lemma Finite_cons_lemma: "Finite(C) \ C\Fin(cons(c, C))" apply (drule Finite_into_Fin) apply (blast intro: Fin_mono [THEN subsetD]) done @@ -243,7 +243,7 @@ by (simp add: setsum_def) lemma setsum_cons [simp]: - "Finite(C) ==> + "Finite(C) \ setsum(g, cons(c,C)) = (if c \ C then setsum(g,C) else g(c) $+ setsum(g,C))" apply (auto simp add: setsum_def Finite_cons cons_absorb) @@ -259,8 +259,8 @@ (*The reversed orientation looks more natural, but LOOPS as a simprule!*) lemma setsum_Un_Int: - "[| Finite(C); Finite(D) |] - ==> setsum(g, C \ D) $+ setsum(g, C \ D) + "\Finite(C); Finite(D)\ + \ setsum(g, C \ D) $+ setsum(g, C \ D) = setsum(g, C) $+ setsum(g, D)" apply (erule Finite_induct) apply (simp_all add: Int_cons_right cons_absorb Un_cons Int_commute Finite_Un @@ -274,21 +274,21 @@ done lemma setsum_Un_disjoint: - "[| Finite(C); Finite(D); C \ D = 0 |] - ==> setsum(g, C \ D) = setsum(g, C) $+ setsum(g,D)" + "\Finite(C); Finite(D); C \ D = 0\ + \ setsum(g, C \ D) = setsum(g, C) $+ setsum(g,D)" apply (subst setsum_Un_Int [symmetric]) apply (subgoal_tac [3] "Finite (C \ D) ") apply (auto intro: Finite_Un) done lemma Finite_RepFun [rule_format (no_asm)]: - "Finite(I) ==> (\i\I. Finite(C(i))) \ Finite(RepFun(I, C))" + "Finite(I) \ (\i\I. Finite(C(i))) \ Finite(RepFun(I, C))" apply (erule Finite_induct, auto) done lemma setsum_UN_disjoint [rule_format (no_asm)]: "Finite(I) - ==> (\i\I. Finite(C(i))) \ + \ (\i\I. Finite(C(i))) \ (\i\I. \j\I. i\j \ C(i) \ C(j) = 0) \ setsum(f, \i\I. C(i)) = setsum (%i. setsum(f, C(i)), I)" apply (erule Finite_induct, auto) @@ -310,15 +310,15 @@ lemma fold_set_cong: - "[| A=A'; B=B'; e=e'; (\x\A'. \y\B'. f(x,y) = f'(x,y)) |] - ==> fold_set(A,B,f,e) = fold_set(A',B',f',e')" + "\A=A'; B=B'; e=e'; (\x\A'. \y\B'. f(x,y) = f'(x,y))\ + \ fold_set(A,B,f,e) = fold_set(A',B',f',e')" apply (simp add: fold_set_def) apply (intro refl iff_refl lfp_cong Collect_cong disj_cong ex_cong, auto) done lemma fold_cong: -"[| B=B'; A=A'; e=e'; - !!x y. [|x\A'; y\B'|] ==> f(x,y) = f'(x,y) |] ==> +"\B=B'; A=A'; e=e'; + \x y. \x\A'; y\B'\ \ f(x,y) = f'(x,y)\ \ fold[B](f,e,A) = fold[B'](f', e', A')" apply (simp add: fold_def) apply (subst fold_set_cong) @@ -326,27 +326,27 @@ done lemma setsum_cong: - "[| A=B; !!x. x\B ==> f(x) = g(x) |] ==> + "\A=B; \x. x\B \ f(x) = g(x)\ \ setsum(f, A) = setsum(g, B)" by (simp add: setsum_def cong add: fold_cong) lemma setsum_Un: - "[| Finite(A); Finite(B) |] - ==> setsum(f, A \ B) = + "\Finite(A); Finite(B)\ + \ setsum(f, A \ B) = setsum(f, A) $+ setsum(f, B) $- setsum(f, A \ B)" apply (subst setsum_Un_Int [symmetric], auto) done lemma setsum_zneg_or_0 [rule_format (no_asm)]: - "Finite(A) ==> (\x\A. g(x) $\ #0) \ setsum(g, A) $\ #0" + "Finite(A) \ (\x\A. g(x) $\ #0) \ setsum(g, A) $\ #0" apply (erule Finite_induct) apply (auto intro: zneg_or_0_add_zneg_or_0_imp_zneg_or_0) done lemma setsum_succD_lemma [rule_format]: "Finite(A) - ==> \n\nat. setsum(f,A) = $# succ(n) \ (\a\A. #0 $< f(a))" + \ \n\nat. setsum(f,A) = $# succ(n) \ (\a\A. #0 $< f(a))" apply (erule Finite_induct) apply (auto simp del: int_of_0 int_of_succ simp add: not_zless_iff_zle int_of_0 [symmetric]) apply (subgoal_tac "setsum (f, B) $\ #0") @@ -360,7 +360,7 @@ done lemma setsum_succD: - "[| setsum(f, A) = $# succ(n); n\nat |]==> \a\A. #0 $< f(a)" + "\setsum(f, A) = $# succ(n); n\nat\\ \a\A. #0 $< f(a)" apply (case_tac "Finite (A) ") apply (blast intro: setsum_succD_lemma) apply (unfold setsum_def) @@ -368,27 +368,27 @@ done lemma g_zpos_imp_setsum_zpos [rule_format]: - "Finite(A) ==> (\x\A. #0 $\ g(x)) \ #0 $\ setsum(g, A)" + "Finite(A) \ (\x\A. #0 $\ g(x)) \ #0 $\ setsum(g, A)" apply (erule Finite_induct) apply (simp (no_asm)) apply (auto intro: zpos_add_zpos_imp_zpos) done lemma g_zpos_imp_setsum_zpos2 [rule_format]: - "[| Finite(A); \x. #0 $\ g(x) |] ==> #0 $\ setsum(g, A)" + "\Finite(A); \x. #0 $\ g(x)\ \ #0 $\ setsum(g, A)" apply (erule Finite_induct) apply (auto intro: zpos_add_zpos_imp_zpos) done lemma g_zspos_imp_setsum_zspos [rule_format]: "Finite(A) - ==> (\x\A. #0 $< g(x)) \ A \ 0 \ (#0 $< setsum(g, A))" + \ (\x\A. #0 $< g(x)) \ A \ 0 \ (#0 $< setsum(g, A))" apply (erule Finite_induct) apply (auto intro: zspos_add_zspos_imp_zspos) done lemma setsum_Diff [rule_format]: - "Finite(A) ==> \a. M(a) = #0 \ setsum(M, A) = setsum(M, A-{a})" + "Finite(A) \ \a. M(a) = #0 \ setsum(M, A) = setsum(M, A-{a})" apply (erule Finite_induct) apply (simp_all add: Diff_cons_eq Finite_Diff) done diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/ListN.thy --- a/src/ZF/Induct/ListN.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/ListN.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,11 +17,11 @@ domains "listn(A)" \ "nat \ list(A)" intros NilI: "<0,Nil> \ listn(A)" - ConsI: "[| a \ A; \ listn(A) |] ==> \ listn(A)" + ConsI: "\a \ A; \ listn(A)\ \ \ listn(A)" type_intros nat_typechecks list.intros -lemma list_into_listn: "l \ list(A) ==> \ listn(A)" +lemma list_into_listn: "l \ list(A) \ \ listn(A)" by (induct set: list) (simp_all add: listn.intros) lemma listn_iff: " \ listn(A) \ l \ list(A) & length(l)=n" @@ -36,7 +36,7 @@ apply (simp add: listn_iff separation image_singleton_iff) done -lemma listn_mono: "A \ B ==> listn(A) \ listn(B)" +lemma listn_mono: "A \ B \ listn(A) \ listn(B)" apply (unfold listn.defs) apply (rule lfp_mono) apply (rule listn.bnd_mono)+ @@ -44,7 +44,7 @@ done lemma listn_append: - "[| \ listn(A); \ listn(A) |] ==> \ listn(A)" + "\ \ listn(A); \ listn(A)\ \ \ listn(A)" apply (erule listn.induct) apply (frule listn.dom_subset [THEN subsetD]) apply (simp_all add: listn.intros) diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Multiset.thy Tue Sep 27 16:51:35 2022 +0100 @@ -14,64 +14,64 @@ abbreviation (input) \ \Short cut for multiset space\ Mult :: "i=>i" where - "Mult(A) == A -||> nat-{0}" + "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" + "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)" + "multiset(M) \ \A. M \ A -> nat-{0} & Finite(A)" definition mset_of :: "i=>i" where - "mset_of(M) == domain(M)" + "mset_of(M) \ domain(M)" definition munion :: "[i, i] => i" (infixl \+#\ 65) where - "M +# N == \x \ mset_of(M) \ mset_of(N). + "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) == + "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). + "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#} \ {}" definition MCollect :: "[i, i=>o] => i" (*comprehension*) where - "MCollect(M, P) == funrestrict(M, {x \ mset_of(M). P(x)})" + "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" + "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))" + "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)" + "a :# M \ a \ mset_of(M)" syntax "_MColl" :: "[pttrn, i, o] => i" (\(1{# _ \ _./ _#})\) @@ -84,43 +84,43 @@ (* 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) == + "multirel1(A, r) \ { \ Mult(A)*Mult(A). \a \ A. \M0 \ Mult(A). \K \ Mult(A). N=M0 +# {#a#} & M=M0 +# K & (\b \ mset_of(K). \ r)}" definition multirel :: "[i, i] => i" where - "multirel(A, r) == multirel1(A, r)^+" + "multirel(A, r) \ multirel1(A, r)^+" (* ordinal multiset orderings *) definition omultiset :: "i => o" where - "omultiset(M) == \i. Ord(i) & M \ Mult(field(Memrel(i)))" + "omultiset(M) \ \i. Ord(i) & M \ Mult(field(Memrel(i)))" definition mless :: "[i, i] => o" (infixl \<#\ 50) where - "M <# N == \i. Ord(i) & \ multirel(field(Memrel(i)), Memrel(i))" + "M <# N \ \i. Ord(i) & \ multirel(field(Memrel(i)), Memrel(i))" definition mle :: "[i, i] => o" (infixl \<#=\ 50) where - "M <#= N == (omultiset(M) & M = N) | M <# N" + "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" +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)" + "\\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)" +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" +lemma funrestrict [simp]: "a \ A \ funrestrict(f,A) ` a = f`a" by (simp add: funrestrict_def) lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0" @@ -130,7 +130,7 @@ by (auto simp add: funrestrict_def lam_def) lemma fun_cons_funrestrict_eq: - "f \ cons(a, b) -> B ==> f = cons(, funrestrict(f, b))" + "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) @@ -151,7 +151,7 @@ done (** The multiset space **) -lemma multiset_into_Mult: "[| multiset(M); mset_of(M)\A |] ==> M \ Mult(A)" +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) @@ -159,7 +159,7 @@ apply (simp_all (no_asm_simp) add: multiset_fun_iff) done -lemma Mult_into_multiset: "M \ Mult(A) ==> multiset(M) & mset_of(M)\A" +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) @@ -185,13 +185,13 @@ text\The \<^term>\mset_of\ operator\ -lemma multiset_set_of_Finite [simp]: "multiset(M) ==> Finite(mset_of(M))" +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" +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}" @@ -200,7 +200,7 @@ 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" +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 *) @@ -230,7 +230,7 @@ apply (auto intro!: lam_type simp add: Collect_Finite) done -lemma normalize_multiset [simp]: "multiset(M) ==> normalize(M) = M" +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))" @@ -244,7 +244,7 @@ (* union *) -lemma munion_multiset [simp]: "[| multiset(M); multiset(N) |] ==> multiset(M +# N)" +lemma munion_multiset [simp]: "\multiset(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) @@ -259,7 +259,7 @@ (* Union *) -lemma munion_0 [simp]: "multiset(M) ==> M +# 0 = M & 0 +# M = M" +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 @@ -287,10 +287,10 @@ lemma mdiff_0 [simp]: "0 -# M = 0" by (simp add: mdiff_def normalize_def) -lemma mdiff_0_right [simp]: "multiset(M) ==> M -# 0 = M" +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" +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) 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) @@ -303,7 +303,7 @@ (** Count of elements **) -lemma mcount_type [simp,TC]: "multiset(M) ==> mcount(M, a) \ nat" +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" @@ -312,13 +312,13 @@ 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)" +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)" + "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) @@ -326,7 +326,7 @@ apply (force intro!: lam_type) done -lemma mcount_elem: "[| multiset(M); a \ mset_of(M) |] ==> 0 < mcount(M, a)" +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) @@ -343,23 +343,23 @@ lemma msize_type [simp,TC]: "msize(M) \ int" by (simp add: msize_def) -lemma msize_zpositive: "multiset(M)==> #0 $\ msize(M)" +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" +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)" + "\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" +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 @@ -376,7 +376,7 @@ done lemma setsum_mcount_Int: - "Finite(A) ==> setsum(%a. $# mcount(N, a), A \ mset_of(N)) + "Finite(A) \ setsum(%a. $# mcount(N, a), A \ mset_of(N)) = setsum(%a. $# mcount(N, a), A)" apply (induct rule: Finite_induct) apply auto @@ -386,13 +386,13 @@ done lemma msize_union [simp]: - "[| multiset(M); multiset(N) |] ==> msize(M +# N) = msize(M) $+ msize(N)" + "\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)" +lemma msize_eq_succ_imp_elem: "\msize(M)= $# succ(n); n \ nat\ \ \a. a \ mset_of(M)" apply (unfold msize_def) apply (blast dest: setsum_succD) done @@ -400,8 +400,8 @@ (** Equality of multisets **) lemma equality_lemma: - "[| multiset(M); multiset(N); \a. mcount(M, a)=mcount(N, a) |] - ==> mset_of(M)=mset_of(N)" + "\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) @@ -410,7 +410,7 @@ done lemma multiset_equality: - "[| multiset(M); multiset(N) |]==> M=N\(\a. mcount(M, a)=mcount(N, a))" + "\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) @@ -424,28 +424,28 @@ (** More algebraic properties of multisets **) -lemma munion_eq_0_iff [simp]: "[|multiset(M); multiset(N)|]==>(M +# N =0) \ (M=0 & N=0)" +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)" +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)" + "\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)" + "\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)" +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#})" + "\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 @@ -464,8 +464,8 @@ 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)" +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 @@ -476,7 +476,7 @@ lemma setsum_decr: "Finite(A) - ==> (\M. multiset(M) \ + \ (\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))))" @@ -492,7 +492,7 @@ lemma setsum_decr2: "Finite(A) - ==> \M. multiset(M) \ (\a \ mset_of(M). + \ \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)))" @@ -501,8 +501,8 @@ 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}) = +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) ") @@ -512,10 +512,10 @@ apply (simp add: mcount_def mset_of_def) done -lemma nat_le_1_cases: "n \ nat ==> n \ 1 \ (n=0 | n=1)" +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" +lemma succ_pred_eq_self: "\0 nat\ \ succ(n #- 1) = n" apply (subgoal_tac "1 \ n") apply (drule add_diff_inverse2, auto) done @@ -530,11 +530,11 @@ done lemma multiset_induct_aux: - assumes prem1: "!!M a. [| multiset(M); a\mset_of(M); P(M) |] ==> P(cons(, M))" - and prem2: "!!M b. [| multiset(M); b \ mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))" + assumes prem1: "\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(, 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)\ + "\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) @@ -600,10 +600,10 @@ done lemma multiset_induct2: - "[| multiset(M); P(0); - (!!M a. [| multiset(M); a\mset_of(M); P(M) |] ==> P(cons(, M))); - (!!M b. [| multiset(M); b \ mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))) |] - ==> P(M)" + "\multiset(M); P(0); + (\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(, 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) @@ -614,7 +614,7 @@ done lemma munion_single_case1: - "[| multiset(M); a \mset_of(M) |] ==> M +# {#a#} = cons(, M)" + "\multiset(M); a \mset_of(M)\ \ M +# {#a#} = cons(, M)" apply (simp add: multiset_def msingle_def) apply (auto simp add: munion_def) apply (unfold mset_of_def, simp) @@ -625,7 +625,7 @@ done lemma munion_single_case2: - "[| multiset(M); a \ mset_of(M) |] ==> M +# {#a#} = M(a:=M`a #+ 1)" + "\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) @@ -639,7 +639,7 @@ lemma multiset_induct: assumes M: "multiset(M)" and P0: "P(0)" - and step: "!!M a. [| multiset(M); P(M) |] ==> P(M +# {#a#})" + 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) @@ -651,7 +651,7 @@ (** MCollect **) lemma MCollect_multiset [simp]: - "multiset(M) ==> multiset({# x \ M. P(x)#})" + "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] @@ -659,7 +659,7 @@ done lemma mset_of_MCollect [simp]: - "multiset(M) ==> mset_of({# x \ M. P(x) #}) \ mset_of(M)" + "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]: @@ -670,17 +670,17 @@ "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) #}" +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" + "\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 | +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) @@ -694,7 +694,7 @@ done lemma melem_diff_single: -"multiset(M) ==> +"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) @@ -706,8 +706,8 @@ done lemma munion_eq_conv_exist: -"[| M \ Mult(A); N \ Mult(A) |] - ==> (M +# {#a#} = N +# {#b#}) \ +"\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) @@ -735,7 +735,7 @@ text\Monotonicity of \<^term>\multirel1\\ -lemma multirel1_mono1: "A\B ==> multirel1(A, r)\multirel1(B, r)" +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) @@ -744,7 +744,7 @@ apply (auto simp add: Mult_iff_multiset) done -lemma multirel1_mono2: "r\s ==> multirel1(A,r)\multirel1(A, s)" +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) @@ -754,7 +754,7 @@ done lemma multirel1_mono: - "[| A\B; r\s |] ==> multirel1(A, r) \ multirel1(B, s)" + "\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) @@ -765,7 +765,7 @@ lemma not_less_0 [iff]: " \ multirel1(A, r)" by (auto simp add: multirel1_def Mult_iff_multiset) -lemma less_munion: "[| \ multirel1(A, r); M0 \ Mult(A) |] ==> +lemma less_munion: "\ \ multirel1(A, r); M0 \ Mult(A)\ \ (\M. \ multirel1(A, r) & N = M +# {#a#}) | (\K. K \ Mult(A) & (\b \ mset_of(K). \ r) & N = M0 +# K)" apply (frule multirel1_type [THEN subsetD]) @@ -776,7 +776,7 @@ apply (auto simp add: munion_commute) done -lemma multirel1_base: "[| M \ Mult(A); a \ A |] ==> \ multirel1(A, r)" +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) @@ -787,11 +787,11 @@ lemma acc_0: "acc(0)=0" by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD]) -lemma lemma1: "[| \b \ A. \ r \ +lemma lemma1: "\\b \ A. \ r \ (\M \ acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r))); M0 \ acc(multirel1(A, r)); a \ A; - \M. \ multirel1(A, r) \ M +# {#a#} \ acc(multirel1(A, r)) |] - ==> M0 +# {#a#} \ acc(multirel1(A, r))" + \M. \ 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) @@ -819,20 +819,20 @@ apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset) done -lemma lemma2: "[| \b \ A. \ r +lemma lemma2: "\\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))" + 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))" +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 \ +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 *) @@ -852,14 +852,14 @@ apply (auto simp add: Mult_iff_multiset) done -lemma all_accessible: "[| wf[A](r); M \ Mult(A); A \ 0|] ==> M \ acc(multirel1(A, r))" +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))" +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) @@ -870,7 +870,7 @@ apply (blast intro: all_accessible) done -lemma wf_multirel1: "wf(r) ==>wf(multirel1(field(r), r))" +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) @@ -889,7 +889,7 @@ (* Monotonicity of multirel *) lemma multirel_mono: - "[| A\B; r\s |] ==> multirel(A, r)\multirel(B,s)" + "\A\B; r\s\ \ multirel(A, r)\multirel(B,s)" apply (simp add: multirel_def) apply (rule trancl_mono) apply (rule multirel1_mono, auto) @@ -897,24 +897,24 @@ (* Equivalence of multirel with the usual (closure-free) definition *) -lemma add_diff_eq: "k \ nat ==> 0 < k \ n #+ k #- 1 = n #+ (k #- 1)" +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#})" +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" +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: -" \ multirel(A, r) ==> +" \ multirel(A, r) \ trans[A](r) \ (\I J K. I \ Mult(A) & J \ Mult(A) & K \ Mult(A) & @@ -964,12 +964,12 @@ apply (drule_tac [2] sym, auto) done -lemma melem_imp_eq_diff_union [simp]: "[| a \ mset_of(M); multiset(M) |] ==> M -# {#a#} +# {#a#} = M" +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" + "\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) @@ -981,7 +981,7 @@ (* The second direction *) lemma one_step_implies_multirel_lemma [rule_format (no_asm)]: -"n \ nat ==> +"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). \ r)) @@ -1024,9 +1024,9 @@ done lemma one_step_implies_multirel: - "[| J \ 0; \k \ mset_of(K). \j \ mset_of(J). \ r; - I \ Mult(A); J \ Mult(A); K \ Mult(A) |] - ==> \ multirel(A, r)" + "\J \ 0; \k \ mset_of(K). \j \ mset_of(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) @@ -1038,7 +1038,7 @@ (*irreflexivity*) lemma multirel_irrefl_lemma: - "Finite(A) ==> part_ord(A, r) \ (\x \ A. \y \ A. \ r) \A=0" + "Finite(A) \ part_ord(A, r) \ (\x \ A. \y \ A. \ 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) @@ -1047,7 +1047,7 @@ done lemma irrefl_on_multirel: - "part_ord(A, r) ==> irrefl(Mult(A), multirel(A, r))" + "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) @@ -1066,7 +1066,7 @@ done lemma multirel_trans: - "[| \ multirel(A, r); \ multirel(A, r) |] ==> \ multirel(A,r)" + "\ \ multirel(A, r); \ multirel(A, r)\ \ \ multirel(A,r)" apply (simp add: multirel_def) apply (blast intro: trancl_trans) done @@ -1076,7 +1076,7 @@ apply (rule trans_trancl) done -lemma part_ord_multirel: "part_ord(A,r) ==> part_ord(Mult(A), multirel(A, r))" +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 @@ -1084,7 +1084,7 @@ (** Monotonicity of multiset union **) lemma munion_multirel1_mono: -"[| \ multirel1(A, r); K \ Mult(A) |] ==> \ multirel1(A, r)" +"\ \ 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) @@ -1096,7 +1096,7 @@ done lemma munion_multirel_mono2: - "[| \ multirel(A, r); K \ Mult(A) |]==> \ multirel(A, r)" + "\ \ 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 @@ -1115,7 +1115,7 @@ done lemma munion_multirel_mono1: - "[| \ multirel(A, r); K \ Mult(A)|] ==> \ multirel(A, r)" + "\ \ multirel(A, r); K \ Mult(A)\ \ \ multirel(A, r)" apply (frule multirel_type [THEN subsetD]) apply (rule_tac P = "%x. \ multirel(A, r)" for u in munion_commute [THEN subst]) apply (subst munion_commute [of N]) @@ -1124,8 +1124,8 @@ done lemma munion_multirel_mono: - "[| \ multirel(A, r); \ multirel(A, r)|] - ==> \ multirel(A, r)" + "\ \ multirel(A, r); \ 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) @@ -1134,29 +1134,29 @@ subsection\Ordinal Multisets\ -(* A \ B ==> field(Memrel(A)) \ field(Memrel(B)) *) +(* A \ B \ field(Memrel(A)) \ field(Memrel(B)) *) lemmas field_Memrel_mono = Memrel_mono [THEN field_mono] (* -[| Aa \ Ba; A \ B |] ==> +\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)" +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)" +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)" +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) @@ -1165,7 +1165,7 @@ (** Proving that Memrel is a partial order **) -lemma irrefl_Memrel: "Ord(i) ==> irrefl(field(Memrel(i)), Memrel(i))" +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) @@ -1175,14 +1175,14 @@ 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))" +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) ==> + Ord(i) \ part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i))) *) @@ -1190,18 +1190,18 @@ (*irreflexivity*) -lemma mless_not_refl: "~(M <# M)" +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 *) +(* N R *) lemmas mless_irrefl = mless_not_refl [THEN notE, elim!] (*transitivity*) -lemma mless_trans: "[| K <# M; M <# N |] ==> K <# N" +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] @@ -1210,27 +1210,27 @@ done (*asymmetry*) -lemma mless_not_sym: "M <# N ==> ~ N <# M" +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" +lemma mless_asym: "\M <# N; \P \ N <# M\ \ P" by (blast dest: mless_not_sym) -lemma mle_refl [simp]: "omultiset(M) ==> M <#= M" +lemma mle_refl [simp]: "omultiset(M) \ M <#= M" by (simp add: mle_def) (*anti-symmetry*) lemma mle_antisym: - "[| M <#= N; N <#= M |] ==> M = N" + "\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" +lemma mle_trans: "\K <#= M; M <#= N\ \ K <#= N" apply (simp add: mle_def) apply (blast intro: mless_trans) done @@ -1240,7 +1240,7 @@ (** Monotonicity of mless **) -lemma munion_less_mono2: "[| M <# N; omultiset(K) |] ==> K +# M <# K +# N" +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) @@ -1250,13 +1250,13 @@ apply (blast intro: field_Memrel_mono [THEN subsetD]) done -lemma munion_less_mono1: "[| M <# N; omultiset(K) |] ==> M +# K <# N +# K" +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)" +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" +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) @@ -1264,10 +1264,10 @@ (* <#= *) -lemma mle_imp_omultiset: "M <#= N ==> omultiset(M) & omultiset(N)" +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" +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) @@ -1276,7 +1276,7 @@ lemma omultiset_0 [iff]: "omultiset(0)" by (auto simp add: omultiset_def Mult_iff_multiset) -lemma empty_leI [simp]: "omultiset(M) ==> 0 <#= M" +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) @@ -1286,7 +1286,7 @@ apply (auto simp add: Mult_iff_multiset) done -lemma munion_upper1: "[| omultiset(M); omultiset(N) |] ==> M <#= M +# N" +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 diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Mutil.thy --- a/src/ZF/Induct/Mutil.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Mutil.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,21 +19,21 @@ inductive domains "domino" \ "Pow(nat \ nat)" intros - horiz: "[| i \ nat; j \ nat |] ==> {, } \ domino" - vertl: "[| i \ nat; j \ nat |] ==> {, } \ domino" + horiz: "\i \ nat; j \ nat\ \ {, } \ domino" + vertl: "\i \ nat; j \ nat\ \ {, } \ domino" type_intros empty_subsetI cons_subsetI PowI SigmaI nat_succI inductive domains "tiling(A)" \ "Pow(\(A))" intros empty: "0 \ tiling(A)" - Un: "[| a \ A; t \ tiling(A); a \ t = 0 |] ==> a \ t \ tiling(A)" + Un: "\a \ A; t \ tiling(A); a \ t = 0\ \ a \ t \ tiling(A)" type_intros empty_subsetI Union_upper Un_least PowI type_elims PowD [elim_format] definition evnodd :: "[i, i] => i" where - "evnodd(A,b) == {z \ A. \i j. z = \ (i #+ j) mod 2 = b}" + "evnodd(A,b) \ {z \ A. \i j. z = \ (i #+ j) mod 2 = b}" subsection \Basic properties of evnodd\ @@ -44,7 +44,7 @@ lemma evnodd_subset: "evnodd(A, b) \ A" by (unfold evnodd_def) blast -lemma Finite_evnodd: "Finite(X) ==> Finite(evnodd(X,b))" +lemma Finite_evnodd: "Finite(X) \ Finite(evnodd(X,b))" by (rule lepoll_Finite, rule subset_imp_lepoll, rule evnodd_subset) lemma evnodd_Un: "evnodd(A \ B, b) = evnodd(A,b) \ evnodd(B,b)" @@ -64,11 +64,11 @@ subsection \Dominoes\ -lemma domino_Finite: "d \ domino ==> Finite(d)" +lemma domino_Finite: "d \ domino \ Finite(d)" by (blast intro!: Finite_cons Finite_0 elim: domino.cases) lemma domino_singleton: - "[| d \ domino; b<2 |] ==> \i' j'. evnodd(d,b) = {}" + "\d \ domino; b<2\ \ \i' j'. evnodd(d,b) = {}" apply (erule domino.cases) apply (rule_tac [2] k1 = "i#+j" in mod2_cases [THEN disjE]) apply (rule_tac k1 = "i#+j" in mod2_cases [THEN disjE]) @@ -83,20 +83,20 @@ text \The union of two disjoint tilings is a tiling\ lemma tiling_UnI: - "t \ tiling(A) ==> u \ tiling(A) ==> t \ u = 0 ==> t \ u \ tiling(A)" + "t \ tiling(A) \ u \ tiling(A) \ t \ u = 0 \ t \ u \ tiling(A)" apply (induct set: tiling) apply (simp add: tiling.intros) apply (simp add: Un_assoc subset_empty_iff [THEN iff_sym]) apply (blast intro: tiling.intros) done -lemma tiling_domino_Finite: "t \ tiling(domino) ==> Finite(t)" +lemma tiling_domino_Finite: "t \ tiling(domino) \ Finite(t)" apply (induct set: tiling) apply (rule Finite_0) apply (blast intro!: Finite_Un intro: domino_Finite) done -lemma tiling_domino_0_1: "t \ tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|" +lemma tiling_domino_0_1: "t \ tiling(domino) \ |evnodd(t,0)| = |evnodd(t,1)|" apply (induct set: tiling) apply (simp add: evnodd_def) apply (rule_tac b1 = 0 in domino_singleton [THEN exE]) @@ -115,7 +115,7 @@ done lemma dominoes_tile_row: - "[| i \ nat; n \ nat |] ==> {i} * (n #+ n) \ tiling(domino)" + "\i \ nat; n \ nat\ \ {i} * (n #+ n) \ tiling(domino)" apply (induct_tac n) apply (simp add: tiling.intros) apply (simp add: Un_assoc [symmetric] Sigma_succ2) @@ -131,20 +131,20 @@ done lemma dominoes_tile_matrix: - "[| m \ nat; n \ nat |] ==> m * (n #+ n) \ tiling(domino)" + "\m \ nat; n \ nat\ \ m * (n #+ n) \ tiling(domino)" apply (induct_tac m) apply (simp add: tiling.intros) apply (simp add: Sigma_succ1) apply (blast intro: tiling_UnI dominoes_tile_row elim: mem_irrefl) done -lemma eq_lt_E: "[| x=y; x P" +lemma eq_lt_E: "\x=y; x \ P" by auto -theorem mutil_not_tiling: "[| m \ nat; n \ nat; +theorem mutil_not_tiling: "\m \ nat; n \ nat; t = (succ(m)#+succ(m))*(succ(n)#+succ(n)); - t' = t - {<0,0>} - {} |] - ==> t' \ tiling(domino)" + t' = t - {<0,0>} - {}\ + \ t' \ tiling(domino)" apply (rule notI) apply (drule tiling_domino_0_1) apply (erule_tac x = "|A|" for A in eq_lt_E) diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Ntree.thy --- a/src/ZF/Induct/Ntree.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Ntree.thy Tue Sep 27 16:51:35 2022 +0100 @@ -32,12 +32,12 @@ definition ntree_rec :: "[[i, i, i] => i, i] => i" where - "ntree_rec(b) == + "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)" + "ntree_copy(z) \ ntree_rec(\x h r. Branch(x,r), z)" text \ @@ -50,8 +50,8 @@ 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))" + 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 @@ -66,7 +66,7 @@ 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 |] ==> + 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\ @@ -84,7 +84,7 @@ type definitions. \ -lemma ntree_mono: "A \ B ==> ntree(A) \ ntree(B)" +lemma ntree_mono: "A \ B \ ntree(A) \ ntree(B)" apply (unfold ntree.defs) apply (rule lfp_mono) apply (rule ntree.bnd_mono)+ @@ -99,7 +99,7 @@ apply (blast intro: Pair_in_univ nat_fun_univ [THEN subsetD]) done -lemma ntree_subset_univ: "A \ univ(B) ==> ntree(A) \ univ(B)" +lemma ntree_subset_univ: "A \ univ(B) \ ntree(A) \ univ(B)" by (rule subset_trans [OF ntree_mono ntree_univ]) @@ -108,18 +108,18 @@ \ lemma ntree_rec_Branch: - "function(h) ==> + "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) ==> + "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" +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) @@ -134,9 +134,9 @@ 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); + and step: "\x n h. \x \ A; h \ maptree(A) -||> maptree(A); \y \ field(h). P(y) - |] ==> P(Sons(x,h))" +\ \ P(Sons(x,h))" shows "P(t)" \ \A nicer induction rule than the standard one.\ using t @@ -159,8 +159,8 @@ 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))" + 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 diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Primrec.thy --- a/src/ZF/Induct/Primrec.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Primrec.thy Tue Sep 27 16:51:35 2022 +0100 @@ -18,23 +18,23 @@ definition SC :: "i" where - "SC == \l \ list(nat). list_case(0, \x xs. succ(x), l)" + "SC \ \l \ list(nat). list_case(0, \x xs. succ(x), l)" definition CONSTANT :: "i=>i" where - "CONSTANT(k) == \l \ list(nat). k" + "CONSTANT(k) \ \l \ list(nat). k" definition PROJ :: "i=>i" where - "PROJ(i) == \l \ list(nat). list_case(0, \x xs. x, drop(i,l))" + "PROJ(i) \ \l \ list(nat). list_case(0, \x xs. x, drop(i,l))" definition COMP :: "[i,i]=>i" where - "COMP(g,fs) == \l \ list(nat). g ` map(\f. f`l, fs)" + "COMP(g,fs) \ \l \ list(nat). g ` map(\f. f`l, fs)" definition PREC :: "[i,i]=>i" where - "PREC(f,g) == + "PREC(f,g) \ \l \ list(nat). list_case(0, \x xs. rec(x, f`xs, \y r. g ` Cons(r, Cons(y, xs))), l)" \ \Note that \g\ is applied first to \<^term>\PREC(f,g)`y\ and then to \y\!\ @@ -47,31 +47,31 @@ abbreviation ack :: "[i,i]=>i" where - "ack(x,y) == ACK(x) ` [y]" + "ack(x,y) \ ACK(x) ` [y]" text \ \medskip Useful special cases of evaluation. \ -lemma SC: "[| x \ nat; l \ list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)" +lemma SC: "\x \ nat; l \ list(nat)\ \ SC ` (Cons(x,l)) = succ(x)" by (simp add: SC_def) -lemma CONSTANT: "l \ list(nat) ==> CONSTANT(k) ` l = k" +lemma CONSTANT: "l \ list(nat) \ CONSTANT(k) ` l = k" by (simp add: CONSTANT_def) -lemma PROJ_0: "[| x \ nat; l \ list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x" +lemma PROJ_0: "\x \ nat; l \ list(nat)\ \ PROJ(0) ` (Cons(x,l)) = x" by (simp add: PROJ_def) -lemma COMP_1: "l \ list(nat) ==> COMP(g,[f]) ` l = g` [f`l]" +lemma COMP_1: "l \ list(nat) \ COMP(g,[f]) ` l = g` [f`l]" by (simp add: COMP_def) -lemma PREC_0: "l \ list(nat) ==> PREC(f,g) ` (Cons(0,l)) = f`l" +lemma PREC_0: "l \ list(nat) \ PREC(f,g) ` (Cons(0,l)) = f`l" by (simp add: PREC_def) lemma PREC_succ: - "[| x \ nat; l \ list(nat) |] - ==> PREC(f,g) ` (Cons(succ(x),l)) = + "\x \ nat; l \ list(nat)\ + \ PREC(f,g) ` (Cons(succ(x),l)) = g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))" by (simp add: PREC_def) @@ -85,10 +85,10 @@ domains prim_rec \ "list(nat)->nat" intros "SC \ prim_rec" - "k \ nat ==> CONSTANT(k) \ prim_rec" - "i \ nat ==> PROJ(i) \ prim_rec" - "[| g \ prim_rec; fs\list(prim_rec) |] ==> COMP(g,fs) \ prim_rec" - "[| f \ prim_rec; g \ prim_rec |] ==> PREC(f,g) \ prim_rec" + "k \ nat \ CONSTANT(k) \ prim_rec" + "i \ nat \ PROJ(i) \ prim_rec" + "\g \ prim_rec; fs\list(prim_rec)\ \ COMP(g,fs) \ prim_rec" + "\f \ prim_rec; g \ prim_rec\ \ PREC(f,g) \ prim_rec" monos list_mono con_defs SC_def CONSTANT_def PROJ_def COMP_def PREC_def type_intros nat_typechecks list.intros @@ -96,7 +96,7 @@ apply_type rec_type -lemma prim_rec_into_fun [TC]: "c \ prim_rec ==> c \ list(nat) -> nat" +lemma prim_rec_into_fun [TC]: "c \ prim_rec \ c \ list(nat) -> nat" by (erule subsetD [OF prim_rec.dom_subset]) lemmas [TC] = apply_type [OF prim_rec_into_fun] @@ -105,16 +105,16 @@ declare nat_into_Ord [TC] declare rec_type [TC] -lemma ACK_in_prim_rec [TC]: "i \ nat ==> ACK(i) \ prim_rec" +lemma ACK_in_prim_rec [TC]: "i \ nat \ ACK(i) \ prim_rec" by (induct set: nat) simp_all -lemma ack_type [TC]: "[| i \ nat; j \ nat |] ==> ack(i,j) \ nat" +lemma ack_type [TC]: "\i \ nat; j \ nat\ \ ack(i,j) \ nat" by auto subsection \Ackermann's function cases\ -lemma ack_0: "j \ nat ==> ack(0,j) = succ(j)" +lemma ack_0: "j \ nat \ ack(0,j) = succ(j)" \ \PROPERTY A 1\ by (simp add: SC) @@ -123,7 +123,7 @@ by (simp add: CONSTANT PREC_0) lemma ack_succ_succ: - "[| i\nat; j\nat |] ==> ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))" + "\i\nat; j\nat\ \ ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))" \ \PROPERTY A 3\ by (simp add: CONSTANT PREC_succ COMP_1 PROJ_0) @@ -131,7 +131,7 @@ and [simp del] = ACK.simps -lemma lt_ack2: "i \ nat ==> j \ nat ==> j < ack(i,j)" +lemma lt_ack2: "i \ nat \ j \ nat \ j < ack(i,j)" \ \PROPERTY A 4\ apply (induct i arbitrary: j set: nat) apply simp @@ -141,11 +141,11 @@ apply auto done -lemma ack_lt_ack_succ2: "[|i\nat; j\nat|] ==> ack(i,j) < ack(i, succ(j))" +lemma ack_lt_ack_succ2: "\i\nat; j\nat\ \ ack(i,j) < ack(i, succ(j))" \ \PROPERTY A 5-, the single-step lemma\ by (induct set: nat) (simp_all add: lt_ack2) -lemma ack_lt_mono2: "[| j nat; k \ nat |] ==> ack(i,j) < ack(i,k)" +lemma ack_lt_mono2: "\j nat; k \ nat\ \ ack(i,j) < ack(i,k)" \ \PROPERTY A 5, monotonicity for \<\\ apply (frule lt_nat_in_nat, assumption) apply (erule succ_lt_induct) @@ -154,14 +154,14 @@ apply (auto intro: ack_lt_ack_succ2) done -lemma ack_le_mono2: "[|j\k; i\nat; k\nat|] ==> ack(i,j) \ ack(i,k)" +lemma ack_le_mono2: "\j\k; i\nat; k\nat\ \ ack(i,j) \ ack(i,k)" \ \PROPERTY A 5', monotonicity for \\\\ apply (rule_tac f = "\j. ack (i,j) " in Ord_lt_mono_imp_le_mono) apply (assumption | rule ack_lt_mono2 ack_type [THEN nat_into_Ord])+ done lemma ack2_le_ack1: - "[| i\nat; j\nat |] ==> ack(i, succ(j)) \ ack(succ(i), j)" + "\i\nat; j\nat\ \ ack(i, succ(j)) \ ack(succ(i), j)" \ \PROPERTY A 6\ apply (induct_tac j) apply simp_all @@ -170,14 +170,14 @@ apply auto done -lemma ack_lt_ack_succ1: "[| i \ nat; j \ nat |] ==> ack(i,j) < ack(succ(i),j)" +lemma ack_lt_ack_succ1: "\i \ nat; j \ nat\ \ ack(i,j) < ack(succ(i),j)" \ \PROPERTY A 7-, the single-step lemma\ apply (rule ack_lt_mono2 [THEN lt_trans2]) apply (rule_tac [4] ack2_le_ack1) apply auto done -lemma ack_lt_mono1: "[| i nat; k \ nat |] ==> ack(i,k) < ack(j,k)" +lemma ack_lt_mono1: "\i nat; k \ nat\ \ ack(i,k) < ack(j,k)" \ \PROPERTY A 7, monotonicity for \<\\ apply (frule lt_nat_in_nat, assumption) apply (erule succ_lt_induct) @@ -186,23 +186,23 @@ apply (auto intro: ack_lt_ack_succ1) done -lemma ack_le_mono1: "[| i\j; j \ nat; k \ nat |] ==> ack(i,k) \ ack(j,k)" +lemma ack_le_mono1: "\i\j; j \ nat; k \ nat\ \ ack(i,k) \ ack(j,k)" \ \PROPERTY A 7', monotonicity for \\\\ apply (rule_tac f = "\j. ack (j,k) " in Ord_lt_mono_imp_le_mono) apply (assumption | rule ack_lt_mono1 ack_type [THEN nat_into_Ord])+ done -lemma ack_1: "j \ nat ==> ack(1,j) = succ(succ(j))" +lemma ack_1: "j \ nat \ ack(1,j) = succ(succ(j))" \ \PROPERTY A 8\ by (induct set: nat) simp_all -lemma ack_2: "j \ nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))" +lemma ack_2: "j \ nat \ ack(succ(1),j) = succ(succ(succ(j#+j)))" \ \PROPERTY A 9\ by (induct set: nat) (simp_all add: ack_1) lemma ack_nest_bound: - "[| i1 \ nat; i2 \ nat; j \ nat |] - ==> ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)" + "\i1 \ nat; i2 \ nat; j \ nat\ + \ ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)" \ \PROPERTY A 10\ apply (rule lt_trans2 [OF _ ack2_le_ack1]) apply simp @@ -212,8 +212,8 @@ done lemma ack_add_bound: - "[| i1 \ nat; i2 \ nat; j \ nat |] - ==> ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)" + "\i1 \ nat; i2 \ nat; j \ nat\ + \ ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)" \ \PROPERTY A 11\ apply (rule_tac j = "ack (succ (1), ack (i1 #+ i2, j))" in lt_trans) apply (simp add: ack_2) @@ -223,8 +223,8 @@ done lemma ack_add_bound2: - "[| i < ack(k,j); j \ nat; k \ nat |] - ==> i#+j < ack(succ(succ(succ(succ(k)))), j)" + "\i < ack(k,j); j \ nat; k \ nat\ + \ i#+j < ack(succ(succ(succ(succ(k)))), j)" \ \PROPERTY A 12.\ \ \Article uses existential quantifier but the ALF proof used \<^term>\k#+#4\.\ \ \Quantified version must be nested \\k'. \i,j \\.\ @@ -239,14 +239,14 @@ declare list_add_type [simp] -lemma SC_case: "l \ list(nat) ==> SC ` l < ack(1, list_add(l))" +lemma SC_case: "l \ list(nat) \ SC ` l < ack(1, list_add(l))" apply (unfold SC_def) apply (erule list.cases) apply (simp add: succ_iff) apply (simp add: ack_1 add_le_self) done -lemma lt_ack1: "[| i \ nat; j \ nat |] ==> i < ack(i,j)" +lemma lt_ack1: "\i \ nat; j \ nat\ \ i < ack(i,j)" \ \PROPERTY A 4'? Extra lemma needed for \CONSTANT\ case, constant functions.\ apply (induct_tac i) apply (simp add: nat_0_le) @@ -255,11 +255,11 @@ done lemma CONSTANT_case: - "[| l \ list(nat); k \ nat |] ==> CONSTANT(k) ` l < ack(k, list_add(l))" + "\l \ list(nat); k \ nat\ \ CONSTANT(k) ` l < ack(k, list_add(l))" by (simp add: CONSTANT_def lt_ack1) lemma PROJ_case [rule_format]: - "l \ list(nat) ==> \i \ nat. PROJ(i) ` l < ack(0, list_add(l))" + "l \ list(nat) \ \i \ nat. PROJ(i) ` l < ack(0, list_add(l))" apply (unfold PROJ_def) apply simp apply (erule list.induct) @@ -280,7 +280,7 @@ lemma COMP_map_lemma: "fs \ list({f \ prim_rec. \kf \ nat. \l \ list(nat). f`l < ack(kf, list_add(l))}) - ==> \k \ nat. \l \ list(nat). + \ \k \ nat. \l \ list(nat). list_add(map(\f. f ` l, fs)) < ack(k, list_add(l))" apply (induct set: list) apply (rule_tac x = 0 in bexI) @@ -294,12 +294,12 @@ done lemma COMP_case: - "[| kg\nat; + "\kg\nat; \l \ list(nat). g`l < ack(kg, list_add(l)); fs \ list({f \ prim_rec . \kf \ nat. \l \ list(nat). - f`l < ack(kf, list_add(l))}) |] - ==> \k \ nat. \l \ list(nat). COMP(g,fs)`l < ack(k, list_add(l))" + f`l < ack(kf, list_add(l))})\ + \ \k \ nat. \l \ list(nat). COMP(g,fs)`l < ack(k, list_add(l))" apply (simp add: COMP_def) apply (frule list_CollectD) apply (erule COMP_map_lemma [THEN bexE]) @@ -316,12 +316,12 @@ \ lemma PREC_case_lemma: - "[| \l \ list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); + "\\l \ list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \l \ list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); f \ prim_rec; kf\nat; g \ prim_rec; kg\nat; - l \ list(nat) |] - ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))" + l \ list(nat)\ + \ PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))" apply (unfold PREC_def) apply (erule list.cases) apply (simp add: lt_trans [OF nat_le_refl lt_ack2]) @@ -347,11 +347,11 @@ done lemma PREC_case: - "[| f \ prim_rec; kf\nat; + "\f \ prim_rec; kf\nat; g \ prim_rec; kg\nat; \l \ list(nat). f`l < ack(kf, list_add(l)); - \l \ list(nat). g`l < ack(kg, list_add(l)) |] - ==> \k \ nat. \l \ list(nat). PREC(f,g)`l< ack(k, list_add(l))" + \l \ list(nat). g`l < ack(kg, list_add(l))\ + \ \k \ nat. \l \ list(nat). PREC(f,g)`l< ack(k, list_add(l))" apply (rule ballI [THEN bexI]) apply (rule lt_trans1 [OF add_le_self PREC_case_lemma]) apply typecheck @@ -359,7 +359,7 @@ done lemma ack_bounds_prim_rec: - "f \ prim_rec ==> \k \ nat. \l \ list(nat). f`l < ack(k, list_add(l))" + "f \ prim_rec \ \k \ nat. \l \ list(nat). f`l < ack(k, list_add(l))" apply (induct set: prim_rec) apply (auto intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case) done diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/PropLog.thy --- a/src/ZF/Induct/PropLog.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/PropLog.thy Tue Sep 27 16:51:35 2022 +0100 @@ -36,17 +36,17 @@ abbreviation thms_syntax :: "[i,i] => o" (infixl \|-\ 50) - where "H |- p == p \ thms(H)" + 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" + 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] @@ -65,7 +65,7 @@ definition is_true :: "[i,i] => o" where - "is_true(p,t) == is_true_fun(p,t) = 1" + "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" @@ -87,7 +87,7 @@ definition logcon :: "[i,i] => o" (infixl \|=\ 50) where - "H |= p == \t. (\q \ H. is_true(q,t)) \ is_true(p,t)" + "H |= p \ \t. (\q \ H. is_true(q,t)) \ is_true(p,t)" text \ @@ -106,7 +106,7 @@ subsection \Proof theory of propositional logic\ -lemma thms_mono: "G \ H ==> thms(G) \ thms(H)" +lemma thms_mono: "G \ H \ thms(G) \ thms(H)" apply (unfold thms.defs) apply (rule lfp_mono) apply (rule thms.bnd_mono)+ @@ -117,13 +117,13 @@ inductive_cases ImpE: "p=>q \ propn" -lemma thms_MP: "[| H |- p=>q; H |- p |] ==> H |- q" +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" +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) @@ -134,23 +134,23 @@ subsubsection \Weakening, left and right\ -lemma weaken_left: "[| G \ H; G|-p |] ==> H|-p" +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" +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" +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" +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]) @@ -162,24 +162,24 @@ subsubsection \The cut rule\ -lemma cut: "[| H|-p; cons(p,H) |- q |] ==> H |- q" +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" +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" +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" +theorem soundness: "H |- p \ H |= p" apply (unfold logcon_def) apply (induct set: thms) apply auto @@ -190,14 +190,14 @@ subsubsection \Towards the completeness proof\ -lemma Fls_Imp: "[| H |- p=>Fls; q \ propn |] ==> H |- p=>q" +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" +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) @@ -207,7 +207,7 @@ done lemma hyps_thms_if: - "p \ propn ==> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)" + "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) @@ -216,7 +216,7 @@ 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" +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) @@ -234,14 +234,14 @@ \ lemma thms_excluded_middle: - "[| p \ propn; q \ propn |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q" + "\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" + "\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)+ @@ -255,7 +255,7 @@ \ lemma hyps_Diff: - "p \ propn ==> hyps(p, t-{v}) \ cons(#v=>Fls, hyps(p,t)-{#v})" + "p \ propn \ hyps(p, t-{v}) \ cons(#v=>Fls, hyps(p,t)-{#v})" by (induct set: propn) auto text \ @@ -264,7 +264,7 @@ \ lemma hyps_cons: - "p \ propn ==> hyps(p, cons(v,t)) \ cons(#v, hyps(p,t)-{#v=>Fls})" + "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\\ @@ -281,7 +281,7 @@ \<^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})" +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] @@ -292,7 +292,7 @@ \ lemma completeness_0_lemma [rule_format]: - "[| p \ propn; 0 |= p |] ==> \t. hyps(p,t) - hyps(p,t0) |- p" + "\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) @@ -315,18 +315,18 @@ subsubsection \Completeness theorem\ -lemma completeness_0: "[| p \ propn; 0 |= p |] ==> 0 |- p" +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" +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" + "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]) @@ -334,7 +334,7 @@ apply (blast intro: propn_Is) done -theorem thms_iff: "H \ Fin(propn) ==> H |- p \ H |= p \ p \ propn" +theorem thms_iff: "H \ Fin(propn) \ H |- p \ H |= p \ p \ propn" by (blast intro: soundness completeness thms_in_pl) end diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Rmap.thy --- a/src/ZF/Induct/Rmap.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Rmap.thy Tue Sep 27 16:51:35 2022 +0100 @@ -15,12 +15,12 @@ intros NilI: " \ rmap(r)" - ConsI: "[| : r; \ rmap(r) |] - ==> \ rmap(r)" + ConsI: "\: r; \ rmap(r)\ + \ \ rmap(r)" type_intros domainI rangeI list.intros -lemma rmap_mono: "r \ s ==> rmap(r) \ rmap(s)" +lemma rmap_mono: "r \ s \ rmap(r) \ rmap(s)" apply (unfold rmap.defs) apply (rule lfp_mono) apply (rule rmap.bnd_mono)+ @@ -33,19 +33,19 @@ declare rmap.intros [intro] -lemma rmap_rel_type: "r \ A \ B ==> rmap(r) \ list(A) \ list(B)" +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))" +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))" +lemma rmap_functional: "function(r) \ function(rmap(r))" apply (unfold function_def) apply (rule impI [THEN allI, THEN allI]) apply (erule rmap.induct) @@ -57,14 +57,14 @@ as expected. \ -lemma rmap_fun_type: "f \ A->B ==> rmap(f): list(A)->list(B)" +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)" +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 -r f2094906e491 -r e44d86131648 src/ZF/Induct/Term.thy --- a/src/ZF/Induct/Term.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Term.thy Tue Sep 27 16:51:35 2022 +0100 @@ -22,39 +22,39 @@ definition term_rec :: "[i, [i, i, i] => i] => i" where - "term_rec(t,d) == + "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))" + "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)))" + "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)))" + "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)))" + "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])" + "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)" + "\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) @@ -62,10 +62,10 @@ 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) |] ==> + "\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)" +\ \ 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) @@ -76,7 +76,7 @@ type definitions. \ -lemma term_mono: "A \ B ==> term(A) \ term(B)" +lemma term_mono: "A \ B \ term(A) \ term(B)" apply (unfold term.defs) apply (rule lfp_mono) apply (rule term.bnd_mono)+ @@ -92,21 +92,21 @@ apply (assumption | rule Pair_in_univ list_univ [THEN subsetD])+ done -lemma term_subset_univ: "A \ univ(B) ==> term(A) \ univ(B)" +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)" +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)" +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 @@ -116,7 +116,7 @@ apply (blast dest: rank_rls [THEN lt_trans]) done -lemma term_rec [simp]: "ts \ list(A) ==> +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]) @@ -126,9 +126,9 @@ 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))" + 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 @@ -141,17 +141,17 @@ done lemma def_term_rec: - "[| !!t. j(t)==term_rec(t,d); ts: list(A) |] ==> + "\\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" + "\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 @@ -163,19 +163,19 @@ \ lemma term_map [simp]: - "ts \ list(A) ==> + "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)" + "\t \ term(A); \x. x \ A \ f(x): B\ \ term_map(f,t) \ term(B)" apply (unfold 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})" + "t \ term(A) \ term_map(f,t) \ term({f(u). u \ A})" apply (erule term_map_type) apply (erule RepFunI) done @@ -185,10 +185,10 @@ \ lemma term_size [simp]: - "ts \ list(A) ==> term_size(Apply(a, ts)) = succ(list_add(map(term_size, ts)))" + "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" +lemma term_size_type [TC]: "t \ term(A) \ term_size(t) \ nat" by (auto simp add: term_size_def) @@ -197,10 +197,10 @@ \ lemma reflect [simp]: - "ts \ list(A) ==> reflect(Apply(a, ts)) = Apply(a, rev(map(reflect, ts)))" + "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)" +lemma reflect_type [TC]: "t \ term(A) \ reflect(t) \ term(A)" by (auto simp add: reflect_def) @@ -209,10 +209,10 @@ \ lemma preorder [simp]: - "ts \ list(A) ==> preorder(Apply(a, ts)) = Cons(a, flat(map(preorder, ts)))" + "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)" +lemma preorder_type [TC]: "t \ term(A) \ preorder(t) \ list(A)" by (simp add: preorder_def) @@ -221,10 +221,10 @@ \ lemma postorder [simp]: - "ts \ list(A) ==> postorder(Apply(a, ts)) = flat(map(postorder, ts)) @ [a]" + "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)" +lemma postorder_type [TC]: "t \ term(A) \ postorder(t) \ list(A)" by (simp add: postorder_def) @@ -234,15 +234,15 @@ declare map_compose [simp] -lemma term_map_ident: "t \ term(A) ==> term_map(\u. u, t) = t" +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)" + "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))" + "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]) @@ -250,13 +250,13 @@ \medskip Theorems about \term_size\. \ -lemma term_size_term_map: "t \ term(A) ==> term_size(term_map(f,t)) = term_size(t)" +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)" +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))" +lemma term_size_length: "t \ term(A) \ term_size(t) = length(preorder(t))" by (induct rule: term_induct_eqn) (simp add: length_flat) @@ -264,7 +264,7 @@ \medskip Theorems about \reflect\. \ -lemma reflect_reflect_ident: "t \ term(A) ==> reflect(reflect(t)) = t" +lemma reflect_reflect_ident: "t \ term(A) \ reflect(reflect(t)) = t" by (induct rule: term_induct_eqn) (simp add: rev_map_distrib) @@ -273,11 +273,11 @@ \ lemma preorder_term_map: - "t \ term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))" + "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))" + "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]) diff -r f2094906e491 -r e44d86131648 src/ZF/Induct/Tree_Forest.thy --- a/src/ZF/Induct/Tree_Forest.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Induct/Tree_Forest.thy Tue Sep 27 16:51:35 2022 +0100 @@ -26,10 +26,10 @@ declare tree_forest.intros [simp, TC] -lemma tree_def: "tree(A) == Part(tree_forest(A), Inl)" +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)" +lemma forest_def: "forest(A) \ Part(tree_forest(A), Inr)" by (simp only: tree_forest.defs) @@ -43,7 +43,7 @@ apply (rule Part_subset) done -lemma treeI [TC]: "x \ tree(A) ==> x \ tree_forest(A)" +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)" @@ -51,7 +51,7 @@ apply (rule Part_subset) done -lemma treeI' [TC]: "x \ forest(A) ==> x \ tree_forest(A)" +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)" @@ -91,22 +91,22 @@ \ 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)); + "\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)" + \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)); + "\\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)) \ + \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.\ apply (unfold Ball_def) @@ -161,10 +161,10 @@ \ lemma list_of_TF_type [TC]: - "z \ tree_forest(A) ==> list_of_TF(z) \ list(tree(A))" + "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)" +lemma of_list_type [TC]: "l \ list(tree(A)) \ of_list(l) \ forest(A)" by (induct set: list) simp_all text \ @@ -172,9 +172,9 @@ \ 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)" + 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 @@ -182,7 +182,7 @@ \medskip \size\. \ -lemma size_type [TC]: "z \ tree_forest(A) ==> size(z) \ nat" +lemma size_type [TC]: "z \ tree_forest(A) \ size(z) \ nat" by (induct set: tree_forest) simp_all @@ -190,7 +190,7 @@ \medskip \preorder\. \ -lemma preorder_type [TC]: "z \ tree_forest(A) ==> preorder(z) \ list(A)" +lemma preorder_type [TC]: "z \ tree_forest(A) \ preorder(z) \ list(A)" by (induct set: tree_forest) simp_all @@ -199,10 +199,10 @@ \ lemma forest_induct [consumes 1, case_names Fnil Fcons]: - "[| f \ forest(A); + "\f \ forest(A); R(Fnil); - !!t f. [| t \ tree(A); f \ forest(A); R(f) |] ==> R(Fcons(t,f)) - |] ==> R(f)" + \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]) @@ -211,10 +211,10 @@ apply simp done -lemma forest_iso: "f \ forest(A) ==> of_list(list_of_TF(f)) = f" +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" +lemma tree_list_iso: "ts: list(tree(A)) \ list_of_TF(of_list(ts)) = ts" by (induct set: list) simp_all @@ -222,11 +222,11 @@ \medskip Theorems about \map\. \ -lemma map_ident: "z \ tree_forest(A) ==> map(\u. u, z) = z" +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)" + "z \ tree_forest(A) \ map(h, map(j,z)) = map(\u. h(j(u)), z)" by (induct set: tree_forest) simp_all @@ -234,10 +234,10 @@ \medskip Theorems about \size\. \ -lemma size_map: "z \ tree_forest(A) ==> size(map(h,z)) = size(z)" +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))" +lemma size_length: "z \ tree_forest(A) \ size(z) = length(preorder(z))" by (induct set: tree_forest) (simp_all add: length_app) text \ @@ -245,7 +245,7 @@ \ lemma preorder_map: - "z \ tree_forest(A) ==> preorder(map(h,z)) = List.map(h, preorder(z))" + "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 -r f2094906e491 -r e44d86131648 src/ZF/Inductive.thy --- a/src/ZF/Inductive.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Inductive.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,13 +19,13 @@ "elimination" "induction" "case_eqns" "recursor_eqns" :: quasi_command begin -lemma def_swap_iff: "a == b ==> a = c \ c = b" +lemma def_swap_iff: "a \ b \ a = c \ c = b" by blast -lemma def_trans: "f == g ==> g(a) = b ==> f(a) = b" +lemma def_trans: "f \ g \ g(a) = b \ f(a) = b" by simp -lemma refl_thin: "!!P. a = a ==> P ==> P" . +lemma refl_thin: "\P. a = a \ P \ P" . ML_file \ind_syntax.ML\ ML_file \Tools/ind_cases.ML\ diff -r f2094906e491 -r e44d86131648 src/ZF/InfDatatype.thy --- a/src/ZF/InfDatatype.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/InfDatatype.thy Tue Sep 27 16:51:35 2022 +0100 @@ -45,13 +45,13 @@ qed lemma subset_Vcsucc: - "[| D \ Vfrom(A,csucc(K)); |D| \ K; InfCard(K) |] - ==> \j. D \ Vfrom(A,j) & j < csucc(K)" + "\D \ Vfrom(A,csucc(K)); |D| \ K; InfCard(K)\ + \ \j. D \ Vfrom(A,j) & j < csucc(K)" by (simp add: subset_iff_id fun_Vcsucc_lemma) (*Version for arbitrary index sets*) lemma fun_Vcsucc: - "[| |D| \ K; InfCard(K); D \ Vfrom(A,csucc(K)) |] ==> + "\|D| \ K; InfCard(K); D \ Vfrom(A,csucc(K))\ \ D -> Vfrom(A,csucc(K)) \ Vfrom(A,csucc(K))" apply (safe dest!: fun_Vcsucc_lemma subset_Vcsucc) apply (rule Vfrom [THEN ssubst]) @@ -65,9 +65,9 @@ done lemma fun_in_Vcsucc: - "[| f: D -> Vfrom(A, csucc(K)); |D| \ K; InfCard(K); - D \ Vfrom(A,csucc(K)) |] - ==> f: Vfrom(A,csucc(K))" + "\f: D -> Vfrom(A, csucc(K)); |D| \ K; InfCard(K); + D \ Vfrom(A,csucc(K))\ + \ f: Vfrom(A,csucc(K))" by (blast intro: fun_Vcsucc [THEN subsetD]) text\Remove \\\ from the rule above\ @@ -76,7 +76,7 @@ (** Version where K itself is the index set **) lemma Card_fun_Vcsucc: - "InfCard(K) ==> K -> Vfrom(A,csucc(K)) \ Vfrom(A,csucc(K))" + "InfCard(K) \ K -> Vfrom(A,csucc(K)) \ Vfrom(A,csucc(K))" apply (frule InfCard_is_Card [THEN Card_is_Ord]) apply (blast del: subsetI intro: fun_Vcsucc Ord_cardinal_le i_subset_Vfrom @@ -84,10 +84,10 @@ done lemma Card_fun_in_Vcsucc: - "[| f: K -> Vfrom(A, csucc(K)); InfCard(K) |] ==> f: Vfrom(A,csucc(K))" + "\f: K -> Vfrom(A, csucc(K)); InfCard(K)\ \ f: Vfrom(A,csucc(K))" by (blast intro: Card_fun_Vcsucc [THEN subsetD]) -lemma Limit_csucc: "InfCard(K) ==> Limit(csucc(K))" +lemma Limit_csucc: "InfCard(K) \ Limit(csucc(K))" by (erule InfCard_csucc [THEN InfCard_is_Limit]) lemmas Pair_in_Vcsucc = Pair_in_VLimit [OF _ _ Limit_csucc] diff -r f2094906e491 -r e44d86131648 src/ZF/Int.thy --- a/src/ZF/Int.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Int.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,50 +9,50 @@ definition intrel :: i where - "intrel == {p \ (nat*nat)*(nat*nat). + "intrel \ {p \ (nat*nat)*(nat*nat). \x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" definition int :: i where - "int == (nat*nat)//intrel" + "int \ (nat*nat)//intrel" definition int_of :: "i=>i" \ \coercion from nat to int\ (\$# _\ [80] 80) where - "$# m == intrel `` {}" + "$# m \ intrel `` {}" definition intify :: "i=>i" \ \coercion from ANYTHING to int\ where - "intify(m) == if m \ int then m else $#0" + "intify(m) \ if m \ int then m else $#0" definition raw_zminus :: "i=>i" where - "raw_zminus(z) == \\z. intrel``{}" + "raw_zminus(z) \ \\z. intrel``{}" definition zminus :: "i=>i" (\$- _\ [80] 80) where - "$- z == raw_zminus (intify(z))" + "$- z \ raw_zminus (intify(z))" definition znegative :: "i=>o" where - "znegative(z) == \x y. xnat & \z" + "znegative(z) \ \x y. xnat & \z" definition iszero :: "i=>o" where - "iszero(z) == z = $# 0" + "iszero(z) \ z = $# 0" definition raw_nat_of :: "i=>i" where - "raw_nat_of(z) == natify (\\z. x#-y)" + "raw_nat_of(z) \ natify (\\z. x#-y)" definition nat_of :: "i=>i" where - "nat_of(z) == raw_nat_of (intify(z))" + "nat_of(z) \ raw_nat_of (intify(z))" definition zmagnitude :: "i=>i" where \ \could be replaced by an absolute value function from int to int?\ - "zmagnitude(z) == - THE m. m\nat & ((~ znegative(z) & z = $# m) | + "zmagnitude(z) \ + THE m. m\nat & ((\ znegative(z) & z = $# m) | (znegative(z) & $- z = $# m))" definition @@ -60,35 +60,35 @@ (*Cannot use UN here or in zadd because of the form of congruent2. Perhaps a "curried" or even polymorphic congruent predicate would be better.*) - "raw_zmult(z1,z2) == + "raw_zmult(z1,z2) \ \p1\z1. \p2\z2. split(%x1 y1. split(%x2 y2. intrel``{}, p2), p1)" definition zmult :: "[i,i]=>i" (infixl \$*\ 70) where - "z1 $* z2 == raw_zmult (intify(z1),intify(z2))" + "z1 $* z2 \ raw_zmult (intify(z1),intify(z2))" definition raw_zadd :: "[i,i]=>i" where - "raw_zadd (z1, z2) == + "raw_zadd (z1, z2) \ \z1\z1. \z2\z2. let =z1; =z2 in intrel``{}" definition zadd :: "[i,i]=>i" (infixl \$+\ 65) where - "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))" + "z1 $+ z2 \ raw_zadd (intify(z1),intify(z2))" definition zdiff :: "[i,i]=>i" (infixl \$-\ 65) where - "z1 $- z2 == z1 $+ zminus(z2)" + "z1 $- z2 \ z1 $+ zminus(z2)" definition zless :: "[i,i]=>o" (infixl \$<\ 50) where - "z1 $< z2 == znegative(z1 $- z2)" + "z1 $< z2 \ znegative(z1 $- z2)" definition zle :: "[i,i]=>o" (infixl \$\\ 50) where - "z1 $\ z2 == z1 $< z2 | intify(z1)=intify(z2)" + "z1 $\ z2 \ z1 $< z2 | intify(z1)=intify(z2)" declare quotientE [elim!] @@ -103,19 +103,19 @@ by (simp add: intrel_def) lemma intrelI [intro!]: - "[| x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat |] - ==> <,>: intrel" + "\x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat\ + \ <,>: intrel" by (simp add: intrel_def) lemma intrelE [elim!]: - "[| p \ intrel; - !!x1 y1 x2 y2. [| p = <,>; x1#+y2 = x2#+y1; - x1\nat; y1\nat; x2\nat; y2\nat |] ==> Q |] - ==> Q" + "\p \ intrel; + \x1 y1 x2 y2. \p = <,>; x1#+y2 = x2#+y1; + x1\nat; y1\nat; x2\nat; y2\nat\ \ Q\ + \ Q" by (simp add: intrel_def, blast) lemma int_trans_lemma: - "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1" + "\x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2\ \ x1 #+ y3 = x3 #+ y1" apply (rule sym) apply (erule add_left_cancel)+ apply (simp_all (no_asm_simp)) @@ -126,7 +126,7 @@ apply (fast elim!: sym int_trans_lemma) done -lemma image_intrel_int: "[| m\nat; n\nat |] ==> intrel `` {} \ int" +lemma image_intrel_int: "\m\nat; n\nat\ \ intrel `` {} \ int" by (simp add: int_def) declare equiv_intrel [THEN eq_equiv_class_iff, simp] @@ -142,7 +142,7 @@ lemma int_of_eq [iff]: "($# m = $# n) \ natify(m)=natify(n)" by (simp add: int_of_def) -lemma int_of_inject: "[| $#m = $#n; m\nat; n\nat |] ==> m=n" +lemma int_of_inject: "\$#m = $#n; m\nat; n\nat\ \ m=n" by (drule int_of_eq [THEN iffD1], auto) @@ -151,7 +151,7 @@ lemma intify_in_int [iff,TC]: "intify(x) \ int" by (simp add: intify_def) -lemma intify_ident [simp]: "n \ int ==> intify(n) = n" +lemma intify_ident [simp]: "n \ int \ intify(n) = n" by (simp add: intify_def) @@ -211,7 +211,7 @@ lemma zminus_congruent: "(%. intrel``{}) respects intrel" by (auto simp add: congruent_def add_ac) -lemma raw_zminus_type: "z \ int ==> raw_zminus(z) \ int" +lemma raw_zminus_type: "z \ int \ raw_zminus(z) \ int" apply (simp add: int_def raw_zminus_def) apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent]) done @@ -220,31 +220,31 @@ by (simp add: zminus_def raw_zminus_type) lemma raw_zminus_inject: - "[| raw_zminus(z) = raw_zminus(w); z \ int; w \ int |] ==> z=w" + "\raw_zminus(z) = raw_zminus(w); z \ int; w \ int\ \ z=w" apply (simp add: int_def raw_zminus_def) apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe) apply (auto dest: eq_intrelD simp add: add_ac) done -lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)" +lemma zminus_inject_intify [dest!]: "$-z = $-w \ intify(z) = intify(w)" apply (simp add: zminus_def) apply (blast dest!: raw_zminus_inject) done -lemma zminus_inject: "[| $-z = $-w; z \ int; w \ int |] ==> z=w" +lemma zminus_inject: "\$-z = $-w; z \ int; w \ int\ \ z=w" by auto lemma raw_zminus: - "[| x\nat; y\nat |] ==> raw_zminus(intrel``{}) = intrel `` {}" + "\x\nat; y\nat\ \ raw_zminus(intrel``{}) = intrel `` {}" apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) done lemma zminus: - "[| x\nat; y\nat |] - ==> $- (intrel``{}) = intrel `` {}" + "\x\nat; y\nat\ + \ $- (intrel``{}) = intrel `` {}" by (simp add: zminus_def raw_zminus image_intrel_int) -lemma raw_zminus_zminus: "z \ int ==> raw_zminus (raw_zminus(z)) = z" +lemma raw_zminus_zminus: "z \ int \ raw_zminus (raw_zminus(z)) = z" by (auto simp add: int_def raw_zminus) lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)" @@ -253,13 +253,13 @@ lemma zminus_int0 [simp]: "$- ($#0) = $#0" by (simp add: int_of_def zminus) -lemma zminus_zminus: "z \ int ==> $- ($- z) = z" +lemma zminus_zminus: "z \ int \ $- ($- z) = z" by simp subsection\\<^term>\znegative\: the test for negative integers\ -lemma znegative: "[| x\nat; y\nat |] ==> znegative(intrel``{}) \ xx\nat; y\nat\ \ znegative(intrel``{}) \ x znegative($# n)" by (simp add: znegative int_of_def) lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))" by (simp add: znegative int_of_def zminus natify_succ) -lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> natify(n)=0" +lemma not_znegative_imp_zero: "\ znegative($- $# n) \ natify(n)=0" by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym]) @@ -286,7 +286,7 @@ by (auto simp add: congruent_def split: nat_diff_split) lemma raw_nat_of: - "[| x\nat; y\nat |] ==> raw_nat_of(intrel``{}) = x#-y" + "\x\nat; y\nat\ \ raw_nat_of(intrel``{}) = x#-y" by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)" @@ -306,7 +306,7 @@ lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)" by (auto simp add: zmagnitude_def int_of_eq) -lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n" +lemma natify_int_of_eq: "natify(x)=n \ $#x = $# n" apply (drule sym) apply (simp (no_asm_simp) add: int_of_eq) done @@ -324,7 +324,7 @@ done lemma not_zneg_int_of: - "[| z \ int; ~ znegative(z) |] ==> \n\nat. z = $# n" + "\z \ int; \ znegative(z)\ \ \n\nat. z = $# n" apply (auto simp add: int_def znegative int_of_def not_lt_iff_le) apply (rename_tac x y) apply (rule_tac x="x#-y" in bexI) @@ -332,38 +332,38 @@ done lemma not_zneg_mag [simp]: - "[| z \ int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" + "\z \ int; \ znegative(z)\ \ $# (zmagnitude(z)) = z" by (drule not_zneg_int_of, auto) lemma zneg_int_of: - "[| znegative(z); z \ int |] ==> \n\nat. z = $- ($# succ(n))" + "\znegative(z); z \ int\ \ \n\nat. z = $- ($# succ(n))" by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add) lemma zneg_mag [simp]: - "[| znegative(z); z \ int |] ==> $# (zmagnitude(z)) = $- z" + "\znegative(z); z \ int\ \ $# (zmagnitude(z)) = $- z" by (drule zneg_int_of, auto) -lemma int_cases: "z \ int ==> \n\nat. z = $# n | z = $- ($# succ(n))" +lemma int_cases: "z \ int \ \n\nat. z = $# n | z = $- ($# succ(n))" apply (case_tac "znegative (z) ") prefer 2 apply (blast dest: not_zneg_mag sym) apply (blast dest: zneg_int_of) done lemma not_zneg_raw_nat_of: - "[| ~ znegative(z); z \ int |] ==> $# (raw_nat_of(z)) = z" + "\\ znegative(z); z \ int\ \ $# (raw_nat_of(z)) = z" apply (drule not_zneg_int_of) apply (auto simp add: raw_nat_of_type raw_nat_of_int_of) done lemma not_zneg_nat_of_intify: - "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)" + "\ znegative(intify(z)) \ $# (nat_of(z)) = intify(z)" by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of) -lemma not_zneg_nat_of: "[| ~ znegative(z); z \ int |] ==> $# (nat_of(z)) = z" +lemma not_zneg_nat_of: "\\ znegative(z); z \ int\ \ $# (nat_of(z)) = z" apply (simp (no_asm_simp) add: not_zneg_nat_of_intify) done -lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0" +lemma zneg_nat_of [simp]: "znegative(intify(z)) \ nat_of(z) = 0" apply (subgoal_tac "intify(z) \ int") apply (simp add: int_def) apply (auto simp add: znegative nat_of_def raw_nat_of @@ -389,7 +389,7 @@ apply (simp (no_asm_simp) add: add_assoc [symmetric]) done -lemma raw_zadd_type: "[| z \ int; w \ int |] ==> raw_zadd(z,w) \ int" +lemma raw_zadd_type: "\z \ int; w \ int\ \ raw_zadd(z,w) \ int" apply (simp add: int_def raw_zadd_def) apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+) apply (simp add: Let_def) @@ -399,8 +399,8 @@ by (simp add: zadd_def raw_zadd_type) lemma raw_zadd: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> raw_zadd (intrel``{}, intrel``{}) = + "\x1\nat; y1\nat; x2\nat; y2\nat\ + \ raw_zadd (intrel``{}, intrel``{}) = intrel `` {}" apply (simp add: raw_zadd_def UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) @@ -408,37 +408,37 @@ done lemma zadd: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> (intrel``{}) $+ (intrel``{}) = + "\x1\nat; y1\nat; x2\nat; y2\nat\ + \ (intrel``{}) $+ (intrel``{}) = intrel `` {}" by (simp add: zadd_def raw_zadd image_intrel_int) -lemma raw_zadd_int0: "z \ int ==> raw_zadd ($#0,z) = z" +lemma raw_zadd_int0: "z \ int \ raw_zadd ($#0,z) = z" by (auto simp add: int_def int_of_def raw_zadd) lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)" by (simp add: zadd_def raw_zadd_int0) -lemma zadd_int0: "z \ int ==> $#0 $+ z = z" +lemma zadd_int0: "z \ int \ $#0 $+ z = z" by simp lemma raw_zminus_zadd_distrib: - "[| z \ int; w \ int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" + "\z \ int; w \ int\ \ $- raw_zadd(z,w) = raw_zadd($- z, $- w)" by (auto simp add: zminus raw_zadd int_def) lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w" by (simp add: zadd_def raw_zminus_zadd_distrib) lemma raw_zadd_commute: - "[| z \ int; w \ int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" + "\z \ int; w \ int\ \ raw_zadd(z,w) = raw_zadd(w,z)" by (auto simp add: raw_zadd add_ac int_def) lemma zadd_commute: "z $+ w = w $+ z" by (simp add: zadd_def raw_zadd_commute) lemma raw_zadd_assoc: - "[| z1: int; z2: int; z3: int |] - ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" + "\z1: int; z2: int; z3: int\ + \ raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" by (auto simp add: int_def raw_zadd add_assoc) lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)" @@ -460,13 +460,13 @@ by (simp add: int_of_add [symmetric] natify_succ) lemma int_of_diff: - "[| m\nat; n \ m |] ==> $# (m #- n) = ($#m) $- ($#n)" + "\m\nat; n \ m\ \ $# (m #- n) = ($#m) $- ($#n)" apply (simp add: int_of_def zdiff_def) apply (frule lt_nat_in_nat) apply (simp_all add: zadd zminus add_diff_inverse2) done -lemma raw_zadd_zminus_inverse: "z \ int ==> raw_zadd (z, $- z) = $#0" +lemma raw_zadd_zminus_inverse: "z \ int \ raw_zadd (z, $- z) = $#0" by (auto simp add: int_def int_of_def zminus raw_zadd add_commute) lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0" @@ -481,7 +481,7 @@ lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)" by (rule trans [OF zadd_commute zadd_int0_intify]) -lemma zadd_int0_right: "z \ int ==> z $+ $#0 = z" +lemma zadd_int0_right: "z \ int \ z $+ $#0 = z" by simp @@ -502,7 +502,7 @@ done -lemma raw_zmult_type: "[| z \ int; w \ int |] ==> raw_zmult(z,w) \ int" +lemma raw_zmult_type: "\z \ int; w \ int\ \ raw_zmult(z,w) \ int" apply (simp add: int_def raw_zmult_def) apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+) apply (simp add: Let_def) @@ -512,42 +512,42 @@ by (simp add: zmult_def raw_zmult_type) lemma raw_zmult: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> raw_zmult(intrel``{}, intrel``{}) = + "\x1\nat; y1\nat; x2\nat; y2\nat\ + \ raw_zmult(intrel``{}, intrel``{}) = intrel `` {}" by (simp add: raw_zmult_def UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) lemma zmult: - "[| x1\nat; y1\nat; x2\nat; y2\nat |] - ==> (intrel``{}) $* (intrel``{}) = + "\x1\nat; y1\nat; x2\nat; y2\nat\ + \ (intrel``{}) $* (intrel``{}) = intrel `` {}" by (simp add: zmult_def raw_zmult image_intrel_int) -lemma raw_zmult_int0: "z \ int ==> raw_zmult ($#0,z) = $#0" +lemma raw_zmult_int0: "z \ int \ raw_zmult ($#0,z) = $#0" by (auto simp add: int_def int_of_def raw_zmult) lemma zmult_int0 [simp]: "$#0 $* z = $#0" by (simp add: zmult_def raw_zmult_int0) -lemma raw_zmult_int1: "z \ int ==> raw_zmult ($#1,z) = z" +lemma raw_zmult_int1: "z \ int \ raw_zmult ($#1,z) = z" by (auto simp add: int_def int_of_def raw_zmult) lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)" by (simp add: zmult_def raw_zmult_int1) -lemma zmult_int1: "z \ int ==> $#1 $* z = z" +lemma zmult_int1: "z \ int \ $#1 $* z = z" by simp lemma raw_zmult_commute: - "[| z \ int; w \ int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" + "\z \ int; w \ int\ \ raw_zmult(z,w) = raw_zmult(w,z)" by (auto simp add: int_def raw_zmult add_ac mult_ac) lemma zmult_commute: "z $* w = w $* z" by (simp add: zmult_def raw_zmult_commute) lemma raw_zmult_zminus: - "[| z \ int; w \ int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" + "\z \ int; w \ int\ \ raw_zmult($- z, w) = $- raw_zmult(z, w)" by (auto simp add: int_def zminus raw_zmult add_ac) lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)" @@ -559,8 +559,8 @@ by (simp add: zmult_commute [of w]) lemma raw_zmult_assoc: - "[| z1: int; z2: int; z3: int |] - ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" + "\z1: int; z2: int; z3: int\ + \ raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac) lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)" @@ -576,8 +576,8 @@ lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute lemma raw_zadd_zmult_distrib: - "[| z1: int; z2: int; w \ int |] - ==> raw_zmult(raw_zadd(z1,z2), w) = + "\z1: int; z2: int; w \ int\ + \ raw_zmult(raw_zadd(z1,z2), w) = raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))" by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac) @@ -620,7 +620,7 @@ (*"Less than" is a linear ordering*) lemma zless_linear_lemma: - "[| z \ int; w \ int |] ==> z$z \ int; w \ int\ \ z$ (z$ int; y \ int |] ==> (x \ y) \ (x $< y | y $< x)" +lemma neq_iff_zless: "\x \ int; y \ int\ \ (x \ y) \ (x $< y | y $< x)" by (cut_tac z = x and w = y in zless_linear, auto) -lemma zless_imp_intify_neq: "w $< z ==> intify(w) \ intify(z)" +lemma zless_imp_intify_neq: "w $< z \ intify(w) \ intify(z)" apply auto -apply (subgoal_tac "~ (intify (w) $< intify (z))") +apply (subgoal_tac "\ (intify (w) $< intify (z))") apply (erule_tac [2] ssubst) apply (simp (no_asm_use)) apply auto @@ -648,7 +648,7 @@ (*This lemma allows direct proofs of other <-properties*) lemma zless_imp_succ_zadd_lemma: - "[| w $< z; w \ int; z \ int |] ==> (\n\nat. z = w $+ $#(succ(n)))" + "\w $< z; w \ int; z \ int\ \ (\n\nat. z = w $+ $#(succ(n)))" apply (simp add: zless_def znegative_def zdiff_def int_def) apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def) apply (rule_tac x = k in bexI) @@ -656,14 +656,14 @@ done lemma zless_imp_succ_zadd: - "w $< z ==> (\n\nat. w $+ $#(succ(n)) = intify(z))" + "w $< z \ (\n\nat. w $+ $#(succ(n)) = intify(z))" apply (subgoal_tac "intify (w) $< intify (z) ") apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma) apply auto done lemma zless_succ_zadd_lemma: - "w \ int ==> w $< w $+ $# succ(n)" + "w \ int \ w $< w $+ $# succ(n)" apply (simp add: zless_def znegative_def zdiff_def int_def) apply (auto simp add: zadd zminus int_of_def image_iff) apply (rule_tac x = 0 in exI, auto) @@ -680,13 +680,13 @@ apply (cut_tac w = w and n = n in zless_succ_zadd, auto) done -lemma zless_int_of [simp]: "[| m\nat; n\nat |] ==> ($#m $< $#n) \ (mm\nat; n\nat\ \ ($#m $< $#n) \ (m int; y \ int; z \ int |] ==> x $< z" + "\x $< y; y $< z; x \ int; y \ int; z \ int\ \ x $< z" apply (simp add: zless_def znegative_def zdiff_def int_def) apply (auto simp add: zadd zminus image_iff) apply (rename_tac x1 x2 y1 y2) @@ -699,19 +699,19 @@ apply auto done -lemma zless_trans [trans]: "[| x $< y; y $< z |] ==> x $< z" +lemma zless_trans [trans]: "\x $< y; y $< z\ \ x $< z" apply (subgoal_tac "intify (x) $< intify (z) ") apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma) apply auto done -lemma zless_not_sym: "z $< w ==> ~ (w $< z)" +lemma zless_not_sym: "z $< w \ \ (w $< z)" by (blast dest: zless_trans) -(* [| z $< w; ~ P ==> w $< z |] ==> P *) +(* \z $< w; \ P \ w $< z\ \ P *) lemmas zless_asym = zless_not_sym [THEN swap] -lemma zless_imp_zle: "z $< w ==> z $\ w" +lemma zless_imp_zle: "z $< w \ z $\ w" by (simp add: zle_def) lemma zle_linear: "z $\ w | w $\ z" @@ -725,48 +725,48 @@ lemma zle_refl: "z $\ z" by (simp add: zle_def) -lemma zle_eq_refl: "x=y ==> x $\ y" +lemma zle_eq_refl: "x=y \ x $\ y" by (simp add: zle_refl) -lemma zle_anti_sym_intify: "[| x $\ y; y $\ x |] ==> intify(x) = intify(y)" +lemma zle_anti_sym_intify: "\x $\ y; y $\ x\ \ intify(x) = intify(y)" apply (simp add: zle_def, auto) apply (blast dest: zless_trans) done -lemma zle_anti_sym: "[| x $\ y; y $\ x; x \ int; y \ int |] ==> x=y" +lemma zle_anti_sym: "\x $\ y; y $\ x; x \ int; y \ int\ \ x=y" by (drule zle_anti_sym_intify, auto) lemma zle_trans_lemma: - "[| x \ int; y \ int; z \ int; x $\ y; y $\ z |] ==> x $\ z" + "\x \ int; y \ int; z \ int; x $\ y; y $\ z\ \ x $\ z" apply (simp add: zle_def, auto) apply (blast intro: zless_trans) done -lemma zle_trans [trans]: "[| x $\ y; y $\ z |] ==> x $\ z" +lemma zle_trans [trans]: "\x $\ y; y $\ z\ \ x $\ z" apply (subgoal_tac "intify (x) $\ intify (z) ") apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma) apply auto done -lemma zle_zless_trans [trans]: "[| i $\ j; j $< k |] ==> i $< k" +lemma zle_zless_trans [trans]: "\i $\ j; j $< k\ \ i $< k" apply (auto simp add: zle_def) apply (blast intro: zless_trans) apply (simp add: zless_def zdiff_def zadd_def) done -lemma zless_zle_trans [trans]: "[| i $< j; j $\ k |] ==> i $< k" +lemma zless_zle_trans [trans]: "\i $< j; j $\ k\ \ i $< k" apply (auto simp add: zle_def) apply (blast intro: zless_trans) apply (simp add: zless_def zdiff_def zminus_def) done -lemma not_zless_iff_zle: "~ (z $< w) \ (w $\ z)" +lemma not_zless_iff_zle: "\ (z $< w) \ (w $\ z)" apply (cut_tac z = z and w = w in zless_linear) apply (auto dest: zless_trans simp add: zle_def) apply (auto dest!: zless_imp_intify_neq) done -lemma not_zle_iff_zless: "~ (z $\ w) \ (w $< z)" +lemma not_zle_iff_zless: "\ (z $\ w) \ (w $< z)" by (simp add: not_zless_iff_zle [THEN iff_sym]) @@ -784,21 +784,21 @@ lemma zless_zdiff_iff: "(x $< z$-y) \ (x $+ y $< z)" by (simp add: zless_def zdiff_def zadd_ac) -lemma zdiff_eq_iff: "[| x \ int; z \ int |] ==> (x$-y = z) \ (x = z $+ y)" +lemma zdiff_eq_iff: "\x \ int; z \ int\ \ (x$-y = z) \ (x = z $+ y)" by (auto simp add: zdiff_def zadd_assoc) -lemma eq_zdiff_iff: "[| x \ int; z \ int |] ==> (x = z$-y) \ (x $+ y = z)" +lemma eq_zdiff_iff: "\x \ int; z \ int\ \ (x = z$-y) \ (x $+ y = z)" by (auto simp add: zdiff_def zadd_assoc) lemma zdiff_zle_iff_lemma: - "[| x \ int; z \ int |] ==> (x$-y $\ z) \ (x $\ z $+ y)" + "\x \ int; z \ int\ \ (x$-y $\ z) \ (x $\ z $+ y)" by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff) lemma zdiff_zle_iff: "(x$-y $\ z) \ (x $\ z $+ y)" by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp) lemma zle_zdiff_iff_lemma: - "[| x \ int; z \ int |] ==>(x $\ z$-y) \ (x $+ y $\ z)" + "\x \ int; z \ int\ \(x $\ z$-y) \ (x $+ y $\ z)" apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff) apply (auto simp add: zdiff_def zadd_assoc) done @@ -820,7 +820,7 @@ of the CancelNumerals Simprocs\ lemma zadd_left_cancel: - "[| w \ int; w': int |] ==> (z $+ w' = z $+ w) \ (w' = w)" + "\w \ int; w': int\ \ (z $+ w' = z $+ w) \ (w' = w)" apply safe apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) apply (simp add: zadd_ac) @@ -833,7 +833,7 @@ done lemma zadd_right_cancel: - "[| w \ int; w': int |] ==> (w' $+ z = w $+ z) \ (w' = w)" + "\w \ int; w': int\ \ (w' $+ z = w $+ z) \ (w' = w)" apply safe apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) apply (simp add: zadd_ac) @@ -858,22 +858,22 @@ by (simp add: zadd_commute [of z] zadd_right_cancel_zle) -(*"v $\ w ==> v$+z $\ w$+z"*) +(*"v $\ w \ v$+z $\ w$+z"*) lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2] -(*"v $\ w ==> z$+v $\ z$+w"*) +(*"v $\ w \ z$+v $\ z$+w"*) lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2] -(*"v $\ w ==> v$+z $\ w$+z"*) +(*"v $\ w \ v$+z $\ w$+z"*) lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2] -(*"v $\ w ==> z$+v $\ z$+w"*) +(*"v $\ w \ z$+v $\ z$+w"*) lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2] -lemma zadd_zle_mono: "[| w' $\ w; z' $\ z |] ==> w' $+ z' $\ w $+ z" +lemma zadd_zle_mono: "\w' $\ w; z' $\ z\ \ w' $+ z' $\ w $+ z" by (erule zadd_zle_mono1 [THEN zle_trans], simp) -lemma zadd_zless_mono: "[| w' $< w; z' $\ z |] ==> w' $+ z' $< w $+ z" +lemma zadd_zless_mono: "\w' $< w; z' $\ z\ \ w' $+ z' $< w $+ z" by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp) @@ -887,10 +887,10 @@ subsubsection\More inequality lemmas\ -lemma equation_zminus: "[| x \ int; y \ int |] ==> (x = $- y) \ (y = $- x)" +lemma equation_zminus: "\x \ int; y \ int\ \ (x = $- y) \ (y = $- x)" by auto -lemma zminus_equation: "[| x \ int; y \ int |] ==> ($- x = y) \ ($- y = x)" +lemma zminus_equation: "\x \ int; y \ int\ \ ($- x = y) \ ($- y = x)" by auto lemma equation_zminus_intify: "(intify(x) = $- y) \ (intify(y) = $- x)" diff -r f2094906e491 -r e44d86131648 src/ZF/IntDiv.thy --- a/src/ZF/IntDiv.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/IntDiv.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,20 +11,20 @@ end fun negDivAlg (a,b) = - if 0<=a+b then (~1,a+b) + if 0<=a+b then (\1,a+b) else let val (q,r) = negDivAlg(a, 2*b) in if 0<=r-b then (2*q+1, r-b) else (2*q, r) end; - fun negateSnd (q,r:int) = (q,~r); + fun negateSnd (q,r:int) = (q,\r); fun divAlg (a,b) = if 0<=a then if b>0 then posDivAlg (a,b) else if a=0 then (0,0) - else negateSnd (negDivAlg (~a,~b)) + else negateSnd (negDivAlg (\a,\b)) else if 0a,\b)); *) section\The Division Operators Div and Mod\ @@ -35,13 +35,13 @@ definition quorem :: "[i,i] => o" where - "quorem == % . + "quorem \ % . a = b$*q $+ r & - (#0$r & r$ #0)" + (#0$r & r$(#0$ #0)" definition adjust :: "[i,i] => i" where - "adjust(b) == %. if #0 $\ r$-b then <#2$*q $+ #1,r$-b> + "adjust(b) \ %. if #0 $\ r$-b then <#2$*q $+ #1,r$-b> else <#2$*q,r>" @@ -51,7 +51,7 @@ posDivAlg :: "i => i" where (*for the case a>=0, b>0*) (*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) - "posDivAlg(ab) == + "posDivAlg(ab) \ wfrec(measure(int*int, %. nat_of (a $- b $+ #1)), ab, % f. if (a$#0) then <#0,a> @@ -62,7 +62,7 @@ definition negDivAlg :: "i => i" where (*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) - "negDivAlg(ab) == + "negDivAlg(ab) \ wfrec(measure(int*int, %. nat_of ($- a $- b)), ab, % f. if (#0 $\ a$+b | b$\#0) then <#-1,a$+b> @@ -72,13 +72,13 @@ definition negateSnd :: "i => i" where - "negateSnd == %. " + "negateSnd \ %. " (*The full division algorithm considers all possible signs for a, b including the special case a=0, b<0, because negDivAlg requires a<0*) definition divAlg :: "i => i" where - "divAlg == + "divAlg \ %. if #0 $\ a then if #0 $\ b then posDivAlg () else if a=#0 then <#0,#0> @@ -89,28 +89,28 @@ definition zdiv :: "[i,i]=>i" (infixl \zdiv\ 70) where - "a zdiv b == fst (divAlg ())" + "a zdiv b \ fst (divAlg ())" definition zmod :: "[i,i]=>i" (infixl \zmod\ 70) where - "a zmod b == snd (divAlg ())" + "a zmod b \ snd (divAlg ())" (** Some basic laws by Sidi Ehmety (need linear arithmetic!) **) -lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y" +lemma zspos_add_zspos_imp_zspos: "\#0 $< x; #0 $< y\ \ #0 $< x $+ y" apply (rule_tac y = "y" in zless_trans) apply (rule_tac [2] zdiff_zless_iff [THEN iffD1]) apply auto done -lemma zpos_add_zpos_imp_zpos: "[| #0 $\ x; #0 $\ y |] ==> #0 $\ x $+ y" +lemma zpos_add_zpos_imp_zpos: "\#0 $\ x; #0 $\ y\ \ #0 $\ x $+ y" apply (rule_tac y = "y" in zle_trans) apply (rule_tac [2] zdiff_zle_iff [THEN iffD1]) apply auto done -lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0" +lemma zneg_add_zneg_imp_zneg: "\x $< #0; y $< #0\ \ x $+ y $< #0" apply (rule_tac y = "y" in zless_trans) apply (rule zless_zdiff_iff [THEN iffD1]) apply auto @@ -118,13 +118,13 @@ (* this theorem is used below *) lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0: - "[| x $\ #0; y $\ #0 |] ==> x $+ y $\ #0" + "\x $\ #0; y $\ #0\ \ x $+ y $\ #0" apply (rule_tac y = "y" in zle_trans) apply (rule zle_zdiff_iff [THEN iffD1]) apply auto done -lemma zero_lt_zmagnitude: "[| #0 $< k; k \ int |] ==> 0 < zmagnitude(k)" +lemma zero_lt_zmagnitude: "\#0 $< k; k \ int\ \ 0 < zmagnitude(k)" apply (drule zero_zless_imp_znegative_zminus) apply (drule_tac [2] zneg_int_of) apply (auto simp add: zminus_equation [of k]) @@ -151,7 +151,7 @@ done lemma zadd_succ_lemma: - "z \ int ==> (w $+ $# succ(m) $\ z) \ (w $+ $#m $< z)" + "z \ int \ (w $+ $# succ(m) $\ z) \ (w $+ $#m $< z)" apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff) apply (auto intro: zle_anti_sym elim: zless_asym simp add: zless_imp_zle not_zless_iff_zle) @@ -184,13 +184,13 @@ (*** Monotonicity of Multiplication ***) -lemma zmult_mono_lemma: "k \ nat ==> i $\ j ==> i $* $#k $\ j $* $#k" +lemma zmult_mono_lemma: "k \ nat \ i $\ j \ i $* $#k $\ j $* $#k" apply (induct_tac "k") prefer 2 apply (subst int_succ_int_1) apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono) done -lemma zmult_zle_mono1: "[| i $\ j; #0 $\ k |] ==> i$*k $\ j$*k" +lemma zmult_zle_mono1: "\i $\ j; #0 $\ k\ \ i$*k $\ j$*k" apply (subgoal_tac "i $* intify (k) $\ j $* intify (k) ") apply (simp (no_asm_use)) apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) @@ -199,25 +199,25 @@ apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym]) done -lemma zmult_zle_mono1_neg: "[| i $\ j; k $\ #0 |] ==> j$*k $\ i$*k" +lemma zmult_zle_mono1_neg: "\i $\ j; k $\ #0\ \ j$*k $\ i$*k" apply (rule zminus_zle_zminus [THEN iffD1]) apply (simp del: zmult_zminus_right add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus) done -lemma zmult_zle_mono2: "[| i $\ j; #0 $\ k |] ==> k$*i $\ k$*j" +lemma zmult_zle_mono2: "\i $\ j; #0 $\ k\ \ k$*i $\ k$*j" apply (drule zmult_zle_mono1) apply (simp_all add: zmult_commute) done -lemma zmult_zle_mono2_neg: "[| i $\ j; k $\ #0 |] ==> k$*j $\ k$*i" +lemma zmult_zle_mono2_neg: "\i $\ j; k $\ #0\ \ k$*j $\ k$*i" apply (drule zmult_zle_mono1_neg) apply (simp_all add: zmult_commute) done (* $\ monotonicity, BOTH arguments*) lemma zmult_zle_mono: - "[| i $\ j; k $\ l; #0 $\ j; #0 $\ k |] ==> i$*k $\ j$*l" + "\i $\ j; k $\ l; #0 $\ j; #0 $\ k\ \ i$*k $\ j$*l" apply (erule zmult_zle_mono1 [THEN zle_trans]) apply assumption apply (erule zmult_zle_mono2) @@ -228,7 +228,7 @@ (** strict, in 1st argument; proof is by induction on k>0 **) lemma zmult_zless_mono2_lemma [rule_format]: - "[| i$ nat |] ==> 0 $#k $* i $< $#k $* j" + "\i$ nat\ \ 0 $#k $* i $< $#k $* j" apply (induct_tac "k") prefer 2 apply (subst int_succ_int_1) @@ -241,7 +241,7 @@ apply (simp_all (no_asm_simp) add: zle_def) done -lemma zmult_zless_mono2: "[| i$ k$*i $< k$*j" +lemma zmult_zless_mono2: "\i$ \ k$*i $< k$*j" apply (subgoal_tac "intify (k) $* i $< intify (k) $* j") apply (simp (no_asm_use)) apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) @@ -252,27 +252,27 @@ apply (auto simp add: zero_lt_zmagnitude) done -lemma zmult_zless_mono1: "[| i$ i$*k $< j$*k" +lemma zmult_zless_mono1: "\i$ \ i$*k $< j$*k" apply (drule zmult_zless_mono2) apply (simp_all add: zmult_commute) done (* < monotonicity, BOTH arguments*) lemma zmult_zless_mono: - "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l" + "\i $< j; k $< l; #0 $< j; #0 $< k\ \ i$*k $< j$*l" apply (erule zmult_zless_mono1 [THEN zless_trans]) apply assumption apply (erule zmult_zless_mono2) apply assumption done -lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k" +lemma zmult_zless_mono1_neg: "\i $< j; k $< #0\ \ j$*k $< i$*k" apply (rule zminus_zless_zminus [THEN iffD1]) apply (simp del: zmult_zminus_right add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus) done -lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i" +lemma zmult_zless_mono2_neg: "\i $< j; k $< #0\ \ k$*j $< k$*i" apply (rule zminus_zless_zminus [THEN iffD1]) apply (simp del: zmult_zminus add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus) @@ -282,7 +282,7 @@ (** Products of zeroes **) lemma zmult_eq_lemma: - "[| m \ int; n \ int |] ==> (m = #0 | n = #0) \ (m$*n = #0)" + "\m \ int; n \ int\ \ (m = #0 | n = #0) \ (m$*n = #0)" apply (case_tac "m $< #0") apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless) apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+ @@ -297,8 +297,8 @@ but not (yet?) for k*m < n*k. **) lemma zmult_zless_lemma: - "[| k \ int; m \ int; n \ int |] - ==> (m$*k $< n$*k) \ ((#0 $< k & m$k \ int; m \ int; n \ int\ + \ (m$*k $< n$*k) \ ((#0 $< k & m$ k$*n) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1) -lemma int_eq_iff_zle: "[| m \ int; n \ int |] ==> m=n \ (m $\ n & n $\ m)" +lemma int_eq_iff_zle: "\m \ int; n \ int\ \ m=n \ (m $\ n & n $\ m)" apply (blast intro: zle_refl zle_anti_sym) done lemma zmult_cancel2_lemma: - "[| k \ int; m \ int; n \ int |] ==> (m$*k = n$*k) \ (k=#0 | m=n)" + "\k \ int; m \ int; n \ int\ \ (m$*k = n$*k) \ (k=#0 | m=n)" apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m]) apply (auto simp add: zmult_zle_cancel2 neq_iff_zless) done @@ -352,8 +352,8 @@ subsection\Uniqueness and monotonicity of quotients and remainders\ lemma unique_quotient_lemma: - "[| b$*q' $+ r' $\ b$*q $+ r; #0 $\ r'; #0 $< b; r $< b |] - ==> q' $\ q" + "\b$*q' $+ r' $\ b$*q $+ r; #0 $\ r'; #0 $< b; r $< b\ + \ q' $\ q" apply (subgoal_tac "r' $+ b $* (q'$-q) $\ r") prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls) apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ") @@ -370,8 +370,8 @@ done lemma unique_quotient_lemma_neg: - "[| b$*q' $+ r' $\ b$*q $+ r; r $\ #0; b $< #0; b $< r' |] - ==> q $\ q'" + "\b$*q' $+ r' $\ b$*q $+ r; r $\ #0; b $< #0; b $< r'\ + \ q $\ q'" apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" in unique_quotient_lemma) apply (auto simp del: zminus_zadd_distrib @@ -380,8 +380,8 @@ lemma unique_quotient: - "[| quorem (, ); quorem (, ); b \ int; b \ #0; - q \ int; q' \ int |] ==> q = q'" + "\quorem (, ); quorem (, ); b \ int; b \ #0; + q \ int; q' \ int\ \ q = q'" apply (simp add: split_ifs quorem_def neq_iff_zless) apply safe apply simp_all @@ -391,9 +391,9 @@ done lemma unique_remainder: - "[| quorem (, ); quorem (, ); b \ int; b \ #0; + "\quorem (, ); quorem (, ); b \ int; b \ #0; q \ int; q' \ int; - r \ int; r' \ int |] ==> r = r'" + r \ int; r' \ int\ \ r = r'" apply (subgoal_tac "q = q'") prefer 2 apply (blast intro: unique_quotient) apply (simp add: quorem_def) @@ -411,8 +411,8 @@ lemma posDivAlg_termination: - "[| #0 $< b; ~ a $< b |] - ==> nat_of(a $- #2 $* b $+ #1) < nat_of(a $- b $+ #1)" + "\#0 $< b; \ a $< b\ + \ nat_of(a $- #2 $* b $+ #1) < nat_of(a $- b $+ #1)" apply (simp (no_asm) add: zless_nat_conj) apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls) done @@ -420,7 +420,7 @@ lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure] lemma posDivAlg_eqn: - "[| #0 $< b; a \ int; b \ int |] ==> + "\#0 $< b; a \ int; b \ int\ \ posDivAlg() = (if a$ else adjust(b, posDivAlg ()))" apply (rule posDivAlg_unfold [THEN trans]) @@ -430,8 +430,8 @@ lemma posDivAlg_induct_lemma [rule_format]: assumes prem: - "!!a b. [| a \ int; b \ int; - ~ (a $< b | b $\ #0) \ P() |] ==> P()" + "\a b. \a \ int; b \ int; + \ (a $< b | b $\ #0) \ P()\ \ P()" shows " \ int*int \ P()" using wf_measure [where A = "int*int" and f = "%.nat_of (a $- b $+ #1)"] proof (induct "" arbitrary: u v rule: wf_induct) @@ -449,8 +449,8 @@ lemma posDivAlg_induct [consumes 2]: assumes u_int: "u \ int" and v_int: "v \ int" - and ih: "!!a b. [| a \ int; b \ int; - ~ (a $< b | b $\ #0) \ P(a, #2 $* b) |] ==> P(a,b)" + and ih: "\a b. \a \ int; b \ int; + \ (a $< b | b $\ #0) \ P(a, #2 $* b)\ \ P(a,b)" shows "P(u,v)" apply (subgoal_tac "(%. P (x,y)) ()") apply simp @@ -461,28 +461,28 @@ done (*FIXME: use intify in integ_of so that we always have @{term"integ_of w \ int"}. - then this rewrite can work for all constants!!*) + then this rewrite can work for all constants\*) lemma intify_eq_0_iff_zle: "intify(m) = #0 \ (m $\ #0 & #0 $\ m)" by (simp add: int_eq_iff_zle) subsection\Some convenient biconditionals for products of signs\ -lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" +lemma zmult_pos: "\#0 $< i; #0 $< j\ \ #0 $< i $* j" by (drule zmult_zless_mono1, auto) -lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" +lemma zmult_neg: "\i $< #0; j $< #0\ \ #0 $< i $* j" by (drule zmult_zless_mono1_neg, auto) -lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" +lemma zmult_pos_neg: "\#0 $< i; j $< #0\ \ i $* j $< #0" by (drule zmult_zless_mono1_neg, auto) (** Inequality reasoning **) lemma int_0_less_lemma: - "[| x \ int; y \ int |] - ==> (#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" + "\x \ int; y \ int\ + \ (#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg) apply (rule ccontr) apply (rule_tac [2] ccontr) @@ -502,8 +502,8 @@ done lemma int_0_le_lemma: - "[| x \ int; y \ int |] - ==> (#0 $\ x $* y) \ (#0 $\ x & #0 $\ y | x $\ #0 & y $\ #0)" + "\x \ int; y \ int\ + \ (#0 $\ x $* y) \ (#0 $\ x & #0 $\ y | x $\ #0 & y $\ #0)" by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff) lemma int_0_le_mult_iff: @@ -526,7 +526,7 @@ (*Typechecking for posDivAlg*) lemma posDivAlg_type [rule_format]: - "[| a \ int; b \ int |] ==> posDivAlg() \ int * int" + "\a \ int; b \ int\ \ posDivAlg() \ int * int" apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) apply assumption+ apply (case_tac "#0 $< ba") @@ -541,8 +541,8 @@ (*Correctness of posDivAlg: it computes quotients correctly*) lemma posDivAlg_correct [rule_format]: - "[| a \ int; b \ int |] - ==> #0 $\ a \ #0 $< b \ quorem (, posDivAlg())" + "\a \ int; b \ int\ + \ #0 $\ a \ #0 $< b \ quorem (, posDivAlg())" apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) apply auto apply (simp_all add: quorem_def) @@ -565,8 +565,8 @@ subsection\Correctness of negDivAlg, the division algorithm for a<0 and b>0\ lemma negDivAlg_termination: - "[| #0 $< b; a $+ b $< #0 |] - ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)" + "\#0 $< b; a $+ b $< #0\ + \ nat_of($- a $- #2 $* b) < nat_of($- a $- b)" apply (simp (no_asm) add: zless_nat_conj) apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym] zless_zminus) @@ -575,7 +575,7 @@ lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure] lemma negDivAlg_eqn: - "[| #0 $< b; a \ int; b \ int |] ==> + "\#0 $< b; a \ int; b \ int\ \ negDivAlg() = (if #0 $\ a$+b then <#-1,a$+b> else adjust(b, negDivAlg ()))" @@ -586,9 +586,9 @@ lemma negDivAlg_induct_lemma [rule_format]: assumes prem: - "!!a b. [| a \ int; b \ int; - ~ (#0 $\ a $+ b | b $\ #0) \ P() |] - ==> P()" + "\a b. \a \ int; b \ int; + \ (#0 $\ a $+ b | b $\ #0) \ P()\ + \ P()" shows " \ int*int \ P()" using wf_measure [where A = "int*int" and f = "%.nat_of ($- a $- b)"] proof (induct "" arbitrary: u v rule: wf_induct) @@ -605,9 +605,9 @@ lemma negDivAlg_induct [consumes 2]: assumes u_int: "u \ int" and v_int: "v \ int" - and ih: "!!a b. [| a \ int; b \ int; - ~ (#0 $\ a $+ b | b $\ #0) \ P(a, #2 $* b) |] - ==> P(a,b)" + and ih: "\a b. \a \ int; b \ int; + \ (#0 $\ a $+ b | b $\ #0) \ P(a, #2 $* b)\ + \ P(a,b)" shows "P(u,v)" apply (subgoal_tac " (%. P (x,y)) ()") apply simp @@ -620,7 +620,7 @@ (*Typechecking for negDivAlg*) lemma negDivAlg_type: - "[| a \ int; b \ int |] ==> negDivAlg() \ int * int" + "\a \ int; b \ int\ \ negDivAlg() \ int * int" apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) apply assumption+ apply (case_tac "#0 $< ba") @@ -637,8 +637,8 @@ (*Correctness of negDivAlg: it computes quotients correctly It doesn't work if a=0 because the 0/b=0 rather than -1*) lemma negDivAlg_correct [rule_format]: - "[| a \ int; b \ int |] - ==> a $< #0 \ #0 $< b \ quorem (, negDivAlg())" + "\a \ int; b \ int\ + \ a $< #0 \ #0 $< b \ quorem (, negDivAlg())" apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) apply auto apply (simp_all add: quorem_def) @@ -661,7 +661,7 @@ subsection\Existence shown by proving the division algorithm to be correct\ (*the case a=0*) -lemma quorem_0: "[|b \ #0; b \ int|] ==> quorem (<#0,b>, <#0,#0>)" +lemma quorem_0: "\b \ #0; b \ int\ \ quorem (<#0,b>, <#0,#0>)" by (force simp add: quorem_def neq_iff_zless) lemma posDivAlg_zero_divisor: "posDivAlg() = <#0,a>" @@ -676,7 +676,7 @@ (*Needed below. Actually it's an equivalence.*) -lemma linear_arith_lemma: "~ (#0 $\ #-1 $+ b) ==> (b $\ #0)" +lemma linear_arith_lemma: "\ (#0 $\ #-1 $+ b) \ (b $\ #0)" apply (simp add: not_zle_iff_zless) apply (drule zminus_zless_zminus [THEN iffD2]) apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus) @@ -692,14 +692,14 @@ apply auto done -lemma negateSnd_type: "qr \ int * int ==> negateSnd (qr) \ int * int" +lemma negateSnd_type: "qr \ int * int \ negateSnd (qr) \ int * int" apply (unfold negateSnd_def) apply auto done lemma quorem_neg: - "[|quorem (<$-a,$-b>, qr); a \ int; b \ int; qr \ int * int|] - ==> quorem (, negateSnd(qr))" + "\quorem (<$-a,$-b>, qr); a \ int; b \ int; qr \ int * int\ + \ quorem (, negateSnd(qr))" apply clarify apply (auto elim: zless_asym simp add: quorem_def zless_zminus) txt\linear arithmetic from here on\ @@ -711,7 +711,7 @@ done lemma divAlg_correct: - "[|b \ #0; a \ int; b \ int|] ==> quorem (, divAlg())" + "\b \ #0; a \ int; b \ int\ \ quorem (, divAlg())" apply (auto simp add: quorem_0 divAlg_def) apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct posDivAlg_type negDivAlg_type) @@ -720,7 +720,7 @@ apply (auto simp add: zle_def) done -lemma divAlg_type: "[|a \ int; b \ int|] ==> divAlg() \ int * int" +lemma divAlg_type: "\a \ int; b \ int\ \ divAlg() \ int * int" apply (auto simp add: divAlg_def) apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) done @@ -765,7 +765,7 @@ (** Basic laws about division and remainder **) lemma raw_zmod_zdiv_equality: - "[| a \ int; b \ int |] ==> a = b $* (a zdiv b) $+ (a zmod b)" + "\a \ int; b \ int\ \ a = b $* (a zdiv b) $+ (a zmod b)" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (cut_tac a = "a" and b = "b" in divAlg_correct) @@ -778,7 +778,7 @@ apply auto done -lemma pos_mod: "#0 $< b ==> #0 $\ a zmod b & a zmod b $< b" +lemma pos_mod: "#0 $< b \ #0 $\ a zmod b & a zmod b $< b" apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) apply (blast dest: zle_zless_trans)+ @@ -787,7 +787,7 @@ lemmas pos_mod_sign = pos_mod [THEN conjunct1] and pos_mod_bound = pos_mod [THEN conjunct2] -lemma neg_mod: "b $< #0 ==> a zmod b $\ #0 & b $< a zmod b" +lemma neg_mod: "b $< #0 \ a zmod b $\ #0 & b $< a zmod b" apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) apply (blast dest: zle_zless_trans) @@ -801,8 +801,8 @@ (** proving general properties of zdiv and zmod **) lemma quorem_div_mod: - "[|b \ #0; a \ int; b \ int |] - ==> quorem (, )" + "\b \ #0; a \ int; b \ int\ + \ quorem (, )" apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) @@ -810,58 +810,58 @@ (*Surely quorem(,) implies @{term"a \ int"}, but it doesn't matter*) lemma quorem_div: - "[| quorem(,); b \ #0; a \ int; b \ int; q \ int |] - ==> a zdiv b = q" + "\quorem(,); b \ #0; a \ int; b \ int; q \ int\ + \ a zdiv b = q" by (blast intro: quorem_div_mod [THEN unique_quotient]) lemma quorem_mod: - "[| quorem(,); b \ #0; a \ int; b \ int; q \ int; r \ int |] - ==> a zmod b = r" + "\quorem(,); b \ #0; a \ int; b \ int; q \ int; r \ int\ + \ a zmod b = r" by (blast intro: quorem_div_mod [THEN unique_remainder]) lemma zdiv_pos_pos_trivial_raw: - "[| a \ int; b \ int; #0 $\ a; a $< b |] ==> a zdiv b = #0" + "\a \ int; b \ int; #0 $\ a; a $< b\ \ a zdiv b = #0" apply (rule quorem_div) apply (auto simp add: quorem_def) (*linear arithmetic*) apply (blast dest: zle_zless_trans)+ done -lemma zdiv_pos_pos_trivial: "[| #0 $\ a; a $< b |] ==> a zdiv b = #0" +lemma zdiv_pos_pos_trivial: "\#0 $\ a; a $< b\ \ a zdiv b = #0" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_pos_pos_trivial_raw) apply auto done lemma zdiv_neg_neg_trivial_raw: - "[| a \ int; b \ int; a $\ #0; b $< a |] ==> a zdiv b = #0" + "\a \ int; b \ int; a $\ #0; b $< a\ \ a zdiv b = #0" apply (rule_tac r = "a" in quorem_div) apply (auto simp add: quorem_def) (*linear arithmetic*) apply (blast dest: zle_zless_trans zless_trans)+ done -lemma zdiv_neg_neg_trivial: "[| a $\ #0; b $< a |] ==> a zdiv b = #0" +lemma zdiv_neg_neg_trivial: "\a $\ #0; b $< a\ \ a zdiv b = #0" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_neg_neg_trivial_raw) apply auto done -lemma zadd_le_0_lemma: "[| a$+b $\ #0; #0 $< a; #0 $< b |] ==> False" +lemma zadd_le_0_lemma: "\a$+b $\ #0; #0 $< a; #0 $< b\ \ False" apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono) apply (auto simp add: zle_def) apply (blast dest: zless_trans) done lemma zdiv_pos_neg_trivial_raw: - "[| a \ int; b \ int; #0 $< a; a$+b $\ #0 |] ==> a zdiv b = #-1" + "\a \ int; b \ int; #0 $< a; a$+b $\ #0\ \ a zdiv b = #-1" apply (rule_tac r = "a $+ b" in quorem_div) apply (auto simp add: quorem_def) (*linear arithmetic*) apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ done -lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $\ #0 |] ==> a zdiv b = #-1" +lemma zdiv_pos_neg_trivial: "\#0 $< a; a$+b $\ #0\ \ a zdiv b = #-1" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_pos_neg_trivial_raw) apply auto @@ -871,42 +871,42 @@ lemma zmod_pos_pos_trivial_raw: - "[| a \ int; b \ int; #0 $\ a; a $< b |] ==> a zmod b = a" + "\a \ int; b \ int; #0 $\ a; a $< b\ \ a zmod b = a" apply (rule_tac q = "#0" in quorem_mod) apply (auto simp add: quorem_def) (*linear arithmetic*) apply (blast dest: zle_zless_trans)+ done -lemma zmod_pos_pos_trivial: "[| #0 $\ a; a $< b |] ==> a zmod b = intify(a)" +lemma zmod_pos_pos_trivial: "\#0 $\ a; a $< b\ \ a zmod b = intify(a)" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_pos_pos_trivial_raw) apply auto done lemma zmod_neg_neg_trivial_raw: - "[| a \ int; b \ int; a $\ #0; b $< a |] ==> a zmod b = a" + "\a \ int; b \ int; a $\ #0; b $< a\ \ a zmod b = a" apply (rule_tac q = "#0" in quorem_mod) apply (auto simp add: quorem_def) (*linear arithmetic*) apply (blast dest: zle_zless_trans zless_trans)+ done -lemma zmod_neg_neg_trivial: "[| a $\ #0; b $< a |] ==> a zmod b = intify(a)" +lemma zmod_neg_neg_trivial: "\a $\ #0; b $< a\ \ a zmod b = intify(a)" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_neg_neg_trivial_raw) apply auto done lemma zmod_pos_neg_trivial_raw: - "[| a \ int; b \ int; #0 $< a; a$+b $\ #0 |] ==> a zmod b = a$+b" + "\a \ int; b \ int; #0 $< a; a$+b $\ #0\ \ a zmod b = a$+b" apply (rule_tac q = "#-1" in quorem_mod) apply (auto simp add: quorem_def) (*linear arithmetic*) apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ done -lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $\ #0 |] ==> a zmod b = a$+b" +lemma zmod_pos_neg_trivial: "\#0 $< a; a$+b $\ #0\ \ a zmod b = a$+b" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_pos_neg_trivial_raw) apply auto @@ -918,7 +918,7 @@ (*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*) lemma zdiv_zminus_zminus_raw: - "[|a \ int; b \ int|] ==> ($-a) zdiv ($-b) = a zdiv b" + "\a \ int; b \ int\ \ ($-a) zdiv ($-b) = a zdiv b" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div]) @@ -932,7 +932,7 @@ (*Simpler laws such as -a zmod b = -(a zmod b) FAIL*) lemma zmod_zminus_zminus_raw: - "[|a \ int; b \ int|] ==> ($-a) zmod ($-b) = $- (a zmod b)" + "\a \ int; b \ int\ \ ($-a) zmod ($-b) = $- (a zmod b)" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod]) @@ -947,7 +947,7 @@ subsection\division of a number by itself\ -lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $\ q" +lemma self_quotient_aux1: "\#0 $< a; a = r $+ a$*q; r $< a\ \ #1 $\ q" apply (subgoal_tac "#0 $< a$*q") apply (cut_tac w = "#0" and z = "q" in add1_zle_iff) apply (simp add: int_0_less_mult_iff) @@ -958,7 +958,7 @@ apply (simp add: zcompare_rls) done -lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $\ r |] ==> q $\ #1" +lemma self_quotient_aux2: "\#0 $< a; a = r $+ a$*q; #0 $\ r\ \ q $\ #1" apply (subgoal_tac "#0 $\ a$* (#1$-q)") apply (simp add: int_0_le_mult_iff zcompare_rls) apply (blast dest: zle_zless_trans) @@ -968,7 +968,7 @@ done lemma self_quotient: - "[| quorem(,); a \ int; q \ int; a \ #0|] ==> q = #1" + "\quorem(,); a \ int; q \ int; a \ #0\ \ q = #1" apply (simp add: split_ifs quorem_def neq_iff_zless) apply (rule zle_anti_sym) apply safe @@ -984,22 +984,22 @@ done lemma self_remainder: - "[|quorem(,); a \ int; q \ int; r \ int; a \ #0|] ==> r = #0" + "\quorem(,); a \ int; q \ int; r \ int; a \ #0\ \ r = #0" apply (frule self_quotient) apply (auto simp add: quorem_def) done -lemma zdiv_self_raw: "[|a \ #0; a \ int|] ==> a zdiv a = #1" +lemma zdiv_self_raw: "\a \ #0; a \ int\ \ a zdiv a = #1" apply (blast intro: quorem_div_mod [THEN self_quotient]) done -lemma zdiv_self [simp]: "intify(a) \ #0 ==> a zdiv a = #1" +lemma zdiv_self [simp]: "intify(a) \ #0 \ a zdiv a = #1" apply (drule zdiv_self_raw) apply auto done (*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *) -lemma zmod_self_raw: "a \ int ==> a zmod a = #0" +lemma zmod_self_raw: "a \ int \ a zmod a = #0" apply (case_tac "a = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (blast intro: quorem_div_mod [THEN self_remainder]) @@ -1016,29 +1016,29 @@ lemma zdiv_zero [simp]: "#0 zdiv b = #0" by (simp add: zdiv_def divAlg_def) -lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" +lemma zdiv_eq_minus1: "#0 $< b \ #-1 zdiv b = #-1" by (simp (no_asm_simp) add: zdiv_def divAlg_def) lemma zmod_zero [simp]: "#0 zmod b = #0" by (simp add: zmod_def divAlg_def) -lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" +lemma zdiv_minus1: "#0 $< b \ #-1 zdiv b = #-1" by (simp add: zdiv_def divAlg_def) -lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" +lemma zmod_minus1: "#0 $< b \ #-1 zmod b = b $- #1" by (simp add: zmod_def divAlg_def) (** a positive, b positive **) -lemma zdiv_pos_pos: "[| #0 $< a; #0 $\ b |] - ==> a zdiv b = fst (posDivAlg())" +lemma zdiv_pos_pos: "\#0 $< a; #0 $\ b\ + \ a zdiv b = fst (posDivAlg())" apply (simp (no_asm_simp) add: zdiv_def divAlg_def) apply (auto simp add: zle_def) done lemma zmod_pos_pos: - "[| #0 $< a; #0 $\ b |] - ==> a zmod b = snd (posDivAlg())" + "\#0 $< a; #0 $\ b\ + \ a zmod b = snd (posDivAlg())" apply (simp (no_asm_simp) add: zmod_def divAlg_def) apply (auto simp add: zle_def) done @@ -1046,15 +1046,15 @@ (** a negative, b positive **) lemma zdiv_neg_pos: - "[| a $< #0; #0 $< b |] - ==> a zdiv b = fst (negDivAlg())" + "\a $< #0; #0 $< b\ + \ a zdiv b = fst (negDivAlg())" apply (simp (no_asm_simp) add: zdiv_def divAlg_def) apply (blast dest: zle_zless_trans) done lemma zmod_neg_pos: - "[| a $< #0; #0 $< b |] - ==> a zmod b = snd (negDivAlg())" + "\a $< #0; #0 $< b\ + \ a zmod b = snd (negDivAlg())" apply (simp (no_asm_simp) add: zmod_def divAlg_def) apply (blast dest: zle_zless_trans) done @@ -1062,8 +1062,8 @@ (** a positive, b negative **) lemma zdiv_pos_neg: - "[| #0 $< a; b $< #0 |] - ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" + "\#0 $< a; b $< #0\ + \ a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle) apply auto apply (blast dest: zle_zless_trans)+ @@ -1072,8 +1072,8 @@ done lemma zmod_pos_neg: - "[| #0 $< a; b $< #0 |] - ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" + "\#0 $< a; b $< #0\ + \ a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle) apply auto apply (blast dest: zle_zless_trans)+ @@ -1084,16 +1084,16 @@ (** a negative, b negative **) lemma zdiv_neg_neg: - "[| a $< #0; b $\ #0 |] - ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" + "\a $< #0; b $\ #0\ + \ a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" apply (simp (no_asm_simp) add: zdiv_def divAlg_def) apply auto apply (blast dest!: zle_zless_trans)+ done lemma zmod_neg_neg: - "[| a $< #0; b $\ #0 |] - ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" + "\a $< #0; b $\ #0\ + \ a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" apply (simp (no_asm_simp) add: zmod_def divAlg_def) apply auto apply (blast dest!: zle_zless_trans)+ @@ -1138,7 +1138,7 @@ apply auto done -lemma zdiv_minus1_right_raw: "a \ int ==> a zdiv #-1 = $-a" +lemma zdiv_minus1_right_raw: "a \ int \ a zdiv #-1 = $-a" apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality) apply auto apply (rule equation_zminus [THEN iffD2]) @@ -1154,7 +1154,7 @@ subsection\Monotonicity in the first argument (divisor)\ -lemma zdiv_mono1: "[| a $\ a'; #0 $< b |] ==> a zdiv b $\ a' zdiv b" +lemma zdiv_mono1: "\a $\ a'; #0 $< b\ \ a zdiv b $\ a' zdiv b" apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) apply (rule unique_quotient_lemma) @@ -1163,7 +1163,7 @@ apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound) done -lemma zdiv_mono1_neg: "[| a $\ a'; b $< #0 |] ==> a' zdiv b $\ a zdiv b" +lemma zdiv_mono1_neg: "\a $\ a'; b $< #0\ \ a' zdiv b $\ a zdiv b" apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) apply (rule unique_quotient_lemma_neg) @@ -1176,7 +1176,7 @@ subsection\Monotonicity in the second argument (dividend)\ lemma q_pos_lemma: - "[| #0 $\ b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $\ q'" + "\#0 $\ b'$*q' $+ r'; r' $< b'; #0 $< b'\ \ #0 $\ q'" apply (subgoal_tac "#0 $< b'$* (q' $+ #1)") apply (simp add: int_0_less_mult_iff) apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1]) @@ -1186,9 +1186,9 @@ done lemma zdiv_mono2_lemma: - "[| b$*q $+ r = b'$*q' $+ r'; #0 $\ b'$*q' $+ r'; - r' $< b'; #0 $\ r; #0 $< b'; b' $\ b |] - ==> q $\ q'" + "\b$*q $+ r = b'$*q' $+ r'; #0 $\ b'$*q' $+ r'; + r' $< b'; #0 $\ r; #0 $< b'; b' $\ b\ + \ q $\ q'" apply (frule q_pos_lemma, assumption+) apply (subgoal_tac "b$*q $< b$* (q' $+ #1)") apply (simp add: zmult_zless_cancel1) @@ -1207,8 +1207,8 @@ lemma zdiv_mono2_raw: - "[| #0 $\ a; #0 $< b'; b' $\ b; a \ int |] - ==> a zdiv b $\ a zdiv b'" + "\#0 $\ a; #0 $< b'; b' $\ b; a \ int\ + \ a zdiv b $\ a zdiv b'" apply (subgoal_tac "#0 $< b") prefer 2 apply (blast dest: zless_zle_trans) apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) @@ -1220,14 +1220,14 @@ done lemma zdiv_mono2: - "[| #0 $\ a; #0 $< b'; b' $\ b |] - ==> a zdiv b $\ a zdiv b'" + "\#0 $\ a; #0 $< b'; b' $\ b\ + \ a zdiv b $\ a zdiv b'" apply (cut_tac a = "intify (a)" in zdiv_mono2_raw) apply auto done lemma q_neg_lemma: - "[| b'$*q' $+ r' $< #0; #0 $\ r'; #0 $< b' |] ==> q' $< #0" + "\b'$*q' $+ r' $< #0; #0 $\ r'; #0 $< b'\ \ q' $< #0" apply (subgoal_tac "b'$*q' $< #0") prefer 2 apply (force intro: zle_zless_trans) apply (simp add: zmult_less_0_iff) @@ -1237,9 +1237,9 @@ lemma zdiv_mono2_neg_lemma: - "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; - r $< b; #0 $\ r'; #0 $< b'; b' $\ b |] - ==> q' $\ q" + "\b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; + r $< b; #0 $\ r'; #0 $< b'; b' $\ b\ + \ q' $\ q" apply (subgoal_tac "#0 $< b") prefer 2 apply (blast dest: zless_zle_trans) apply (frule q_neg_lemma, assumption+) @@ -1266,8 +1266,8 @@ done lemma zdiv_mono2_neg_raw: - "[| a $< #0; #0 $< b'; b' $\ b; a \ int |] - ==> a zdiv b' $\ a zdiv b" + "\a $< #0; #0 $< b'; b' $\ b; a \ int\ + \ a zdiv b' $\ a zdiv b" apply (subgoal_tac "#0 $< b") prefer 2 apply (blast dest: zless_zle_trans) apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) @@ -1278,8 +1278,8 @@ apply (simp_all add: pos_mod_sign pos_mod_bound) done -lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $\ b |] - ==> a zdiv b' $\ a zdiv b" +lemma zdiv_mono2_neg: "\a $< #0; #0 $< b'; b' $\ b\ + \ a zdiv b' $\ a zdiv b" apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw) apply auto done @@ -1291,16 +1291,16 @@ (** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) lemma zmult1_lemma: - "[| quorem(, ); c \ int; c \ #0 |] - ==> quorem (, )" + "\quorem(, ); c \ int; c \ #0\ + \ quorem (, )" apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) apply (auto intro: raw_zmod_zdiv_equality) done lemma zdiv_zmult1_eq_raw: - "[|b \ int; c \ int|] - ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" + "\b \ int; c \ int\ + \ (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" apply (case_tac "c = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div]) @@ -1313,7 +1313,7 @@ done lemma zmod_zmult1_eq_raw: - "[|b \ int; c \ int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c" + "\b \ int; c \ int\ \ (a$*b) zmod c = a$*(b zmod c) zmod c" apply (case_tac "c = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod]) @@ -1337,10 +1337,10 @@ apply (rule zmod_zmult1_eq) done -lemma zdiv_zmult_self1 [simp]: "intify(b) \ #0 ==> (a$*b) zdiv b = intify(a)" +lemma zdiv_zmult_self1 [simp]: "intify(b) \ #0 \ (a$*b) zdiv b = intify(a)" by (simp add: zdiv_zmult1_eq) -lemma zdiv_zmult_self2 [simp]: "intify(b) \ #0 ==> (b$*a) zdiv b = intify(a)" +lemma zdiv_zmult_self2 [simp]: "intify(b) \ #0 \ (b$*a) zdiv b = intify(a)" by (simp add: zmult_commute) lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" @@ -1354,9 +1354,9 @@ a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) lemma zadd1_lemma: - "[| quorem(, ); quorem(, ); - c \ int; c \ #0 |] - ==> quorem (, )" + "\quorem(, ); quorem(, ); + c \ int; c \ #0\ + \ quorem (, )" apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) apply (auto intro: raw_zmod_zdiv_equality) @@ -1364,7 +1364,7 @@ (*NOT suitable for rewriting: the RHS has an instance of the LHS*) lemma zdiv_zadd1_eq_raw: - "[|a \ int; b \ int; c \ int|] ==> + "\a \ int; b \ int; c \ int\ \ (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" apply (case_tac "c = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) @@ -1380,8 +1380,8 @@ done lemma zmod_zadd1_eq_raw: - "[|a \ int; b \ int; c \ int|] - ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" + "\a \ int; b \ int; c \ int\ + \ (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" apply (case_tac "c = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, @@ -1395,7 +1395,7 @@ done lemma zmod_div_trivial_raw: - "[|a \ int; b \ int|] ==> (a zmod b) zdiv b = #0" + "\a \ int; b \ int\ \ (a zmod b) zdiv b = #0" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound @@ -1408,7 +1408,7 @@ done lemma zmod_mod_trivial_raw: - "[|a \ int; b \ int|] ==> (a zmod b) zmod b = a zmod b" + "\a \ int; b \ int\ \ (a zmod b) zmod b = a zmod b" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound @@ -1436,11 +1436,11 @@ lemma zdiv_zadd_self1 [simp]: - "intify(a) \ #0 ==> (a$+b) zdiv a = b zdiv a $+ #1" + "intify(a) \ #0 \ (a$+b) zdiv a = b zdiv a $+ #1" by (simp (no_asm_simp) add: zdiv_zadd1_eq) lemma zdiv_zadd_self2 [simp]: - "intify(a) \ #0 ==> (b$+a) zdiv a = b zdiv a $+ #1" + "intify(a) \ #0 \ (b$+a) zdiv a = b zdiv a $+ #1" by (simp (no_asm_simp) add: zdiv_zadd1_eq) lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a" @@ -1458,14 +1458,14 @@ subsection\proving a zdiv (b*c) = (a zdiv b) zdiv c\ -(*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but - 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems +(*The condition c>0 seems necessary. Consider that 7 zdiv \6 = \2 but + 7 zdiv 2 zdiv \3 = 3 zdiv \3 = \1. The subcase (a zdiv b) zmod c = 0 seems to cause particular problems.*) (** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) lemma zdiv_zmult2_aux1: - "[| #0 $< c; b $< r; r $\ #0 |] ==> b$*c $< b$*(q zmod c) $+ r" + "\#0 $< c; b $< r; r $\ #0\ \ b$*c $< b$*(q zmod c) $+ r" apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1") apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) apply (rule zle_zless_trans) @@ -1476,7 +1476,7 @@ done lemma zdiv_zmult2_aux2: - "[| #0 $< c; b $< r; r $\ #0 |] ==> b $* (q zmod c) $+ r $\ #0" + "\#0 $< c; b $< r; r $\ #0\ \ b $* (q zmod c) $+ r $\ #0" apply (subgoal_tac "b $* (q zmod c) $\ #0") prefer 2 apply (simp add: zmult_le_0_iff pos_mod_sign) @@ -1488,7 +1488,7 @@ done lemma zdiv_zmult2_aux3: - "[| #0 $< c; #0 $\ r; r $< b |] ==> #0 $\ b $* (q zmod c) $+ r" + "\#0 $< c; #0 $\ r; r $< b\ \ #0 $\ b $* (q zmod c) $+ r" apply (subgoal_tac "#0 $\ b $* (q zmod c)") prefer 2 apply (simp add: int_0_le_mult_iff pos_mod_sign) @@ -1500,7 +1500,7 @@ done lemma zdiv_zmult2_aux4: - "[| #0 $< c; #0 $\ r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c" + "\#0 $< c; #0 $\ r; r $< b\ \ b $* (q zmod c) $+ r $< b $* c" apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)") apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) apply (rule zless_zle_trans) @@ -1511,8 +1511,8 @@ done lemma zdiv_zmult2_lemma: - "[| quorem (, ); a \ int; b \ int; b \ #0; #0 $< c |] - ==> quorem (, )" + "\quorem (, ); a \ int; b \ int; b \ #0; #0 $< c\ + \ quorem (, )" apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def neq_iff_zless int_0_less_mult_iff zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2 @@ -1521,7 +1521,7 @@ done lemma zdiv_zmult2_eq_raw: - "[|#0 $< c; a \ int; b \ int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c" + "\#0 $< c; a \ int; b \ int\ \ a zdiv (b$*c) = (a zdiv b) zdiv c" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div]) @@ -1529,14 +1529,14 @@ apply (blast dest: zle_zless_trans) done -lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c" +lemma zdiv_zmult2_eq: "#0 $< c \ a zdiv (b$*c) = (a zdiv b) zdiv c" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw) apply auto done lemma zmod_zmult2_eq_raw: - "[|#0 $< c; a \ int; b \ int|] - ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" + "\#0 $< c; a \ int; b \ int\ + \ a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod]) @@ -1545,7 +1545,7 @@ done lemma zmod_zmult2_eq: - "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" + "#0 $< c \ a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw) apply auto done @@ -1553,32 +1553,32 @@ subsection\Cancellation of common factors in "zdiv"\ lemma zdiv_zmult_zmult1_aux1: - "[| #0 $< b; intify(c) \ #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" + "\#0 $< b; intify(c) \ #0\ \ (c$*a) zdiv (c$*b) = a zdiv b" apply (subst zdiv_zmult2_eq) apply auto done lemma zdiv_zmult_zmult1_aux2: - "[| b $< #0; intify(c) \ #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" + "\b $< #0; intify(c) \ #0\ \ (c$*a) zdiv (c$*b) = a zdiv b" apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)") apply (rule_tac [2] zdiv_zmult_zmult1_aux1) apply auto done lemma zdiv_zmult_zmult1_raw: - "[|intify(c) \ #0; b \ int|] ==> (c$*a) zdiv (c$*b) = a zdiv b" + "\intify(c) \ #0; b \ int\ \ (c$*a) zdiv (c$*b) = a zdiv b" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (auto simp add: neq_iff_zless [of b] zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) done -lemma zdiv_zmult_zmult1: "intify(c) \ #0 ==> (c$*a) zdiv (c$*b) = a zdiv b" +lemma zdiv_zmult_zmult1: "intify(c) \ #0 \ (c$*a) zdiv (c$*b) = a zdiv b" apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw) apply auto done -lemma zdiv_zmult_zmult2: "intify(c) \ #0 ==> (a$*c) zdiv (b$*c) = a zdiv b" +lemma zdiv_zmult_zmult2: "intify(c) \ #0 \ (a$*c) zdiv (b$*c) = a zdiv b" apply (drule zdiv_zmult_zmult1) apply (auto simp add: zmult_commute) done @@ -1587,22 +1587,22 @@ subsection\Distribution of factors over "zmod"\ lemma zmod_zmult_zmult1_aux1: - "[| #0 $< b; intify(c) \ #0 |] - ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" + "\#0 $< b; intify(c) \ #0\ + \ (c$*a) zmod (c$*b) = c $* (a zmod b)" apply (subst zmod_zmult2_eq) apply auto done lemma zmod_zmult_zmult1_aux2: - "[| b $< #0; intify(c) \ #0 |] - ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" + "\b $< #0; intify(c) \ #0\ + \ (c$*a) zmod (c$*b) = c $* (a zmod b)" apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))") apply (rule_tac [2] zmod_zmult_zmult1_aux1) apply auto done lemma zmod_zmult_zmult1_raw: - "[|b \ int; c \ int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" + "\b \ int; c \ int\ \ (c$*a) zmod (c$*b) = c $* (a zmod b)" apply (case_tac "b = #0") apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) apply (case_tac "c = #0") @@ -1624,7 +1624,7 @@ (** Quotients of signs **) -lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0" +lemma zdiv_neg_pos_less0: "\a $< #0; #0 $< b\ \ a zdiv b $< #0" apply (subgoal_tac "a zdiv b $\ #-1") apply (erule zle_zless_trans) apply (simp (no_asm)) @@ -1635,12 +1635,12 @@ apply (auto simp add: zdiv_minus1) done -lemma zdiv_nonneg_neg_le0: "[| #0 $\ a; b $< #0 |] ==> a zdiv b $\ #0" +lemma zdiv_nonneg_neg_le0: "\#0 $\ a; b $< #0\ \ a zdiv b $\ #0" apply (drule zdiv_mono1_neg) apply auto done -lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $\ a zdiv b) \ (#0 $\ a)" +lemma pos_imp_zdiv_nonneg_iff: "#0 $< b \ (#0 $\ a zdiv b) \ (#0 $\ a)" apply auto apply (drule_tac [2] zdiv_mono1) apply (auto simp add: neq_iff_zless) @@ -1648,7 +1648,7 @@ apply (blast intro: zdiv_neg_pos_less0) done -lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $\ a zdiv b) \ (a $\ #0)" +lemma neg_imp_zdiv_nonneg_iff: "b $< #0 \ (#0 $\ a zdiv b) \ (a $\ #0)" apply (subst zdiv_zminus_zminus [symmetric]) apply (rule iff_trans) apply (rule pos_imp_zdiv_nonneg_iff) @@ -1656,13 +1656,13 @@ done (*But not (a zdiv b $\ 0 iff a$\0); consider a=1, b=2 when a zdiv b = 0.*) -lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) \ (a $< #0)" +lemma pos_imp_zdiv_neg_iff: "#0 $< b \ (a zdiv b $< #0) \ (a $< #0)" apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) apply (erule pos_imp_zdiv_nonneg_iff) done (*Again the law fails for $\: consider a = -1, b = -2 when a zdiv b = 0*) -lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) \ (#0 $< a)" +lemma neg_imp_zdiv_neg_iff: "b $< #0 \ (a zdiv b $< #0) \ (#0 $< a)" apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) apply (erule neg_imp_zdiv_nonneg_iff) done @@ -1674,7 +1674,7 @@ (** computing "zdiv" by shifting **) - lemma pos_zdiv_mult_2: "#0 $\ a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" + lemma pos_zdiv_mult_2: "#0 $\ a \ (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" apply (case_tac "a = #0") apply (subgoal_tac "#1 $\ a") apply (arith_tac 2) @@ -1694,7 +1694,7 @@ done - lemma neg_zdiv_mult_2: "a $\ #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) \ (b$+#1) zdiv a" + lemma neg_zdiv_mult_2: "a $\ #0 \ (#1 $+ #2$*b) zdiv (#2$*a) \ (b$+#1) zdiv a" apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) \ ($-b-#1) zdiv ($-a)") apply (rule_tac [2] pos_zdiv_mult_2) apply (auto simp add: zmult_zminus_right) @@ -1706,12 +1706,12 @@ (*Not clear why this must be proved separately; probably integ_of causes simplification problems*) - lemma lemma: "~ #0 $\ x ==> x $\ #0" + lemma lemma: "\ #0 $\ x \ x $\ #0" apply auto done lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) = - (if ~b | #0 $\ integ_of w + (if \b | #0 $\ integ_of w then integ_of v zdiv (integ_of w) else (integ_of v $+ #1) zdiv (integ_of w))" apply (simp_tac (simpset_of @{theory_context Int} add: zadd_assoc integ_of_BIT) @@ -1723,7 +1723,7 @@ (** computing "zmod" by shifting (proofs resemble those for "zdiv") **) - lemma pos_zmod_mult_2: "#0 $\ a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" + lemma pos_zmod_mult_2: "#0 $\ a \ (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" apply (case_tac "a = #0") apply (subgoal_tac "#1 $\ a") apply (arith_tac 2) @@ -1743,7 +1743,7 @@ done - lemma neg_zmod_mult_2: "a $\ #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" + lemma neg_zmod_mult_2: "a $\ #0 \ (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))") apply (rule_tac [2] pos_zmod_mult_2) apply (auto simp add: zmult_zminus_right) diff -r f2094906e491 -r e44d86131648 src/ZF/List.thy --- a/src/ZF/List.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/List.thy Tue Sep 27 16:51:35 2022 +0100 @@ -93,19 +93,19 @@ definition (* Function `take' returns the first n elements of a list *) take :: "[i,i]=>i" where - "take(n, as) == list_rec(\n\nat. [], + "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, + "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, + "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 @@ -123,11 +123,11 @@ definition min :: "[i,i] =>i" where - "min(x, y) == (if x \ y then x else y)" + "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)" + "max(x, y) \ (if x \ y then y else x)" (*** Aspects of the datatype definition ***) @@ -143,7 +143,7 @@ lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' & l=l'" by auto -lemma Nil_Cons_iff: "~ Nil=Cons(a,l)" +lemma Nil_Cons_iff: "\ Nil=Cons(a,l)" by auto lemma list_unfold: "list(A) = {0} + (A * list(A))" @@ -153,7 +153,7 @@ (** Lemmas to justify using "list" in other recursive type definitions **) -lemma list_mono: "A<=B ==> list(A) \ list(B)" +lemma list_mono: "A<=B \ list(A) \ list(B)" apply (unfold list.defs ) apply (rule lfp_mono) apply (simp_all add: list.bnd_mono) @@ -171,14 +171,14 @@ (*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)" +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); + "\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)" + \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}" @@ -189,26 +189,26 @@ (*** List functions ***) -lemma tl_type: "l \ list(A) ==> tl(l) \ list(A)" +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" +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)" +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)" +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 @@ -219,60 +219,60 @@ (** Type checking -- proved by induction, as usual **) lemma list_rec_type [TC]: - "[| l \ list(A); + "\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)" + \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)" + "\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})" +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" +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" + "\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)" +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)" +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)" +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)" +lemma set_of_list_type [TC]: "l \ list(A) \ set_of_list(l) \ Pow(A)" apply (unfold 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)" + "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 @@ -280,34 +280,34 @@ (** list_add **) -lemma list_add_type [TC]: "xs: list(nat) ==> list_add(xs) \ nat" +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" +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)" +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)" +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))" +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) ==> + "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") @@ -316,31 +316,31 @@ (** theorems about list(Collect(A,P)) -- used in Induct/Term.thy **) -(* @{term"c \ list(Collect(B,P)) ==> c \ list"} *) +(* @{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)" +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)" +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)" + "\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)" +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))" + "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 @@ -349,12 +349,12 @@ (*Lemma for the inductive step of drop_length*) lemma drop_length_Cons [rule_format]: - "xs: list(A) ==> + "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))" + "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) @@ -365,20 +365,20 @@ (*** theorems about app ***) -lemma app_right_Nil [simp]: "xs: list(A) ==> xs@Nil=xs" +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)" +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)" +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))" +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 @@ -388,17 +388,17 @@ 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)" + "\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" +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)))" +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 @@ -407,18 +407,18 @@ (*** theorems about list_add ***) lemma list_add_app: - "[| xs: list(nat); ys: list(nat) |] - ==> list_add(xs@ys) = list_add(ys) #+ list_add(xs)" + "\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)" +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))" + "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 @@ -426,30 +426,30 @@ (** New induction rules **) lemma list_append_induct [case_names Nil snoc, consumes 1]: - "[| l \ list(A); + "\l \ list(A); P(Nil); - !!x y. [| x \ A; y \ list(A); P(y) |] ==> P(y @ [x]) - |] ==> P(l)" + \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)" + "\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)" + "\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 @@ -462,31 +462,31 @@ (** 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)" +lemma min_sym: "\i \ nat; j \ nat\ \ min(i,j)=min(j,i)" apply (unfold 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" +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" +lemma min_0 [simp]: "i \ nat \ min(0,i) = 0" apply (unfold min_def) apply (auto dest: not_lt_imp_le) done -lemma min_02 [simp]: "i \ nat ==> min(i, 0) = 0" +lemma min_02 [simp]: "i \ nat \ min(i, 0) = 0" apply (unfold min_def) apply (auto dest: not_lt_imp_le) done -lemma lt_min_iff: "[| i \ nat; j \ nat; k \ nat |] ==> i ii \ nat; j \ nat; k \ nat\ \ i i nat; j \ nat |] ==> min(succ(i), succ(j))= succ(min(i, j))" + "\i \ nat; j \ nat\ \ min(succ(i), succ(j))= succ(min(i, j))" apply (unfold min_def, auto) done @@ -495,65 +495,65 @@ (** filter **) lemma filter_append [simp]: - "xs:list(A) ==> filter(P, xs@ys) = filter(P, xs) @ filter(P, ys)" + "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)" +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)" +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)" +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" +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" +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" +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" +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" +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" +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)" +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)" + "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)" + "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)" + "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)" + "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) |] ==> + "\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) @@ -561,14 +561,14 @@ (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff2 [rule_format]: - "[| xs:list(A); ys:list(A); zs:list(A) |] ==> + "\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). + "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)) @@ -578,7 +578,7 @@ declare append_eq_append_iff [simp] lemma append_eq_append [rule_format]: - "xs:list(A) ==> + "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") @@ -589,24 +589,24 @@ 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)" + "\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" + "\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" + "\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)" + "xs:list(A) \ \ys \ list(A). xs@[x] = ys@[y] \ (xs = ys & x=y)" apply (erule list.induct) apply clarify apply (erule list.cases) @@ -618,41 +618,41 @@ declare append1_eq_iff [simp] lemma append_right_is_self_iff [simp]: - "[| xs:list(A); ys:list(A) |] ==> (xs@ys = ys) \ (xs=Nil)" + "\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)" + "\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)" + "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" + "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)" +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)" +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" + "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:list(A) \ (xs=Nil \ P) \ (\ys \ list(A). \y \ A. xs =ys@[y] \P)\P" by (erule list_append_induct, auto) @@ -660,63 +660,63 @@ (** more theorems about drop **) lemma length_drop [rule_format]: - "n \ nat ==> \xs \ list(A). length(drop(n, xs)) = length(xs) #- n" + "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" + "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 ==> + "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)" + "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" +lemma take_0 [simp]: "xs:list(A) \ take(0, xs) = Nil" apply (unfold 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))" + "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" +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" + "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)" + "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) ==> + "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) @@ -725,7 +725,7 @@ declare take_append [simp] lemma take_take [rule_format]: - "m \ nat ==> + "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) @@ -739,14 +739,14 @@ 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)" +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" + "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) @@ -754,13 +754,13 @@ declare nth_type [simp,TC] lemma nth_eq_0 [rule_format]: - "xs:list(A) ==> \n \ nat. length(xs) \ n \ nth(n,xs) = 0" + "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) ==> + "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) @@ -769,7 +769,7 @@ lemma set_of_list_conv_nth: "xs:list(A) - ==> set_of_list(xs) = {x \ A. \i\nat. i set_of_list(xs) = {x \ A. \i\nat. i nat ==> + "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") @@ -797,9 +797,9 @@ 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" + "\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) @@ -808,8 +808,8 @@ (*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" + "\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) @@ -821,14 +821,14 @@ done lemma nth_drop [rule_format]: - "n \ nat ==> \i \ nat. \xs \ list(A). nth(i, drop(n, xs)) = nth(n #+ i, xs)" + "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)]" + \ \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) @@ -836,14 +836,14 @@ 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))" + "\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))" + "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 @@ -865,29 +865,29 @@ definition zip :: "[i, i]=>i" where - "zip(xs, ys) == zip_aux(set_of_list(ys),xs)`ys" + "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))" +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" +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" +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" + "\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) @@ -896,7 +896,7 @@ done lemma zip_Cons_Cons [simp]: - "[| xs:list(A); ys:list(B); x \ A; y \ B |] ==> + "\xs:list(A); ys:list(B); x \ A; y \ B\ \ zip(Cons(x,xs), Cons(y, ys)) = Cons(, zip(xs, ys))" apply (simp add: zip_def, auto) apply (rule zip_aux_unique, auto) @@ -905,7 +905,7 @@ done lemma zip_type [rule_format]: - "xs:list(A) ==> \ys \ list(B). zip(xs, ys):list(A*B)" + "xs:list(A) \ \ys \ list(B). zip(xs, ys):list(A*B)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify @@ -915,7 +915,7 @@ (* zip length *) lemma length_zip [rule_format]: - "xs:list(A) ==> \ys \ list(B). length(zip(xs,ys)) = + "xs:list(A) \ \ys \ list(B). length(zip(xs,ys)) = min(length(xs), length(ys))" apply (unfold min_def) apply (induct_tac "xs", simp_all, clarify) @@ -924,7 +924,7 @@ declare length_zip [simp] lemma zip_append1 [rule_format]: - "[| ys:list(A); zs:list(B) |] ==> + "\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) @@ -932,21 +932,21 @@ done lemma zip_append2 [rule_format]: - "[| xs:list(A); zs:list(B) |] ==> \ys \ list(B). zip(xs, ys@zs) = + "\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)" + "\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). + "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) @@ -955,7 +955,7 @@ declare zip_rev [simp] lemma nth_zip [rule_format]: - "ys:list(B) ==> \i \ nat. \xs \ list(A). + "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) @@ -965,28 +965,28 @@ 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)) = + "\xs:list(A); ys:list(B); i \ nat\ + \ set_of_list(zip(xs, ys)) = {: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" +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 ==> + "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)" + "\xs:list(A); v \ A\ \ \n \ nat. list_update(xs, n, v):list(A)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify @@ -995,7 +995,7 @@ 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)" + "xs:list(A) \ \i \ nat. length(list_update(xs, i, v))=length(xs)" apply (induct_tac "xs") apply (simp (no_asm)) apply clarify @@ -1004,7 +1004,7 @@ declare length_list_update [simp] lemma nth_list_update [rule_format]: - "[| xs:list(A) |] ==> \i \ nat. \j \ nat. i < length(xs) \ + "\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 @@ -1016,12 +1016,12 @@ done lemma nth_list_update_eq [simp]: - "[| i < length(xs); xs:list(A) |] ==> nth(i, list_update(xs, i,x)) = x" + "\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) ==> + "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)) @@ -1033,7 +1033,7 @@ declare nth_list_update_neq [simp] lemma list_update_overwrite [rule_format]: - "xs:list(A) ==> \i \ nat. i < length(xs) + "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)) @@ -1043,7 +1043,7 @@ declare list_update_overwrite [simp] lemma list_update_same_conv [rule_format]: - "xs:list(A) ==> + "xs:list(A) \ \i \ nat. i < length(xs) \ (list_update(xs, i, x) = xs) \ (nth(i, xs) = x)" apply (induct_tac "xs") @@ -1053,7 +1053,7 @@ done lemma update_zip [rule_format]: - "ys:list(B) ==> + "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)), @@ -1065,7 +1065,7 @@ done lemma set_update_subset_cons [rule_format]: - "xs:list(A) ==> + "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 @@ -1074,8 +1074,8 @@ 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" + "\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 @@ -1083,13 +1083,13 @@ (** upt **) lemma upt_rec: - "j \ nat ==> upt(i,j) = (if i nat \ upt(i,j) = (if i i; j \ nat |] ==> upt(i,j) = Nil" +lemma upt_conv_Nil [simp]: "\j \ 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) @@ -1097,34 +1097,34 @@ (*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]" + "\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))" + "\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)" +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)" + "\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" +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" + "\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 @@ -1132,7 +1132,7 @@ done lemma take_upt [rule_format]: - "[| m \ nat; n \ nat |] ==> + "\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) @@ -1147,13 +1147,13 @@ declare take_upt [simp] lemma map_succ_upt: - "[| m \ nat; n \ nat |] ==> map(succ, upt(m,n))= upt(succ(m), succ(n))" + "\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) ==> + "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) @@ -1162,7 +1162,7 @@ declare nth_map [simp] lemma nth_map_upt [rule_format]: - "[| m \ nat; n \ nat |] ==> + "\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) @@ -1180,17 +1180,17 @@ definition sublist :: "[i, i] => i" where - "sublist(xs, A)== + "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" +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 |] ==> + "\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) @@ -1199,18 +1199,18 @@ done lemma sublist_type [simp,TC]: - "xs:list(B) ==> sublist(xs, A):list(B)" + "xs:list(B) \ sublist(xs, A):list(B)" apply (unfold 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)" + "\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) |] ==> + "\xs:list(B); ys:list(B)\ \ sublist(xs@ys, A) = sublist(xs, A) @ sublist(ys, {j \ nat. j #+ length(xs): A})" apply (unfold sublist_def) apply (erule_tac l = ys in list_append_induct, simp) @@ -1221,7 +1221,7 @@ lemma sublist_Cons: - "[| xs:list(B); x \ B |] ==> + "\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) @@ -1234,7 +1234,7 @@ by (simp add: sublist_Cons) lemma sublist_upt_eq_take [rule_format]: - "xs:list(A) ==> \n\nat. sublist(xs,n) = take(n,xs)" + "xs:list(A) \ \n\nat. sublist(xs,n) = take(n,xs)" apply (erule list.induct, simp) apply (clarify ) apply (erule natE) @@ -1243,7 +1243,7 @@ declare sublist_upt_eq_take [simp] lemma sublist_Int_eq: - "xs \ list(B) ==> sublist(xs, A \ nat) = sublist(xs, A)" + "xs \ list(B) \ sublist(xs, A \ nat) = sublist(xs, A)" apply (erule list.induct) apply (simp_all add: sublist_Cons) done @@ -1256,15 +1256,15 @@ "repeat(a,succ(n)) = Cons(a,repeat(a,n))" -lemma length_repeat: "n \ nat ==> length(repeat(a,n)) = 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]" +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)" +lemma repeat_type [TC]: "\a \ A; n \ nat\ \ repeat(a,n) \ list(A)" by (induct_tac n, auto) end diff -r f2094906e491 -r e44d86131648 src/ZF/Nat.thy --- a/src/ZF/Nat.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Nat.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,43 +9,43 @@ definition nat :: i where - "nat == lfp(Inf, %X. {0} \ {succ(i). i \ X})" + "nat \ lfp(Inf, %X. {0} \ {succ(i). i \ X})" definition quasinat :: "i => o" where - "quasinat(n) == n=0 | (\m. n = succ(m))" + "quasinat(n) \ n=0 | (\m. n = succ(m))" definition (*Has an unconditional succ case, which is used in "recursor" below.*) nat_case :: "[i, i=>i, i]=>i" where - "nat_case(a,b,k) == THE y. k=0 & y=a | (\x. k=succ(x) & y=b(x))" + "nat_case(a,b,k) \ THE y. k=0 & y=a | (\x. k=succ(x) & y=b(x))" definition nat_rec :: "[i, i, [i,i]=>i]=>i" where - "nat_rec(k,a,b) == + "nat_rec(k,a,b) \ wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))" (*Internalized relations on the naturals*) definition Le :: i where - "Le == {:nat*nat. x \ y}" + "Le \ {:nat*nat. x \ y}" definition Lt :: i where - "Lt == {:nat*nat. x < y}" + "Lt \ {:nat*nat. x < y}" definition Ge :: i where - "Ge == {:nat*nat. y \ x}" + "Ge \ {:nat*nat. y \ x}" definition Gt :: i where - "Gt == {:nat*nat. y < x}" + "Gt \ {:nat*nat. y < x}" definition greater_than :: "i=>i" where - "greater_than(n) == {i \ nat. n < i}" + "greater_than(n) \ {i \ nat. n < i}" text\No need for a less-than operator: a natural number is its list of predecessors!\ @@ -66,7 +66,7 @@ apply (rule singletonI [THEN UnI1]) done -lemma nat_succI [intro!,TC]: "n \ nat ==> succ(n) \ nat" +lemma nat_succI [intro!,TC]: "n \ nat \ succ(n) \ nat" apply (subst nat_unfold) apply (erule RepFunI [THEN UnI2]) done @@ -87,7 +87,7 @@ (*Mathematical induction*) lemma nat_induct [case_names 0 succ, induct set: nat]: - "[| n \ nat; P(0); !!x. [| x \ nat; P(x) |] ==> P(succ(x)) |] ==> P(n)" + "\n \ nat; P(0); \x. \x \ nat; P(x)\ \ P(succ(x))\ \ P(n)" by (erule def_induct [OF nat_def nat_bnd_mono], blast) lemma natE: @@ -96,13 +96,13 @@ using assms by (rule nat_unfold [THEN equalityD1, THEN subsetD, THEN UnE]) auto -lemma nat_into_Ord [simp]: "n \ nat ==> Ord(n)" +lemma nat_into_Ord [simp]: "n \ nat \ Ord(n)" by (erule nat_induct, auto) -(* @{term"i \ nat ==> 0 \ i"}; same thing as @{term"0 nat \ 0 \ i"}; same thing as @{term"0 nat ==> i \ i"}; same thing as @{term"i nat \ i \ i"}; same thing as @{term"i nat ==> ~ Limit(a)" +lemma naturals_not_limit: "a \ nat \ \ Limit(a)" by (induct a rule: nat_induct, auto) -lemma succ_natD: "succ(i): nat ==> i \ nat" +lemma succ_natD: "succ(i): nat \ i \ nat" by (rule Ord_trans [OF succI1], auto) lemma nat_succ_iff [iff]: "succ(n): nat \ n \ nat" by (blast dest!: succ_natD) -lemma nat_le_Limit: "Limit(i) ==> nat \ i" +lemma nat_le_Limit: "Limit(i) \ nat \ i" apply (rule subset_imp_le) apply (simp_all add: Limit_is_Ord) apply (rule subsetI) @@ -137,15 +137,15 @@ apply (blast intro: Limit_has_succ [THEN ltD] ltI Limit_is_Ord) done -(* [| succ(i): k; k \ nat |] ==> i \ k *) +(* \succ(i): k; k \ nat\ \ i \ k *) lemmas succ_in_naturalD = Ord_trans [OF succI1 _ nat_into_Ord] -lemma lt_nat_in_nat: "[| m nat |] ==> m \ nat" +lemma lt_nat_in_nat: "\m nat\ \ m \ nat" apply (erule ltE) apply (erule Ord_trans, assumption, simp) done -lemma le_in_nat: "[| m \ n; n \ nat |] ==> m \ nat" +lemma le_in_nat: "\m \ n; n \ nat\ \ m \ nat" by (blast dest!: lt_nat_in_nat) @@ -163,7 +163,7 @@ lemma nat_induct_from: assumes "m \ n" "m \ nat" "n \ nat" and "P(m)" - and "!!x. [| x \ nat; m \ x; P(x) |] ==> P(succ(x))" + and "\x. \x \ nat; m \ x; P(x)\ \ P(succ(x))" shows "P(n)" proof - from assms(3) have "m \ n \ P(m) \ P(n)" @@ -173,11 +173,11 @@ (*Induction suitable for subtraction and less-than*) lemma diff_induct [case_names 0 0_succ succ_succ, consumes 2]: - "[| m \ nat; n \ nat; - !!x. x \ nat ==> P(x,0); - !!y. y \ nat ==> P(0,succ(y)); - !!x y. [| x \ nat; y \ nat; P(x,y) |] ==> P(succ(x),succ(y)) |] - ==> P(m,n)" + "\m \ nat; n \ nat; + \x. x \ nat \ P(x,0); + \y. y \ nat \ P(0,succ(y)); + \x y. \x \ nat; y \ nat; P(x,y)\ \ P(succ(x),succ(y))\ + \ P(m,n)" apply (erule_tac x = m in rev_bspec) apply (erule nat_induct, simp) apply (rule ballI) @@ -189,7 +189,7 @@ (** Induction principle analogous to trancl_induct **) lemma succ_lt_induct_lemma [rule_format]: - "m \ nat ==> P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ + "m \ nat \ P(m,succ(m)) \ (\x\nat. P(m,x) \ P(m,succ(x))) \ (\n\nat. m P(m,n))" apply (erule nat_induct) apply (intro impI, rule nat_induct [THEN ballI]) @@ -198,10 +198,10 @@ done lemma succ_lt_induct: - "[| m nat; + "\m nat; P(m,succ(m)); - !!x. [| x \ nat; P(m,x) |] ==> P(m,succ(x)) |] - ==> P(m,n)" + \x. \x \ nat; P(m,x)\ \ P(m,succ(x))\ + \ P(m,n)" by (blast intro: succ_lt_induct_lemma lt_nat_in_nat) subsection\quasinat: to allow a case-split rule for \<^term>\nat_case\\ @@ -213,20 +213,20 @@ lemma [iff]: "quasinat(succ(x))" by (simp add: quasinat_def) -lemma nat_imp_quasinat: "n \ nat ==> quasinat(n)" +lemma nat_imp_quasinat: "n \ nat \ quasinat(n)" by (erule natE, simp_all) -lemma non_nat_case: "~ quasinat(x) ==> nat_case(a,b,x) = 0" +lemma non_nat_case: "\ quasinat(x) \ nat_case(a,b,x) = 0" by (simp add: quasinat_def nat_case_def) -lemma nat_cases_disj: "k=0 | (\y. k = succ(y)) | ~ quasinat(k)" +lemma nat_cases_disj: "k=0 | (\y. k = succ(y)) | \ quasinat(k)" apply (case_tac "k=0", simp) apply (case_tac "\m. k = succ(m)") apply (simp_all add: quasinat_def) done lemma nat_cases: - "[|k=0 ==> P; !!y. k = succ(y) ==> P; ~ quasinat(k) ==> P|] ==> P" + "\k=0 \ P; \y. k = succ(y) \ P; \ quasinat(k) \ P\ \ P" by (insert nat_cases_disj [of k], blast) (** nat_case **) @@ -238,13 +238,13 @@ by (simp add: nat_case_def) lemma nat_case_type [TC]: - "[| n \ nat; a \ C(0); !!m. m \ nat ==> b(m): C(succ(m)) |] - ==> nat_case(a,b,n) \ C(n)" + "\n \ nat; a \ C(0); \m. m \ nat \ b(m): C(succ(m))\ + \ nat_case(a,b,n) \ C(n)" by (erule nat_induct, auto) lemma split_nat_case: "P(nat_case(a,b,k)) \ - ((k=0 \ P(a)) & (\x. k=succ(x) \ P(b(x))) & (~ quasinat(k) \ P(0)))" + ((k=0 \ P(a)) & (\x. k=succ(x) \ P(b(x))) & (\ quasinat(k) \ P(0)))" apply (rule nat_cases [of k]) apply (auto simp add: non_nat_case) done @@ -261,7 +261,7 @@ apply (rule nat_case_0) done -lemma nat_rec_succ: "m \ nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" +lemma nat_rec_succ: "m \ nat \ nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))" apply (rule nat_rec_def [THEN def_wfrec, THEN trans]) apply (rule wf_Memrel) apply (simp add: vimage_singleton_iff) @@ -269,12 +269,12 @@ (** The union of two natural numbers is a natural number -- their maximum **) -lemma Un_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" +lemma Un_nat_type [TC]: "\i \ nat; j \ nat\ \ i \ j \ nat" apply (rule Un_least_lt [THEN ltD]) apply (simp_all add: lt_def) done -lemma Int_nat_type [TC]: "[| i \ nat; j \ nat |] ==> i \ j \ nat" +lemma Int_nat_type [TC]: "\i \ nat; j \ nat\ \ i \ j \ nat" apply (rule Int_greatest_lt [THEN ltD]) apply (simp_all add: lt_def) done @@ -284,7 +284,7 @@ by blast text\A natural number is the set of its predecessors\ -lemma nat_eq_Collect_lt: "i \ nat ==> {j\nat. j nat \ {j\nat. j o] => o" where - "oall(A, P) == \x. x P(x)" + "oall(A, P) \ \x. x P(x)" definition oex :: "[i, i => o] => o" where - "oex(A, P) == \x. x \x. x i] => i" where - "OUnion(i,B) == {z: \x\i. B(x). Ord(i)}" + "OUnion(i,B) \ {z: \x\i. B(x). Ord(i)}" syntax "_oall" :: "[idt, i, o] => o" (\(3\_<_./ _)\ 10) @@ -37,11 +37,11 @@ (*MOST IMPORTANT that this is added to the simpset BEFORE Ord_atomize is proved. Ord_atomize would convert this rule to - x < 0 ==> P(x) == True, which causes dire effects!*) + x < 0 \ P(x) \ True, which causes dire effects!*) lemma [simp]: "(\x<0. P(x))" by (simp add: oall_def) -lemma [simp]: "~(\x<0. P(x))" +lemma [simp]: "\(\x<0. P(x))" by (simp add: oex_def) lemma [simp]: "(\x (Ord(i) \ P(i) & (\xUnion over ordinals\ lemma Ord_OUN [intro,simp]: - "[| !!x. x Ord(B(x)) |] ==> Ord(\x\x. x Ord(B(x))\ \ Ord(\xx i < (\xax \ i < (\xb(a); Ord(\x i \ (\xab(a); Ord(\x \ i \ (\x (\x (\xi\nat.i)=nat"}! *) lemma OUN_least: - "(!!x. x B(x) \ C) ==> (\x C" + "(\x. x B(x) \ C) \ (\x C" by (simp add: OUnion_def UN_least ltI) lemma OUN_least_le: - "[| Ord(i); !!x. x b(x) \ i |] ==> (\x i" + "\Ord(i); \x. x b(x) \ i\ \ (\x i" by (simp add: OUnion_def UN_least_le ltI Ord_0_le) lemma le_implies_OUN_le_OUN: - "[| !!x. x c(x) \ d(x) |] ==> (\x (\x\x. x c(x) \ d(x)\ \ (\x (\x A ==> Ord(B(x))) - ==> (\z < (\x\A. B(x)). C(z)) = (\x\A. \z < B(x). C(z))" + "(\x. x \ A \ Ord(B(x))) + \ (\z < (\x\A. B(x)). C(z)) = (\x\A. \z < B(x). C(z))" by (simp add: OUnion_def) lemma OUN_Union_eq: - "(!!x. x \ X ==> Ord(x)) - ==> (\z < \(X). C(z)) = (\x\X. \z < x. C(z))" + "(\x. x \ X \ Ord(x)) + \ (\z < \(X). C(z)) = (\x\X. \z < x. C(z))" by (simp add: OUnion_def) (*So that rule_format will get rid of this quantifier...*) lemma atomize_oall [symmetric, rulify]: - "(!!x. x P(x)) == Trueprop (\xx. x P(x)) \ Trueprop (\xuniversal quantifier for ordinals\ lemma oallI [intro!]: - "[| !!x. x P(x) |] ==> \x\x. x P(x)\ \ \xx P(x)" +lemma ospec: "\\x \ P(x)" by (simp add: oall_def) lemma oallE: - "[| \x Q; ~x Q |] ==> Q" + "\\x Q; \x Q\ \ Q" by (simp add: oall_def, blast) lemma rev_oallE [elim]: - "[| \x Q; P(x) ==> Q |] ==> Q" + "\\xx Q; P(x) \ Q\ \ Q" by (simp add: oall_def, blast) @@ -126,43 +126,43 @@ (*Congruence rule for rewriting*) lemma oall_cong [cong]: - "[| a=a'; !!x. x P(x) <-> P'(x) |] - ==> oall(a, %x. P(x)) <-> oall(a', %x. P'(x))" + "\a=a'; \x. x P(x) <-> P'(x)\ + \ oall(a, %x. P(x)) <-> oall(a', %x. P'(x))" by (simp add: oall_def) subsubsection \existential quantifier for ordinals\ lemma oexI [intro]: - "[| P(x); x \xP(x); x \ \xx P(a); a \x\xP(x) \ P(a); a \ \xx Q |] ==> Q" + "\\xx. \x \ Q\ \ Q" apply (simp add: oex_def, blast) done lemma oex_cong [cong]: - "[| a=a'; !!x. x P(x) <-> P'(x) |] - ==> oex(a, %x. P(x)) <-> oex(a', %x. P'(x))" + "\a=a'; \x. x P(x) <-> P'(x)\ + \ oex(a, %x. P(x)) <-> oex(a', %x. P'(x))" apply (simp add: oex_def cong add: conj_cong) done subsubsection \Rules for Ordinal-Indexed Unions\ -lemma OUN_I [intro]: "[| a B(a) |] ==> b: (\za B(a)\ \ b: (\z (\z B(a); a R |] ==> R" + "\b \ (\za.\b \ B(a); a \ R\ \ R" apply (unfold OUnion_def lt_def, blast) done @@ -170,11 +170,11 @@ by (unfold OUnion_def oex_def lt_def, blast) lemma OUN_cong [cong]: - "[| i=j; !!x. x C(x)=D(x) |] ==> (\xxi=j; \x. x C(x)=D(x)\ \ (\xxy P(x) |] ==> P(i)" + "\ix.\xy \ P(x)\ \ P(i)" apply (simp add: lt_def oall_def) apply (erule conjE) apply (erule Ord_induct, assumption, blast) @@ -185,11 +185,11 @@ definition "rall" :: "[i=>o, i=>o] => o" where - "rall(M, P) == \x. M(x) \ P(x)" + "rall(M, P) \ \x. M(x) \ P(x)" definition "rex" :: "[i=>o, i=>o] => o" where - "rex(M, P) == \x. M(x) & P(x)" + "rex(M, P) \ \x. M(x) & P(x)" syntax "_rall" :: "[pttrn, i=>o, o] => o" (\(3\_[_]./ _)\ 10) @@ -201,18 +201,18 @@ subsubsection\Relativized universal quantifier\ -lemma rallI [intro!]: "[| !!x. M(x) ==> P(x) |] ==> \x[M]. P(x)" +lemma rallI [intro!]: "\\x. M(x) \ P(x)\ \ \x[M]. P(x)" by (simp add: rall_def) -lemma rspec: "[| \x[M]. P(x); M(x) |] ==> P(x)" +lemma rspec: "\\x[M]. P(x); M(x)\ \ P(x)" by (simp add: rall_def) (*Instantiates x first: better for automatic theorem proving?*) lemma rev_rallE [elim]: - "[| \x[M]. P(x); ~ M(x) ==> Q; P(x) ==> Q |] ==> Q" + "\\x[M]. P(x); \ M(x) \ Q; P(x) \ Q\ \ Q" by (simp add: rall_def, blast) -lemma rallE: "[| \x[M]. P(x); P(x) ==> Q; ~ M(x) ==> Q |] ==> Q" +lemma rallE: "\\x[M]. P(x); P(x) \ Q; \ M(x) \ Q\ \ Q" by blast (*Trival rewrite rule; (\x[M].P)<->P holds only if A is nonempty!*) @@ -221,32 +221,32 @@ (*Congruence rule for rewriting*) lemma rall_cong [cong]: - "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (\x[M]. P(x)) <-> (\x[M]. P'(x))" + "(\x. M(x) \ P(x) <-> P'(x)) \ (\x[M]. P(x)) <-> (\x[M]. P'(x))" by (simp add: rall_def) subsubsection\Relativized existential quantifier\ -lemma rexI [intro]: "[| P(x); M(x) |] ==> \x[M]. P(x)" +lemma rexI [intro]: "\P(x); M(x)\ \ \x[M]. P(x)" by (simp add: rex_def, blast) (*The best argument order when there is only one M(x)*) -lemma rev_rexI: "[| M(x); P(x) |] ==> \x[M]. P(x)" +lemma rev_rexI: "\M(x); P(x)\ \ \x[M]. P(x)" by blast (*Not of the general form for such rules... *) -lemma rexCI: "[| \x[M]. ~P(x) ==> P(a); M(a) |] ==> \x[M]. P(x)" +lemma rexCI: "\\x[M]. \P(x) \ P(a); M(a)\ \ \x[M]. P(x)" by blast -lemma rexE [elim!]: "[| \x[M]. P(x); !!x. [| M(x); P(x) |] ==> Q |] ==> Q" +lemma rexE [elim!]: "\\x[M]. P(x); \x. \M(x); P(x)\ \ Q\ \ Q" by (simp add: rex_def, blast) -(*We do not even have (\x[M]. True) <-> True unless A is nonempty!!*) +(*We do not even have (\x[M]. True) <-> True unless A is nonempty\*) lemma rex_triv [simp]: "(\x[M]. P) \ ((\x. M(x)) \ P)" by (simp add: rex_def) lemma rex_cong [cong]: - "(!!x. M(x) ==> P(x) <-> P'(x)) ==> (\x[M]. P(x)) <-> (\x[M]. P'(x))" + "(\x. M(x) \ P(x) <-> P'(x)) \ (\x[M]. P(x)) <-> (\x[M]. P'(x))" by (simp add: rex_def cong: conj_cong) lemma rall_is_ball [simp]: "(\x[%z. z\A]. P(x)) <-> (\x\A. P(x))" @@ -255,7 +255,7 @@ lemma rex_is_bex [simp]: "(\x[%z. z\A]. P(x)) <-> (\x\A. P(x))" by blast -lemma atomize_rall: "(!!x. M(x) ==> P(x)) == Trueprop (\x[M]. P(x))" +lemma atomize_rall: "(\x. M(x) \ P(x)) \ Trueprop (\x[M]. P(x))" by (simp add: rall_def atomize_all atomize_imp) declare atomize_rall [symmetric, rulify] @@ -264,7 +264,7 @@ "(\x[M]. P(x) & Q) <-> (\x[M]. P(x)) & ((\x[M]. False) | Q)" "(\x[M]. P(x) | Q) <-> ((\x[M]. P(x)) | Q)" "(\x[M]. P(x) \ Q) <-> ((\x[M]. P(x)) \ Q)" - "(~(\x[M]. P(x))) <-> (\x[M]. ~P(x))" + "(\(\x[M]. P(x))) <-> (\x[M]. \P(x))" by blast+ lemma rall_simps2: @@ -283,7 +283,7 @@ "(\x[M]. P(x) & Q) <-> ((\x[M]. P(x)) & Q)" "(\x[M]. P(x) | Q) <-> (\x[M]. P(x)) | ((\x[M]. True) & Q)" "(\x[M]. P(x) \ Q) <-> ((\x[M]. P(x)) \ ((\x[M]. True) & Q))" - "(~(\x[M]. P(x))) <-> (\x[M]. ~P(x))" + "(\(\x[M]. P(x))) <-> (\x[M]. \P(x))" by blast+ lemma rex_simps2: @@ -324,7 +324,7 @@ definition setclass :: "[i,i] => o" (\##_\ [40] 40) where - "setclass(A) == %x. x \ A" + "setclass(A) \ %x. x \ A" lemma setclass_iff [simp]: "setclass(A,x) <-> x \ A" by (simp add: setclass_def) diff -r f2094906e491 -r e44d86131648 src/ZF/Order.thy --- a/src/ZF/Order.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Order.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,15 +17,15 @@ definition part_ord :: "[i,i]=>o" (*Strict partial ordering*) where - "part_ord(A,r) == irrefl(A,r) & trans[A](r)" + "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. :r | x=y | :r)" + "linear(A,r) \ (\x\A. \y\A. :r | x=y | :r)" definition tot_ord :: "[i,i]=>o" (*Strict total ordering*) where - "tot_ord(A,r) == part_ord(A,r) & linear(A,r)" + "tot_ord(A,r) \ part_ord(A,r) & linear(A,r)" definition "preorder_on(A, r) \ refl(A, r) \ trans[A](r)" @@ -41,62 +41,62 @@ definition well_ord :: "[i,i]=>o" (*Well-ordering*) where - "well_ord(A,r) == tot_ord(A,r) & wf[A](r)" + "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) == + "mono_map(A,r,B,s) \ {f \ A->B. \x\A. \y\A. :r \ :s}" definition ord_iso :: "[i,i,i,i]=>i" (\(\_, _\ \/ \_, _\)\ 51) (*Order isomorphisms*) where - "\A,r\ \ \B,s\ == + "\A,r\ \ \B,s\ \ {f \ bij(A,B). \x\A. \y\A. :r \ :s}" definition pred :: "[i,i,i]=>i" (*Set of predecessors*) where - "pred(A,x,r) == {y \ A. :r}" + "pred(A,x,r) \ {y \ A. :r}" definition ord_iso_map :: "[i,i,i,i]=>i" (*Construction for linearity theorem*) where - "ord_iso_map(A,r,B,s) == + "ord_iso_map(A,r,B,s) \ \x\A. \y\B. \f \ ord_iso(pred(A,x,r), r, pred(B,y,s), s). {}" definition first :: "[i, i, i] => o" where - "first(u, X, R) == u \ X & (\v\X. v\u \ \ R)" + "first(u, X, R) \ u \ X & (\v\X. v\u \ \ R)" subsection\Immediate Consequences of the Definitions\ lemma part_ord_Imp_asym: - "part_ord(A,r) ==> asym(r \ A*A)" + "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; - :r ==> P; x=y ==> P; :r ==> P |] - ==> P" + "\linear(A,r); x \ A; y \ A; + :r \ P; x=y \ P; :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)" + "\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)" + "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)" + "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)" +lemma well_ord_is_linear: "well_ord(A,r) \ linear(A,r)" by (unfold well_ord_def tot_ord_def, blast) @@ -107,7 +107,7 @@ lemmas predI = conjI [THEN pred_iff [THEN iffD2]] -lemma predE: "[| y \ pred(A,x,r); [| y \ A; :r |] ==> P |] ==> P" +lemma predE: "\y \ pred(A,x,r); \y \ A; :r\ \ P\ \ P" by (simp add: pred_def) lemma pred_subset_under: "pred(A,x,r) \ r -`` {x}" @@ -121,8 +121,8 @@ by (simp add: pred_def, blast) lemma trans_pred_pred_eq: - "[| trans[A](r); :r; x \ A; y \ A |] - ==> pred(pred(A,x,r), y, r) = pred(A,y,r)" + "\trans[A](r); :r; x \ A; y \ A\ + \ pred(pred(A,x,r), y, r) = pred(A,y,r)" by (unfold trans_on_def pred_def, blast) @@ -133,21 +133,21 @@ (*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)" + "\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)" + "\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)" + "\tot_ord(A,r); B<=A\ \ tot_ord(B,r)" apply (unfold 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)" + "\well_ord(A,r); B<=A\ \ well_ord(B,r)" apply (unfold well_ord_def) apply (fast elim!: tot_ord_subset wf_on_subset_A) done @@ -239,40 +239,40 @@ (** Order-preserving (monotone) maps **) -lemma mono_map_is_fun: "f \ mono_map(A,r,B,s) ==> f \ A->B" +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)" + "\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 |] ==> \ r \ \ s |] - ==> f \ ord_iso(A,r,B,s)" + "\f \ bij(A, B); + \x y. \x \ A; y \ A\ \ \ 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)" + "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)" + "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); : r; x \ A; y \ A |] ==> \ s" + "\f \ ord_iso(A,r,B,s); : r; x \ A; y \ A\ \ \ s" by (simp add: ord_iso_def) lemma ord_iso_converse: - "[| f \ ord_iso(A,r,B,s); : s; x \ B; y \ B |] - ==> \ r" + "\f \ ord_iso(A,r,B,s); : 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])+ @@ -287,7 +287,7 @@ 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)" +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]) @@ -295,16 +295,16 @@ (*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)" + "\g \ mono_map(A,r,B,s); f \ mono_map(B,s,C,t)\ + \ (f O g): mono_map(A,r,C,t)" apply (unfold 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)" + "\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]) @@ -314,8 +314,8 @@ (** 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)" + "\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") @@ -324,9 +324,9 @@ 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)" + "\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) @@ -338,13 +338,13 @@ (** 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)" + "\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)" + "\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]) @@ -353,7 +353,7 @@ done lemma wf_on_ord_iso: - "[| wf[B](s); f \ ord_iso(A,r,B,s) |] ==> wf[A](r)" + "\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) @@ -361,7 +361,7 @@ done lemma well_ord_ord_iso: - "[| well_ord(B,s); f \ ord_iso(A,r,B,s) |] ==> well_ord(A,r)" + "\well_ord(B,s); f \ ord_iso(A,r,B,s)\ \ well_ord(A,r)" apply (unfold well_ord_def tot_ord_def) apply (fast elim!: part_ord_ord_iso linear_ord_iso wf_on_ord_iso) done @@ -372,8 +372,8 @@ (*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" + "\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+) @@ -383,7 +383,7 @@ (*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" + "\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 *) @@ -394,8 +394,8 @@ (*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" + "\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+) @@ -408,7 +408,7 @@ (*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)" + "\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) apply (erule CollectE) apply (simp (no_asm_simp) add: image_fun [OF bij_is_fun Collect_subset]) @@ -420,8 +420,8 @@ 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)" + "\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 @@ -429,18 +429,18 @@ (*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)" + "\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); : r; + "\well_ord(A,r); well_ord(B,s); : 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 |] ==> : s" + a \ A; c \ A; b \ B; d \ B\ \ : 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)) @@ -452,9 +452,9 @@ (*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" + "\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 @@ -470,8 +470,8 @@ (*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" +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)") @@ -499,19 +499,19 @@ done lemma function_ord_iso_map: - "well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))" + "well_ord(B,s) \ function(ord_iso_map(A,r,B,s))" apply (unfold 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) +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) + "\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)" apply (unfold mono_map_def) @@ -524,7 +524,7 @@ done lemma ord_iso_map_ord_iso: - "[| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) + "\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) @@ -539,9 +539,9 @@ (*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)" + "\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)" apply (unfold ord_iso_map_def) apply (safe intro!: predI) (*Case analysis on xa vs a in r *) @@ -563,8 +563,8 @@ (*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 | + "\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) @@ -582,8 +582,8 @@ (*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 | + "\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) @@ -591,8 +591,8 @@ 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) | + "\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) @@ -614,21 +614,21 @@ (** Properties of converse(r) **) -lemma irrefl_converse: "irrefl(A,r) ==> irrefl(A,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))" +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))" +lemma part_ord_converse: "part_ord(A,r) \ part_ord(A,converse(r))" apply (unfold part_ord_def) apply (blast intro!: irrefl_converse trans_on_converse) done -lemma linear_converse: "linear(A,r) ==> linear(A,converse(r))" +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))" +lemma tot_ord_converse: "tot_ord(A,r) \ tot_ord(A,converse(r))" apply (unfold tot_ord_def) apply (blast intro!: part_ord_converse linear_converse) done @@ -637,11 +637,11 @@ (** By Krzysztof Grabczewski. Lemmas involving the first element of a well ordered set **) -lemma first_is_elem: "first(b,B,r) ==> b \ B" +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))" + "\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) apply (elim conjE allE disjE, blast) apply (erule bexE) @@ -650,7 +650,7 @@ done lemma the_first_in: - "[| well_ord(A,r); B<=A; B\0 |] ==> (THE b. first(b,B,r)) \ B" + "\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) @@ -660,7 +660,7 @@ subsection \Lemmas for the Reflexive Orders\ lemma subset_vimage_vimage_iff: - "[| Preorder(r); A \ field(r); B \ field(r) |] ==> + "\Preorder(r); A \ field(r); B \ field(r)\ \ r -`` A \ r -`` B \ (\a\A. \b\B. \ r)" apply (auto simp: subset_def preorder_on_def refl_def vimage_def image_def) apply blast @@ -673,12 +673,12 @@ done lemma subset_vimage1_vimage1_iff: - "[| Preorder(r); a \ field(r); b \ field(r) |] ==> + "\Preorder(r); a \ field(r); b \ field(r)\ \ r -`` {a} \ r -`` {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) |] ==> + "\refl(field(r), r); antisym(r); a \ field(r); b \ field(r)\ \ r `` {a} = r `` {b} \ a = b" apply rule apply (frule equality_iffD) @@ -689,13 +689,13 @@ done lemma Partial_order_eq_Image1_Image1_iff: - "[| Partial_order(r); a \ field(r); b \ field(r) |] ==> + "\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) |] ==> + "\refl(field(r), r); antisym(r); a \ field(r); b \ field(r)\ \ r -`` {a} = r -`` {b} \ a = b" apply rule apply (frule equality_iffD) @@ -706,7 +706,7 @@ done lemma Partial_order_eq_vimage1_vimage1_iff: - "[| Partial_order(r); a \ field(r); b \ field(r) |] ==> + "\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) diff -r f2094906e491 -r e44d86131648 src/ZF/OrderArith.thy --- a/src/ZF/OrderArith.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/OrderArith.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,7 +10,7 @@ definition (*disjoint sum of two relations; underlies ordinal addition*) radd :: "[i,i,i,i]=>i" where - "radd(A,r,B,s) == + "radd(A,r,B,s) \ {z: (A+B) * (A+B). (\x y. z = ) | (\x' x. z = & :r) | @@ -19,7 +19,7 @@ definition (*lexicographic product of two relations; underlies ordinal multiplication*) rmult :: "[i,i,i,i]=>i" where - "rmult(A,r,B,s) == + "rmult(A,r,B,s) \ {z: (A*B) * (A*B). \x' y' x y. z = <, > & (: r | (x'=x & : s))}" @@ -27,11 +27,11 @@ definition (*inverse image of a relation*) rvimage :: "[i,i,i]=>i" where - "rvimage(A,f,r) == {z \ A*A. \x y. z = & : r}" + "rvimage(A,f,r) \ {z \ A*A. \x y. z = & : r}" definition measure :: "[i, i\i] \ i" where - "measure(A,f) == {: A*A. f(x) < f(y)}" + "measure(A,f) \ {: A*A. f(x) < f(y)}" subsection\Addition of Relations -- Disjoint Sum\ @@ -59,11 +59,11 @@ subsubsection\Elimination Rule\ lemma raddE: - "[| \ 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); : r; x':A; x \ A |] ==> Q; - !!y' y. [| p'=Inr(y'); p=Inr(y); : s; y':B; y \ B |] ==> Q - |] ==> Q" + "\ \ 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); : r; x':A; x \ A\ \ Q; + \y' y. \p'=Inr(y'); p=Inr(y); : s; y':B; y \ B\ \ Q +\ \ Q" by (unfold radd_def, blast) subsubsection\Type checking\ @@ -78,13 +78,13 @@ subsubsection\Linearity\ lemma linear_radd: - "[| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))" + "\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))" +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!\ @@ -99,14 +99,14 @@ 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))" +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))" + "\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) @@ -115,8 +115,8 @@ 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)" + "\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) @@ -124,7 +124,7 @@ done lemma sum_ord_iso_cong: - "[| f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s') |] ==> + "\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'))" apply (unfold ord_iso_def) @@ -135,7 +135,7 @@ (*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 ==> +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 @@ -170,10 +170,10 @@ by (unfold rmult_def, blast) lemma rmultE: - "[| <, > \ rmult(A,r,B,s); - [| : r; a':A; a \ A; b':B; b \ B |] ==> Q; - [| : s; a \ A; a'=a; b':B; b \ B |] ==> Q - |] ==> Q" + "\<, > \ rmult(A,r,B,s); + \: r; a':A; a \ A; b':B; b \ B\ \ Q; + \: s; a \ A; a'=a; b':B; b \ B\ \ Q +\ \ Q" by blast subsubsection\Type checking\ @@ -186,12 +186,12 @@ subsubsection\Linearity\ lemma linear_rmult: - "[| linear(A,r); linear(B,s) |] ==> linear(A*B,rmult(A,r,B,s))" + "\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))" +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) @@ -203,14 +203,14 @@ done -lemma wf_rmult: "[| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))" +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))" + "\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) @@ -220,8 +220,8 @@ subsubsection\An \<^term>\ord_iso\ congruence law\ lemma prod_bij: - "[| f \ bij(A,C); g \ bij(B,D) |] - ==> (lam :A*B. ) \ bij(A*B, C*D)" + "\f \ bij(A,C); g \ bij(B,D)\ + \ (lam :A*B. ) \ bij(A*B, C*D)" apply (rule_tac d = "%. " in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) @@ -229,8 +229,8 @@ done lemma prod_ord_iso_cong: - "[| f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s') |] - ==> (lam :A*B. ) + "\f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s')\ + \ (lam :A*B. ) \ ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))" apply (unfold ord_iso_def) apply (safe intro!: prod_bij) @@ -243,7 +243,7 @@ (*Used??*) lemma singleton_prod_ord_iso: - "well_ord({x},xr) ==> + "well_ord({x},xr) \ (\z\A. ) \ ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))" apply (rule singleton_prod_bij [THEN ord_isoI]) apply (simp (no_asm_simp)) @@ -253,7 +253,7 @@ (*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 ==> + "a\C \ (\x\C*B + D. case(%x. x, %y., x)) \ bij(C*B + D, C*B \ {a}*D)" apply (rule subst_elem) @@ -267,7 +267,7 @@ done lemma prod_sum_singleton_ord_iso: - "[| a \ A; well_ord(A,r) |] ==> + "\a \ A; well_ord(A,r)\ \ (\x\pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y., x)) \ ord_iso(pred(A,a,r)*B + pred(B,b,s), radd(A*B, rmult(A,r,B,s), B, s), @@ -324,19 +324,19 @@ subsubsection\Partial Ordering Properties\ lemma irrefl_rvimage: - "[| f \ inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))" + "\f \ inj(A,B); irrefl(B,r)\ \ irrefl(A, rvimage(A,f,r))" apply (unfold 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))" + "\f \ inj(A,B); trans[B](r)\ \ trans[A](rvimage(A,f,r))" apply (unfold 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))" + "\f \ inj(A,B); part_ord(B,r)\ \ part_ord(A, rvimage(A,f,r))" apply (unfold part_ord_def) apply (blast intro!: irrefl_rvimage trans_on_rvimage) done @@ -344,13 +344,13 @@ subsubsection\Linearity\ lemma linear_rvimage: - "[| f \ inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))" + "\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))" + "\f \ inj(A,B); tot_ord(B,r)\ \ tot_ord(A, rvimage(A,f,r))" apply (unfold tot_ord_def) apply (blast intro!: part_ord_rvimage linear_rvimage) done @@ -358,7 +358,7 @@ subsubsection\Well-foundedness\ -lemma wf_rvimage [intro!]: "wf(r) ==> wf(rvimage(A,f,r))" +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) }") @@ -370,8 +370,8 @@ 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))" + \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 @@ -382,7 +382,7 @@ (*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))" + "\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) apply (blast intro!: wf_on_rvimage inj_is_fun) @@ -390,13 +390,13 @@ done lemma ord_iso_rvimage: - "f \ bij(A,B) ==> f \ ord_iso(A, rvimage(A,f,s), B, s)" + "f \ bij(A,B) \ f \ ord_iso(A, rvimage(A,f,s), B, s)" apply (unfold 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" + "f \ ord_iso(A,r, B,s) \ rvimage(A,f,s) = r \ A*A" by (unfold ord_iso_def rvimage_def, blast) @@ -409,28 +409,28 @@ definition wfrank :: "[i,i]=>i" where - "wfrank(r,a) == wfrec(r, a, %x f. \y \ r-``{x}. succ(f`y))" + "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))" + "wftype(r) \ \y \ range(r). succ(wfrank(r,y))" -lemma wfrank: "wf(r) ==> wfrank(r,a) = (\y \ r-``{a}. 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))" +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); \ r|] ==> wfrank(r,a) < wfrank(r,b)" +lemma wfrank_lt: "\wf(r); \ 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))" +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)" @@ -440,7 +440,7 @@ lemma wf_imp_subset_rvimage: - "[|wf(r); r \ A*A|] ==> \i f. Ord(i) & r \ rvimage(A, f, Memrel(i))" + "\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) @@ -450,19 +450,19 @@ done theorem wf_iff_subset_rvimage: - "relation(r) ==> wf(r) \ (\i f A. Ord(i) & r \ rvimage(A, f, Memrel(i)))" + "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)" +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)" + "\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 @@ -500,8 +500,8 @@ by (simp (no_asm) add: measure_def) lemma linear_measure: - assumes Ordf: "!!x. x \ A ==> Ord(f(x))" - and inj: "!!x y. [|x \ A; y \ A; f(x) = f(y) |] ==> x=y" + assumes Ordf: "\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) @@ -513,8 +513,8 @@ 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" + 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) @@ -528,9 +528,9 @@ lemma wf_on_Union: assumes wfA: "wf[A](r)" - and wfB: "!!a. a\A ==> wf[B(a)](s)" - and ok: "!!a u v. [| \ s; v \ B(a); a \ A|] - ==> (\a'\A. \ r & u \ B(a')) | u \ B(a)" + and wfB: "\a. a\A \ wf[B(a)](s)" + and ok: "\a u v. \ \ s; v \ B(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) diff -r f2094906e491 -r e44d86131648 src/ZF/OrderType.thy --- a/src/ZF/OrderType.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/OrderType.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,45 +13,45 @@ definition ordermap :: "[i,i]=>i" where - "ordermap(A,r) == \x\A. wfrec[A](r, x, %x f. f `` pred(A,x,r))" + "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" + "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)))" + "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" + "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)))" + "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)))" + "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))" + "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))" + "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))" +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 @@ -60,23 +60,23 @@ apply (blast intro: Ord_in_Ord lt_Ord)+ done -(*"Ord(i) ==> well_ord(i, Memrel(i))"*) +(*"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" + "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" + "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" + "\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) @@ -88,8 +88,8 @@ (*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" + "\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 @@ -108,8 +108,8 @@ (*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)" + "\wf[A](r); x \ A\ + \ ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)" apply (unfold ordermap_def pred_def) apply (simp (no_asm_simp)) apply (erule wfrec_on [THEN trans], assumption) @@ -118,8 +118,8 @@ (*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)}" + "\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!*) @@ -127,8 +127,8 @@ (*The theorem above is -[| wf[A](r); x \ A |] -==> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \ A . \ r}} +\wf[A](r); x \ A\ +\ ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \ A . \ 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 @@ -144,7 +144,7 @@ subsubsection\Showing that ordermap, ordertype yield ordinals\ lemma Ord_ordermap: - "[| well_ord(A,r); x \ A |] ==> Ord(ordermap(A,r) ` x)" + "\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) @@ -155,7 +155,7 @@ done lemma Ord_ordertype: - "well_ord(A,r) ==> Ord(ordertype(A,r))" + "well_ord(A,r) \ Ord(ordertype(A,r))" apply (unfold ordertype_def) apply (subst image_fun [OF ordermap_type subset_refl]) apply (rule OrdI [OF _ Ord_is_Transset]) @@ -169,15 +169,15 @@ subsubsection\ordermap preserves the orderings in both directions\ lemma ordermap_mono: - "[| : r; wf[A](r); w \ A; x \ A |] - ==> ordermap(A,r)`w \ ordermap(A,r)`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 |] - ==> : r" + "\ordermap(A,r)`w \ ordermap(A,r)`x; well_ord(A,r); w \ A; x \ A\ + \ : 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]) @@ -189,7 +189,7 @@ by (rule surj_image) (rule ordermap_type) lemma ordermap_bij: - "well_ord(A,r) ==> ordermap(A,r) \ bij(A, ordertype(A,r))" + "well_ord(A,r) \ ordermap(A,r) \ bij(A, ordertype(A,r))" apply (unfold well_ord_def tot_ord_def bij_def inj_def) apply (force intro!: ordermap_type ordermap_surj elim: linearE dest: ordermap_mono @@ -200,7 +200,7 @@ lemma ordertype_ord_iso: "well_ord(A,r) - ==> ordermap(A,r) \ ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))" + \ ordermap(A,r) \ ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))" apply (unfold ord_iso_def) apply (safe elim!: well_ord_is_wf intro!: ordermap_type [THEN apply_type] ordermap_mono ordermap_bij) @@ -208,16 +208,16 @@ done lemma ordertype_eq: - "[| f \ ord_iso(A,r,B,s); well_ord(B,s) |] - ==> ordertype(A,r) = ordertype(B,s)" + "\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)" + "\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) @@ -227,7 +227,7 @@ subsubsection\Basic equalities for ordertype\ (*Ordertype of Memrel*) -lemma le_ordertype_Memrel: "j \ i ==> ordertype(j,Memrel(i)) = j" +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) @@ -238,7 +238,7 @@ apply (fast elim: ltE Ord_in_Ord Ord_trans) done -(*"Ord(i) ==> ordertype(i, Memrel(i)) = i"*) +(*"Ord(i) \ ordertype(i, Memrel(i)) = i"*) lemmas ordertype_Memrel = le_refl [THEN le_ordertype_Memrel] lemma ordertype_0 [simp]: "ordertype(0,r) = 0" @@ -248,7 +248,7 @@ apply (rule Ord_0 [THEN ordertype_Memrel]) done -(*Ordertype of rvimage: [| f \ bij(A,B); well_ord(B,s) |] ==> +(*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] @@ -256,8 +256,8 @@ (*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" + "\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) @@ -278,15 +278,15 @@ text\Theorems by Krzysztof Grabczewski; proofs simplified by lcp\ -lemma ordertype_pred_subset: "[| well_ord(A,r); x \ A |] ==> +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)" + "\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]) @@ -298,7 +298,7 @@ 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}" + \ 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] @@ -310,7 +310,7 @@ subsection\Alternative definition of ordinal\ (*proof by Krzysztof Grabczewski*) -lemma Ord_is_Ord_alt: "Ord(i) ==> Ord_alt(i)" +lemma Ord_is_Ord_alt: "Ord(i) \ Ord_alt(i)" apply (unfold Ord_alt_def) apply (rule conjI) apply (erule well_ord_Memrel) @@ -319,7 +319,7 @@ (*proof by lcp*) lemma Ord_alt_is_Ord: - "Ord_alt(i) ==> Ord(i)" + "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) @@ -339,7 +339,7 @@ done lemma ordertype_sum_0_eq: - "well_ord(A,r) ==> ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)" + "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 @@ -351,7 +351,7 @@ done lemma ordertype_0_sum_eq: - "well_ord(A,r) ==> ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)" + "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 @@ -361,7 +361,7 @@ (*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)) + "a \ A \ (\x\pred(A,a,r). Inl(x)) \ bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))" apply (unfold pred_def) apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective) @@ -369,8 +369,8 @@ 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)) = + "\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]) @@ -378,7 +378,7 @@ done lemma pred_Inr_bij: - "b \ B ==> + "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) @@ -386,8 +386,8 @@ 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)) = + "\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) @@ -407,7 +407,7 @@ subsubsection\Basic laws for ordinal addition\ -lemma Ord_raw_oadd: "[|Ord(i); Ord(j)|] ==> Ord(raw_oadd(i,j))" +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) @@ -417,19 +417,19 @@ text\Ordinal addition with zero\ -lemma raw_oadd_0: "Ord(i) ==> raw_oadd(i,0) = i" +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" +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" +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" +lemma oadd_0_left [simp]: "Ord(i) \ 0++i = i" by (simp add: oadd_def raw_oadd_0_left ordify_def) @@ -438,14 +438,14 @@ 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" +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" +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) @@ -455,31 +455,31 @@ 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" +(*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))" +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)" + "\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)" + "\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) @@ -488,15 +488,15 @@ 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)) + "\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))) = + "\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]) @@ -504,7 +504,7 @@ apply (simp_all add: well_ord_radd well_ord_Memrel) done -lemma oadd_lt_mono2: "k i++k < i++j" +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) @@ -516,7 +516,7 @@ leI [THEN le_ordertype_Memrel] ordertype_sum_Memrel) done -lemma oadd_lt_cancel2: "[| i++j < i++k; Ord(j) |] ==> ji++j < i++k; Ord(j)\ \ j i++j < i++k \ j i++j < i++k \ j j=k" +lemma oadd_inject: "\i++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 )" +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 @@ -564,7 +564,7 @@ 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})" +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) @@ -572,12 +572,12 @@ apply (force intro: lt_oadd1 oadd_lt_mono2 simp add: Ord_mem_iff_lt) done -lemma oadd_1: "Ord(i) ==> i++1 = succ(i)" +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)" +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] @@ -588,28 +588,28 @@ 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))" + "\\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)" +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" +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); Ord(j)\ \ 0 < (i ++ j) \ 0 Limit(i ++ j)" +lemma oadd_LimitI: "\Ord(i); Limit(j)\ \ Limit(i ++ j)" apply (simp add: oadd_Limit) apply (frule Limit_has_1 [THEN ltD]) apply (rule increasing_LimitI) @@ -624,7 +624,7 @@ text\Order/monotonicity properties of ordinal addition\ -lemma oadd_le_self2: "Ord(i) ==> i \ j++i" +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 @@ -639,7 +639,7 @@ thus ?case using limit.hyps by (simp add: oadd_Limit) qed -lemma oadd_le_mono1: "k \ j ==> k++i \ j++i" +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) @@ -651,16 +651,16 @@ apply (rule le_implies_UN_le_UN, blast) done -lemma oadd_lt_mono: "[| i' \ i; j' i'++j' < i++j" +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" +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" +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" +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]) @@ -668,12 +668,12 @@ done text\Every ordinal is exceeded by some limit ordinal.\ -lemma Ord_imp_greater_Limit: "Ord(i) ==> \k. i \k. i \k. iOrd(i); Ord(j)\ \ \k. i lemma bij_sum_Diff: - "A<=B ==> (\y\B. if(y \ A, Inl(y), Inr(y))) \ bij(B, A+(B-A))" + "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) @@ -694,7 +694,7 @@ done lemma ordertype_sum_Diff: - "i \ j ==> + "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]) @@ -708,7 +708,7 @@ done lemma Ord_odiff [simp,TC]: - "[| Ord(i); Ord(j) |] ==> Ord(i--j)" + "\Ord(i); Ord(j)\ \ Ord(i--j)" apply (unfold odiff_def) apply (blast intro: Ord_ordertype Diff_subset well_ord_subset well_ord_Memrel) done @@ -716,7 +716,7 @@ lemma raw_oadd_ordertype_Diff: "i \ j - ==> raw_oadd(i,j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(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]) @@ -725,19 +725,19 @@ 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" +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" + 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" +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) @@ -749,7 +749,7 @@ subsection\Ordinal Multiplication\ lemma Ord_omult [simp,TC]: - "[| Ord(i); Ord(j) |] ==> Ord(i**j)" + "\Ord(i); Ord(j)\ \ Ord(i**j)" apply (unfold omult_def) apply (blast intro: Ord_ordertype well_ord_rmult well_ord_Memrel) done @@ -757,13 +757,13 @@ subsubsection\A useful unfolding law\ lemma pred_Pair_eq: - "[| a \ A; b \ B |] ==> pred(A*B, , rmult(A,r,B,s)) = + "\a \ A; b \ B\ \ pred(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) |] ==> + "\a \ A; b \ B; well_ord(A,r); well_ord(B,s)\ \ ordertype(pred(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))" @@ -776,8 +776,8 @@ done lemma ordertype_pred_Pair_lemma: - "[| i' ordertype(pred(i*j, , rmult(i,Memrel(i),j,Memrel(j))), + "\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) @@ -795,8 +795,8 @@ done lemma lt_omult: - "[| Ord(i); Ord(j); k \j' i'. k = j**i' ++ j' & j'Ord(i); Ord(j); k + \ \j' i'. k = j**i' ++ j' & j' j**i' ++ j' < j**i" + "\j' \ j**i' ++ j' < j**i" apply (unfold omult_def) apply (rule ltI) prefer 2 @@ -820,7 +820,7 @@ done lemma omult_unfold: - "[| Ord(i); Ord(j) |] ==> j**i = (\j'\j. \i'\i. {j**i' ++ j'})" + "\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) @@ -845,7 +845,7 @@ text\Ordinal multiplication by 1\ -lemma omult_1 [simp]: "Ord(i) ==> i**1 = i" +lemma omult_1 [simp]: "Ord(i) \ i**1 = i" apply (unfold omult_def) apply (rule_tac s1="Memrel(i)" in ord_isoI [THEN ordertype_eq, THEN trans]) @@ -853,7 +853,7 @@ apply (auto elim!: snd_type well_ord_Memrel ordertype_Memrel) done -lemma omult_1_left [simp]: "Ord(i) ==> 1**i = i" +lemma omult_1_left [simp]: "Ord(i) \ 1**i = i" apply (unfold omult_def) apply (rule_tac s1="Memrel(i)" in ord_isoI [THEN ordertype_eq, THEN trans]) @@ -864,7 +864,7 @@ 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)" + "\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]) @@ -879,13 +879,13 @@ Ord_ordertype) done -lemma omult_succ: "[| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i" +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)" + "\Ord(i); Ord(j); Ord(k)\ \ (i**j)**k = i**(j**k)" apply (unfold omult_def) apply (rule ordertype_eq [THEN trans]) apply (rule prod_ord_iso_cong [OF ord_iso_refl @@ -902,24 +902,24 @@ 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))" + "\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)" +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" +(*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" +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: @@ -940,24 +940,24 @@ qed qed -lemma omult_lt_mono2: "[| k i**k < i**j" +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" +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" +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" +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: @@ -988,7 +988,7 @@ text\Further properties of ordinal multiplication\ -lemma omult_inject: "[| i**j = i**k; 0 j=k" +lemma omult_inject: "\i**j = i**k; 0 \ j=k" apply (rule Ord_linear_lt) prefer 4 apply assumption apply auto diff -r f2094906e491 -r e44d86131648 src/ZF/Ordinal.thy --- a/src/ZF/Ordinal.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Ordinal.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,27 +9,27 @@ definition Memrel :: "i=>i" where - "Memrel(A) == {z\A*A . \x y. z= & x\y }" + "Memrel(A) \ {z\A*A . \x y. z= & x\y }" definition Transset :: "i=>o" where - "Transset(i) == \x\i. x<=i" + "Transset(i) \ \x\i. x<=i" definition Ord :: "i=>o" where - "Ord(i) == Transset(i) & (\x\i. Transset(x))" + "Ord(i) \ Transset(i) & (\x\i. Transset(x))" definition lt :: "[i,i] => o" (infixl \<\ 50) (*less-than on ordinals*) where - "ij & Ord(j)" + "i i\j & Ord(j)" definition Limit :: "i=>o" where - "Limit(i) == Ord(i) & 0y. y succ(y) Ord(i) & 0y. y succ(y)\\ 50) where - "x \ y == x < succ(y)" + "x \ y \ x < succ(y)" subsection\Rules for Transset\ @@ -50,21 +50,21 @@ subsubsection\Consequences of Downwards Closure\ lemma Transset_doubleton_D: - "[| Transset(C); {a,b}: C |] ==> a\C & b\C" + "\Transset(C); {a,b}: C\ \ a\C & b\C" by (unfold Transset_def, blast) lemma Transset_Pair_D: - "[| Transset(C); \C |] ==> a\C & b\C" + "\Transset(C); \C\ \ a\C & b\C" apply (simp add: Pair_def) apply (blast dest: Transset_doubleton_D) done lemma Transset_includes_domain: - "[| Transset(C); A*B \ C; b \ B |] ==> A \ C" + "\Transset(C); A*B \ C; b \ B\ \ A \ C" by (blast dest: Transset_Pair_D) lemma Transset_includes_range: - "[| Transset(C); A*B \ C; a \ A |] ==> B \ C" + "\Transset(C); A*B \ C; a \ A\ \ B \ C" by (blast dest: Transset_Pair_D) subsubsection\Closure Properties\ @@ -73,73 +73,73 @@ by (unfold Transset_def, blast) lemma Transset_Un: - "[| Transset(i); Transset(j) |] ==> Transset(i \ j)" + "\Transset(i); Transset(j)\ \ Transset(i \ j)" by (unfold Transset_def, blast) lemma Transset_Int: - "[| Transset(i); Transset(j) |] ==> Transset(i \ j)" + "\Transset(i); Transset(j)\ \ Transset(i \ j)" by (unfold Transset_def, blast) -lemma Transset_succ: "Transset(i) ==> Transset(succ(i))" +lemma Transset_succ: "Transset(i) \ Transset(succ(i))" by (unfold Transset_def, blast) -lemma Transset_Pow: "Transset(i) ==> Transset(Pow(i))" +lemma Transset_Pow: "Transset(i) \ Transset(Pow(i))" by (unfold Transset_def, blast) -lemma Transset_Union: "Transset(A) ==> Transset(\(A))" +lemma Transset_Union: "Transset(A) \ Transset(\(A))" by (unfold Transset_def, blast) lemma Transset_Union_family: - "[| !!i. i\A ==> Transset(i) |] ==> Transset(\(A))" + "\\i. i\A \ Transset(i)\ \ Transset(\(A))" by (unfold Transset_def, blast) lemma Transset_Inter_family: - "[| !!i. i\A ==> Transset(i) |] ==> Transset(\(A))" + "\\i. i\A \ Transset(i)\ \ Transset(\(A))" by (unfold Inter_def Transset_def, blast) lemma Transset_UN: - "(!!x. x \ A ==> Transset(B(x))) ==> Transset (\x\A. B(x))" + "(\x. x \ A \ Transset(B(x))) \ Transset (\x\A. B(x))" by (rule Transset_Union_family, auto) lemma Transset_INT: - "(!!x. x \ A ==> Transset(B(x))) ==> Transset (\x\A. B(x))" + "(\x. x \ A \ Transset(B(x))) \ Transset (\x\A. B(x))" by (rule Transset_Inter_family, auto) subsection\Lemmas for Ordinals\ lemma OrdI: - "[| Transset(i); !!x. x\i ==> Transset(x) |] ==> Ord(i)" + "\Transset(i); \x. x\i \ Transset(x)\ \ Ord(i)" by (simp add: Ord_def) -lemma Ord_is_Transset: "Ord(i) ==> Transset(i)" +lemma Ord_is_Transset: "Ord(i) \ Transset(i)" by (simp add: Ord_def) lemma Ord_contains_Transset: - "[| Ord(i); j\i |] ==> Transset(j) " + "\Ord(i); j\i\ \ Transset(j) " by (unfold Ord_def, blast) -lemma Ord_in_Ord: "[| Ord(i); j\i |] ==> Ord(j)" +lemma Ord_in_Ord: "\Ord(i); j\i\ \ Ord(j)" by (unfold Ord_def Transset_def, blast) (*suitable for rewriting PROVIDED i has been fixed*) -lemma Ord_in_Ord': "[| j\i; Ord(i) |] ==> Ord(j)" +lemma Ord_in_Ord': "\j\i; Ord(i)\ \ Ord(j)" by (blast intro: Ord_in_Ord) -(* Ord(succ(j)) ==> Ord(j) *) +(* Ord(succ(j)) \ Ord(j) *) lemmas Ord_succD = Ord_in_Ord [OF _ succI1] -lemma Ord_subset_Ord: "[| Ord(i); Transset(j); j<=i |] ==> Ord(j)" +lemma Ord_subset_Ord: "\Ord(i); Transset(j); j<=i\ \ Ord(j)" by (simp add: Ord_def Transset_def, blast) -lemma OrdmemD: "[| j\i; Ord(i) |] ==> j<=i" +lemma OrdmemD: "\j\i; Ord(i)\ \ j<=i" by (unfold Ord_def Transset_def, blast) -lemma Ord_trans: "[| i\j; j\k; Ord(k) |] ==> i\k" +lemma Ord_trans: "\i\j; j\k; Ord(k)\ \ i\k" by (blast dest: OrdmemD) -lemma Ord_succ_subsetI: "[| i\j; Ord(j) |] ==> succ(i) \ j" +lemma Ord_succ_subsetI: "\i\j; Ord(j)\ \ succ(i) \ j" by (blast dest: OrdmemD) @@ -148,7 +148,7 @@ lemma Ord_0 [iff,TC]: "Ord(0)" by (blast intro: OrdI Transset_0) -lemma Ord_succ [TC]: "Ord(i) ==> Ord(succ(i))" +lemma Ord_succ [TC]: "Ord(i) \ Ord(succ(i))" by (blast intro: OrdI Transset_succ Ord_is_Transset Ord_contains_Transset) lemmas Ord_1 = Ord_0 [THEN Ord_succ] @@ -156,18 +156,18 @@ lemma Ord_succ_iff [iff]: "Ord(succ(i)) <-> Ord(i)" by (blast intro: Ord_succ dest!: Ord_succD) -lemma Ord_Un [intro,simp,TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \ j)" +lemma Ord_Un [intro,simp,TC]: "\Ord(i); Ord(j)\ \ Ord(i \ j)" apply (unfold Ord_def) apply (blast intro!: Transset_Un) done -lemma Ord_Int [TC]: "[| Ord(i); Ord(j) |] ==> Ord(i \ j)" +lemma Ord_Int [TC]: "\Ord(i); Ord(j)\ \ Ord(i \ j)" apply (unfold Ord_def) apply (blast intro!: Transset_Int) done text\There is no set of all ordinals, for then it would contain itself\ -lemma ON_class: "~ (\i. i\X <-> Ord(i))" +lemma ON_class: "\ (\i. i\X <-> Ord(i))" proof (rule notI) assume X: "\i. i \ X \ Ord(i)" have "\x y. x\X \ y\x \ y\X" @@ -183,63 +183,63 @@ subsection\< is 'less Than' for Ordinals\ -lemma ltI: "[| i\j; Ord(j) |] ==> ii\j; Ord(j)\ \ ij; Ord(i); Ord(j) |] ==> P |] ==> P" + "\ii\j; Ord(i); Ord(j)\ \ P\ \ P" apply (unfold lt_def) apply (blast intro: Ord_in_Ord) done -lemma ltD: "i i\j" +lemma ltD: "i i\j" by (erule ltE, assumption) -lemma not_lt0 [simp]: "~ i<0" +lemma not_lt0 [simp]: "\ i<0" by (unfold lt_def, blast) -lemma lt_Ord: "j Ord(j)" +lemma lt_Ord: "j Ord(j)" by (erule ltE, assumption) -lemma lt_Ord2: "j Ord(i)" +lemma lt_Ord2: "j Ord(i)" by (erule ltE, assumption) -(* @{term"ja \ j ==> Ord(j)"} *) +(* @{term"ja \ j \ Ord(j)"} *) lemmas le_Ord2 = lt_Ord2 [THEN Ord_succD] -(* i<0 ==> R *) +(* i<0 \ R *) lemmas lt0E = not_lt0 [THEN notE, elim!] -lemma lt_trans [trans]: "[| i ii \ i ~ (j \ (j j P *) +(* \iP \ j \ P *) lemmas lt_asym = lt_not_sym [THEN swap] -lemma lt_irrefl [elim!]: "i P" +lemma lt_irrefl [elim!]: "i P" by (blast intro: lt_asym) -lemma lt_not_refl: "~ i iRecall that \<^term>\i \ j\ abbreviates \<^term>\i !!\ +text\Recall that \<^term>\i \ j\ abbreviates \<^term>\i \\ lemma le_iff: "i \ j <-> i i < succ(j)*) -lemma leI: "i i \ j" +(*Equivalently, i i < succ(j)*) +lemma leI: "i i \ j" by (simp add: le_iff) -lemma le_eqI: "[| i=j; Ord(j) |] ==> i \ j" +lemma le_eqI: "\i=j; Ord(j)\ \ i \ j" by (simp add: le_iff) lemmas le_refl = refl [THEN le_eqI] @@ -247,14 +247,14 @@ lemma le_refl_iff [iff]: "i \ i <-> Ord(i)" by (simp (no_asm_simp) add: lt_not_refl le_iff) -lemma leCI: "(~ (i=j & Ord(j)) ==> i i \ j" +lemma leCI: "(\ (i=j & Ord(j)) \ i i \ j" by (simp add: le_iff, blast) lemma leE: - "[| i \ j; i P; [| i=j; Ord(j) |] ==> P |] ==> P" + "\i \ j; i P; \i=j; Ord(j)\ \ P\ \ P" by (simp add: le_iff, blast) -lemma le_anti_sym: "[| i \ j; j \ i |] ==> i=j" +lemma le_anti_sym: "\i \ j; j \ i\ \ i=j" apply (simp add: le_iff) apply (blast elim: lt_asym) done @@ -270,19 +270,19 @@ lemma Memrel_iff [simp]: " \ Memrel(A) <-> a\b & a\A & b\A" by (unfold Memrel_def, blast) -lemma MemrelI [intro!]: "[| a \ b; a \ A; b \ A |] ==> \ Memrel(A)" +lemma MemrelI [intro!]: "\a \ b; a \ A; b \ A\ \ \ Memrel(A)" by auto lemma MemrelE [elim!]: - "[| \ Memrel(A); - [| a \ A; b \ A; a\b |] ==> P |] - ==> P" + "\ \ Memrel(A); + \a \ A; b \ A; a\b\ \ P\ + \ P" by auto lemma Memrel_type: "Memrel(A) \ A*A" by (unfold Memrel_def, blast) -lemma Memrel_mono: "A<=B ==> Memrel(A) \ Memrel(B)" +lemma Memrel_mono: "A<=B \ Memrel(A) \ Memrel(B)" by (unfold Memrel_def, blast) lemma Memrel_0 [simp]: "Memrel(0) = 0" @@ -303,17 +303,17 @@ text\The premise \<^term>\Ord(i)\ does not suffice.\ lemma trans_Memrel: - "Ord(i) ==> trans(Memrel(i))" + "Ord(i) \ trans(Memrel(i))" by (unfold Ord_def Transset_def trans_def, blast) text\However, the following premise is strong enough.\ lemma Transset_trans_Memrel: - "\j\i. Transset(j) ==> trans(Memrel(i))" + "\j\i. Transset(j) \ trans(Memrel(i))" by (unfold Transset_def trans_def, blast) (*If Transset(A) then Memrel(A) internalizes the membership relation below A*) lemma Transset_Memrel_iff: - "Transset(A) ==> \ Memrel(A) <-> a\b & b\A" + "Transset(A) \ \ Memrel(A) <-> a\b & b\A" by (unfold Transset_def, blast) @@ -321,9 +321,9 @@ (*Epsilon induction over a transitive set*) lemma Transset_induct: - "[| i \ k; Transset(k); - !!x.[| x \ k; \y\x. P(y) |] ==> P(x) |] - ==> P(i)" + "\i \ k; Transset(k); + \x.\x \ k; \y\x. P(y)\ \ P(x)\ + \ P(i)" apply (simp add: Transset_def) apply (erule wf_Memrel [THEN wf_induct2], blast+) done @@ -382,34 +382,34 @@ apply (blast intro: leI le_eqI o) + done -lemma le_imp_not_lt: "j \ i ==> ~ i i \ \ i j \ i" +lemma not_lt_imp_le: "\\ i \ j \ i" by (rule_tac i = i and j = j in Ord_linear2, auto) subsubsection \Some Rewrite Rules for \<\, \\\\ -lemma Ord_mem_iff_lt: "Ord(j) ==> i\j <-> i i\j <-> i ~ i j \ i" +lemma not_lt_iff_le: "\Ord(i); Ord(j)\ \ \ i j \ i" by (blast dest: le_imp_not_lt not_lt_imp_le) -lemma not_le_iff_lt: "[| Ord(i); Ord(j) |] ==> ~ i \ j <-> jOrd(i); Ord(j)\ \ \ i \ j <-> j 0 \ i" +lemma Ord_0_le: "Ord(i) \ 0 \ i" by (erule not_lt_iff_le [THEN iffD1], auto) -lemma Ord_0_lt: "[| Ord(i); i\0 |] ==> 0Ord(i); i\0\ \ 0 i\0 <-> 0 i\0 <-> 0 succ(x) <-> Ord(x)" by (blast intro: Ord_0_le elim: ltE) -lemma subset_imp_le: "[| j<=i; Ord(i); Ord(j) |] ==> j \ i" +lemma subset_imp_le: "\j<=i; Ord(i); Ord(j)\ \ j \ i" apply (rule not_lt_iff_le [THEN iffD1], assumption+) apply (blast elim: ltE mem_irrefl) done -lemma le_imp_subset: "i \ j ==> i<=j" +lemma le_imp_subset: "i \ j \ i<=j" by (blast dest: OrdmemD elim: ltE leE) lemma le_subset_iff: "j \ i <-> j<=i & Ord(i) & Ord(j)" @@ -437,27 +437,27 @@ done (*Just a variant of subset_imp_le*) -lemma all_lt_imp_le: "[| Ord(i); Ord(j); !!x. x x j \ i" +lemma all_lt_imp_le: "\Ord(i); Ord(j); \x. x x \ j \ i" by (blast intro: not_lt_imp_le dest: lt_irrefl) subsubsection\Transitivity Laws\ -lemma lt_trans1: "[| i \ j; j ii \ j; j \ i k |] ==> ii k\ \ i j; j \ k |] ==> i \ k" +lemma le_trans: "\i \ j; j \ k\ \ i \ k" by (blast intro: lt_trans1) -lemma succ_leI: "i succ(i) \ j" +lemma succ_leI: "i succ(i) \ j" apply (rule not_lt_iff_le [THEN iffD1]) apply (blast elim: ltE leE lt_asym)+ done -(*Identical to succ(i) < succ(j) ==> i j ==> i i j \ i j <-> i succ(j) ==> i \ j" +lemma succ_le_imp_le: "succ(i) \ succ(j) \ i \ j" by (blast dest!: succ_leE) -lemma lt_subset_trans: "[| i \ j; j ii \ j; j \ i 0 0 i j" @@ -487,86 +487,86 @@ apply (blast dest: lt_asym) done -lemma Ord_succ_mem_iff: "Ord(j) ==> succ(i) \ succ(j) <-> i\j" +lemma Ord_succ_mem_iff: "Ord(j) \ succ(i) \ succ(j) <-> i\j" apply (insert succ_le_iff [of i j]) apply (simp add: lt_def) done subsubsection\Union and Intersection\ -lemma Un_upper1_le: "[| Ord(i); Ord(j) |] ==> i \ i \ j" +lemma Un_upper1_le: "\Ord(i); Ord(j)\ \ i \ i \ j" by (rule Un_upper1 [THEN subset_imp_le], auto) -lemma Un_upper2_le: "[| Ord(i); Ord(j) |] ==> j \ i \ j" +lemma Un_upper2_le: "\Ord(i); Ord(j)\ \ j \ i \ j" by (rule Un_upper2 [THEN subset_imp_le], auto) (*Replacing k by succ(k') yields the similar rule for le!*) -lemma Un_least_lt: "[| i i \ j < k" +lemma Un_least_lt: "\i \ i \ j < k" apply (rule_tac i = i and j = j in Ord_linear_le) apply (auto simp add: Un_commute le_subset_iff subset_Un_iff lt_Ord) done -lemma Un_least_lt_iff: "[| Ord(i); Ord(j) |] ==> i \ j < k <-> iOrd(i); Ord(j)\ \ i \ j < k <-> i i \ j \ k <-> i\k & j\k" + "\Ord(i); Ord(j); Ord(k)\ \ i \ j \ k <-> i\k & j\k" apply (insert Un_least_lt_iff [of i j k]) apply (simp add: lt_def) done (*Replacing k by succ(k') yields the similar rule for le!*) -lemma Int_greatest_lt: "[| i i \ j < k" +lemma Int_greatest_lt: "\i \ i \ j < k" apply (rule_tac i = i and j = j in Ord_linear_le) apply (auto simp add: Int_commute le_subset_iff subset_Int_iff lt_Ord) done lemma Ord_Un_if: - "[| Ord(i); Ord(j) |] ==> i \ j = (if jOrd(i); Ord(j)\ \ i \ j = (if j succ(i \ j) = succ(i) \ succ(j)" + "\Ord(i); Ord(j)\ \ succ(i \ j) = succ(i) \ succ(j)" by (simp add: Ord_Un_if lt_Ord le_Ord2) lemma lt_Un_iff: - "[| Ord(i); Ord(j) |] ==> k < i \ j <-> k < i | k < j" + "\Ord(i); Ord(j)\ \ k < i \ j <-> k < i | k < j" apply (simp add: Ord_Un_if not_lt_iff_le) apply (blast intro: leI lt_trans2)+ done lemma le_Un_iff: - "[| Ord(i); Ord(j) |] ==> k \ i \ j <-> k \ i | k \ j" + "\Ord(i); Ord(j)\ \ k \ i \ j <-> k \ i | k \ j" by (simp add: succ_Un_distrib lt_Un_iff [symmetric]) -lemma Un_upper1_lt: "[|k < i; Ord(j)|] ==> k < i \ j" +lemma Un_upper1_lt: "\k < i; Ord(j)\ \ k < i \ j" by (simp add: lt_Un_iff lt_Ord2) -lemma Un_upper2_lt: "[|k < j; Ord(i)|] ==> k < i \ j" +lemma Un_upper2_lt: "\k < j; Ord(i)\ \ k < i \ j" by (simp add: lt_Un_iff lt_Ord2) (*See also Transset_iff_Union_succ*) -lemma Ord_Union_succ_eq: "Ord(i) ==> \(succ(i)) = i" +lemma Ord_Union_succ_eq: "Ord(i) \ \(succ(i)) = i" by (blast intro: Ord_trans) subsection\Results about Limits\ -lemma Ord_Union [intro,simp,TC]: "[| !!i. i\A ==> Ord(i) |] ==> Ord(\(A))" +lemma Ord_Union [intro,simp,TC]: "\\i. i\A \ Ord(i)\ \ Ord(\(A))" apply (rule Ord_is_Transset [THEN Transset_Union_family, THEN OrdI]) apply (blast intro: Ord_contains_Transset)+ done lemma Ord_UN [intro,simp,TC]: - "[| !!x. x\A ==> Ord(B(x)) |] ==> Ord(\x\A. B(x))" + "\\x. x\A \ Ord(B(x))\ \ Ord(\x\A. B(x))" by (rule Ord_Union, blast) lemma Ord_Inter [intro,simp,TC]: - "[| !!i. i\A ==> Ord(i) |] ==> Ord(\(A))" + "\\i. i\A \ Ord(i)\ \ Ord(\(A))" apply (rule Transset_Inter_family [THEN OrdI]) apply (blast intro: Ord_is_Transset) apply (simp add: Inter_def) @@ -574,95 +574,95 @@ done lemma Ord_INT [intro,simp,TC]: - "[| !!x. x\A ==> Ord(B(x)) |] ==> Ord(\x\A. B(x))" + "\\x. x\A \ Ord(B(x))\ \ Ord(\x\A. B(x))" by (rule Ord_Inter, blast) (* No < version of this theorem: consider that @{term"(\i\nat.i)=nat"}! *) lemma UN_least_le: - "[| Ord(i); !!x. x\A ==> b(x) \ i |] ==> (\x\A. b(x)) \ i" + "\Ord(i); \x. x\A \ b(x) \ i\ \ (\x\A. b(x)) \ i" apply (rule le_imp_subset [THEN UN_least, THEN subset_imp_le]) apply (blast intro: Ord_UN elim: ltE)+ done lemma UN_succ_least_lt: - "[| jA ==> b(x) (\x\A. succ(b(x))) < i" + "\jx. x\A \ b(x) \ (\x\A. succ(b(x))) < i" apply (rule ltE, assumption) apply (rule UN_least_le [THEN lt_trans2]) apply (blast intro: succ_leI)+ done lemma UN_upper_lt: - "[| a\A; i < b(a); Ord(\x\A. b(x)) |] ==> i < (\x\A. b(x))" + "\a\A; i < b(a); Ord(\x\A. b(x))\ \ i < (\x\A. b(x))" by (unfold lt_def, blast) lemma UN_upper_le: - "[| a \ A; i \ b(a); Ord(\x\A. b(x)) |] ==> i \ (\x\A. b(x))" + "\a \ A; i \ b(a); Ord(\x\A. b(x))\ \ i \ (\x\A. b(x))" apply (frule ltD) apply (rule le_imp_subset [THEN subset_trans, THEN subset_imp_le]) apply (blast intro: lt_Ord UN_upper)+ done -lemma lt_Union_iff: "\i\A. Ord(i) ==> (j < \(A)) <-> (\i\A. ji\A. Ord(i) \ (j < \(A)) <-> (\i\A. j J; i\j; Ord(\(J)) |] ==> i \ \J" + "\j \ J; i\j; Ord(\(J))\ \ i \ \J" apply (subst Union_eq_UN) apply (rule UN_upper_le, auto) done lemma le_implies_UN_le_UN: - "[| !!x. x\A ==> c(x) \ d(x) |] ==> (\x\A. c(x)) \ (\x\A. d(x))" + "\\x. x\A \ c(x) \ d(x)\ \ (\x\A. c(x)) \ (\x\A. d(x))" apply (rule UN_least_le) apply (rule_tac [2] UN_upper_le) apply (blast intro: Ord_UN le_Ord2)+ done -lemma Ord_equality: "Ord(i) ==> (\y\i. succ(y)) = i" +lemma Ord_equality: "Ord(i) \ (\y\i. succ(y)) = i" by (blast intro: Ord_trans) (*Holds for all transitive sets, not just ordinals*) -lemma Ord_Union_subset: "Ord(i) ==> \(i) \ i" +lemma Ord_Union_subset: "Ord(i) \ \(i) \ i" by (blast intro: Ord_trans) subsection\Limit Ordinals -- General Properties\ -lemma Limit_Union_eq: "Limit(i) ==> \(i) = i" +lemma Limit_Union_eq: "Limit(i) \ \(i) = i" apply (unfold Limit_def) apply (fast intro!: ltI elim!: ltE elim: Ord_trans) done -lemma Limit_is_Ord: "Limit(i) ==> Ord(i)" +lemma Limit_is_Ord: "Limit(i) \ Ord(i)" apply (unfold Limit_def) apply (erule conjunct1) done -lemma Limit_has_0: "Limit(i) ==> 0 < i" +lemma Limit_has_0: "Limit(i) \ 0 < i" apply (unfold Limit_def) apply (erule conjunct2 [THEN conjunct1]) done -lemma Limit_nonzero: "Limit(i) ==> i \ 0" +lemma Limit_nonzero: "Limit(i) \ i \ 0" by (drule Limit_has_0, blast) -lemma Limit_has_succ: "[| Limit(i); j succ(j) < i" +lemma Limit_has_succ: "\Limit(i); j \ succ(j) < i" by (unfold Limit_def, blast) -lemma Limit_succ_lt_iff [simp]: "Limit(i) ==> succ(j) < i <-> (j succ(j) < i <-> (j Limit(0)" by (simp add: Limit_def) -lemma Limit_has_1: "Limit(i) ==> 1 < i" +lemma Limit_has_1: "Limit(i) \ 1 < i" by (blast intro: Limit_has_0 Limit_has_succ) -lemma increasing_LimitI: "[| 0x\l. \y\l. x Limit(l)" +lemma increasing_LimitI: "\0x\l. \y\l. x \ Limit(l)" apply (unfold Limit_def, simp add: lt_Ord2, clarify) apply (drule_tac i=y in ltD) apply (blast intro: lt_trans1 [OF _ ltI] lt_Ord2) @@ -676,28 +676,28 @@ { fix y assume yi: "y y" using yi by (blast dest: le_imp_not_lt) + have "\ i \ y" using yi by (blast dest: le_imp_not_lt) hence "succ(y) < i" using nsucc [of y] by (blast intro: Ord_linear_lt [OF Osy Oi]) } thus ?thesis using i Oi by (auto simp add: Limit_def) qed -lemma succ_LimitE [elim!]: "Limit(succ(i)) ==> P" +lemma succ_LimitE [elim!]: "Limit(succ(i)) \ P" apply (rule lt_irrefl) apply (rule Limit_has_succ, assumption) apply (erule Limit_is_Ord [THEN Ord_succD, THEN le_refl]) done -lemma not_succ_Limit [simp]: "~ Limit(succ(i))" +lemma not_succ_Limit [simp]: "\ Limit(succ(i))" by blast -lemma Limit_le_succD: "[| Limit(i); i \ succ(j) |] ==> i \ j" +lemma Limit_le_succD: "\Limit(i); i \ succ(j)\ \ i \ j" by (blast elim!: leE) subsubsection\Traditional 3-Way Case Analysis on Ordinals\ -lemma Ord_cases_disj: "Ord(i) ==> i=0 | (\j. Ord(j) & i=succ(j)) | Limit(i)" +lemma Ord_cases_disj: "Ord(i) \ i=0 | (\j. Ord(j) & i=succ(j)) | Limit(i)" by (blast intro!: non_succ_LimitI Ord_0_lt) lemma Ord_cases: @@ -706,11 +706,11 @@ by (insert Ord_cases_disj [OF i], auto) lemma trans_induct3_raw: - "[| Ord(i); + "\Ord(i); P(0); - !!x. [| Ord(x); P(x) |] ==> P(succ(x)); - !!x. [| Limit(x); \y\x. P(y) |] ==> P(x) - |] ==> P(i)" + \x. \Ord(x); P(x)\ \ P(succ(x)); + \x. \Limit(x); \y\x. P(y)\ \ P(x) +\ \ P(i)" apply (erule trans_induct) apply (erule Ord_cases, blast+) done @@ -722,7 +722,7 @@ text\A set of ordinals is either empty, contains its own union, or its union is a limit ordinal.\ -lemma Union_le: "[| !!x. x\I ==> x\j; Ord(j) |] ==> \(I) \ j" +lemma Union_le: "\\x. x\I \ x\j; Ord(j)\ \ \(I) \ j" by (auto simp add: le_subset_iff Union_least) lemma Ord_set_cases: @@ -754,10 +754,10 @@ qed text\If the union of a set of ordinals is a successor, then it is an element of that set.\ -lemma Ord_Union_eq_succD: "[|\x\X. Ord(x); \X = succ(j)|] ==> succ(j) \ X" +lemma Ord_Union_eq_succD: "\\x\X. Ord(x); \X = succ(j)\ \ succ(j) \ X" by (drule Ord_set_cases, auto) -lemma Limit_Union [rule_format]: "[| I \ 0; (\i. i\I \ Limit(i)) |] ==> Limit(\I)" +lemma Limit_Union [rule_format]: "\I \ 0; (\i. i\I \ Limit(i))\ \ Limit(\I)" apply (simp add: Limit_def lt_def) apply (blast intro!: equalityI) done diff -r f2094906e491 -r e44d86131648 src/ZF/Perm.thy --- a/src/ZF/Perm.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Perm.thy Tue Sep 27 16:51:35 2022 +0100 @@ -15,43 +15,43 @@ 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) . + "r O s \ {xz \ domain(s)*range(r) . \x y z. xz= & :s & :r}" definition (*the identity function for A*) id :: "i=>i" where - "id(A) == (\x\A. x)" + "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}" + "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}" + "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)" + "bij(A,B) \ inj(A,B) \ surj(A,B)" subsection\Surjective Function Space\ -lemma surj_is_fun: "f \ surj(A,B) ==> f \ A->B" +lemma surj_is_fun: "f \ surj(A,B) \ f \ A->B" apply (unfold surj_def) apply (erule CollectD1) done -lemma fun_is_surj: "f \ Pi(A,B) ==> f \ surj(A,range(f))" +lemma fun_is_surj: "f \ Pi(A,B) \ f \ surj(A,range(f))" apply (unfold surj_def) apply (blast intro: apply_equality range_of_fun domain_type) done -lemma surj_range: "f \ surj(A,B) ==> range(f)=B" +lemma surj_range: "f \ surj(A,B) \ range(f)=B" apply (unfold surj_def) apply (best intro: apply_Pair elim: range_type) done @@ -59,15 +59,15 @@ 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)" + "\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)" + "\\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 @@ -82,63 +82,63 @@ subsection\Injective Function Space\ -lemma inj_is_fun: "f \ inj(A,B) ==> f \ A->B" +lemma inj_is_fun: "f \ inj(A,B) \ f \ A->B" apply (unfold 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: - "[| :f; :f; f \ inj(A,B) |] ==> a=c" + "\:f; :f; f \ inj(A,B)\ \ a=c" apply (unfold 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" +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)" +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)" + "\\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)" +lemma bij_is_inj: "f \ bij(A,B) \ f \ inj(A,B)" apply (unfold bij_def) apply (erule IntD1) done -lemma bij_is_surj: "f \ bij(A,B) ==> f \ surj(A,B)" +lemma bij_is_surj: "f \ bij(A,B) \ f \ surj(A,B)" apply (unfold bij_def) apply (erule IntD2) done -lemma bij_is_fun: "f \ bij(A,B) ==> f \ A->B" +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)" + "\\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)" apply (unfold 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)" + \ (\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 @@ -146,12 +146,12 @@ subsection\Identity Function\ -lemma idI [intro!]: "a \ A ==> \ id(A)" +lemma idI [intro!]: "a \ A \ \ id(A)" apply (unfold id_def) apply (erule lamI) done -lemma idE [elim!]: "[| p \ id(A); !!x.[| x \ A; p= |] ==> P |] ==> P" +lemma idE [elim!]: "\p \ id(A); \x.\x \ A; p=\ \ P\ \ P" by (simp add: id_def lam_def, blast) lemma id_type: "id(A) \ A->A" @@ -159,17 +159,17 @@ apply (rule lam_type, assumption) done -lemma id_conv [simp]: "x \ A ==> id(A)`x = x" +lemma id_conv [simp]: "x \ A \ id(A)`x = x" apply (unfold id_def) apply (simp (no_asm_simp)) done -lemma id_mono: "A<=B ==> id(A) \ id(B)" +lemma id_mono: "A<=B \ id(A) \ id(B)" apply (unfold id_def) apply (erule lam_mono) done -lemma id_subset_inj: "A<=B ==> id(A): inj(A,B)" +lemma id_subset_inj: "A<=B \ id(A): inj(A,B)" apply (simp add: inj_def id_def) apply (blast intro: lam_type) done @@ -198,7 +198,7 @@ subsection\Converse of a Function\ -lemma inj_converse_fun: "f \ inj(A,B) ==> converse(f) \ range(f)->A" +lemma inj_converse_fun: "f \ inj(A,B) \ converse(f) \ range(f)->A" apply (unfold inj_def) apply (simp (no_asm_simp) add: Pi_iff function_def) apply (erule CollectE) @@ -210,34 +210,34 @@ 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" + "\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" +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" + "\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" + "\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" + "\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" +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)" +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) @@ -245,12 +245,12 @@ apply blast done -lemma inj_converse_surj: "f \ inj(A,B) ==> converse(f): surj(range(f), A)" +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)" +lemma bij_converse_bij [TC]: "f \ bij(A,B) \ converse(f): bij(B,A)" apply (unfold bij_def) apply (fast elim: surj_range [THEN subst] inj_converse_inj inj_converse_surj) done @@ -261,19 +261,19 @@ text\The inductive definition package could derive these theorems for \<^term>\r O s\\ -lemma compI [intro]: "[| :s; :r |] ==> \ r O s" +lemma compI [intro]: "\:s; :r\ \ \ r O s" by (unfold comp_def, blast) lemma compE [elim!]: - "[| xz \ r O s; - !!x y z. [| xz=; :s; :r |] ==> P |] - ==> P" + "\xz \ r O s; + \x y z. \xz=; :s; :r\ \ P\ + \ P" by (unfold comp_def, blast) lemma compEpair: - "[| \ r O s; - !!y. [| :s; :r |] ==> P |] - ==> P" + "\ \ r O s; + \y. \:s; :r\ \ P\ + \ P" by (erule compE, simp) lemma converse_comp: "converse(R O S) = converse(S) O converse(R)" @@ -286,32 +286,32 @@ 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)" +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)" +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))" +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))" +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)" +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" +lemma comp_rel: "\s<=A*B; r<=B*C\ \ (r O s) \ A*C" by blast text\associative law for composition\ @@ -320,31 +320,31 @@ (*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" + 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" + 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)" +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" +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)" + "\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) @@ -352,8 +352,8 @@ 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)))" + "\\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) @@ -363,7 +363,7 @@ done lemma comp_inj: - "[| g \ inj(A,B); f \ inj(B,C) |] ==> (f O g) \ inj(A,C)" + "\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) @@ -371,13 +371,13 @@ done lemma comp_surj: - "[| g \ surj(A,B); f \ surj(B,C) |] ==> (f O g) \ surj(A,C)" + "\g \ surj(A,B); f \ surj(B,C)\ \ (f O g) \ surj(A,C)" apply (unfold 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)" + "\g \ bij(A,B); f \ bij(B,C)\ \ (f O g) \ bij(A,C)" apply (unfold bij_def) apply (blast intro: comp_inj comp_surj) done @@ -390,11 +390,11 @@ 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)" + "\(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)" + "\(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) @@ -404,14 +404,14 @@ done lemma comp_mem_surjD1: - "[| (f O g): surj(A,C); g \ A->B; f \ B->C |] ==> f \ surj(B,C)" + "\(f O g): surj(A,C); g \ A->B; f \ B->C\ \ f \ surj(B,C)" apply (unfold 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)" + "\(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) @@ -420,17 +420,17 @@ 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)" + \<^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)\\ + \<^term>\f \ A->B \ f O converse(f) \ id(B)\\ lemma right_comp_inverse: - "f \ surj(A,B) ==> f O converse(f) = id(B)" + "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) @@ -441,7 +441,7 @@ 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)" + "\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 @@ -451,17 +451,17 @@ 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)" + "\f \ A->B; g \ B->A; f O g = id(B); g O f = id(A)\ \ f \ bij(A,B)" apply (unfold 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)" +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)" + "\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) @@ -471,16 +471,16 @@ 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)" + "\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)" + "\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) @@ -489,8 +489,8 @@ 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)" + "\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) @@ -500,30 +500,30 @@ subsubsection\Restrictions as Surjections and Bijections\ lemma surj_image: - "f \ Pi(A,B) ==> f \ surj(A, f``A)" + "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" +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)" + "\f \ inj(A,B); C<=A\ \ restrict(f,C): inj(C,B)" apply (unfold 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)" +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)" + "\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 @@ -531,13 +531,13 @@ subsubsection\Lemmas for Ramsey's Theorem\ -lemma inj_weaken_type: "[| f \ inj(A,B); B<=D |] ==> f \ inj(A,D)" +lemma inj_weaken_type: "\f \ inj(A,B); B<=D\ \ f \ inj(A,D)" apply (unfold 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})" + "\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) apply (unfold inj_def) apply (fast elim: range_type mem_irrefl dest: apply_equality) @@ -545,8 +545,8 @@ lemma inj_extend: - "[| f \ inj(A,B); a\A; b\B |] - ==> cons(,f) \ inj(cons(a,A), cons(b,B))" + "\f \ inj(A,B); a\A; b\B\ + \ cons(,f) \ inj(cons(a,A), cons(b,B))" apply (unfold inj_def) apply (force intro: apply_type simp add: fun_extend) done diff -r f2094906e491 -r e44d86131648 src/ZF/QPair.thy --- a/src/ZF/QPair.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/QPair.thy Tue Sep 27 16:51:35 2022 +0100 @@ -22,27 +22,27 @@ definition QPair :: "[i, i] => i" (\<(_;/ _)>\) where - " == a+b" + " \ a+b" definition qfst :: "i => i" where - "qfst(p) == THE a. \b. p=" + "qfst(p) \ THE a. \b. p=" definition qsnd :: "i => i" where - "qsnd(p) == THE b. \a. p=" + "qsnd(p) \ THE b. \a. p=" definition qsplit :: "[[i, i] => 'a, i] => 'a::{}" (*for pattern-matching*) where - "qsplit(c,p) == c(qfst(p), qsnd(p))" + "qsplit(c,p) \ c(qfst(p), qsnd(p))" definition qconverse :: "i => i" where - "qconverse(r) == {z. w \ r, \x y. w= & z=}" + "qconverse(r) \ {z. w \ r, \x y. w= & z=}" definition QSigma :: "[i, i => i] => i" where - "QSigma(A,B) == \x\A. \y\B(x). {}" + "QSigma(A,B) \ \x\A. \y\B(x). {}" syntax "_QSUM" :: "[idt, i, i] => i" (\(3QSUM _ \ _./ _)\ 10) @@ -51,23 +51,23 @@ abbreviation qprod (infixr \<*>\ 80) where - "A <*> B == QSigma(A, %_. B)" + "A <*> B \ QSigma(A, %_. B)" definition qsum :: "[i,i]=>i" (infixr \<+>\ 65) where - "A <+> B == ({0} <*> A) \ ({1} <*> B)" + "A <+> B \ ({0} <*> A) \ ({1} <*> B)" definition QInl :: "i=>i" where - "QInl(a) == <0;a>" + "QInl(a) \ <0;a>" definition QInr :: "i=>i" where - "QInr(b) == <1;b>" + "QInr(b) \ <1;b>" definition qcase :: "[i=>i, i=>i, i]=>i" where - "qcase(c,d) == qsplit(%y z. cond(y, d(z), c(z)))" + "qcase(c,d) \ qsplit(%y z. cond(y, d(z), c(z)))" subsection\Quine ordered pairing\ @@ -84,40 +84,40 @@ lemmas QPair_inject = QPair_iff [THEN iffD1, THEN conjE, elim!] -lemma QPair_inject1: " = ==> a=c" +lemma QPair_inject1: " = \ a=c" by blast -lemma QPair_inject2: " = ==> b=d" +lemma QPair_inject2: " = \ b=d" by blast subsubsection\QSigma: Disjoint union of a family of sets Generalizes Cartesian product\ -lemma QSigmaI [intro!]: "[| a \ A; b \ B(a) |] ==> \ QSigma(A,B)" +lemma QSigmaI [intro!]: "\a \ A; b \ B(a)\ \ \ QSigma(A,B)" by (simp add: QSigma_def) (** Elimination rules for :A*B -- introducing no eigenvariables **) lemma QSigmaE [elim!]: - "[| c \ QSigma(A,B); - !!x y.[| x \ A; y \ B(x); c= |] ==> P - |] ==> P" + "\c \ QSigma(A,B); + \x y.\x \ A; y \ B(x); c=\ \ P +\ \ P" by (simp add: QSigma_def, blast) lemma QSigmaE2 [elim!]: - "[| : QSigma(A,B); [| a \ A; b \ B(a) |] ==> P |] ==> P" + "\: QSigma(A,B); \a \ A; b \ B(a)\ \ P\ \ P" by (simp add: QSigma_def) -lemma QSigmaD1: " \ QSigma(A,B) ==> a \ A" +lemma QSigmaD1: " \ QSigma(A,B) \ a \ A" by blast -lemma QSigmaD2: " \ QSigma(A,B) ==> b \ B(a)" +lemma QSigmaD2: " \ QSigma(A,B) \ b \ B(a)" by blast lemma QSigma_cong: - "[| A=A'; !!x. x \ A' ==> B(x)=B'(x) |] ==> + "\A=A'; \x. x \ A' \ B(x)=B'(x)\ \ QSigma(A,B) = QSigma(A',B')" by (simp add: QSigma_def) @@ -136,69 +136,69 @@ lemma qsnd_conv [simp]: "qsnd() = b" by (simp add: qsnd_def) -lemma qfst_type [TC]: "p \ QSigma(A,B) ==> qfst(p) \ A" +lemma qfst_type [TC]: "p \ QSigma(A,B) \ qfst(p) \ A" by auto -lemma qsnd_type [TC]: "p \ QSigma(A,B) ==> qsnd(p) \ B(qfst(p))" +lemma qsnd_type [TC]: "p \ QSigma(A,B) \ qsnd(p) \ B(qfst(p))" by auto -lemma QPair_qfst_qsnd_eq: "a \ QSigma(A,B) ==> = a" +lemma QPair_qfst_qsnd_eq: "a \ QSigma(A,B) \ = a" by auto subsubsection\Eliminator: qsplit\ (*A META-equality, so that it applies to higher types as well...*) -lemma qsplit [simp]: "qsplit(%x y. c(x,y), ) == c(a,b)" +lemma qsplit [simp]: "qsplit(%x y. c(x,y), ) \ c(a,b)" by (simp add: qsplit_def) lemma qsplit_type [elim!]: - "[| p \ QSigma(A,B); - !!x y.[| x \ A; y \ B(x) |] ==> c(x,y):C() - |] ==> qsplit(%x y. c(x,y), p) \ C(p)" + "\p \ QSigma(A,B); + \x y.\x \ A; y \ B(x)\ \ c(x,y):C() +\ \ qsplit(%x y. c(x,y), p) \ C(p)" by auto lemma expand_qsplit: - "u \ A<*>B ==> R(qsplit(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" + "u \ A<*>B \ R(qsplit(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" apply (simp add: qsplit_def, auto) done subsubsection\qsplit for predicates: result type o\ -lemma qsplitI: "R(a,b) ==> qsplit(R, )" +lemma qsplitI: "R(a,b) \ qsplit(R, )" by (simp add: qsplit_def) lemma qsplitE: - "[| qsplit(R,z); z \ QSigma(A,B); - !!x y. [| z = ; R(x,y) |] ==> P - |] ==> P" + "\qsplit(R,z); z \ QSigma(A,B); + \x y. \z = ; R(x,y)\ \ P +\ \ P" by (simp add: qsplit_def, auto) -lemma qsplitD: "qsplit(R,) ==> R(a,b)" +lemma qsplitD: "qsplit(R,) \ R(a,b)" by (simp add: qsplit_def) subsubsection\qconverse\ -lemma qconverseI [intro!]: ":r ==> :qconverse(r)" +lemma qconverseI [intro!]: ":r \ :qconverse(r)" by (simp add: qconverse_def, blast) -lemma qconverseD [elim!]: " \ qconverse(r) ==> \ r" +lemma qconverseD [elim!]: " \ qconverse(r) \ \ r" by (simp add: qconverse_def, blast) lemma qconverseE [elim!]: - "[| yx \ qconverse(r); - !!x y. [| yx=; :r |] ==> P - |] ==> P" + "\yx \ qconverse(r); + \x y. \yx=; :r\ \ P +\ \ P" by (simp add: qconverse_def, blast) -lemma qconverse_qconverse: "r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" +lemma qconverse_qconverse: "r<=QSigma(A,B) \ qconverse(qconverse(r)) = r" by blast -lemma qconverse_type: "r \ A <*> B ==> qconverse(r) \ B <*> A" +lemma qconverse_type: "r \ A <*> B \ qconverse(r) \ B <*> A" by blast lemma qconverse_prod: "qconverse(A <*> B) = B <*> A" @@ -214,19 +214,19 @@ (** Introduction rules for the injections **) -lemma QInlI [intro!]: "a \ A ==> QInl(a) \ A <+> B" +lemma QInlI [intro!]: "a \ A \ QInl(a) \ A <+> B" by (simp add: qsum_defs, blast) -lemma QInrI [intro!]: "b \ B ==> QInr(b) \ A <+> B" +lemma QInrI [intro!]: "b \ B \ QInr(b) \ A <+> B" by (simp add: qsum_defs, blast) (** Elimination rules **) lemma qsumE [elim!]: - "[| u \ A <+> B; - !!x. [| x \ A; u=QInl(x) |] ==> P; - !!y. [| y \ B; u=QInr(y) |] ==> P - |] ==> P" + "\u \ A <+> B; + \x. \x \ A; u=QInl(x)\ \ P; + \y. \y \ B; u=QInr(y)\ \ P +\ \ P" by (simp add: qsum_defs, blast) @@ -254,10 +254,10 @@ lemmas QInl_neq_QInr = QInl_QInr_iff [THEN iffD1, THEN FalseE, elim!] lemmas QInr_neq_QInl = QInr_QInl_iff [THEN iffD1, THEN FalseE, elim!] -lemma QInlD: "QInl(a): A<+>B ==> a \ A" +lemma QInlD: "QInl(a): A<+>B \ a \ A" by blast -lemma QInrD: "QInr(b): A<+>B ==> b \ B" +lemma QInrD: "QInr(b): A<+>B \ b \ B" by blast (** <+> is itself injective... who cares?? **) @@ -284,10 +284,10 @@ by (simp add: qsum_defs ) lemma qcase_type: - "[| u \ A <+> B; - !!x. x \ A ==> c(x): C(QInl(x)); - !!y. y \ B ==> d(y): C(QInr(y)) - |] ==> qcase(c,d,u) \ C(u)" + "\u \ A <+> B; + \x. x \ A \ c(x): C(QInl(x)); + \y. y \ B \ d(y): C(QInr(y)) +\ \ qcase(c,d,u) \ C(u)" by (simp add: qsum_defs, auto) (** Rules for the Part primitive **) @@ -301,26 +301,26 @@ lemma Part_QInr2: "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y \ Part(B,h)}" by blast -lemma Part_qsum_equality: "C \ A <+> B ==> Part(C,QInl) \ Part(C,QInr) = C" +lemma Part_qsum_equality: "C \ A <+> B \ Part(C,QInl) \ Part(C,QInr) = C" by blast subsubsection\Monotonicity\ -lemma QPair_mono: "[| a<=c; b<=d |] ==> \ " +lemma QPair_mono: "\a<=c; b<=d\ \ \ " by (simp add: QPair_def sum_mono) lemma QSigma_mono [rule_format]: - "[| A<=C; \x\A. B(x) \ D(x) |] ==> QSigma(A,B) \ QSigma(C,D)" + "\A<=C; \x\A. B(x) \ D(x)\ \ QSigma(A,B) \ QSigma(C,D)" by blast -lemma QInl_mono: "a<=b ==> QInl(a) \ QInl(b)" +lemma QInl_mono: "a<=b \ QInl(a) \ QInl(b)" by (simp add: QInl_def subset_refl [THEN QPair_mono]) -lemma QInr_mono: "a<=b ==> QInr(a) \ QInr(b)" +lemma QInr_mono: "a<=b \ QInr(a) \ QInr(b)" by (simp add: QInr_def subset_refl [THEN QPair_mono]) -lemma qsum_mono: "[| A<=C; B<=D |] ==> A <+> B \ C <+> D" +lemma qsum_mono: "\A<=C; B<=D\ \ A <+> B \ C <+> D" by blast end diff -r f2094906e491 -r e44d86131648 src/ZF/QUniv.thy --- a/src/ZF/QUniv.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/QUniv.thy Tue Sep 27 16:51:35 2022 +0100 @@ -21,32 +21,32 @@ definition quniv :: "i => i" where - "quniv(A) == Pow(univ(eclose(A)))" + "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" + "\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)" + "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)" +lemma qunivI: "X \ univ(eclose(A)) \ X \ quniv(A)" by (simp add: quniv_def) -lemma qunivD: "X \ quniv(A) ==> X \ univ(eclose(A))" +lemma qunivD: "X \ quniv(A) \ X \ univ(eclose(A))" by (simp add: quniv_def) -lemma quniv_mono: "A<=B ==> quniv(A) \ quniv(B)" +lemma quniv_mono: "A<=B \ quniv(A) \ quniv(B)" apply (unfold quniv_def) apply (erule eclose_mono [THEN univ_mono, THEN Pow_mono]) done @@ -86,12 +86,12 @@ (*Quine ordered pairs*) lemma QPair_subset_univ: - "[| a \ univ(A); b \ univ(A) |] ==> \ univ(A)" + "\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)" +lemma QInl_subset_univ: "a \ univ(A) \ QInl(a) \ univ(A)" apply (unfold QInl_def) apply (erule empty_subsetI [THEN QPair_subset_univ]) done @@ -102,7 +102,7 @@ lemmas naturals_subset_univ = subset_trans [OF naturals_subset_nat nat_subset_univ] -lemma QInr_subset_univ: "a \ univ(A) ==> QInr(a) \ univ(A)" +lemma QInr_subset_univ: "a \ univ(A) \ QInr(a) \ univ(A)" apply (unfold QInr_def) apply (erule nat_1I [THEN naturals_subset_univ, THEN QPair_subset_univ]) done @@ -111,7 +111,7 @@ (*Quine ordered pairs*) lemma QPair_in_quniv: - "[| a: quniv(A); b: quniv(A) |] ==> \ quniv(A)" + "\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)" @@ -121,7 +121,7 @@ (*The opposite inclusion*) lemma quniv_QPair_D: - " \ quniv(A) ==> a: quniv(A) & b: quniv(A)" + " \ quniv(A) \ a: quniv(A) & b: quniv(A)" apply (unfold quniv_def QPair_def) apply (rule Transset_includes_summands [THEN conjE]) apply (rule Transset_eclose [THEN Transset_univ]) @@ -136,10 +136,10 @@ subsection\Quine Disjoint Sum\ -lemma QInl_in_quniv: "a: quniv(A) ==> QInl(a) \ quniv(A)" +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)" +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)" @@ -152,7 +152,7 @@ lemmas nat_subset_quniv = subset_trans [OF nat_subset_univ univ_subset_quniv] -(* n:nat ==> n:quniv(A) *) +(* 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] @@ -163,7 +163,7 @@ (*Intersecting with Vfrom...*) lemma QPair_Int_Vfrom_succ_subset: - "Transset(X) ==> + "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] @@ -176,18 +176,18 @@ (*Rule for level i -- preserving the level, not decreasing it*) lemma QPair_Int_Vfrom_subset: - "Transset(X) ==> + "Transset(X) \ \ Vfrom(X,i) \ Vfrom(X,i); b \ Vfrom(X,i)>" apply (unfold QPair_def) apply (erule Transset_Vfrom [THEN Transset_sum_Int_subset]) done -(*@{term"[| a \ Vset(i) \ c; b \ Vset(i) \ d |] ==> \ Vset(i) \ "}*) +(*@{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)>)" + "Ord(i) \ \ Vset(i) \ (\j\i. Vset(j); b \ Vset(j)>)" apply (erule Ord_cases) (*0 case*) apply (simp add: Vfrom_0) diff -r f2094906e491 -r e44d86131648 src/ZF/Resid/Confluence.thy --- a/src/ZF/Resid/Confluence.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Resid/Confluence.thy Tue Sep 27 16:51:35 2022 +0100 @@ -7,13 +7,13 @@ definition confluence :: "i=>o" where - "confluence(R) == + "confluence(R) \ \x y. \ R \ (\z. \ R \ (\u. \ R & \ R))" definition strip :: "o" where - "strip == \x y. (x ===> y) \ - (\z.(x =1=> z) \ (\u.(y =1=> u) & (z===>u)))" + "strip \ \x y. (x =\ y) \ + (\z.(x =1=> z) \ (\u.(y =1=> u) & (z=\u)))" (* ------------------------------------------------------------------------- *) @@ -21,7 +21,7 @@ (* ------------------------------------------------------------------------- *) lemma strip_lemma_r: - "[|confluence(Spar_red1)|]==> strip" + "\confluence(Spar_red1)\\ strip" apply (unfold confluence_def strip_def) apply (rule impI [THEN allI, THEN allI]) apply (erule Spar_red.induct, fast) @@ -30,7 +30,7 @@ lemma strip_lemma_l: - "strip==> confluence(Spar_red)" + "strip\ confluence(Spar_red)" apply (unfold confluence_def strip_def) apply (rule impI [THEN allI, THEN allI]) apply (erule Spar_red.induct, blast) @@ -54,7 +54,7 @@ lemmas confluence_parallel_reduction = parallel_moves [THEN strip_lemma_r, THEN strip_lemma_l] -lemma lemma1: "[|confluence(Spar_red)|]==> confluence(Sred)" +lemma lemma1: "\confluence(Spar_red)\\ confluence(Sred)" by (unfold confluence_def, blast intro: par_red_red red_par_red) lemmas confluence_beta_reduction = @@ -69,17 +69,17 @@ abbreviation Sconv1_rel (infixl \<-1->\ 50) where - "a<-1->b == \ Sconv1" + "a<-1->b \ \ Sconv1" abbreviation Sconv_rel (infixl \<-\\ 50) where - "a<-\b == \ Sconv" + "a<-\b \ \ Sconv" inductive domains "Sconv1" \ "lambda*lambda" intros - red1: "m -1-> n ==> m<-1->n" - expl: "n -1-> m ==> m<-1->n" + 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] @@ -87,14 +87,14 @@ 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" + 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" +lemma conv_sym: "m<-\n \ n<-\m" apply (erule Sconv.induct) apply (erule Sconv1.induct, blast+) done @@ -103,7 +103,7 @@ (* Church_Rosser Theorem *) (* ------------------------------------------------------------------------- *) -lemma Church_Rosser: "m<-\n ==> \p.(m -\p) & (n -\ p)" +lemma Church_Rosser: "m<-\n \ \p.(m -\p) & (n -\ p)" apply (erule Sconv.induct) apply (erule Sconv1.induct) apply (blast intro: red1D1 redD2) diff -r f2094906e491 -r e44d86131648 src/ZF/Resid/Redex.thy --- a/src/ZF/Resid/Redex.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Resid/Redex.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,14 +19,14 @@ abbreviation Ssub_rel (infixl \\\ 70) where - "a \ b == \ Ssub" + "a \ b \ \ Ssub" abbreviation Scomp_rel (infixl \\\ 70) where - "a \ b == \ Scomp" + "a \ b \ \ Scomp" abbreviation - "regular(a) == a \ Sreg" + "regular(a) \ a \ Sreg" consts union_aux :: "i=>i" primrec (*explicit lambda is required because both arguments of "\" vary*) @@ -44,35 +44,35 @@ definition union (infixl \\\ 70) where - "u \ v == union_aux(u)`v" + "u \ v \ union_aux(u)`v" inductive domains "Ssub" \ "redexes*redexes" intros - Sub_Var: "n \ nat ==> Var(n) \ Var(n)" - Sub_Fun: "[|u \ v|]==> Fun(u) \ Fun(v)" - Sub_App1: "[|u1 \ v1; u2 \ v2; b \ bool|]==> + Sub_Var: "n \ nat \ Var(n) \ Var(n)" + Sub_Fun: "\u \ v\\ Fun(u) \ Fun(v)" + Sub_App1: "\u1 \ v1; u2 \ v2; b \ bool\\ App(0,u1,u2) \ App(b,v1,v2)" - Sub_App2: "[|u1 \ v1; u2 \ v2|]==> App(1,u1,u2) \ App(1,v1,v2)" + Sub_App2: "\u1 \ v1; u2 \ v2\\ App(1,u1,u2) \ App(1,v1,v2)" type_intros redexes.intros bool_typechecks inductive domains "Scomp" \ "redexes*redexes" intros - Comp_Var: "n \ nat ==> Var(n) \ Var(n)" - Comp_Fun: "[|u \ v|]==> Fun(u) \ Fun(v)" - Comp_App: "[|u1 \ v1; u2 \ v2; b1 \ bool; b2 \ bool|] - ==> App(b1,u1,u2) \ App(b2,v1,v2)" + Comp_Var: "n \ nat \ Var(n) \ Var(n)" + Comp_Fun: "\u \ v\\ Fun(u) \ Fun(v)" + Comp_App: "\u1 \ v1; u2 \ v2; b1 \ bool; b2 \ bool\ + \ App(b1,u1,u2) \ App(b2,v1,v2)" type_intros redexes.intros bool_typechecks inductive domains "Sreg" \ redexes intros - Reg_Var: "n \ nat ==> regular(Var(n))" - Reg_Fun: "[|regular(u)|]==> regular(Fun(u))" - Reg_App1: "[|regular(Fun(u)); regular(v) |]==>regular(App(1,Fun(u),v))" - Reg_App2: "[|regular(u); regular(v) |]==>regular(App(0,u,v))" + Reg_Var: "n \ nat \ regular(Var(n))" + Reg_Fun: "\regular(u)\\ regular(Fun(u))" + Reg_App1: "\regular(Fun(u)); regular(v)\\regular(App(1,Fun(u),v))" + Reg_App2: "\regular(u); regular(v)\\regular(App(0,u,v))" type_intros redexes.intros bool_typechecks @@ -90,15 +90,15 @@ (* Equality rules for union *) (* ------------------------------------------------------------------------- *) -lemma union_Var [simp]: "n \ nat ==> Var(n) \ Var(n)=Var(n)" +lemma union_Var [simp]: "n \ nat \ Var(n) \ Var(n)=Var(n)" by (simp add: union_def) -lemma union_Fun [simp]: "v \ redexes ==> Fun(u) \ Fun(v) = Fun(u \ v)" +lemma union_Fun [simp]: "v \ redexes \ Fun(u) \ Fun(v) = Fun(u \ v)" by (simp add: union_def) lemma union_App [simp]: - "[|b2 \ bool; u2 \ redexes; v2 \ redexes|] - ==> App(b1,u1,v1) \ App(b2,u2,v2)=App(b1 or b2,u1 \ u2,v1 \ v2)" + "\b2 \ bool; u2 \ redexes; v2 \ redexes\ + \ App(b1,u1,v1) \ App(b2,u2,v2)=App(b1 or b2,u1 \ u2,v1 \ v2)" by (simp add: union_def) @@ -135,16 +135,16 @@ (* comp proofs *) (* ------------------------------------------------------------------------- *) -lemma comp_refl [simp]: "u \ redexes ==> u \ u" +lemma comp_refl [simp]: "u \ redexes \ u \ u" by (erule redexes.induct, blast+) -lemma comp_sym: "u \ v ==> v \ u" +lemma comp_sym: "u \ v \ v \ u" by (erule Scomp.induct, blast+) lemma comp_sym_iff: "u \ v \ v \ u" by (blast intro: comp_sym) -lemma comp_trans [rule_format]: "u \ v ==> \w. v \ w\u \ w" +lemma comp_trans [rule_format]: "u \ v \ \w. v \ w\u \ w" by (erule Scomp.induct, blast+) (* ------------------------------------------------------------------------- *) diff -r f2094906e491 -r e44d86131648 src/ZF/Resid/Reduction.thy --- a/src/ZF/Resid/Reduction.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Resid/Reduction.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,14 +13,14 @@ abbreviation Apl :: "[i,i]=>i" where - "Apl(n,m) == App(0,n,m)" + "Apl(n,m) \ App(0,n,m)" inductive domains "lambda" \ redexes intros - Lambda_Var: " n \ nat ==> Var(n) \ lambda" - Lambda_Fun: " u \ lambda ==> Fun(u) \ lambda" - Lambda_App: "[|u \ lambda; v \ lambda|] ==> Apl(u,v) \ lambda" + Lambda_Var: " n \ nat \ Var(n) \ lambda" + Lambda_Fun: " u \ lambda \ Fun(u) \ lambda" + Lambda_App: "\u \ lambda; v \ lambda\ \ Apl(u,v) \ lambda" type_intros redexes.intros bool_typechecks declare lambda.intros [intro] @@ -40,10 +40,10 @@ (* ------------------------------------------------------------------------- *) lemma unmark_type [intro, simp]: - "u \ redexes ==> unmark(u) \ lambda" + "u \ redexes \ unmark(u) \ lambda" by (erule redexes.induct, simp_all) -lemma lambda_unmark: "u \ lambda ==> unmark(u) = u" +lemma lambda_unmark: "u \ lambda \ unmark(u) = u" by (erule lambda.induct, simp_all) @@ -52,11 +52,11 @@ (* ------------------------------------------------------------------------- *) lemma liftL_type [rule_format]: - "v \ lambda ==> \k \ nat. lift_rec(v,k) \ lambda" + "v \ lambda \ \k \ nat. lift_rec(v,k) \ lambda" by (erule lambda.induct, simp_all add: lift_rec_Var) lemma substL_type [rule_format, simp]: - "v \ lambda ==> \n \ nat. \u \ lambda. subst_rec(u,v,n) \ lambda" + "v \ lambda \ \n \ nat. \u \ lambda. subst_rec(u,v,n) \ lambda" by (erule lambda.induct, simp_all add: liftL_type subst_Var) @@ -75,28 +75,28 @@ abbreviation Sred1_rel (infixl \-1->\ 50) where - "a -1-> b == \ Sred1" + "a -1-> b \ \ Sred1" abbreviation Sred_rel (infixl \-\\ 50) where - "a -\ b == \ Sred" + "a -\ b \ \ Sred" abbreviation Spar_red1_rel (infixl \=1=>\ 50) where - "a =1=> b == \ Spar_red1" + "a =1=> b \ \ Spar_red1" abbreviation - Spar_red_rel (infixl \===>\ 50) where - "a ===> b == \ Spar_red" + Spar_red_rel (infixl \=\\ 50) where + "a =\ b \ \ Spar_red" inductive domains "Sred1" \ "lambda*lambda" intros - beta: "[|m \ lambda; n \ lambda|] ==> Apl(Fun(m),n) -1-> n/m" - rfun: "[|m -1-> n|] ==> Fun(m) -1-> Fun(n)" - apl_l: "[|m2 \ lambda; m1 -1-> n1|] ==> Apl(m1,m2) -1-> Apl(n1,m2)" - apl_r: "[|m1 \ lambda; m2 -1-> n2|] ==> Apl(m1,m2) -1-> Apl(m1,n2)" + beta: "\m \ lambda; n \ lambda\ \ Apl(Fun(m),n) -1-> n/m" + rfun: "\m -1-> n\ \ Fun(m) -1-> Fun(n)" + apl_l: "\m2 \ lambda; m1 -1-> n1\ \ Apl(m1,m2) -1-> Apl(n1,m2)" + apl_r: "\m1 \ lambda; m2 -1-> n2\ \ Apl(m1,m2) -1-> Apl(m1,n2)" type_intros red_typechecks declare Sred1.intros [intro, simp] @@ -104,9 +104,9 @@ inductive domains "Sred" \ "lambda*lambda" intros - one_step: "m-1->n ==> m-\n" - refl: "m \ lambda==>m -\m" - trans: "[|m-\n; n-\p|] ==>m-\p" + one_step: "m-1->n \ m-\n" + refl: "m \ lambda\m -\m" + trans: "\m-\n; n-\p\ \m-\p" type_intros Sred1.dom_subset [THEN subsetD] red_typechecks declare Sred.one_step [intro, simp] @@ -115,10 +115,10 @@ inductive domains "Spar_red1" \ "lambda*lambda" intros - beta: "[|m =1=> m'; n =1=> n'|] ==> Apl(Fun(m),n) =1=> n'/m'" - rvar: "n \ nat ==> Var(n) =1=> Var(n)" - rfun: "m =1=> m' ==> Fun(m) =1=> Fun(m')" - rapl: "[|m =1=> m'; n =1=> n'|] ==> Apl(m,n) =1=> Apl(m',n')" + beta: "\m =1=> m'; n =1=> n'\ \ Apl(Fun(m),n) =1=> n'/m'" + rvar: "n \ nat \ Var(n) =1=> Var(n)" + rfun: "m =1=> m' \ Fun(m) =1=> Fun(m')" + rapl: "\m =1=> m'; n =1=> n'\ \ Apl(m,n) =1=> Apl(m',n')" type_intros red_typechecks declare Spar_red1.intros [intro, simp] @@ -126,8 +126,8 @@ inductive domains "Spar_red" \ "lambda*lambda" intros - one_step: "m =1=> n ==> m ===> n" - trans: "[|m===>n; n===>p|] ==> m===>p" + one_step: "m =1=> n \ m =\ n" + trans: "\m=\n; n=\p\ \ m=\p" type_intros Spar_red1.dom_subset [THEN subsetD] red_typechecks declare Spar_red.one_step [intro, simp] @@ -158,27 +158,27 @@ (* Lemmas for reduction *) (* ------------------------------------------------------------------------- *) -lemma red_Fun: "m-\n ==> Fun(m) -\ Fun(n)" +lemma red_Fun: "m-\n \ Fun(m) -\ Fun(n)" apply (erule Sred.induct) apply (rule_tac [3] Sred.trans, simp_all) done -lemma red_Apll: "[|n \ lambda; m -\ m'|] ==> Apl(m,n)-\Apl(m',n)" +lemma red_Apll: "\n \ lambda; m -\ m'\ \ Apl(m,n)-\Apl(m',n)" apply (erule Sred.induct) apply (rule_tac [3] Sred.trans, simp_all) done -lemma red_Aplr: "[|n \ lambda; m -\ m'|] ==> Apl(n,m)-\Apl(n,m')" +lemma red_Aplr: "\n \ lambda; m -\ m'\ \ Apl(n,m)-\Apl(n,m')" apply (erule Sred.induct) apply (rule_tac [3] Sred.trans, simp_all) done -lemma red_Apl: "[|m -\ m'; n-\n'|] ==> Apl(m,n)-\Apl(m',n')" +lemma red_Apl: "\m -\ m'; n-\n'\ \ Apl(m,n)-\Apl(m',n')" apply (rule_tac n = "Apl (m',n) " in Sred.trans) apply (simp_all add: red_Apll red_Aplr) done -lemma red_beta: "[|m \ lambda; m':lambda; n \ lambda; n':lambda; m -\ m'; n-\n'|] ==> +lemma red_beta: "\m \ lambda; m':lambda; n \ lambda; n':lambda; m -\ m'; n-\n'\ \ Apl(Fun(m),n)-\ n'/m'" apply (rule_tac n = "Apl (Fun (m'),n') " in Sred.trans) apply (simp_all add: red_Apl red_Fun) @@ -190,19 +190,19 @@ (* ------------------------------------------------------------------------- *) -lemma refl_par_red1: "m \ lambda==> m =1=> m" +lemma refl_par_red1: "m \ lambda\ m =1=> m" by (erule lambda.induct, simp_all) -lemma red1_par_red1: "m-1->n ==> m=1=>n" +lemma red1_par_red1: "m-1->n \ m=1=>n" by (erule Sred1.induct, simp_all add: refl_par_red1) -lemma red_par_red: "m-\n ==> m===>n" +lemma red_par_red: "m-\n \ m=\n" apply (erule Sred.induct) apply (rule_tac [3] Spar_red.trans) apply (simp_all add: refl_par_red1 red1_par_red1) done -lemma par_red_red: "m===>n ==> m-\n" +lemma par_red_red: "m=\n \ m-\n" apply (erule Spar_red.induct) apply (erule Spar_red1.induct) apply (rule_tac [5] Sred.trans) @@ -214,7 +214,7 @@ (* Simulation *) (* ------------------------------------------------------------------------- *) -lemma simulation: "m=1=>n ==> \v. m|>v = n & m \ v & regular(v)" +lemma simulation: "m=1=>n \ \v. m|>v = n & m \ v & regular(v)" by (erule Spar_red1.induct, force+) @@ -223,11 +223,11 @@ (* ------------------------------------------------------------------------- *) lemma unmmark_lift_rec: - "u \ redexes ==> \k \ nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)" + "u \ redexes \ \k \ nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)" by (erule redexes.induct, simp_all add: lift_rec_Var) lemma unmmark_subst_rec: - "v \ redexes ==> \k \ nat. \u \ redexes. + "v \ redexes \ \k \ nat. \u \ redexes. unmark(subst_rec(u,v,k)) = subst_rec(unmark(u),unmark(v),k)" by (erule redexes.induct, simp_all add: unmmark_lift_rec subst_Var) @@ -237,12 +237,12 @@ (* ------------------------------------------------------------------------- *) lemma completeness_l [rule_format]: - "u \ v ==> regular(v) \ unmark(u) =1=> unmark(u|>v)" + "u \ v \ regular(v) \ unmark(u) =1=> unmark(u|>v)" apply (erule Scomp.induct) apply (auto simp add: unmmark_subst_rec) done -lemma completeness: "[|u \ lambda; u \ v; regular(v)|] ==> u =1=> unmark(u|>v)" +lemma completeness: "\u \ lambda; u \ v; regular(v)\ \ u =1=> unmark(u|>v)" by (drule completeness_l, simp_all add: lambda_unmark) end diff -r f2094906e491 -r e44d86131648 src/ZF/Resid/Residuals.thy --- a/src/ZF/Resid/Residuals.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Resid/Residuals.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,25 +9,25 @@ Sres :: "i" abbreviation - "residuals(u,v,w) == \ Sres" + "residuals(u,v,w) \ \ Sres" inductive domains "Sres" \ "redexes*redexes*redexes" intros - Res_Var: "n \ nat ==> residuals(Var(n),Var(n),Var(n))" - Res_Fun: "[|residuals(u,v,w)|]==> + Res_Var: "n \ nat \ residuals(Var(n),Var(n),Var(n))" + Res_Fun: "\residuals(u,v,w)\\ residuals(Fun(u),Fun(v),Fun(w))" - Res_App: "[|residuals(u1,v1,w1); - residuals(u2,v2,w2); b \ bool|]==> + Res_App: "\residuals(u1,v1,w1); + residuals(u2,v2,w2); b \ bool\\ residuals(App(b,u1,u2),App(0,v1,v2),App(b,w1,w2))" - Res_redex: "[|residuals(u1,v1,w1); - residuals(u2,v2,w2); b \ bool|]==> + Res_redex: "\residuals(u1,v1,w1); + residuals(u2,v2,w2); b \ bool\\ residuals(App(b,Fun(u1),u2),App(1,Fun(v1),v2),w2/w1)" type_intros subst_type nat_typechecks redexes.intros bool_typechecks definition res_func :: "[i,i]=>i" (infixl \|>\ 70) where - "u |> v == THE w. residuals(u,v,w)" + "u |> v \ THE w. residuals(u,v,w)" subsection\Setting up rule lists\ @@ -64,15 +64,15 @@ subsection\residuals is a partial function\ lemma residuals_function [rule_format]: - "residuals(u,v,w) ==> \w1. residuals(u,v,w1) \ w1 = w" + "residuals(u,v,w) \ \w1. residuals(u,v,w1) \ w1 = w" by (erule Sres.induct, force+) lemma residuals_intro [rule_format]: - "u \ v ==> regular(v) \ (\w. residuals(u,v,w))" + "u \ v \ regular(v) \ (\w. residuals(u,v,w))" by (erule Scomp.induct, force+) lemma comp_resfuncD: - "[| u \ v; regular(v) |] ==> residuals(u, v, THE w. residuals(u, v, w))" + "\u \ v; regular(v)\ \ residuals(u, v, THE w. residuals(u, v, w))" apply (frule residuals_intro, assumption, clarify) apply (subst the_equality) apply (blast intro: residuals_function)+ @@ -80,32 +80,32 @@ subsection\Residual function\ -lemma res_Var [simp]: "n \ nat ==> Var(n) |> Var(n) = Var(n)" +lemma res_Var [simp]: "n \ nat \ Var(n) |> Var(n) = Var(n)" by (unfold res_func_def, blast) lemma res_Fun [simp]: - "[|s \ t; regular(t)|]==> Fun(s) |> Fun(t) = Fun(s |> t)" + "\s \ t; regular(t)\\ Fun(s) |> Fun(t) = Fun(s |> t)" apply (unfold res_func_def) apply (blast intro: comp_resfuncD residuals_function) done lemma res_App [simp]: - "[|s \ u; regular(u); t \ v; regular(v); b \ bool|] - ==> App(b,s,t) |> App(0,u,v) = App(b, s |> u, t |> v)" + "\s \ u; regular(u); t \ v; regular(v); b \ bool\ + \ App(b,s,t) |> App(0,u,v) = App(b, s |> u, t |> v)" apply (unfold res_func_def) apply (blast dest!: comp_resfuncD intro: residuals_function) done lemma res_redex [simp]: - "[|s \ u; regular(u); t \ v; regular(v); b \ bool|] - ==> App(b,Fun(s),t) |> App(1,Fun(u),v) = (t |> v)/ (s |> u)" + "\s \ u; regular(u); t \ v; regular(v); b \ bool\ + \ App(b,Fun(s),t) |> App(1,Fun(u),v) = (t |> v)/ (s |> u)" apply (unfold res_func_def) apply (blast elim!: redexes.free_elims dest!: comp_resfuncD intro: residuals_function) done lemma resfunc_type [simp]: - "[|s \ t; regular(t)|]==> regular(t) \ s |> t \ redexes" + "\s \ t; regular(t)\\ regular(t) \ s |> t \ redexes" by (erule Scomp.induct, auto) subsection\Commutation theorem\ @@ -117,14 +117,14 @@ "u \ v \ regular(v) \ regular(u)" by (erule Ssub.induct, auto) -lemma residuals_lift_rec: "[|u \ v; k \ nat|]==> regular(v)\ (\n \ nat. +lemma residuals_lift_rec: "\u \ v; k \ nat\\ regular(v)\ (\n \ nat. lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))" apply (erule Scomp.induct, safe) apply (simp_all add: lift_rec_Var subst_Var lift_subst) done lemma residuals_subst_rec: - "u1 \ u2 ==> \v1 v2. v1 \ v2 \ regular(v2) \ regular(u2) \ + "u1 \ u2 \ \v1 v2. v1 \ v2 \ regular(v2) \ regular(u2) \ (\n \ nat. subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) = subst_rec(v1 |> v2, u1 |> u2,n))" apply (erule Scomp.induct, safe) @@ -135,29 +135,29 @@ lemma commutation [simp]: - "[|u1 \ u2; v1 \ v2; regular(u2); regular(v2)|] - ==> (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)" + "\u1 \ u2; v1 \ v2; regular(u2); regular(v2)\ + \ (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)" by (simp add: residuals_subst_rec) subsection\Residuals are comp and regular\ lemma residuals_preserve_comp [rule_format, simp]: - "u \ v ==> \w. u \ w \ v \ w \ regular(w) \ (u|>w) \ (v|>w)" + "u \ v \ \w. u \ w \ v \ w \ regular(w) \ (u|>w) \ (v|>w)" by (erule Scomp.induct, force+) lemma residuals_preserve_reg [rule_format, simp]: - "u \ v ==> regular(u) \ regular(v) \ regular(u|>v)" + "u \ v \ regular(u) \ regular(v) \ regular(u|>v)" apply (erule Scomp.induct, auto) done subsection\Preservation lemma\ -lemma union_preserve_comp: "u \ v ==> v \ (u \ v)" +lemma union_preserve_comp: "u \ v \ v \ (u \ v)" by (erule Scomp.induct, simp_all) lemma preservation [rule_format]: - "u \ v ==> regular(v) \ u|>v = (u \ v)|>v" + "u \ v \ regular(v) \ u|>v = (u \ v)|>v" apply (erule Scomp.induct, safe) apply (drule_tac [3] psi = "Fun (u) |> v = w" for u v w in asm_rl) apply (auto simp add: union_preserve_comp comp_sym_iff) @@ -175,7 +175,7 @@ w |> u = (w|>v) |> (u|>v))" by (erule Ssub.induct, force+) -lemma prism: "[|v \ u; regular(u); w \ v|] ==> w |> u = (w|>v) |> (u|>v)" +lemma prism: "\v \ u; regular(u); w \ v\ \ w |> u = (w|>v) |> (u|>v)" apply (rule prism_l) apply (rule_tac [4] comp_trans, auto) done @@ -183,7 +183,7 @@ subsection\Levy's Cube Lemma\ -lemma cube: "[|u \ v; regular(v); regular(u); w \ u|]==> +lemma cube: "\u \ v; regular(v); regular(u); w \ u\\ (w|>u) |> (v|>u) = (w|>v) |> (u|>v)" apply (subst preservation [of u], assumption, assumption) apply (subst preservation [of v], erule comp_sym, assumption) @@ -198,7 +198,7 @@ subsection\paving theorem\ -lemma paving: "[|w \ u; w \ v; regular(u); regular(v)|]==> +lemma paving: "\w \ u; w \ v; regular(u); regular(v)\\ \uv vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u) \ vu \ regular(vu) & (w|>v) \ uv \ regular(uv)" apply (subgoal_tac "u \ v") diff -r f2094906e491 -r e44d86131648 src/ZF/Resid/Substitution.thy --- a/src/ZF/Resid/Substitution.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Resid/Substitution.thy Tue Sep 27 16:51:35 2022 +0100 @@ -18,11 +18,11 @@ definition lift_rec :: "[i,i]=> i" where - "lift_rec(r,k) == lift_aux(r)`k" + "lift_rec(r,k) \ lift_aux(r)`k" abbreviation lift :: "i=>i" where - "lift(r) == lift_rec(r,0)" + "lift(r) \ lift_rec(r,0)" consts @@ -39,11 +39,11 @@ definition subst_rec :: "[i,i,i]=> i" (**NOTE THE ARGUMENT ORDER BELOW**) where - "subst_rec(u,r,k) == subst_aux(r)`u`k" + "subst_rec(u,r,k) \ subst_aux(r)`u`k" abbreviation subst :: "[i,i]=>i" (infixl \'/\ 70) where - "u/v == subst_rec(u,v,0)" + "u/v \ subst_rec(u,v,0)" @@ -51,18 +51,18 @@ (* Arithmetic extensions *) (* ------------------------------------------------------------------------- *) -lemma gt_not_eq: "p < n ==> n\p" +lemma gt_not_eq: "p < n \ n\p" by blast -lemma succ_pred [rule_format, simp]: "j \ nat ==> i < j \ succ(j #- 1) = j" +lemma succ_pred [rule_format, simp]: "j \ nat \ i < j \ succ(j #- 1) = j" by (induct_tac "j", auto) -lemma lt_pred: "[|succ(x) nat|] ==> x < n #- 1 " +lemma lt_pred: "\succ(x) nat\ \ x < n #- 1 " apply (rule succ_leE) apply (simp add: succ_pred) done -lemma gt_pred: "[|n < succ(x); p nat|] ==> n #- 1 < x " +lemma gt_pred: "\n < succ(x); p nat\ \ n #- 1 < x " apply (rule succ_leE) apply (simp add: succ_pred) done @@ -75,22 +75,22 @@ (* lift_rec equality rules *) (* ------------------------------------------------------------------------- *) lemma lift_rec_Var: - "n \ nat ==> lift_rec(Var(i),n) = (if i nat \ lift_rec(Var(i),n) = (if i nat; k\i|] ==> lift_rec(Var(i),k) = Var(succ(i))" + "\i \ nat; k\i\ \ lift_rec(Var(i),k) = Var(succ(i))" by (simp add: lift_rec_def le_in_nat) -lemma lift_rec_gt [simp]: "[| k \ nat; i lift_rec(Var(i),k) = Var(i)" +lemma lift_rec_gt [simp]: "\k \ nat; i \ lift_rec(Var(i),k) = Var(i)" by (simp add: lift_rec_def) lemma lift_rec_Fun [simp]: - "k \ nat ==> lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))" + "k \ nat \ lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))" by (simp add: lift_rec_def) lemma lift_rec_App [simp]: - "k \ nat ==> lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))" + "k \ nat \ lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))" by (simp add: lift_rec_def) @@ -99,43 +99,43 @@ (* ------------------------------------------------------------------------- *) lemma subst_Var: - "[|k \ nat; u \ redexes|] - ==> subst_rec(u,Var(i),k) = + "\k \ nat; u \ redexes\ + \ subst_rec(u,Var(i),k) = (if k nat; u \ redexes|] ==> subst_rec(u,Var(n),n) = u" + "\n \ nat; u \ redexes\ \ subst_rec(u,Var(n),n) = u" by (simp add: subst_rec_def) lemma subst_gt [simp]: - "[|u \ redexes; p \ nat; p subst_rec(u,Var(n),p) = Var(n #- 1)" + "\u \ redexes; p \ nat; p \ subst_rec(u,Var(n),p) = Var(n #- 1)" by (simp add: subst_rec_def) lemma subst_lt [simp]: - "[|u \ redexes; p \ nat; n subst_rec(u,Var(n),p) = Var(n)" + "\u \ redexes; p \ nat; n \ subst_rec(u,Var(n),p) = Var(n)" by (simp add: subst_rec_def gt_not_eq leI lt_nat_in_nat) lemma subst_Fun [simp]: - "[|p \ nat; u \ redexes|] - ==> subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) " + "\p \ nat; u \ redexes\ + \ subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) " by (simp add: subst_rec_def) lemma subst_App [simp]: - "[|p \ nat; u \ redexes|] - ==> subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))" + "\p \ nat; u \ redexes\ + \ subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))" by (simp add: subst_rec_def) lemma lift_rec_type [rule_format, simp]: - "u \ redexes ==> \k \ nat. lift_rec(u,k) \ redexes" + "u \ redexes \ \k \ nat. lift_rec(u,k) \ redexes" apply (erule redexes.induct) apply (simp_all add: lift_rec_Var lift_rec_Fun lift_rec_App) done lemma subst_type [rule_format, simp]: - "v \ redexes ==> \n \ nat. \u \ redexes. subst_rec(u,v,n) \ redexes" + "v \ redexes \ \n \ nat. \u \ redexes. subst_rec(u,v,n) \ redexes" apply (erule redexes.induct) apply (simp_all add: subst_Var lift_rec_type) done @@ -148,7 +148,7 @@ (*The i\nat is redundant*) lemma lift_lift_rec [rule_format]: "u \ redexes - ==> \n \ nat. \i \ nat. i\n \ + \ \n \ nat. \i \ nat. i\n \ (lift_rec(lift_rec(u,i),succ(n)) = lift_rec(lift_rec(u,n),i))" apply (erule redexes.induct, auto) apply (case_tac "n < i") @@ -157,15 +157,15 @@ done lemma lift_lift: - "[|u \ redexes; n \ nat|] - ==> lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))" + "\u \ redexes; n \ nat\ + \ lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))" by (simp add: lift_lift_rec) -lemma lt_not_m1_lt: "\m < n; n \ nat; m \ nat\\ ~ n #- 1 < m" +lemma lt_not_m1_lt: "\m < n; n \ nat; m \ nat\\ \ n #- 1 < m" by (erule natE, auto) lemma lift_rec_subst_rec [rule_format]: - "v \ redexes ==> + "v \ redexes \ \n \ nat. \m \ nat. \u \ redexes. n\m\ lift_rec(subst_rec(u,v,n),m) = subst_rec(lift_rec(u,m),lift_rec(v,succ(m)),n)" @@ -181,13 +181,13 @@ lemma lift_subst: - "[|v \ redexes; u \ redexes; n \ nat|] - ==> lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))" + "\v \ redexes; u \ redexes; n \ nat\ + \ lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))" by (simp add: lift_rec_subst_rec) lemma lift_rec_subst_rec_lt [rule_format]: - "v \ redexes ==> + "v \ redexes \ \n \ nat. \m \ nat. \u \ redexes. m\n\ lift_rec(subst_rec(u,v,n),m) = subst_rec(lift_rec(u,m),lift_rec(v,m),succ(n))" @@ -204,14 +204,14 @@ lemma subst_rec_lift_rec [rule_format]: - "u \ redexes ==> + "u \ redexes \ \n \ nat. \v \ redexes. subst_rec(v,lift_rec(u,n),n) = u" apply (erule redexes.induct, auto) apply (case_tac "n < na", auto) done lemma subst_rec_subst_rec [rule_format]: - "v \ redexes ==> + "v \ redexes \ \m \ nat. \n \ nat. \u \ redexes. \w \ redexes. m\n \ subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m) = subst_rec(w,subst_rec(u,v,m),n)" @@ -237,8 +237,8 @@ lemma substitution: - "[|v \ redexes; u \ redexes; w \ redexes; n \ nat|] - ==> subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)" + "\v \ redexes; u \ redexes; w \ redexes; n \ nat\ + \ subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)" by (simp add: subst_rec_subst_rec) @@ -249,21 +249,21 @@ lemma lift_rec_preserve_comp [rule_format, simp]: - "u \ v ==> \m \ nat. lift_rec(u,m) \ lift_rec(v,m)" + "u \ v \ \m \ nat. lift_rec(u,m) \ lift_rec(v,m)" by (erule Scomp.induct, simp_all add: comp_refl) lemma subst_rec_preserve_comp [rule_format, simp]: - "u2 \ v2 ==> \m \ nat. \u1 \ redexes. \v1 \ redexes. + "u2 \ v2 \ \m \ nat. \u1 \ redexes. \v1 \ redexes. u1 \ v1\ subst_rec(u1,u2,m) \ subst_rec(v1,v2,m)" by (erule Scomp.induct, simp_all add: subst_Var lift_rec_preserve_comp comp_refl) lemma lift_rec_preserve_reg [simp]: - "regular(u) ==> \m \ nat. regular(lift_rec(u,m))" + "regular(u) \ \m \ nat. regular(lift_rec(u,m))" by (erule Sreg.induct, simp_all add: lift_rec_Var) lemma subst_rec_preserve_reg [simp]: - "regular(v) ==> + "regular(v) \ \m \ nat. \u \ redexes. regular(u)\regular(subst_rec(u,v,m))" by (erule Sreg.induct, simp_all add: subst_Var lift_rec_preserve_reg) diff -r f2094906e491 -r e44d86131648 src/ZF/Sum.thy --- a/src/ZF/Sum.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Sum.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,20 +10,20 @@ text\And the "Part" primitive for simultaneous recursive type definitions\ definition sum :: "[i,i]=>i" (infixr \+\ 65) where - "A+B == {0}*A \ {1}*B" + "A+B \ {0}*A \ {1}*B" definition Inl :: "i=>i" where - "Inl(a) == <0,a>" + "Inl(a) \ <0,a>" definition Inr :: "i=>i" where - "Inr(b) == <1,b>" + "Inr(b) \ <1,b>" definition "case" :: "[i=>i, i=>i, i]=>i" where - "case(c,d) == (%. cond(y, d(z), c(z)))" + "case(c,d) \ (%. cond(y, d(z), c(z)))" (*operator for selecting out the various summands*) definition Part :: "[i,i=>i] => i" where - "Part(A,h) == {x \ A. \z. x = h(z)}" + "Part(A,h) \ {x \ A. \z. x = h(z)}" subsection\Rules for the \<^term>\Part\ Primitive\ @@ -34,14 +34,14 @@ done lemma Part_eqI [intro]: - "[| a \ A; a=h(b) |] ==> a \ Part(A,h)" + "\a \ A; a=h(b)\ \ a \ Part(A,h)" by (unfold Part_def, blast) lemmas PartI = refl [THEN [2] Part_eqI] lemma PartE [elim!]: - "[| a \ Part(A,h); !!z. [| a \ A; a=h(z) |] ==> P - |] ==> P" + "\a \ Part(A,h); \z. \a \ A; a=h(z)\ \ P +\ \ P" apply (unfold Part_def, blast) done @@ -60,19 +60,19 @@ (** Introduction rules for the injections **) -lemma InlI [intro!,simp,TC]: "a \ A ==> Inl(a) \ A+B" +lemma InlI [intro!,simp,TC]: "a \ A \ Inl(a) \ A+B" by (unfold sum_defs, blast) -lemma InrI [intro!,simp,TC]: "b \ B ==> Inr(b) \ A+B" +lemma InrI [intro!,simp,TC]: "b \ B \ Inr(b) \ A+B" by (unfold sum_defs, blast) (** Elimination rules **) lemma sumE [elim!]: - "[| u \ A+B; - !!x. [| x \ A; u=Inl(x) |] ==> P; - !!y. [| y \ B; u=Inr(y) |] ==> P - |] ==> P" + "\u \ A+B; + \x. \x \ A; u=Inl(x)\ \ P; + \y. \y \ B; u=Inr(y)\ \ P +\ \ P" by (unfold sum_defs, blast) (** Injection and freeness equivalences, for rewriting **) @@ -100,10 +100,10 @@ lemmas Inr_neq_Inl = Inr_Inl_iff [THEN iffD1, THEN FalseE, elim!] -lemma InlD: "Inl(a): A+B ==> a \ A" +lemma InlD: "Inl(a): A+B \ a \ A" by blast -lemma InrD: "Inr(b): A+B ==> b \ B" +lemma InrD: "Inr(b): A+B \ b \ B" by blast lemma sum_iff: "u \ A+B \ (\x. x \ A & u=Inl(x)) | (\y. y \ B & u=Inr(y))" @@ -134,26 +134,26 @@ by (simp add: sum_defs) lemma case_type [TC]: - "[| u \ A+B; - !!x. x \ A ==> c(x): C(Inl(x)); - !!y. y \ B ==> d(y): C(Inr(y)) - |] ==> case(c,d,u) \ C(u)" + "\u \ A+B; + \x. x \ A \ c(x): C(Inl(x)); + \y. y \ B \ d(y): C(Inr(y)) +\ \ case(c,d,u) \ C(u)" by auto -lemma expand_case: "u \ A+B ==> +lemma expand_case: "u \ A+B \ R(case(c,d,u)) \ ((\x\A. u = Inl(x) \ R(c(x))) & (\y\B. u = Inr(y) \ R(d(y))))" by auto lemma case_cong: - "[| z \ A+B; - !!x. x \ A ==> c(x)=c'(x); - !!y. y \ B ==> d(y)=d'(y) - |] ==> case(c,d,z) = case(c',d',z)" + "\z \ A+B; + \x. x \ A \ c(x)=c'(x); + \y. y \ B \ d(y)=d'(y) +\ \ case(c,d,z) = case(c',d',z)" by auto -lemma case_case: "z \ A+B ==> +lemma case_case: "z \ A+B \ case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = case(%x. c(c'(x)), %y. d(d'(y)), z)" by auto @@ -161,7 +161,7 @@ subsection\More Rules for \<^term>\Part(A,h)\\ -lemma Part_mono: "A<=B ==> Part(A,h)<=Part(B,h)" +lemma Part_mono: "A<=B \ Part(A,h)<=Part(B,h)" by blast lemma Part_Collect: "Part(Collect(A,P), h) = Collect(Part(A,h), P)" @@ -176,7 +176,7 @@ lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y \ B}" by blast -lemma PartD1: "a \ Part(A,h) ==> a \ A" +lemma PartD1: "a \ Part(A,h) \ a \ A" by (simp add: Part_def) lemma Part_id: "Part(A,%x. x) = A" @@ -185,7 +185,7 @@ lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y \ Part(B,h)}" by blast -lemma Part_sum_equality: "C \ A+B ==> Part(C,Inl) \ Part(C,Inr) = C" +lemma Part_sum_equality: "C \ A+B \ Part(C,Inl) \ Part(C,Inr) = C" by blast end diff -r f2094906e491 -r e44d86131648 src/ZF/Trancl.thy --- a/src/ZF/Trancl.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Trancl.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,44 +9,44 @@ definition refl :: "[i,i]=>o" where - "refl(A,r) == (\x\A. \ r)" + "refl(A,r) \ (\x\A. \ r)" definition irrefl :: "[i,i]=>o" where - "irrefl(A,r) == \x\A. \ r" + "irrefl(A,r) \ \x\A. \ r" definition sym :: "i=>o" where - "sym(r) == \x y. : r \ : r" + "sym(r) \ \x y. : r \ : r" definition asym :: "i=>o" where - "asym(r) == \x y. :r \ ~ :r" + "asym(r) \ \x y. :r \ \ :r" definition antisym :: "i=>o" where - "antisym(r) == \x y.:r \ :r \ x=y" + "antisym(r) \ \x y.:r \ :r \ x=y" definition trans :: "i=>o" where - "trans(r) == \x y z. : r \ : r \ : r" + "trans(r) \ \x y z. : r \ : r \ : r" definition trans_on :: "[i,i]=>o" (\trans[_]'(_')\) where - "trans[A](r) == \x\A. \y\A. \z\A. + "trans[A](r) \ \x\A. \y\A. \z\A. : r \ : r \ : r" definition rtrancl :: "i=>i" (\(_^*)\ [100] 100) (*refl/transitive closure*) where - "r^* == lfp(field(r)*field(r), %s. id(field(r)) \ (r O s))" + "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^*" + "r^+ \ r O r^*" definition equiv :: "[i,i]=>o" where - "equiv(A,r) == r \ A*A & refl(A,r) & sym(r) & trans(r)" + "equiv(A,r) \ r \ A*A & refl(A,r) & sym(r) & trans(r)" subsection\General properties of relations\ @@ -54,43 +54,43 @@ subsubsection\irreflexivity\ lemma irreflI: - "[| !!x. x \ A ==> \ r |] ==> irrefl(A,r)" + "\\x. x \ A \ \ r\ \ irrefl(A,r)" by (simp add: irrefl_def) -lemma irreflE: "[| irrefl(A,r); x \ A |] ==> \ r" +lemma irreflE: "\irrefl(A,r); x \ A\ \ \ r" by (simp add: irrefl_def) subsubsection\symmetry\ lemma symI: - "[| !!x y.: r ==> : r |] ==> sym(r)" + "\\x y.: r \ : r\ \ sym(r)" by (unfold sym_def, blast) -lemma symE: "[| sym(r); : r |] ==> : r" +lemma symE: "\sym(r); : r\ \ : r" by (unfold sym_def, blast) subsubsection\antisymmetry\ lemma antisymI: - "[| !!x y.[| : r; : r |] ==> x=y |] ==> antisym(r)" + "\\x y.\: r; : r\ \ x=y\ \ antisym(r)" by (simp add: antisym_def, blast) -lemma antisymE: "[| antisym(r); : r; : r |] ==> x=y" +lemma antisymE: "\antisym(r); : r; : r\ \ x=y" by (simp add: antisym_def, blast) subsubsection\transitivity\ -lemma transD: "[| trans(r); :r; :r |] ==> :r" +lemma transD: "\trans(r); :r; :r\ \ :r" by (unfold trans_def, blast) lemma trans_onD: - "[| trans[A](r); :r; :r; a \ A; b \ A; c \ A |] ==> :r" + "\trans[A](r); :r; :r; a \ A; b \ A; c \ A\ \ :r" by (unfold trans_on_def, blast) -lemma trans_imp_trans_on: "trans(r) ==> trans[A](r)" +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)" +lemma trans_on_imp_trans: "\trans[A](r); r \ A*A\ \ trans(r)" by (simp add: trans_on_def trans_def, blast) @@ -100,7 +100,7 @@ "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^*" +lemma rtrancl_mono: "r<=s \ r^* \ s^*" apply (unfold rtrancl_def) apply (rule lfp_mono) apply (rule rtrancl_bnd_mono)+ @@ -122,23 +122,23 @@ done (*Reflexivity of rtrancl*) -lemma rtrancl_refl: "[| a \ field(r) |] ==> \ r^*" +lemma rtrancl_refl: "\a \ field(r)\ \ \ r^*" apply (rule rtrancl_unfold [THEN ssubst]) apply (erule idI [THEN UnI1]) done (*Closure under composition with r *) -lemma rtrancl_into_rtrancl: "[| \ r^*; \ r |] ==> \ r^*" +lemma rtrancl_into_rtrancl: "\ \ r^*; \ r\ \ \ 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: " \ r ==> \ r^*" +lemma r_into_rtrancl: " \ r \ \ 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^*" +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)" @@ -148,20 +148,20 @@ (** standard induction rule **) lemma rtrancl_full_induct [case_names initial step, consumes 1]: - "[| \ r^*; - !!x. x \ field(r) ==> P(); - !!x y z.[| P(); : r^*; : r |] ==> P() |] - ==> P()" + "\ \ r^*; + \x. x \ field(r) \ P(); + \x y z.\P(); : r^*; : r\ \ P()\ + \ P()" 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]: - "[| \ r^*; + "\ \ r^*; P(a); - !!y z.[| \ r^*; \ r; P(y) |] ==> P(z) - |] ==> P(b)" + \y z.\ \ r^*; \ r; P(y)\ \ P(z) +\ \ P(b)" (*by induction on this formula*) apply (subgoal_tac "\y. = \ P (y) ") (*now solve first subgoal: this formula is sufficient*) @@ -170,7 +170,7 @@ apply (erule rtrancl_full_induct, blast+) done -(*transitivity of transitive closure!! -- by induction.*) +(*transitivity of transitive closure\-- by induction.*) lemma trans_rtrancl: "trans(r^*)" apply (unfold trans_def) apply (intro allI impI) @@ -182,9 +182,9 @@ (*elimination of rtrancl -- by induction on a special formula*) lemma rtranclE: - "[| \ r^*; (a=b) ==> P; - !!y.[| \ r^*; \ r |] ==> P |] - ==> P" + "\ \ r^*; (a=b) \ P; + \y.\ \ r^*; \ r\ \ P\ + \ P" apply (subgoal_tac "a = b | (\y. \ r^* & \ r) ") (*see HOL/trancl*) apply blast @@ -207,29 +207,29 @@ (** Conversions between trancl and rtrancl **) -lemma trancl_into_rtrancl: " \ r^+ ==> \ r^*" +lemma trancl_into_rtrancl: " \ r^+ \ \ r^*" apply (unfold trancl_def) apply (blast intro: rtrancl_into_rtrancl) done (*r^+ contains all pairs in r *) -lemma r_into_trancl: " \ r ==> \ r^+" +lemma r_into_trancl: " \ r \ \ r^+" apply (unfold 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^+" +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: "[| \ r^*; \ r |] ==> \ r^+" +lemma rtrancl_into_trancl1: "\ \ r^*; \ r\ \ \ r^+" by (unfold trancl_def, blast) (*intro rule from r and r^* *) lemma rtrancl_into_trancl2: - "[| \ r; \ r^* |] ==> \ r^+" + "\ \ r; \ r^*\ \ \ r^+" apply (erule rtrancl_induct) apply (erule r_into_trancl) apply (blast intro: r_into_trancl trancl_trans) @@ -237,10 +237,10 @@ (*Nice induction rule for trancl*) lemma trancl_induct [case_names initial step, induct set: trancl]: - "[| \ r^+; - !!y. [| \ r |] ==> P(y); - !!y z.[| \ r^+; \ r; P(y) |] ==> P(z) - |] ==> P(b)" + "\ \ r^+; + \y. \ \ r\ \ P(y); + \y z.\ \ r^+; \ r; P(y)\ \ P(z) +\ \ P(b)" apply (rule compEpair) apply (unfold trancl_def, assumption) (*by induction on this formula*) @@ -253,10 +253,10 @@ (*elimination of r^+ -- NOT an induction rule*) lemma tranclE: - "[| \ r^+; - \ r ==> P; - !!y.[| \ r^+; \ r |] ==> P - |] ==> P" + "\ \ r^+; + \ r \ P; + \y.\ \ r^+; \ r\ \ P +\ \ P" apply (subgoal_tac " \ r | (\y. \ r^+ & \ r) ") apply blast apply (rule compEpair) @@ -275,13 +275,13 @@ apply (blast dest: trancl_type [THEN subsetD]) done -lemma trancl_subset_times: "r \ A * A ==> r^+ \ A * A" +lemma trancl_subset_times: "r \ A * A \ r^+ \ A * A" by (insert trancl_type [of r], blast) -lemma trancl_mono: "r<=s ==> r^+ \ s^+" +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" +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) @@ -304,13 +304,13 @@ apply (blast intro: rtrancl_trans) done -lemma rtrancl_subset: "[| R \ S; S \ R^* |] ==> S^* = R^*" +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)^*" + "\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]) @@ -320,7 +320,7 @@ (** rtrancl **) -lemma rtrancl_converseD: ":converse(r)^* ==> :converse(r^*)" +lemma rtrancl_converseD: ":converse(r)^* \ :converse(r^*)" apply (rule converseI) apply (frule rtrancl_type [THEN subsetD]) apply (erule rtrancl_induct) @@ -328,7 +328,7 @@ apply (blast intro: r_into_rtrancl rtrancl_trans) done -lemma rtrancl_converseI: ":converse(r^*) ==> :converse(r)^*" +lemma rtrancl_converseI: ":converse(r^*) \ :converse(r)^*" apply (drule converseD) apply (frule rtrancl_type [THEN subsetD]) apply (erule rtrancl_induct) @@ -344,12 +344,12 @@ (** trancl **) -lemma trancl_converseD: ":converse(r)^+ ==> :converse(r^+)" +lemma trancl_converseD: ":converse(r)^+ \ :converse(r^+)" apply (erule trancl_induct) apply (auto intro: r_into_trancl trancl_trans) done -lemma trancl_converseI: ":converse(r^+) ==> :converse(r)^+" +lemma trancl_converseI: ":converse(r^+) \ :converse(r)^+" apply (drule converseD) apply (erule trancl_induct) apply (auto intro: r_into_trancl trancl_trans) @@ -362,9 +362,9 @@ done lemma converse_trancl_induct [case_names initial step, consumes 1]: -"[| :r^+; !!y. :r ==> P(y); - !!y z. [| \ r; \ r^+; P(z) |] ==> P(y) |] - ==> P(a)" +"\:r^+; \y. :r \ P(y); + \y z. \ \ r; \ r^+; P(z)\ \ P(y)\ + \ P(a)" apply (drule converseI) apply (simp (no_asm_use) add: trancl_converse [symmetric]) apply (erule trancl_induct) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/AllocBase.thy --- a/src/ZF/UNITY/AllocBase.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/AllocBase.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,7 +10,7 @@ abbreviation (input) tokbag :: i (* tokbags could be multisets...or any ordered type?*) where - "tokbag == nat" + "tokbag \ nat" axiomatization NbT :: i and (* Number of tokens in system *) @@ -42,15 +42,15 @@ definition all_distinct :: "i=>o" where - "all_distinct(l) == all_distinct0(l)=1" + "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" + "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" + "lift(x) \ %s. s`x" text\function to show that the set of variables is infinite\ consts @@ -66,7 +66,7 @@ definition nat_var_inj :: "i=>i" where - "nat_var_inj(n) == Var(nat_list_inj(n))" + "nat_var_inj(n) \ Var(nat_list_inj(n))" subsection\Various simple lemmas\ @@ -90,7 +90,7 @@ by (force simp add: INT_iff) lemma setsum_fun_mono [rule_format]: - "n\nat ==> + "n\nat \ (\i\nat. i f(i) $\ g(i)) \ setsum(f, n) $\ setsum(g,n)" apply (induct_tac "n", simp_all) @@ -102,17 +102,17 @@ apply (rule zadd_zle_mono, simp_all) done -lemma tokens_type [simp,TC]: "l\list(A) ==> tokens(l)\nat" +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). \prefix(A) + "xs\list(A) \ \ys\list(A). \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: "\prefix(A) ==> tokens(xs) \ tokens(ys)" +lemma tokens_mono: "\prefix(A) \ tokens(xs) \ tokens(ys)" apply (cut_tac prefix_type) apply (blast intro: tokens_mono_aux) done @@ -123,31 +123,31 @@ done lemma tokens_append [simp]: -"[| xs\list(A); ys\list(A) |] ==> tokens(xs@ys) = tokens(xs) #+ tokens(ys)" +"\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)" +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" + "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)" + "\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). \prefix(A) + "xs\list(A) \ \ys\list(A). \prefix(A) \ \MultLe(A, r)" apply (induct_tac "xs", simp_all, clarify) apply (frule_tac l = ys in bag_of_multiset) @@ -158,8 +158,8 @@ done lemma bag_of_mono [intro]: - "[| \prefix(A); xs\list(A); ys\list(A) |] - ==> \MultLe(A, r)" + "\\prefix(A); xs\list(A); ys\list(A)\ + \ \MultLe(A, r)" apply (blast intro: bag_of_mono_aux) done @@ -172,7 +172,7 @@ lemmas nat_into_Fin = eqpoll_refl [THEN [2] Fin_lemma] -lemma list_Int_length_Fin: "l \ list(A) ==> C \ length(l) \ Fin(length(l))" +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) @@ -182,7 +182,7 @@ lemma mem_Int_imp_lt_length: - "[|xs \ list(A); k \ C \ length(xs)|] ==> k < length(xs)" + "\xs \ list(A); k \ C \ length(xs)\ \ k < length(xs)" by (simp add: ltI) lemma Int_succ_right: @@ -191,8 +191,8 @@ lemma bag_of_sublist_lemma: - "[|C \ nat; x \ A; xs \ list(A)|] - ==> msetsum(\i. {#nth(i, xs @ [x])#}, C \ succ(length(xs)), A) = + "\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))" @@ -209,8 +209,8 @@ done lemma bag_of_sublist_lemma2: - "l\list(A) ==> - C \ nat ==> + "l\list(A) \ + C \ nat \ bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C \ length(l), A)" apply (erule list_append_induct) @@ -219,14 +219,14 @@ done -lemma nat_Int_length_eq: "l \ list(A) ==> nat \ length(l) = length(l)" +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) ==> + "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) @@ -234,7 +234,7 @@ done lemma bag_of_sublist_Un_Int: -"l\list(A) ==> +"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))") @@ -247,16 +247,16 @@ lemma bag_of_sublist_Un_disjoint: - "[| l\list(A); B \ C = 0 |] - ==> bag_of(sublist(l, B \ C)) = + "\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))) = + "\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) @@ -286,7 +286,7 @@ subsubsection\The function \<^term>\state_of\\ -lemma state_of_state: "s\state ==> state_of(s)=s" +lemma state_of_state: "s\state \ state_of(s)=s" by (unfold state_of_def, auto) declare state_of_state [simp] @@ -328,10 +328,10 @@ 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)" +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" +lemma length_nat_list_inj: "n\nat \ length(nat_list_inj(n)) = n" by (induct_tac "n", auto) lemma var_infinite_lemma: @@ -348,20 +348,20 @@ apply (rule var_infinite_lemma) done -lemma var_not_Finite: "~Finite(var)" +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" +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))" + "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 (subgoal_tac "\Finite (var-A) ") apply (drule not_Finite_imp_exist, auto) apply (cut_tac var_not_Finite) apply (erule swap) @@ -369,10 +369,10 @@ done lemma Inter_var_DiffD: - "[| b\\(RepFun(var-A, B)); Finite(A); x\var-A |] ==> b\B(x)" + "\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)) *) +(* \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!] @@ -382,15 +382,15 @@ 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)" + "\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)" + "\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 diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/AllocImpl.thy --- a/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,11 +10,11 @@ abbreviation NbR :: i (*number of consumed messages*) where - "NbR == Var([succ(2)])" + "NbR \ Var([succ(2)])" abbreviation available_tok :: i (*number of free tokens (T in paper)*) where - "available_tok == Var([succ(succ(2))])" + "available_tok \ Var([succ(succ(2))])" axiomatization where alloc_type_assumes [simp]: @@ -24,7 +24,7 @@ "default_val(NbR) = 0 & default_val(available_tok)=0" definition - "alloc_giv_act == + "alloc_giv_act \ { \ state*state. \k. k = length(s`giv) & t = s(giv := s`giv @ [nth(k, s`ask)], @@ -32,7 +32,7 @@ k < length(s`ask) & nth(k, s`ask) \ s`available_tok}" definition - "alloc_rel_act == + "alloc_rel_act \ { \ state*state. t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel), NbR := succ(s`NbR)) & @@ -41,7 +41,7 @@ definition (*The initial condition s`giv=[] is missing from the original definition: S. O. Ehmety *) - "alloc_prog == + "alloc_prog \ mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil}, {alloc_giv_act, alloc_rel_act}, \G \ preserves(lift(available_tok)) \ @@ -49,12 +49,12 @@ preserves(lift(giv)). Acts(G))" -lemma available_tok_value_type [simp,TC]: "s\state ==> s`available_tok \ nat" +lemma available_tok_value_type [simp,TC]: "s\state \ s`available_tok \ nat" apply (unfold state_def) apply (drule_tac a = available_tok in apply_type, auto) done -lemma NbR_value_type [simp,TC]: "s\state ==> s`NbR \ nat" +lemma NbR_value_type [simp,TC]: "s\state \ s`NbR \ nat" apply (unfold state_def) apply (drule_tac a = NbR in apply_type, auto) done @@ -141,8 +141,8 @@ done lemma giv_Bounded_lemma2: -"[| G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(rel)) |] - ==> alloc_prog \ G \ Stable({s\state. s`NbR \ length(s`rel)} \ +"\G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(rel))\ + \ alloc_prog \ G \ Stable({s\state. s`NbR \ length(s`rel)} \ {s\state. s`available_tok #+ tokens(s`giv) = NbT #+ tokens(take(s`NbR, s`rel))})" apply (cut_tac giv_Bounded_lamma1) @@ -214,8 +214,8 @@ subsubsection\First, we lead up to a proof of Lemma 49, page 28.\ lemma alloc_prog_transient_lemma: - "[|G \ program; k\nat|] - ==> alloc_prog \ G \ + "\G \ program; k\nat\ + \ alloc_prog \ G \ transient({s\state. k \ length(s`rel)} \ {s\state. succ(s`NbR) = k})" apply auto @@ -232,8 +232,8 @@ done lemma alloc_prog_rel_Stable_NbR_lemma: - "[| G \ program; alloc_prog ok G; k\nat |] - ==> alloc_prog \ G \ Stable({s\state . k \ succ(s ` NbR)})" + "\G \ program; alloc_prog ok G; k\nat\ + \ alloc_prog \ G \ Stable({s\state . k \ succ(s ` NbR)})" apply (auto intro!: stable_imp_Stable simp add: alloc_prog_ok_iff, safety, auto) apply (blast intro: le_trans leI) apply (drule_tac f = "lift (NbR)" and A = nat in preserves_imp_increasing) @@ -244,9 +244,9 @@ done lemma alloc_prog_NbR_LeadsTo_lemma: - "[| G \ program; alloc_prog ok G; - alloc_prog \ G \ Incr(lift(rel)); k\nat |] - ==> alloc_prog \ G \ + "\G \ program; alloc_prog ok G; + alloc_prog \ G \ Incr(lift(rel)); k\nat\ + \ alloc_prog \ G \ {s\state. k \ length(s`rel)} \ {s\state. succ(s`NbR) = k} \w {s\state. k \ s`NbR}" apply (subgoal_tac "alloc_prog \ G \ Stable ({s\state. k \ length (s`rel)})") @@ -264,9 +264,9 @@ done lemma alloc_prog_NbR_LeadsTo_lemma2 [rule_format]: - "[| G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(rel)); - k\nat; n \ nat; n < k |] - ==> alloc_prog \ G \ + "\G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(rel)); + k\nat; n \ nat; n < k\ + \ alloc_prog \ G \ {s\state . k \ length(s ` rel)} \ {s\state . s ` NbR = n} \w {x \ state. k \ length(x`rel)} \ (\m \ greater_than(n). {x \ state. x ` NbR=m})" @@ -292,15 +292,15 @@ apply (blast intro: lt_trans2) done -lemma Collect_vimage_eq: "u\nat ==> {. s \ A} -`` u = {s\A. f(s) < u}" +lemma Collect_vimage_eq: "u\nat \ {. s \ A} -`` u = {s\A. f(s) < u}" by (force simp add: lt_def) (* Lemma 49, page 28 *) lemma alloc_prog_NbR_LeadsTo_lemma3: - "[|G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(rel)); - k\nat|] - ==> alloc_prog \ G \ + "\G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(rel)); + k\nat\ + \ alloc_prog \ G \ {s\state. k \ length(s`rel)} \w {s\state. k \ s`NbR}" (* Proof by induction over the difference between k and n *) apply (rule_tac f = "\s\state. k #- s`NbR" in LessThan_induct) @@ -322,12 +322,12 @@ subsubsection\Towards proving lemma 50, page 29\ lemma alloc_prog_giv_Ensures_lemma: -"[| G \ program; k\nat; alloc_prog ok G; - alloc_prog \ G \ Incr(lift(ask)) |] ==> +"\G \ program; k\nat; alloc_prog ok G; + alloc_prog \ G \ Incr(lift(ask))\ \ alloc_prog \ G \ {s\state. nth(length(s`giv), s`ask) \ s`available_tok} \ {s\state. k < length(s`ask)} \ {s\state. length(s`giv)=k} - Ensures {s\state. ~ k {s\state. length(s`giv) \ k}" + Ensures {s\state. \ k {s\state. length(s`giv) \ k}" apply (rule EnsuresI, auto) apply (erule_tac [2] V = "G\u" for u in thin_rl) apply (rule_tac [2] act = alloc_giv_act in transientI) @@ -339,7 +339,7 @@ apply (rule_tac [2] ReplaceI) apply (rule_tac [2] x = "x (giv := x ` giv @ [nth (length(x`giv), x ` ask) ], available_tok := x ` available_tok #- nth (length(x`giv), x ` ask))" in exI) apply (auto intro!: state_update_type simp add: app_type) -apply (rule_tac A = "{s\state . nth (length(s ` giv), s ` ask) \ s ` available_tok} \ {s\state . k < length(s ` ask) } \ {s\state. length(s`giv) =k}" and A' = "{s\state . nth (length(s ` giv), s ` ask) \ s ` available_tok} \ {s\state. ~ k < length(s`ask) } \ {s\state . length(s ` giv) \ k}" in Constrains_weaken) +apply (rule_tac A = "{s\state . nth (length(s ` giv), s ` ask) \ s ` available_tok} \ {s\state . k < length(s ` ask) } \ {s\state. length(s`giv) =k}" and A' = "{s\state . nth (length(s ` giv), s ` ask) \ s ` available_tok} \ {s\state. \ k < length(s`ask) } \ {s\state . length(s ` giv) \ k}" in Constrains_weaken) apply (auto dest: ActsD simp add: Constrains_def constrains_def alloc_prog_def [THEN def_prg_Acts] alloc_prog_ok_iff) apply (subgoal_tac "length(xa ` giv @ [nth (length(xa ` giv), xa ` ask) ]) = length(xa ` giv) #+ 1") apply (rule_tac [2] trans) @@ -356,8 +356,8 @@ done lemma alloc_prog_giv_Stable_lemma: -"[| G \ program; alloc_prog ok G; k\nat |] - ==> alloc_prog \ G \ Stable({s\state . k \ length(s`giv)})" +"\G \ program; alloc_prog ok G; k\nat\ + \ alloc_prog \ G \ Stable({s\state . k \ length(s`giv)})" apply (auto intro!: stable_imp_Stable simp add: alloc_prog_ok_iff, safety) apply (auto intro: leI) apply (drule_tac f = "lift (giv)" and g = length in imp_preserves_comp) @@ -370,14 +370,14 @@ (* Lemma 50, page 29 *) lemma alloc_prog_giv_LeadsTo_lemma: -"[| G \ program; alloc_prog ok G; - alloc_prog \ G \ Incr(lift(ask)); k\nat |] - ==> alloc_prog \ G \ +"\G \ program; alloc_prog ok G; + alloc_prog \ G \ Incr(lift(ask)); k\nat\ + \ alloc_prog \ G \ {s\state. nth(length(s`giv), s`ask) \ s`available_tok} \ {s\state. k < length(s`ask)} \ {s\state. length(s`giv) = k} \w {s\state. k < length(s`giv)}" -apply (subgoal_tac "alloc_prog \ G \ {s\state. nth (length(s`giv), s`ask) \ s`available_tok} \ {s\state. k < length(s`ask) } \ {s\state. length(s`giv) = k} \w {s\state. ~ k {s\state. length(s`giv) \ k}") +apply (subgoal_tac "alloc_prog \ G \ {s\state. nth (length(s`giv), s`ask) \ s`available_tok} \ {s\state. k < length(s`ask) } \ {s\state. length(s`giv) = k} \w {s\state. \ k {s\state. length(s`giv) \ k}") prefer 2 apply (blast intro: alloc_prog_giv_Ensures_lemma [THEN LeadsTo_Basis]) apply (subgoal_tac "alloc_prog \ G \ Stable ({s\state. k < length(s`ask) }) ") apply (drule PSP_Stable, assumption) @@ -397,10 +397,10 @@ tokens given does not exceed the number returned, then the upper limit (\<^term>\NbT\) does not exceed the number currently available.\ lemma alloc_prog_Always_lemma: -"[| G \ program; alloc_prog ok G; +"\G \ program; alloc_prog ok G; alloc_prog \ G \ Incr(lift(ask)); - alloc_prog \ G \ Incr(lift(rel)) |] - ==> alloc_prog \ G \ + alloc_prog \ G \ Incr(lift(rel))\ + \ alloc_prog \ G \ Always({s\state. tokens(s`giv) \ tokens(take(s`NbR, s`rel)) \ NbT \ s`available_tok})" apply (subgoal_tac @@ -424,19 +424,19 @@ subsubsection\Main lemmas towards proving property (31)\ lemma LeadsTo_strength_R: - "[| F \ C \w B'; F \ A-C \w B; B'<=B |] ==> F \ A \w B" + "\F \ C \w B'; F \ A-C \w B; B'<=B\ \ F \ A \w B" by (blast intro: LeadsTo_weaken LeadsTo_Un_Un) lemma PSP_StableI: -"[| F \ Stable(C); F \ A - C \w B; - F \ A \ C \w B \ (state - C) |] ==> F \ A \w B" +"\F \ Stable(C); F \ A - C \w B; + F \ A \ C \w B \ (state - C)\ \ F \ A \w B" apply (rule_tac A = " (A-C) \ (A \ C)" in LeadsTo_weaken_L) prefer 2 apply blast apply (rule LeadsTo_Un, assumption) apply (blast intro: LeadsTo_weaken dest: PSP_Stable) done -lemma state_compl_eq [simp]: "state - {s\state. P(s)} = {s\state. ~P(s)}" +lemma state_compl_eq [simp]: "state - {s\state. P(s)} = {s\state. \P(s)}" by auto (*needed?*) @@ -461,7 +461,7 @@ will eventually recognize that they've been released.*) lemma (in alloc_progress) tokens_take_NbR_lemma: "k \ tokbag - ==> alloc_prog \ G \ + \ alloc_prog \ G \ {s\state. k \ tokens(s`rel)} \w {s\state. k \ tokens(take(s`NbR, s`rel))}" apply (rule single_LeadsTo_I, safe) @@ -480,7 +480,7 @@ \w *) lemma (in alloc_progress) tokens_take_NbR_lemma2: "k \ tokbag - ==> alloc_prog \ G \ + \ alloc_prog \ G \ {s\state. tokens(s`giv) = k} \w {s\state. k \ tokens(take(s`NbR, s`rel))}" apply (rule LeadsTo_Trans) @@ -492,8 +492,8 @@ (*Third step in proof of (31): by PSP with the fact that giv increases *) lemma (in alloc_progress) length_giv_disj: - "[| k \ tokbag; n \ nat |] - ==> alloc_prog \ G \ + "\k \ tokbag; n \ nat\ + \ alloc_prog \ G \ {s\state. length(s`giv) = n & tokens(s`giv) = k} \w {s\state. (length(s`giv) = n & tokens(s`giv) = k & @@ -512,8 +512,8 @@ (*Fourth step in proof of (31): we apply lemma (51) *) lemma (in alloc_progress) length_giv_disj2: - "[|k \ tokbag; n \ nat|] - ==> alloc_prog \ G \ + "\k \ tokbag; n \ nat\ + \ alloc_prog \ G \ {s\state. length(s`giv) = n & tokens(s`giv) = k} \w {s\state. (length(s`giv) = n & NbT \ s`available_tok) | @@ -526,7 +526,7 @@ k\nat *) lemma (in alloc_progress) length_giv_disj3: "n \ nat - ==> alloc_prog \ G \ + \ alloc_prog \ G \ {s\state. length(s`giv) = n} \w {s\state. (length(s`giv) = n & NbT \ s`available_tok) | @@ -540,8 +540,8 @@ (*Sixth step in proof of (31): from the fifth step, by PSP with the assumption that ask increases *) lemma (in alloc_progress) length_ask_giv: - "[|k \ nat; n < k|] - ==> alloc_prog \ G \ + "\k \ nat; n < k\ + \ alloc_prog \ G \ {s\state. length(s`ask) = k & length(s`giv) = n} \w {s\state. (NbT \ s`available_tok & length(s`giv) < length(s`ask) & @@ -563,8 +563,8 @@ (*Seventh step in proof of (31): no request (ask[k]) exceeds NbT *) lemma (in alloc_progress) length_ask_giv2: - "[|k \ nat; n < k|] - ==> alloc_prog \ G \ + "\k \ nat; n < k\ + \ alloc_prog \ G \ {s\state. length(s`ask) = k & length(s`giv) = n} \w {s\state. (nth(length(s`giv), s`ask) \ s`available_tok & @@ -580,8 +580,8 @@ (*Eighth step in proof of (31): by 50, we get |giv| > n. *) lemma (in alloc_progress) extend_giv: - "[| k \ nat; n < k|] - ==> alloc_prog \ G \ + "\k \ nat; n < k\ + \ alloc_prog \ G \ {s\state. length(s`ask) = k & length(s`giv) = n} \w {s\state. n < length(s`giv)}" apply (rule LeadsTo_Un_duplicate) @@ -597,7 +597,7 @@ we can't expect |ask| to remain fixed until |giv| increases.*) lemma (in alloc_progress) alloc_prog_ask_LeadsTo_giv: "k \ nat - ==> alloc_prog \ G \ + \ alloc_prog \ G \ {s\state. k \ length(s`ask)} \w {s\state. k \ length(s`giv)}" (* Proof by induction over the difference between k and n *) apply (rule_tac f = "\s\state. k #- length(s`giv)" in LessThan_induct) @@ -622,7 +622,7 @@ (*Final lemma: combine previous result with lemma (30)*) lemma (in alloc_progress) final: "h \ list(tokbag) - ==> alloc_prog \ G + \ alloc_prog \ G \ {s\state. \ prefix(tokbag)} \w {s\state. \ prefix(tokbag)}" apply (rule single_LeadsTo_I) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/ClientImpl.thy --- a/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 16:51:35 2022 +0100 @@ -7,10 +7,10 @@ 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 *) +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: @@ -21,11 +21,11 @@ default_val(rel) = Nil & default_val(tok) = 0" -(*Array indexing is translated to list indexing as A[n] == nth(n-1,A). *) +(*Array indexing is translated to list indexing as A[n] \ nth(n-1,A). *) definition (** Release some client_tokens **) - "client_rel_act == + "client_rel_act \ { \ state*state. \nrel \ nat. nrel = length(s`rel) & t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) & @@ -36,14 +36,14 @@ (** Including t=s suppresses fairness, allowing the non-trivial part of the action to be ignored **) definition - "client_tok_act == { \ state*state. t=s | + "client_tok_act \ { \ state*state. t=s | t = s(tok:=succ(s`tok mod NbT))}" definition - "client_ask_act == { \ state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" + "client_ask_act \ { \ state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" definition - "client_prog == + "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}, @@ -55,22 +55,22 @@ 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)" +lemma ask_value_type [simp,TC]: "s \ state \ s`ask \ list(nat)" apply (unfold state_def) apply (drule_tac a = ask in apply_type, auto) done -lemma giv_value_type [simp,TC]: "s \ state ==> s`giv \ list(nat)" +lemma giv_value_type [simp,TC]: "s \ state \ s`giv \ list(nat)" apply (unfold state_def) apply (drule_tac a = giv in apply_type, auto) done -lemma rel_value_type [simp,TC]: "s \ state ==> s`rel \ list(nat)" +lemma rel_value_type [simp,TC]: "s \ state \ s`rel \ list(nat)" apply (unfold state_def) apply (drule_tac a = rel in apply_type, auto) done -lemma tok_value_type [simp,TC]: "s \ state ==> s`tok \ nat" +lemma tok_value_type [simp,TC]: "s \ state \ s`tok \ nat" apply (unfold state_def) apply (drule_tac a = tok in apply_type, auto) done @@ -105,14 +105,14 @@ lemma preserves_lift_imp_stable: - "G \ preserves(lift(ff)) ==> G \ stable({s \ state. P(s`ff)})" + "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)})" + \ G \ stable({s \ state. \k, s`ff\ \ prefix(nat)})" by (erule preserves_lift_imp_stable) (*Safety property 1 \ ask, rel are increasing: (24) *) @@ -135,8 +135,8 @@ With no Substitution Axiom, we must prove the two invariants simultaneously. *) lemma ask_Bounded_lemma: -"[| client_prog ok G; G \ program |] - ==> client_prog \ G \ +"\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) @@ -166,22 +166,22 @@ 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)})" +"\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)})" + "\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 == { \ state*state. P(s, t)} ==> A={ \ state*state. P(s, t)}" + "A \ { \ state*state. P(s, t)} \ A={ \ state*state. P(s, t)}" by auto -lemma act_subset: "A={ \ state*state. P(s, t)} ==> A<=state*state" +lemma act_subset: "A={ \ state*state. P(s, t)} \ A<=state*state" by auto lemma transient_lemma: @@ -215,8 +215,8 @@ done lemma induct_lemma: -"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] - ==> client_prog \ G \ +"\client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program\ + \ client_prog \ G \ {s \ state. s`rel = k & \ strict_prefix(nat) & \ prefix(nat) & h pfixGe s`ask} \w {s \ state. \ strict_prefix(nat) @@ -244,8 +244,8 @@ done lemma rel_progress_lemma: -"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] - ==> client_prog \ G \ +"\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)}" @@ -273,8 +273,8 @@ done lemma progress_lemma: -"[| client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program |] - ==> client_prog \ G +"\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], diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Comp.thy --- a/src/ZF/UNITY/Comp.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Comp.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,40 +19,40 @@ definition component :: "[i,i]=>o" (infixl \component\ 65) where - "F component H == (\G. F \ G = H)" + "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" + "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)" + "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" + "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})}" + "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. " + "fun_pair(f, g) \ %x. " definition localize :: "[i=>i, i] => i" where - "localize(f,F) == mk_program(Init(F), Acts(F), + "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)" + "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) @@ -60,26 +60,26 @@ done lemma component_eq_subset: - "G \ program ==> (F component G) \ + "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" +lemma component_SKIP [simp]: "F \ program \ SKIP component F" apply (unfold 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" +lemma component_refl [simp]: "F \ program \ F component F" apply (unfold 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" +lemma SKIP_minimal: "F component SKIP \ programify(F) = SKIP" apply (rule program_equalityI) apply (simp_all add: component_eq_subset, blast) done @@ -93,14 +93,14 @@ apply blast done -lemma Join_absorb1: "F component G ==> F \ G = G" +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" +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)" + "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 @@ -110,32 +110,32 @@ apply (blast elim!: not_emptyE)+ done -lemma component_JOIN: "i \ I ==> F(i) component (\i \ I. (F(i)))" +lemma component_JOIN: "i \ I \ F(i) component (\i \ I. (F(i)))" apply (unfold component_def) apply (blast intro: JOIN_absorb) done -lemma component_trans: "[| F component G; G component H |] ==> F component H" +lemma component_trans: "\F component G; G component H\ \ F component H" apply (unfold component_def) apply (blast intro: Join_assoc [symmetric]) done lemma component_antisym: - "[| F \ program; G \ program |] ==>(F component G & G component F) \ F = G" + "\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 ==> + "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" + "\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 @@ -150,13 +150,13 @@ done lemma preservesI [rule_format]: - "\z. F \ stable({s \ state. f(s) = z}) ==> F \ preserves(f)" + "\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); \ act |] ==> f(s) = f(t)" + "\F \ preserves(f); act \ Acts(F); \ act\ \ f(s) = f(t)" apply (unfold preserves_def stable_def constrains_def) apply (subgoal_tac "s \ state & t \ state") prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) @@ -199,7 +199,7 @@ lemma preserves_type: "preserves(v)<=program" by (unfold preserves_def, auto) -lemma preserves_into_program [TC]: "F \ preserves(f) ==> F \ program" +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)" @@ -208,7 +208,7 @@ apply (drule_tac x = act in bspec, auto) done -lemma imp_preserves_comp: "F \ preserves(f) ==> F \ preserves(g comp f)" +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))})" @@ -219,18 +219,18 @@ done lemma preserves_imp_stable: - "F \ preserves(f) ==> F \ stable({s \ state. P(f(s))})" + "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)" + "\F \ preserves(f); \x \ state. f(x):A\ \ F \ Increasing.increasing(A, r, f)" apply (unfold Increasing.increasing_def) apply (auto intro: preserves_into_program) apply (rule_tac P = "%x. :r" in preserves_imp_stable, auto) done lemma preserves_id_subset_stable: - "st_set(A) ==> preserves(%x. x) \ stable(A)" + "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) @@ -238,7 +238,7 @@ done lemma preserves_id_imp_stable: - "[| F \ preserves(%x. x); st_set(A) |] ==> F \ stable(A)" + "\F \ preserves(%x. x); st_set(A)\ \ F \ stable(A)" by (blast intro: preserves_id_subset_stable [THEN subsetD]) (** Added by Sidi **) @@ -246,23 +246,23 @@ (* component_of is stronger than component *) lemma component_of_imp_component: -"F component_of H ==> F component H" +"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" +lemma component_of_refl [simp]: "F \ program \ F component_of F" apply (unfold component_of_def) apply (rule_tac x = SKIP in exI, auto) done -lemma component_of_SKIP [simp]: "F \ program ==>SKIP component_of F" +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" + "\F component_of G; G component_of H\ \ F component_of H" apply (unfold component_of_def) apply (blast intro: Join_assoc [symmetric]) done @@ -285,8 +285,8 @@ (** 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))})" + "\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 @@ -295,10 +295,10 @@ done lemma Increasing_preserves_Stable: - "[| F \ stable({s \ state. :r}); G \ preserves(f); + "\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})" + \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) @@ -315,15 +315,15 @@ (** Lemma used in AllocImpl **) lemma Constrains_UN_left: - "[| \x \ I. F \ A(x) Co B; F \ program |] ==> F:(\x \ I. A(x)) Co B" + "\\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))}); + "\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))})" + 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) 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 diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Constrains.thy Tue Sep 27 16:51:35 2022 +0100 @@ -18,10 +18,10 @@ "(init \ (\act\acts. field(act)))*list(\act\acts. field(act))" intros (*Initial trace is empty*) - Init: "s: init ==> \ traces(init,acts)" + Init: "s: init \ \ traces(init,acts)" - Acts: "[| act:acts; \ traces(init,acts); : act |] - ==> \ traces(init, acts)" + Acts: "\act:acts; \ traces(init,acts); : act\ + \ \ traces(init, acts)" type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1 @@ -31,30 +31,30 @@ domains "reachable(F)" \ "Init(F) \ (\act\Acts(F). field(act))" intros - Init: "s:Init(F) ==> s:reachable(F)" + Init: "s:Init(F) \ s:reachable(F)" - Acts: "[| act: Acts(F); s:reachable(F); : act |] - ==> 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}" + "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)" + "A Unless B \ (A-B) Co (A \ B)" definition Stable :: "i => i" where - "Stable(A) == A Co A" + "Stable(A) \ A Co A" definition (*Always is the weak form of "invariant"*) Always :: "i => i" where - "Always(A) == initially(A) \ Stable(A)" + "Always(A) \ initially(A) \ Stable(A)" (*** traces and reachable ***) @@ -80,7 +80,7 @@ declare state_Int_reachable [iff] lemma reachable_equiv_traces: -"F \ program ==> reachable(F)={s \ state. \evs. :traces(Init(F), Acts(F))}" +"F \ program \ reachable(F)={s \ state. \evs. :traces(Init(F), Acts(F))}" apply (rule equalityI, safe) apply (blast dest: reachable_type [THEN subsetD]) apply (erule_tac [2] traces.induct) @@ -91,8 +91,8 @@ 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))" +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 @@ -102,13 +102,13 @@ (*The set of all reachable states is an invariant...*) lemma invariant_reachable: - "F \ program ==> F \ invariant(reachable(F))" + "F \ program \ F \ invariant(reachable(F))" apply (unfold 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" +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) @@ -120,7 +120,7 @@ (*** Co ***) -lemma constrains_reachable_Int: "F \ B co B'==>F:(reachable(F) \ B) co (reachable(F) \ B')" +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) @@ -136,15 +136,15 @@ lemmas Constrains_def2 = Constrains_eq_constrains [THEN eq_reflection] -lemma constrains_imp_Constrains: "F \ A co A' ==> F \ A Co A'" +lemma constrains_imp_Constrains: "F \ A co A' \ F \ A Co A'" apply (unfold 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'" + "\\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 @@ -166,34 +166,34 @@ declare Constrains_state [iff] lemma Constrains_weaken_R: - "[| F \ A Co A'; A'<=B' |] ==> F \ A Co B'" + "\F \ A Co A'; A'<=B'\ \ F \ A Co B'" apply (unfold Constrains_def2) apply (blast intro: constrains_weaken_R) done lemma Constrains_weaken_L: - "[| F \ A Co A'; B<=A |] ==> F \ B Co A'" + "\F \ A Co A'; B<=A\ \ F \ B Co A'" apply (unfold 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'" + "\F \ A Co A'; B<=A; A'<=B'\ \ F \ B Co B'" apply (unfold 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')" + "\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))" + "\(\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) @@ -201,33 +201,33 @@ (** Intersection **) lemma Constrains_Int: - "[| F \ A Co A'; F \ B Co B'|]==> F:(A \ B) Co (A' \ B')" + "\F \ A Co A'; F \ B Co B'\\ F:(A \ B) Co (A' \ B')" apply (unfold 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))" + "\(\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'" +lemma Constrains_imp_subset: "F \ A Co A' \ reachable(F) \ A \ A'" apply (unfold Constrains_def) apply (blast dest: constrains_imp_subset) done lemma Constrains_trans: - "[| F \ A Co B; F \ B Co C |] ==> F \ A Co C" + "\F \ A Co B; F \ B Co C\ \ F \ A Co C" apply (unfold 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')" +"\F \ A Co (A' \ B); F \ B Co B'\ \ F \ A Co (A' \ B')" apply (unfold Constrains_def2) apply (simp (no_asm_use) add: Int_Un_distrib) apply (blast intro: constrains_cancel) @@ -237,13 +237,13 @@ (* Useful because there's no Stable_weaken. [Tanja Vos] *) lemma stable_imp_Stable: -"F \ stable(A) ==> F \ Stable(A)" +"F \ stable(A) \ F \ Stable(A)" apply (unfold stable_def Stable_def) apply (erule constrains_imp_Constrains) done -lemma Stable_eq: "[| F \ Stable(A); A = B |] ==> F \ Stable(B)" +lemma Stable_eq: "\F \ Stable(A); A = B\ \ F \ Stable(B)" by blast lemma Stable_eq_stable: @@ -251,53 +251,53 @@ apply (auto dest: constrainsD2 simp add: Stable_def stable_def Constrains_def2) done -lemma StableI: "F \ A Co A ==> F \ Stable(A)" +lemma StableI: "F \ A Co A \ F \ Stable(A)" by (unfold Stable_def, assumption) -lemma StableD: "F \ Stable(A) ==> F \ A Co A" +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')" + "\F \ Stable(A); F \ Stable(A')\ \ F \ Stable(A \ A')" apply (unfold Stable_def) apply (blast intro: Constrains_Un) done lemma Stable_Int: - "[| F \ Stable(A); F \ Stable(A') |] ==> F \ Stable (A \ A')" + "\F \ Stable(A); F \ Stable(A')\ \ F \ Stable (A \ A')" apply (unfold 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')" + "\F \ Stable(C); F \ A Co (C \ A')\ + \ F \ (C \ A) Co (C \ A')" apply (unfold 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')" + "\F \ Stable(C); F \ (C \ A) Co A'\ + \ F \ (C \ A) Co (C \ A')" apply (unfold 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))" + "\(\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))" + "\(\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))" +lemma Stable_reachable: "F \ program \F \ Stable (reachable(F))" apply (simp (no_asm_simp) add: Stable_eq_stable Int_absorb) done @@ -307,12 +307,12 @@ 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 + 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))" + "\\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]) @@ -321,8 +321,8 @@ (* 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))" + "\\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 @@ -338,20 +338,20 @@ (** Natural deduction rules for "Always A" **) lemma AlwaysI: -"[| Init(F)<=A; F \ Stable(A) |] ==> F \ Always(A)" +"\Init(F)<=A; F \ Stable(A)\ \ F \ Always(A)" apply (unfold Always_def initially_def) apply (frule Stable_type [THEN subsetD], auto) done -lemma AlwaysD: "F \ Always(A) ==> Init(F)<=A & F \ Stable(A)" +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" +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) @@ -359,7 +359,7 @@ done lemma invariant_imp_Always: - "F \ invariant(A) ==> F \ Always(A)" + "F \ invariant(A) \ F \ Always(A)" apply (unfold Always_def invariant_def Stable_def stable_def) apply (blast intro: constrains_imp_Constrains) done @@ -389,11 +389,11 @@ done declare Always_state_eq [simp] -lemma state_AlwaysI: "F \ program ==> F \ Always(state)" +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))" +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] @@ -401,30 +401,30 @@ dest: invariant_type [THEN subsetD])+ done -lemma Always_weaken: "[| F \ Always(A); A \ B |] ==> F \ Always(B)" +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')" +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')" +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'" +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') *) +(* \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'" +"\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) @@ -435,18 +435,18 @@ 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)))" +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)" +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))" +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 diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Distributor.thy --- a/src/ZF/UNITY/Distributor.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Distributor.thy Tue Sep 27 16:51:35 2022 +0100 @@ -14,7 +14,7 @@ definition distr_follows :: "[i, i, i, i =>i] =>i" where - "distr_follows(A, In, iIn, Out) == + "distr_follows(A, In, iIn, Out) \ (lift(In) IncreasingWrt prefix(A)/list(A)) \ (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \ Always({s \ state. \elt \ set_of_list(s`iIn). elt < Nclients}) @@ -27,13 +27,13 @@ definition distr_allowed_acts :: "[i=>i]=>i" where - "distr_allowed_acts(Out) == + "distr_allowed_acts(Out) \ {D \ program. AllowedActs(D) = cons(id(state), \G \ (\n\nat. preserves(lift(Out(n)))). Acts(G))}" definition distr_spec :: "[i, i, i, i =>i]=>i" where - "distr_spec(A, In, iIn, Out) == + "distr_spec(A, In, iIn, Out) \ distr_follows(A, In, iIn, Out) \ distr_allowed_acts(Out)" locale distr = @@ -54,19 +54,19 @@ and distr_spec: "D \ distr_spec(A, In, iIn, Out)" -lemma (in distr) In_value_type [simp,TC]: "s \ state ==> s`In \ list(A)" +lemma (in distr) In_value_type [simp,TC]: "s \ state \ s`In \ list(A)" apply (unfold state_def) apply (drule_tac a = In in apply_type, auto) done lemma (in distr) iIn_value_type [simp,TC]: - "s \ state ==> s`iIn \ list(nat)" + "s \ state \ s`iIn \ list(nat)" apply (unfold state_def) apply (drule_tac a = iIn in apply_type, auto) done lemma (in distr) Out_value_type [simp,TC]: - "s \ state ==> s`Out(n):list(A)" + "s \ state \ s`Out(n):list(A)" apply (unfold state_def) apply (drule_tac a = "Out (n)" in apply_type) apply auto @@ -78,7 +78,7 @@ done lemma (in distr) D_ok_iff: - "G \ program ==> + "G \ program \ D ok G \ ((\n \ nat. G \ preserves(lift(Out(n)))) & D \ Allowed(G))" apply (cut_tac distr_spec) apply (auto simp add: INT_iff distr_spec_def distr_allowed_acts_def @@ -101,10 +101,10 @@ done lemma (in distr) distr_bag_Follows_lemma: -"[| \n \ nat. G \ preserves(lift(Out(n))); +"\\n \ nat. G \ preserves(lift(Out(n))); D \ G \ Always({s \ state. - \elt \ set_of_list(s`iIn). elt < Nclients}) |] - ==> D \ G \ Always + \elt \ set_of_list(s`iIn). elt < Nclients})\ + \ D \ G \ Always ({s \ state. msetsum(%n. bag_of (sublist(s`In, {k \ nat. k < length(s`iIn) & nth(k, s`iIn)= n})), diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/FP.thy --- a/src/ZF/UNITY/FP.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/FP.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,11 +13,11 @@ definition FP_Orig :: "i=>i" where - "FP_Orig(F) == \({A \ Pow(state). \B. F \ stable(A \ B)})" + "FP_Orig(F) \ \({A \ Pow(state). \B. F \ stable(A \ B)})" definition FP :: "i=>i" where - "FP(F) == {s\state. F \ stable({s})}" + "FP(F) \ {s\state. F \ stable({s})}" lemma FP_Orig_type: "FP_Orig(F) \ state" @@ -36,18 +36,18 @@ apply (rule FP_type) done -lemma stable_FP_Orig_Int: "F \ program ==> F \ stable(FP_Orig(F) \ B)" +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)" + "\\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)" +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) @@ -56,10 +56,10 @@ apply (auto simp add: cons_absorb) done -lemma FP_subset_FP_Orig: "F \ program ==> FP(F) \ FP_Orig(F)" +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)" +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) @@ -67,17 +67,17 @@ apply (auto simp add: cons_absorb st_set_def) done -lemma FP_equivalence: "F \ program ==> FP(F) = FP_Orig(F)" +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)" + "\\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}})" + "\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 -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Follows.thy --- a/src/ZF/UNITY/Follows.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Follows.thy Tue Sep 27 16:51:35 2022 +0100 @@ -11,7 +11,7 @@ definition Follows :: "[i, i, i=>i, i=>i] => i" where - "Follows(A, r, f, g) == + "Follows(A, r, f, g) \ Increasing(A, r, g) Int Increasing(A, r,f) Int Always({s \ state. :r}) Int @@ -19,57 +19,57 @@ abbreviation Incr :: "[i=>i]=>i" where - "Incr(f) == Increasing(list(nat), prefix(nat), f)" + "Incr(f) \ Increasing(list(nat), prefix(nat), f)" abbreviation n_Incr :: "[i=>i]=>i" where - "n_Incr(f) == Increasing(nat, Le, f)" + "n_Incr(f) \ Increasing(nat, Le, f)" abbreviation s_Incr :: "[i=>i]=>i" where - "s_Incr(f) == Increasing(Pow(nat), SetLe(nat), f)" + "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)" + "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)" + "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)" + "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')" + "\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 |] ==> +"\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) 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. \ 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. \ 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})" + \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 @@ -77,8 +77,8 @@ (* 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 |] ==> +"\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})" @@ -97,18 +97,18 @@ done lemma imp_LeadsTo_comp: -"[| F:(\j \ A. {s \ state. \ r} \w {s \ state. \ r}); +"\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 |] ==> + \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); +"\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 |] ==> + \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) apply (rule single_LeadsTo_I, auto) @@ -128,10 +128,10 @@ done lemma imp_LeadsTo_comp_left: -"[| F \ Increasing(A, r, f); +"\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 |] ==> + \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) apply (rule single_LeadsTo_I, auto) @@ -152,12 +152,12 @@ (** This general result is used to prove Follows Un, munion, etc. **) lemma imp_LeadsTo_comp2: -"[| F \ Increasing(A, r, f1) \ Increasing(B, s, g); +"\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}" + \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) @@ -169,25 +169,25 @@ apply (blast dest: Increasing_type [THEN subsetD]) done -lemma Follows_into_program [TC]: "F \ Follows(A, r, f, g) ==> F \ program" +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 \ Follows(A, r, f, g)\ F \ program & (\a. a \ A) & (\x \ state. f(x):A & g(x):A)" apply (unfold 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)" + "\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)" +"\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) @@ -198,8 +198,8 @@ 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)" +"\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 @@ -208,9 +208,9 @@ (* 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)))" +"\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) @@ -232,8 +232,8 @@ lemma Follows_trans: - "[| F \ Follows(A, r, f, g); F \ Follows(A,r, g, h); - trans[A](r) |] ==> F \ Follows(A, r, f, h)" + "\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) @@ -246,35 +246,35 @@ (** Destruction rules for Follows **) lemma Follows_imp_Increasing_left: - "F \ Follows(A, r, f,g) ==> F \ Increasing(A, r, f)" + "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)" + "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})" + "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 \ 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)}" + "\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)}" + "\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)" +"\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) apply (simp add: INT_iff, auto) apply (rule_tac [3] C = "{s \ state. f(s)=g(s)}" @@ -291,8 +291,8 @@ 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)" +"\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) apply (simp add: INT_iff, auto) apply (rule_tac [3] C = "{s \ state. g (s) =h (s) }" @@ -322,27 +322,27 @@ 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))" + "\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))" + "\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)})" + "\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))" +"\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) **) @@ -351,12 +351,12 @@ by (unfold MultLe_def refl_def, auto) lemma MultLe_refl1 [simp]: - "[| multiset(M); mset_of(M)<=A |] ==> \ MultLe(A, r)" + "\multiset(M); mset_of(M)<=A\ \ \ MultLe(A, r)" apply (unfold MultLe_def id_def lam_def) apply (auto simp add: Mult_iff_multiset) done -lemma MultLe_refl2 [simp]: "M \ Mult(A) ==> \ MultLe(A, r)" +lemma MultLe_refl2 [simp]: "M \ Mult(A) \ \ MultLe(A, r)" by (unfold MultLe_def id_def lam_def, auto) @@ -371,14 +371,14 @@ done lemma MultLe_trans: - "[| \ MultLe(A,r); \ MultLe(A,r) |] ==> \ MultLe(A,r)" + "\ \ MultLe(A,r); \ MultLe(A,r)\ \ \ 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))" + "part_order(A, r) \ part_ord(A, r-id(A))" apply (unfold 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) @@ -389,7 +389,7 @@ done lemma antisym_MultLe [simp]: - "part_order(A, r) ==> antisym(MultLe(A,r))" + "part_order(A, r) \ antisym(MultLe(A,r))" apply (unfold MultLe_def antisym_def) apply (drule part_order_imp_part_ord, auto) apply (drule irrefl_on_multirel) @@ -399,13 +399,13 @@ done lemma part_order_MultLe [simp]: - "part_order(A, r) ==> part_order(Mult(A), MultLe(A, r))" + "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)" +"\multiset(M); mset_of(M)<= A\ \ <0, M> \ MultLe(A, r)" apply (unfold MultLe_def) apply (case_tac "M=0") apply (auto simp add: FiniteFun.intros) @@ -414,11 +414,11 @@ apply (auto simp add: Mult_iff_multiset) done -lemma empty_le_MultLe2 [simp]: "M \ Mult(A) ==> <0, M> \ MultLe(A, r)" +lemma empty_le_MultLe2 [simp]: "M \ Mult(A) \ <0, M> \ MultLe(A, r)" by (simp add: Mult_iff_multiset) lemma munion_mono: -"[| \ MultLe(A, r); \ MultLe(A, r) |] ==> +"\ \ MultLe(A, r); \ MultLe(A, r)\ \ \ MultLe(A, r)" apply (unfold MultLe_def) apply (auto intro: munion_multirel_mono1 munion_multirel_mono2 @@ -426,40 +426,40 @@ 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))" + "\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))" + "\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)}); 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)})" + \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))" +"\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)); +"\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), + 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) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/GenPrefix.thy --- a/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,7 +17,7 @@ definition (*really belongs in ZF/Trancl*) part_order :: "[i, i] => o" where - "part_order(A, r) == refl(A,r) & trans[A](r) & antisym(r)" + "part_order(A, r) \ refl(A,r) & trans[A](r) & antisym(r)" consts gen_prefix :: "[i, i] => i" @@ -30,46 +30,46 @@ intros Nil: "<[],[]>:gen_prefix(A, r)" - prepend: "[| :gen_prefix(A, r); :r; x \ A; y \ A |] - ==> : gen_prefix(A, r)" + prepend: "\:gen_prefix(A, r); :r; x \ A; y \ A\ + \ : gen_prefix(A, r)" - append: "[| :gen_prefix(A, r); zs:list(A) |] - ==> :gen_prefix(A, r)" + append: "\: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))" + "prefix(A) \ gen_prefix(A, id(A))" definition strict_prefix :: "i=>i" where - "strict_prefix(A) == prefix(A) - id(list(A))" + "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 == :gen_prefix(nat, Le)" + "xs pfixLe ys \ :gen_prefix(nat, Le)" abbreviation pfixGe :: "[i, i] => o" (infixl \pfixGe\ 50) where - "xs pfixGe ys == :gen_prefix(nat, Ge)" + "xs pfixGe ys \ :gen_prefix(nat, Ge)" lemma reflD: - "[| refl(A, r); x \ A |] ==> :r" + "\refl(A, r); x \ A\ \ :r" apply (unfold refl_def, auto) done (*** preliminary lemmas ***) -lemma Nil_gen_prefix: "xs \ list(A) ==> <[], xs> \ gen_prefix(A, r)" +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: " \ gen_prefix(A, r) ==> length(xs) \ length(ys)" +lemma gen_prefix_length_le: " \ 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] @@ -78,8 +78,8 @@ lemma Cons_gen_prefix_aux: - "[| \ gen_prefix(A, r) |] - ==> (\x xs. x \ A \ xs'= Cons(x,xs) \ + "\ \ gen_prefix(A, r)\ + \ (\x xs. x \ A \ xs'= Cons(x,xs) \ (\y ys. y \ A & ys' = Cons(y,ys) & :r & \ gen_prefix(A, r)))" apply (erule gen_prefix.induct) @@ -87,9 +87,9 @@ done lemma Cons_gen_prefixE: - "[| \ gen_prefix(A, r); - !!y ys. [|zs = Cons(y, ys); y \ A; x \ A; :r; - \ gen_prefix(A, r) |] ==> P |] ==> P" + "\ \ gen_prefix(A, r); + \y ys. \zs = Cons(y, ys); y \ A; x \ A; :r; + \ gen_prefix(A, r)\ \ P\ \ P" apply (frule gen_prefix.dom_subset [THEN subsetD], auto) apply (blast dest: Cons_gen_prefix_aux) done @@ -104,7 +104,7 @@ (** Monotonicity of gen_prefix **) -lemma gen_prefix_mono2: "r<=s ==> gen_prefix(A, r) \ gen_prefix(A, s)" +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) @@ -112,7 +112,7 @@ apply (auto intro: gen_prefix.append) done -lemma gen_prefix_mono1: "A<=B ==>gen_prefix(A, r) \ gen_prefix(B, r)" +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) @@ -126,7 +126,7 @@ 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)" +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) @@ -135,7 +135,7 @@ (*** gen_prefix order ***) (* reflexivity *) -lemma refl_gen_prefix: "refl(A, r) ==> refl(list(A), gen_prefix(A, r))" +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 @@ -144,7 +144,7 @@ (* Transitivity *) (* A lemma for proving gen_prefix_trans_comp *) -lemma append_gen_prefix [rule_format (no_asm)]: "xs \ list(A) ==> +lemma append_gen_prefix [rule_format (no_asm)]: "xs \ list(A) \ \zs. \ gen_prefix(A, r) \ : gen_prefix(A, r)" apply (erule list.induct) apply (auto dest: gen_prefix.dom_subset [THEN subsetD]) @@ -153,7 +153,7 @@ (* Lemma proving transitivity and more*) lemma gen_prefix_trans_comp [rule_format (no_asm)]: - ": gen_prefix(A, r) ==> + ": gen_prefix(A, r) \ (\z \ list(A). \ gen_prefix(A, s)\ \ gen_prefix(A, s O r))" apply (erule gen_prefix.induct) apply (auto elim: ConsE simp add: Nil_gen_prefix) @@ -162,10 +162,10 @@ apply (drule_tac xs = ys and r = s in append_gen_prefix, auto) done -lemma trans_comp_subset: "trans(r) ==> r O r \ r" +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))" +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) @@ -174,14 +174,14 @@ done lemma trans_on_gen_prefix: - "trans(r) ==> trans[list(A)](gen_prefix(A, r))" + "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: - "[| \ prefix(A); \ gen_prefix(A, r); r<=A*A |] - ==> \ gen_prefix(A, r)" + "\ \ prefix(A); \ gen_prefix(A, r); r<=A*A\ + \ \ gen_prefix(A, r)" apply (unfold prefix_def) apply (rule_tac P = "%r. \ gen_prefix (A, r) " in right_comp_id [THEN subst]) apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ @@ -189,8 +189,8 @@ lemma gen_prefix_prefix_trans: -"[| \ gen_prefix(A,r); \ prefix(A); r<=A*A |] - ==> \ gen_prefix(A, r)" +"\ \ gen_prefix(A,r); \ prefix(A); r<=A*A\ + \ \ gen_prefix(A, r)" apply (unfold prefix_def) apply (rule_tac P = "%r. \ gen_prefix (A, r) " in left_comp_id [THEN subst]) apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ @@ -198,10 +198,10 @@ (** Antisymmetry **) -lemma nat_le_lemma [rule_format]: "n \ nat ==> \b \ nat. n #+ b \ n \ b = 0" +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))" +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) @@ -225,12 +225,12 @@ (*** recursion equations ***) -lemma gen_prefix_Nil: "xs \ list(A) ==> \ gen_prefix(A,r) \ (xs = [])" +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) |] ==> + "\refl(A, r); xs \ list(A)\ \ : gen_prefix(A, r) \ \ gen_prefix(A, r)" apply (unfold refl_def) apply (induct_tac "xs") @@ -238,14 +238,14 @@ done declare same_gen_prefix_gen_prefix [simp] -lemma gen_prefix_Cons: "[| xs \ list(A); ys \ list(A); y \ A |] ==> +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 & :r & \ gen_prefix(A,r)))" apply (induct_tac "xs", auto) done -lemma gen_prefix_take_append: "[| refl(A,r); \ gen_prefix(A, r); zs \ list(A) |] - ==> \ gen_prefix(A, r)" +lemma gen_prefix_take_append: "\refl(A,r); \ 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) @@ -254,21 +254,21 @@ apply (simp_all (no_asm_simp) add: diff_is_0_iff [THEN iffD2] take_type) done -lemma gen_prefix_append_both: "[| refl(A, r); \ gen_prefix(A,r); - length(xs) = length(ys); zs \ list(A) |] - ==> \ gen_prefix(A, r)" +lemma gen_prefix_append_both: "\refl(A, r); \ 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" +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]: - "[| \ gen_prefix(A, r); refl(A, r) |] - ==> length(xs) < length(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) @@ -290,15 +290,15 @@ apply (auto elim: ConsE simp add: gen_prefix_append_both) done -lemma append_one_gen_prefix: "[| : gen_prefix(A, r); length(xs) < length(ys); refl(A, r) |] - ==> \ gen_prefix(A, r)" +lemma append_one_gen_prefix: "\: 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) ==> +lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \ list(A) \ \ys \ list(A). \i \ nat. i < length(xs) \ : gen_prefix(A, r) \ :r" apply (induct_tac "xs", simp, clarify) @@ -306,14 +306,14 @@ apply (erule natE, auto) done -lemma gen_prefix_imp_nth: "[| \ gen_prefix(A,r); i < length(xs)|] - ==> :r" +lemma gen_prefix_imp_nth: "\ \ 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) ==> +lemma nth_imp_gen_prefix [rule_format]: "xs \ list(A) \ \ys \ list(A). length(xs) \ length(ys) \ (\i. i < length(xs) \ :r) \ \ gen_prefix(A, r)" @@ -363,7 +363,7 @@ (* Monotonicity of "set" operator WRT prefix *) lemma set_of_list_prefix_mono: -" \ prefix(A) ==> set_of_list(xs) \ set_of_list(ys)" +" \ prefix(A) \ set_of_list(xs) \ set_of_list(ys)" apply (unfold prefix_def) apply (erule gen_prefix.induct) @@ -374,7 +374,7 @@ (** recursion equations **) -lemma Nil_prefix: "xs \ list(A) ==> <[],xs> \ prefix(A)" +lemma Nil_prefix: "xs \ list(A) \ <[],xs> \ prefix(A)" apply (unfold prefix_def) apply (simp (no_asm_simp) add: Nil_gen_prefix) @@ -398,7 +398,7 @@ declare Cons_prefix_Cons [iff] lemma same_prefix_prefix: -"xs \ list(A)==> \ prefix(A) \ ( \ prefix(A))" +"xs \ list(A)\ \ prefix(A) \ ( \ prefix(A))" apply (unfold prefix_def) apply (subgoal_tac "refl (A,id (A))") apply (simp (no_asm_simp)) @@ -406,21 +406,21 @@ done declare same_prefix_prefix [simp] -lemma same_prefix_prefix_Nil: "xs \ list(A) ==> \ prefix(A) \ ( \ prefix(A))" +lemma same_prefix_prefix_Nil: "xs \ list(A) \ \ prefix(A) \ ( \ prefix(A))" apply (rule_tac P = "%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: -"[| \ prefix(A); zs \ list(A) |] ==> \ prefix(A)" +"\ \ prefix(A); zs \ list(A)\ \ \ prefix(A)" apply (unfold prefix_def) apply (erule gen_prefix.append, assumption) done declare prefix_appendI [simp] lemma prefix_Cons: -"[| xs \ list(A); ys \ list(A); y \ A |] ==> +"\xs \ list(A); ys \ list(A); y \ A\ \ \ prefix(A) \ (xs=[] | (\zs. xs=Cons(y,zs) & \ prefix(A)))" apply (unfold prefix_def) @@ -428,8 +428,8 @@ done lemma append_one_prefix: - "[| \ prefix(A); length(xs) < length(ys) |] - ==> \ prefix(A)" + "\ \ prefix(A); length(xs) < length(ys)\ + \ \ prefix(A)" apply (unfold prefix_def) apply (subgoal_tac "refl (A, id (A))") apply (simp (no_asm_simp) add: append_one_gen_prefix) @@ -437,7 +437,7 @@ done lemma prefix_length_le: -" \ prefix(A) ==> length(xs) \ length(ys)" +" \ prefix(A) \ length(xs) \ length(ys)" apply (unfold prefix_def) apply (blast dest: gen_prefix_length_le) done @@ -454,7 +454,7 @@ done lemma strict_prefix_length_lt_aux: - " \ prefix(A) ==> xs\ys \ length(xs) < length(ys)" + " \ prefix(A) \ xs\ys \ length(xs) < length(ys)" apply (unfold prefix_def) apply (erule gen_prefix.induct, clarify) apply (subgoal_tac [!] "ys \ list(A) & xs \ list(A)") @@ -467,7 +467,7 @@ done lemma strict_prefix_length_lt: - ":strict_prefix(A) ==> length(xs) < length(ys)" + ":strict_prefix(A) \ length(xs) < length(ys)" apply (unfold strict_prefix_def) apply (rule strict_prefix_length_lt_aux [THEN mp]) apply (auto dest: prefix_type [THEN subsetD]) @@ -493,7 +493,7 @@ done lemma prefix_snoc: -"[|xs \ list(A); ys \ list(A); y \ A |] ==> +"\xs \ list(A); ys \ list(A); y \ A\ \ \ prefix(A) \ (xs = ys@[y] | \ prefix(A))" apply (simp (no_asm) add: prefix_iff) apply (rule iffI, clarify) @@ -503,7 +503,7 @@ done declare prefix_snoc [simp] -lemma prefix_append_iff [rule_format]: "zs \ list(A) ==> \xs \ list(A). \ys \ list(A). +lemma prefix_append_iff [rule_format]: "zs \ list(A) \ \xs \ list(A). \ys \ list(A). ( \ prefix(A)) \ ( \ prefix(A) | (\us. xs = ys@us & \ prefix(A)))" apply (erule list_append_induct, force, clarify) @@ -518,14 +518,14 @@ (*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) |] - ==> \ prefix(A) \ \ prefix(A) +lemma common_prefix_linear_lemma [rule_format]: "\zs \ list(A); xs \ list(A); ys \ list(A)\ + \ \ prefix(A) \ \ prefix(A) \ \ prefix(A) | \ prefix(A)" apply (erule list_append_induct, auto) done -lemma common_prefix_linear: "[| \ prefix(A); \ prefix(A) |] - ==> \ prefix(A) | \ prefix(A)" +lemma common_prefix_linear: "\ \ prefix(A); \ prefix(A)\ + \ \ prefix(A) | \ prefix(A)" apply (cut_tac prefix_type) apply (blast del: disjCI intro: common_prefix_linear_lemma) done @@ -565,19 +565,19 @@ by (unfold part_order_def, auto) declare part_order_Le [simp] -lemma pfixLe_refl: "x \ list(nat) ==> x pfixLe x" +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" +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" +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: -":prefix(nat)==> xs pfixLe ys" +":prefix(nat)\ xs pfixLe ys" apply (unfold prefix_def) apply (rule gen_prefix_mono [THEN subsetD], auto) @@ -599,25 +599,25 @@ done declare trans_Ge [iff] -lemma pfixGe_refl: "x \ list(nat) ==> x pfixGe x" +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" +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" +lemma pfixGe_antisym: "\x pfixGe y; y pfixGe x\ \ x = y" by (blast intro: antisym_gen_prefix [THEN antisymE]) lemma prefix_imp_pfixGe: - ":prefix(nat) ==> xs pfixGe ys" + ":prefix(nat) \ xs pfixGe ys" apply (unfold prefix_def Ge_def) apply (rule gen_prefix_mono [THEN subsetD], auto) done (* Added by Sidi \ prefix and take *) lemma prefix_imp_take: -" \ prefix(A) ==> xs = take(length(xs), ys)" +" \ prefix(A) \ xs = take(length(xs), ys)" apply (unfold prefix_def) apply (erule gen_prefix.induct) @@ -630,17 +630,17 @@ apply (simp_all (no_asm_simp) add: diff_is_0_iff) done -lemma prefix_length_equal: "[| \ prefix(A); length(xs)=length(ys)|] ==> xs = ys" +lemma prefix_length_equal: "\ \ 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: "[| \ prefix(A); length(ys) \ length(xs)|] ==> xs = ys" +lemma prefix_length_le_equal: "\ \ 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)" +lemma take_prefix [rule_format]: "xs \ list(A) \ \n \ nat. \ prefix(A)" apply (unfold prefix_def) apply (erule list.induct, simp, clarify) apply (erule natE, auto) @@ -654,21 +654,21 @@ apply (blast intro: take_prefix length_type) done -lemma prefix_imp_nth: "[| \ prefix(A); i < length(xs)|] ==> nth(i,xs) = nth(i,ys)" +lemma prefix_imp_nth: "\ \ 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)|] - ==> \ prefix(A)" + "\xs \ list(A); ys \ list(A); length(xs) \ length(ys); + \i. i < length(xs) \ nth(i, xs) = nth(i,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); - \ prefix(A); \ prefix(A)|] ==> \ prefix(A)" +lemma length_le_prefix_imp_prefix: "\length(xs) \ length(ys); + \ prefix(A); \ prefix(A)\ \ \ prefix(A)" apply (cut_tac A = A in prefix_type) apply (rule nth_imp_prefix, blast, blast) apply assumption diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Guar.thy --- a/src/ZF/UNITY/Guar.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Guar.thy Tue Sep 27 16:51:35 2022 +0100 @@ -25,7 +25,7 @@ (* To be moved to theory WFair???? *) -lemma leadsTo_Basis': "[| F \ A co A \ B; F \ transient(A); st_set(B) |] ==> F \ A \ B" +lemma leadsTo_Basis': "\F \ A co A \ B; F \ transient(A); st_set(B)\ \ F \ A \ B" apply (frule constrainsD2) apply (drule_tac B = "A-B" in constrains_weaken_L, blast) apply (drule_tac B = "A-B" in transient_strengthen, blast) @@ -38,63 +38,63 @@ definition ex_prop :: "i => o" where - "ex_prop(X) == X<=program & + "ex_prop(X) \ X<=program & (\F \ program. \G \ program. F ok G \ F \ X | G \ X \ (F \ G) \ X)" definition strict_ex_prop :: "i => o" where - "strict_ex_prop(X) == X<=program & + "strict_ex_prop(X) \ X<=program & (\F \ program. \G \ program. F ok G \ (F \ X | G \ X) \ (F \ G \ X))" definition uv_prop :: "i => o" where - "uv_prop(X) == X<=program & + "uv_prop(X) \ X<=program & (SKIP \ X & (\F \ program. \G \ program. F ok G \ F \ X & G \ X \ (F \ G) \ X))" definition strict_uv_prop :: "i => o" where - "strict_uv_prop(X) == X<=program & + "strict_uv_prop(X) \ X<=program & (SKIP \ X & (\F \ program. \G \ program. F ok G \(F \ X & G \ X) \ (F \ G \ X)))" definition guar :: "[i, i] => i" (infixl \guarantees\ 55) where (*higher than membership, lower than Co*) - "X guarantees Y == {F \ program. \G \ program. F ok G \ F \ G \ X \ F \ G \ Y}" + "X guarantees Y \ {F \ program. \G \ program. F ok G \ F \ G \ X \ F \ G \ Y}" definition (* Weakest guarantees *) wg :: "[i,i] => i" where - "wg(F,Y) == \({X \ Pow(program). F:(X guarantees Y)})" + "wg(F,Y) \ \({X \ Pow(program). F:(X guarantees Y)})" definition (* Weakest existential property stronger than X *) wx :: "i =>i" where - "wx(X) == \({Y \ Pow(program). Y<=X & ex_prop(Y)})" + "wx(X) \ \({Y \ Pow(program). Y<=X & ex_prop(Y)})" definition (*Ill-defined programs can arise through "\"*) welldef :: i where - "welldef == {F \ program. Init(F) \ 0}" + "welldef \ {F \ program. Init(F) \ 0}" definition refines :: "[i, i, i] => o" (\(3_ refines _ wrt _)\ [10,10,10] 10) where - "G refines F wrt X == + "G refines F wrt X \ \H \ program. (F ok H & G ok H & F \ H \ welldef \ X) \ (G \ H \ welldef \ X)" definition iso_refines :: "[i,i, i] => o" (\(3_ iso'_refines _ wrt _)\ [10,10,10] 10) where - "G iso_refines F wrt X == F \ welldef \ X \ G \ welldef \ X" + "G iso_refines F wrt X \ F \ welldef \ X \ G \ welldef \ X" (*** existential properties ***) -lemma ex_imp_subset_program: "ex_prop(X) ==> X\program" +lemma ex_imp_subset_program: "ex_prop(X) \ X\program" by (simp add: ex_prop_def) lemma ex1 [rule_format]: "GG \ Fin(program) - ==> ex_prop(X) \ GG \ X\0 \ OK(GG, (%G. G)) \(\G \ GG. G) \ X" + \ ex_prop(X) \ GG \ X\0 \ OK(GG, (%G. G)) \(\G \ GG. G) \ X" apply (unfold ex_prop_def) apply (erule Fin_induct) apply (simp_all add: OK_cons_iff) @@ -102,7 +102,7 @@ done lemma ex2 [rule_format]: -"X \ program ==> +"X \ program \ (\GG \ Fin(program). GG \ X \ 0 \ OK(GG,(%G. G))\(\G \ GG. G) \ X) \ ex_prop(X)" apply (unfold ex_prop_def, clarify) @@ -131,13 +131,13 @@ (*** universal properties ***) -lemma uv_imp_subset_program: "uv_prop(X)==> X\program" +lemma uv_imp_subset_program: "uv_prop(X)\ X\program" apply (unfold uv_prop_def) apply (simp (no_asm_simp)) done lemma uv1 [rule_format]: - "GG \ Fin(program) ==> + "GG \ Fin(program) \ (uv_prop(X)\ GG \ X & OK(GG, (%G. G)) \ (\G \ GG. G) \ X)" apply (unfold uv_prop_def) apply (erule Fin_induct) @@ -145,7 +145,7 @@ done lemma uv2 [rule_format]: - "X\program ==> + "X\program \ (\GG \ Fin(program). GG \ X & OK(GG,(%G. G)) \ (\G \ GG. G) \ X) \ uv_prop(X)" apply (unfold uv_prop_def, auto) @@ -166,14 +166,14 @@ (*** guarantees ***) lemma guaranteesI: - "[| (!!G. [| F ok G; F \ G \ X; G \ program |] ==> F \ G \ Y); - F \ program |] - ==> F \ X guarantees Y" + "\(\G. \F ok G; F \ G \ X; G \ program\ \ F \ G \ Y); + F \ program\ + \ F \ X guarantees Y" by (simp add: guar_def component_def) lemma guaranteesD: - "[| F \ X guarantees Y; F ok G; F \ G \ X; G \ program |] - ==> F \ G \ Y" + "\F \ X guarantees Y; F ok G; F \ G \ X; G \ program\ + \ F \ G \ Y" by (simp add: guar_def component_def) @@ -181,41 +181,41 @@ The major premise can no longer be F\H since we need to reason about G*) lemma component_guaranteesD: - "[| F \ X guarantees Y; F \ G = H; H \ X; F ok G; G \ program |] - ==> H \ Y" + "\F \ X guarantees Y; F \ G = H; H \ X; F ok G; G \ program\ + \ H \ Y" by (simp add: guar_def, blast) lemma guarantees_weaken: - "[| F \ X guarantees X'; Y \ X; X' \ Y' |] ==> F \ Y guarantees Y'" + "\F \ X guarantees X'; Y \ X; X' \ Y'\ \ F \ Y guarantees Y'" by (simp add: guar_def, auto) lemma subset_imp_guarantees_program: - "X \ Y ==> X guarantees Y = program" + "X \ Y \ X guarantees Y = program" by (unfold guar_def, blast) (*Equivalent to subset_imp_guarantees_UNIV but more intuitive*) lemma subset_imp_guarantees: - "[| X \ Y; F \ program |] ==> F \ X guarantees Y" + "\X \ Y; F \ program\ \ F \ X guarantees Y" by (unfold guar_def, blast) -lemma component_of_Join1: "F ok G ==> F component_of (F \ G)" +lemma component_of_Join1: "F ok G \ F component_of (F \ G)" by (unfold component_of_def, blast) -lemma component_of_Join2: "F ok G ==> G component_of (F \ G)" +lemma component_of_Join2: "F ok G \ G component_of (F \ G)" apply (subst Join_commute) apply (blast intro: ok_sym component_of_Join1) done (*Remark at end of section 4.1 *) lemma ex_prop_imp: - "ex_prop(Y) ==> (Y = (program guarantees Y))" + "ex_prop(Y) \ (Y = (program guarantees Y))" apply (simp (no_asm_use) add: ex_prop_equiv guar_def component_of_def) apply clarify apply (rule equalityI, blast, safe) apply (drule_tac x = x in bspec, assumption, force) done -lemma guarantees_imp: "(Y = program guarantees Y) ==> ex_prop(Y)" +lemma guarantees_imp: "(Y = program guarantees Y) \ ex_prop(Y)" apply (unfold guar_def) apply (simp (no_asm_simp) add: ex_prop_equiv) apply safe @@ -230,7 +230,7 @@ (** Distributive laws. Re-orient to perform miniscoping **) lemma guarantees_UN_left: - "i \ I ==>(\i \ I. X(i)) guarantees Y = (\i \ I. X(i) guarantees Y)" + "i \ I \(\i \ I. X(i)) guarantees Y = (\i \ I. X(i) guarantees Y)" apply (unfold guar_def) apply (rule equalityI, safe) prefer 2 apply force @@ -244,7 +244,7 @@ done lemma guarantees_INT_right: - "i \ I ==> X guarantees (\i \ I. Y(i)) = (\i \ I. X guarantees Y(i))" + "i \ I \ X guarantees (\i \ I. Y(i)) = (\i \ I. X guarantees Y(i))" apply (unfold guar_def) apply (rule equalityI, safe, blast+) done @@ -254,12 +254,12 @@ by (unfold guar_def, blast) lemma guarantees_Int_right_I: - "[| F \ Z guarantees X; F \ Z guarantees Y |] - ==> F \ Z guarantees (X \ Y)" + "\F \ Z guarantees X; F \ Z guarantees Y\ + \ F \ Z guarantees (X \ Y)" by (simp (no_asm_simp) add: guarantees_Int_right) lemma guarantees_INT_right_iff: - "i \ I==> (F \ X guarantees (\i \ I. Y(i))) \ + "i \ I\ (F \ X guarantees (\i \ I. Y(i))) \ (\i \ I. F \ X guarantees Y(i))" by (simp add: guarantees_INT_right INT_iff, blast) @@ -276,36 +276,36 @@ **) lemma combining1: - "[| F \ V guarantees X; F \ (X \ Y) guarantees Z |] - ==> F \ (V \ Y) guarantees Z" + "\F \ V guarantees X; F \ (X \ Y) guarantees Z\ + \ F \ (V \ Y) guarantees Z" by (unfold guar_def, blast) lemma combining2: - "[| F \ V guarantees (X \ Y); F \ Y guarantees Z |] - ==> F \ V guarantees (X \ Z)" + "\F \ V guarantees (X \ Y); F \ Y guarantees Z\ + \ F \ V guarantees (X \ Z)" by (unfold guar_def, blast) (** The following two follow Chandy-Sanders, but the use of object-quantifiers does not suit Isabelle... **) -(*Premise should be (!!i. i \ I ==> F \ X guarantees Y i) *) +(*Premise should be (\i. i \ I \ F \ X guarantees Y i) *) lemma all_guarantees: - "[| \i \ I. F \ X guarantees Y(i); i \ I |] - ==> F \ X guarantees (\i \ I. Y(i))" + "\\i \ I. F \ X guarantees Y(i); i \ I\ + \ F \ X guarantees (\i \ I. Y(i))" by (unfold guar_def, blast) -(*Premises should be [| F \ X guarantees Y i; i \ I |] *) +(*Premises should be \F \ X guarantees Y i; i \ I\ *) lemma ex_guarantees: - "\i \ I. F \ X guarantees Y(i) ==> F \ X guarantees (\i \ I. Y(i))" + "\i \ I. F \ X guarantees Y(i) \ F \ X guarantees (\i \ I. Y(i))" by (unfold guar_def, blast) (*** Additional guarantees laws, by lcp ***) lemma guarantees_Join_Int: - "[| F \ U guarantees V; G \ X guarantees Y; F ok G |] - ==> F \ G: (U \ X) guarantees (V \ Y)" + "\F \ U guarantees V; G \ X guarantees Y; F ok G\ + \ F \ G: (U \ X) guarantees (V \ Y)" apply (unfold guar_def) apply (simp (no_asm)) @@ -317,8 +317,8 @@ done lemma guarantees_Join_Un: - "[| F \ U guarantees V; G \ X guarantees Y; F ok G |] - ==> F \ G: (U \ X) guarantees (V \ Y)" + "\F \ U guarantees V; G \ X guarantees Y; F ok G\ + \ F \ G: (U \ X) guarantees (V \ Y)" apply (unfold guar_def) apply (simp (no_asm)) apply safe @@ -332,8 +332,8 @@ done lemma guarantees_JOIN_INT: - "[| \i \ I. F(i) \ X(i) guarantees Y(i); OK(I,F); i \ I |] - ==> (\i \ I. F(i)) \ (\i \ I. X(i)) guarantees (\i \ I. Y(i))" + "\\i \ I. F(i) \ X(i) guarantees Y(i); OK(I,F); i \ I\ + \ (\i \ I. F(i)) \ (\i \ I. X(i)) guarantees (\i \ I. Y(i))" apply (unfold guar_def, safe) prefer 2 apply blast apply (drule_tac x = xa in bspec) @@ -343,8 +343,8 @@ done lemma guarantees_JOIN_UN: - "[| \i \ I. F(i) \ X(i) guarantees Y(i); OK(I,F) |] - ==> JOIN(I,F) \ (\i \ I. X(i)) guarantees (\i \ I. Y(i))" + "\\i \ I. F(i) \ X(i) guarantees Y(i); OK(I,F)\ + \ JOIN(I,F) \ (\i \ I. X(i)) guarantees (\i \ I. Y(i))" apply (unfold guar_def, auto) apply (drule_tac x = y in bspec, simp_all, safe) apply (rename_tac G y) @@ -355,20 +355,20 @@ (*** guarantees laws for breaking down the program, by lcp ***) lemma guarantees_Join_I1: - "[| F \ X guarantees Y; F ok G |] ==> F \ G \ X guarantees Y" + "\F \ X guarantees Y; F ok G\ \ F \ G \ X guarantees Y" apply (simp add: guar_def, safe) apply (simp add: Join_assoc) done lemma guarantees_Join_I2: - "[| G \ X guarantees Y; F ok G |] ==> F \ G \ X guarantees Y" + "\G \ X guarantees Y; F ok G\ \ F \ G \ X guarantees Y" apply (simp add: Join_commute [of _ G] ok_commute [of _ G]) apply (blast intro: guarantees_Join_I1) done lemma guarantees_JOIN_I: - "[| i \ I; F(i) \ X guarantees Y; OK(I,F) |] - ==> (\i \ I. F(i)) \ X guarantees Y" + "\i \ I; F(i) \ X guarantees Y; OK(I,F)\ + \ (\i \ I. F(i)) \ X guarantees Y" apply (unfold guar_def, safe) apply (drule_tac x = "JOIN (I-{i},F) \ G" in bspec) apply (simp (no_asm)) @@ -377,10 +377,10 @@ (*** well-definedness ***) -lemma Join_welldef_D1: "F \ G \ welldef ==> programify(F) \ welldef" +lemma Join_welldef_D1: "F \ G \ welldef \ programify(F) \ welldef" by (unfold welldef_def, auto) -lemma Join_welldef_D2: "F \ G \ welldef ==> programify(G) \ welldef" +lemma Join_welldef_D2: "F \ G \ welldef \ programify(G) \ welldef" by (unfold welldef_def, auto) (*** refinement ***) @@ -396,7 +396,7 @@ lemma guarantees_type: "X guarantees Y \ program" by (unfold guar_def, auto) -lemma wgD2: "G \ wg(F, X) ==> G \ program & F \ program" +lemma wgD2: "G \ wg(F, X) \ G \ program & F \ program" apply (unfold wg_def, auto) apply (blast dest: guarantees_type [THEN subsetD]) done @@ -407,10 +407,10 @@ by (unfold guar_def component_of_def, force) lemma wg_weakest: - "!!X. [| F \ (X guarantees Y); X \ program |] ==> X \ wg(F,Y)" + "\X. \F \ (X guarantees Y); X \ program\ \ X \ wg(F,Y)" by (unfold wg_def, auto) -lemma wg_guarantees: "F \ program ==> F \ wg(F,Y) guarantees Y" +lemma wg_guarantees: "F \ program \ F \ wg(F,Y) guarantees Y" by (unfold wg_def guar_def, blast) lemma wg_equiv: @@ -423,7 +423,7 @@ done lemma component_of_wg: - "F component_of H ==> H \ wg(F,X) \ (H \ X & F \ program & H \ program)" + "F component_of H \ H \ wg(F,X) \ (H \ X & F \ program & H \ program)" by (simp (no_asm_simp) add: wg_equiv) lemma wg_finite [rule_format]: @@ -439,7 +439,7 @@ done lemma wg_ex_prop: - "ex_prop(X) ==> (F \ X) \ (\H \ program. H \ wg(F,X) & F \ program)" + "ex_prop(X) \ (F \ X) \ (\H \ program. H \ wg(F,X) & F \ program)" apply (simp (no_asm_use) add: ex_prop_equiv wg_equiv) apply blast done @@ -510,7 +510,7 @@ Reasoning About Program composition paper *) lemma stable_guarantees_Always: - "[| Init(F) \ A; F \ program |] ==> F \ stable(A) guarantees Always(A)" + "\Init(F) \ A; F \ program\ \ F \ stable(A) guarantees Always(A)" apply (rule guaranteesI) prefer 2 apply assumption apply (simp (no_asm) add: Join_commute) @@ -520,8 +520,8 @@ done lemma constrains_guarantees_leadsTo: - "[| F \ transient(A); st_set(B) |] - ==> F: (A co A \ B) guarantees (A \ (B-A))" + "\F \ transient(A); st_set(B)\ + \ F: (A co A \ B) guarantees (A \ (B-A))" apply (rule guaranteesI) prefer 2 apply (blast dest: transient_type [THEN subsetD]) apply (rule leadsTo_Basis') diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Increasing.thy --- a/src/ZF/UNITY/Increasing.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Increasing.thy Tue Sep 27 16:51:35 2022 +0100 @@ -12,19 +12,19 @@ definition increasing :: "[i, i, i=>i] => i" (\increasing[_]'(_, _')\) where - "increasing[A](r, f) == + "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) == + "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)" + "f IncreasingWrt r/A \ Increasing[A](r,f)" (** increasing **) @@ -32,15 +32,15 @@ lemma increasing_type: "increasing[A](r, f) \ program" by (unfold increasing_def, blast) -lemma increasing_into_program: "F \ increasing[A](r, f) ==> F \ program" +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})" +"\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)" +"F \ increasing[A](r,f) \ F \ program & (\a. a \ A) & (\s \ state. f(s):A)" apply (unfold increasing_def) apply (subgoal_tac "\x. x \ state") apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state) @@ -54,7 +54,7 @@ done lemma subset_increasing_comp: -"[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==> +"\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) @@ -74,8 +74,8 @@ 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)" + "\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: @@ -92,7 +92,7 @@ (** Increasing **) lemma increasing_imp_Increasing: - "F \ increasing[A](r, f) ==> F \ Increasing[A](r, f)" + "F \ increasing[A](r, f) \ F \ Increasing[A](r, f)" apply (unfold increasing_def Increasing_def) apply (auto intro: stable_imp_Stable) @@ -101,15 +101,15 @@ lemma Increasing_type: "Increasing[A](r, f) \ program" by (unfold Increasing_def, auto) -lemma Increasing_into_program: "F \ Increasing[A](r, f) ==> F \ program" +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})" +"\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)" +"F \ Increasing[A](r, f) \ F \ program & (\a. a \ A) & (\s \ state. f(s):A)" apply (unfold Increasing_def) apply (subgoal_tac "\x. x \ state") apply (auto intro: st0_in_state) @@ -122,7 +122,7 @@ done lemma subset_Increasing_comp: -"[| mono1(A, r, B, s, g); refl(A, r); trans[B](s) |] ==> +"\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) @@ -144,8 +144,8 @@ 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)" + "\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 @@ -161,9 +161,9 @@ (** 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)))" +"\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 @@ -193,8 +193,8 @@ 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[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) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Merge.thy --- a/src/ZF/UNITY/Merge.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Merge.thy Tue Sep 27 16:51:35 2022 +0100 @@ -14,27 +14,27 @@ definition (*spec (10)*) merge_increasing :: "[i, i, i] =>i" where - "merge_increasing(A, Out, iOut) == program guarantees + "merge_increasing(A, Out, iOut) \ program guarantees (lift(Out) IncreasingWrt prefix(A)/list(A)) Int (lift(iOut) IncreasingWrt prefix(nat)/list(nat))" definition (*spec (11)*) merge_eq_Out :: "[i, i] =>i" where - "merge_eq_Out(Out, iOut) == program guarantees + "merge_eq_Out(Out, iOut) \ program guarantees Always({s \ state. length(s`Out) = length(s`iOut)})" definition (*spec (12)*) merge_bounded :: "i=>i" where - "merge_bounded(iOut) == program guarantees + "merge_bounded(iOut) \ program guarantees Always({s \ state. \elt \ set_of_list(s`iOut). elti, i, i] =>i" where - "merge_follows(A, In, Out, iOut) == + "merge_follows(A, In, Out, iOut) \ (\n \ Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A)) guarantees (\n \ Nclients. @@ -45,19 +45,19 @@ definition (*spec: preserves part*) merge_preserves :: "[i=>i] =>i" where - "merge_preserves(In) == \n \ nat. preserves(lift(In(n)))" + "merge_preserves(In) \ \n \ nat. preserves(lift(In(n)))" definition (* environmental constraints*) merge_allowed_acts :: "[i, i] =>i" where - "merge_allowed_acts(Out, iOut) == + "merge_allowed_acts(Out, iOut) \ {F \ program. AllowedActs(F) = cons(id(state), (\G \ preserves(lift(Out)) \ preserves(lift(iOut)). Acts(G)))}" definition merge_spec :: "[i, i =>i, i, i]=>i" where - "merge_spec(A, In, Out, iOut) == + "merge_spec(A, In, Out, iOut) \ merge_increasing(A, Out, iOut) \ merge_eq_Out(Out, iOut) \ merge_bounded(iOut) \ merge_follows(A, In, Out, iOut) \ merge_allowed_acts(Out, iOut) \ merge_preserves(In)" @@ -84,18 +84,18 @@ and merge_spec: "M \ merge_spec(A, In, Out, iOut)" -lemma (in merge) In_value_type [TC,simp]: "s \ state ==> s`In(n) \ list(A)" +lemma (in merge) In_value_type [TC,simp]: "s \ state \ s`In(n) \ list(A)" apply (unfold state_def) apply (drule_tac a = "In (n)" in apply_type) apply auto done -lemma (in merge) Out_value_type [TC,simp]: "s \ state ==> s`Out \ list(A)" +lemma (in merge) Out_value_type [TC,simp]: "s \ state \ s`Out \ list(A)" apply (unfold state_def) apply (drule_tac a = Out in apply_type, auto) done -lemma (in merge) iOut_value_type [TC,simp]: "s \ state ==> s`iOut \ list(nat)" +lemma (in merge) iOut_value_type [TC,simp]: "s \ state \ s`iOut \ list(nat)" apply (unfold state_def) apply (drule_tac a = iOut in apply_type, auto) done @@ -113,7 +113,7 @@ done lemma (in merge) M_ok_iff: - "G \ program ==> + "G \ program \ M ok G \ (G \ preserves(lift(Out)) & G \ preserves(lift(iOut)) & M \ Allowed(G))" apply (cut_tac merge_spec) @@ -121,9 +121,9 @@ done lemma (in merge) merge_Always_Out_eq_iOut: - "[| G \ preserves(lift(Out)); G \ preserves(lift(iOut)); - M \ Allowed(G) |] - ==> M \ G \ Always({s \ state. length(s`Out)=length(s`iOut)})" + "\G \ preserves(lift(Out)); G \ preserves(lift(iOut)); + M \ Allowed(G)\ + \ M \ G \ Always({s \ state. length(s`Out)=length(s`iOut)})" apply (frule preserves_type [THEN subsetD]) apply (subgoal_tac "G \ program") prefer 2 apply assumption @@ -133,8 +133,8 @@ done lemma (in merge) merge_Bounded: - "[| G \ preserves(lift(iOut)); G \ preserves(lift(Out)); - M \ Allowed(G) |] ==> + "\G \ preserves(lift(iOut)); G \ preserves(lift(Out)); + M \ Allowed(G)\ \ M \ G: Always({s \ state. \elt \ set_of_list(s`iOut). elt preserves(lift(iOut)); - G: preserves(lift(Out)); M \ Allowed(G) |] - ==> M \ G \ Always +"\G \ preserves(lift(iOut)); + G: preserves(lift(Out)); M \ Allowed(G)\ + \ M \ G \ Always ({s \ state. msetsum(%i. bag_of(sublist(s`Out, {k \ nat. k < length(s`iOut) & nth(k, s`iOut)=i})), Nclients, A) = bag_of(s`Out)})" diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Monotonicity.thy --- a/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 16:51:35 2022 +0100 @@ -13,14 +13,14 @@ definition mono1 :: "[i, i, i, i, i=>i] => o" where - "mono1(A, r, B, s, f) == + "mono1(A, r, B, s, f) \ (\x \ A. \y \ A. \ 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) == + "mono2(A, r, B, s, C, t, f) \ (\x \ A. \y \ A. \u \ B. \v \ B. \ r & \ s \ \ t) & (\x \ A. \y \ B. f(x,y) \ C)" @@ -29,29 +29,29 @@ definition SetLe :: "i =>i" where - "SetLe(A) == { \ Pow(A)*Pow(A). x \ y}" + "SetLe(A) \ { \ Pow(A)*Pow(A). x \ y}" definition MultLe :: "[i,i] =>i" where - "MultLe(A, r) == multirel(A, r - id(A)) \ id(Mult(A))" + "MultLe(A, r) \ multirel(A, r - id(A)) \ id(Mult(A))" lemma mono1D: - "[| mono1(A, r, B, s, f); \ r; x \ A; y \ A |] ==> \ s" + "\mono1(A, r, B, s, f); \ r; x \ A; y \ A\ \ \ s" by (unfold mono1_def, auto) lemma mono2D: - "[| mono2(A, r, B, s, C, t, f); - \ r; \ s; x \ A; y \ A; u \ B; v \ B |] - ==> \ t" + "\mono2(A, r, B, s, C, t, f); + \ r; \ 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)" + "\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) @@ -67,18 +67,18 @@ done lemma take_mono_left: - "[| i \ j; xs \ list(A); j \ nat |] - ==> \ prefix(A)" + "\i \ j; xs \ list(A); j \ nat\ + \ \ prefix(A)" by (blast intro: le_in_nat take_mono_left_lemma) lemma take_mono_right: - "[| \ prefix(A); i \ nat |] - ==> \ prefix(A)" + "\ \ prefix(A); i \ nat\ + \ \ prefix(A)" by (auto simp add: prefix_iff) lemma take_mono: - "[| i \ j; \ prefix(A); j \ nat |] - ==> \ prefix(A)" + "\i \ j; \ 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 diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/MultisetSum.thy --- a/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,23 +10,23 @@ definition lcomm :: "[i, i, [i,i]=>i]=>o" where - "lcomm(A, B, f) == + "lcomm(A, B, f) \ (\x \ A. \y \ A. \z \ B. f(x, f(y, z))= f(y, f(x, z))) & (\x \ A. \y \ B. f(x, y) \ B)" definition general_setsum :: "[i,i,i, [i,i]=>i, i=>i] =>i" where - "general_setsum(C, B, e, f, g) == + "general_setsum(C, B, e, f, g) \ if Finite(C) then fold[cons(e, B)](%x y. f(g(x), y), e, C) else e" definition msetsum :: "[i=>i, i, i]=>i" where - "msetsum(g, C, B) == normalize(general_setsum(C, Mult(B), 0, (+#), g))" + "msetsum(g, C, B) \ normalize(general_setsum(C, Mult(B), 0, (+#), g))" definition (* sum for naturals *) nsetsum :: "[i=>i, i] =>i" where - "nsetsum(g, C) == general_setsum(C, nat, 0, (#+), g)" + "nsetsum(g, C) \ general_setsum(C, nat, 0, (#+), g)" lemma mset_of_normalize: "mset_of(normalize(M)) \ mset_of(M)" @@ -36,7 +36,7 @@ by (simp add: general_setsum_def) lemma general_setsum_cons [simp]: -"[| C \ Fin(A); a \ A; a\C; e \ B; \x \ A. g(x) \ B; lcomm(B, B, f) |] ==> +"\C \ Fin(A); a \ A; a\C; e \ B; \x \ A. g(x) \ B; lcomm(B, B, f)\ \ general_setsum(cons(a, C), B, e, f, g) = f(g(a), general_setsum(C, B, e, f,g))" apply (simp add: general_setsum_def) @@ -48,7 +48,7 @@ (** lcomm **) -lemma lcomm_mono: "[| C\A; lcomm(A, B, f) |] ==> lcomm(C, B,f)" +lemma lcomm_mono: "\C\A; lcomm(A, B, f)\ \ lcomm(C, B,f)" by (auto simp add: lcomm_def, blast) lemma munion_lcomm [simp]: "lcomm(Mult(A), Mult(A), (+#))" @@ -57,8 +57,8 @@ (** msetsum **) lemma multiset_general_setsum: - "[| C \ Fin(A); \x \ A. multiset(g(x))& mset_of(g(x))\B |] - ==> multiset(general_setsum(C, B -||> nat - {0}, 0, \u v. u +# v, g))" + "\C \ Fin(A); \x \ A. multiset(g(x))& mset_of(g(x))\B\ + \ multiset(general_setsum(C, B -||> nat - {0}, 0, \u v. u +# v, g))" apply (erule Fin_induct, auto) apply (subst general_setsum_cons) apply (auto simp add: Mult_iff_multiset) @@ -68,8 +68,8 @@ by (simp add: msetsum_def) lemma msetsum_cons [simp]: - "[| C \ Fin(A); a\C; a \ A; \x \ A. multiset(g(x)) & mset_of(g(x))\B |] - ==> msetsum(g, cons(a, C), B) = g(a) +# msetsum(g, C, B)" + "\C \ Fin(A); a\C; a \ A; \x \ A. multiset(g(x)) & mset_of(g(x))\B\ + \ msetsum(g, cons(a, C), B) = g(a) +# msetsum(g, C, B)" apply (simp add: msetsum_def) apply (subst general_setsum_cons) apply (auto simp add: multiset_general_setsum Mult_iff_multiset) @@ -81,15 +81,15 @@ by (simp add: msetsum_def) lemma mset_of_msetsum: - "[| C \ Fin(A); \x \ A. multiset(g(x)) & mset_of(g(x))\B |] - ==> mset_of(msetsum(g, C, B))\B" + "\C \ Fin(A); \x \ A. multiset(g(x)) & mset_of(g(x))\B\ + \ mset_of(msetsum(g, C, B))\B" by (erule Fin_induct, auto) (*The reversed orientation looks more natural, but LOOPS as a simprule!*) lemma msetsum_Un_Int: - "[| C \ Fin(A); D \ Fin(A); \x \ A. multiset(g(x)) & mset_of(g(x))\B |] - ==> msetsum(g, C \ D, B) +# msetsum(g, C \ D, B) + "\C \ Fin(A); D \ Fin(A); \x \ A. multiset(g(x)) & mset_of(g(x))\B\ + \ msetsum(g, C \ D, B) +# msetsum(g, C \ D, B) = msetsum(g, C, B) +# msetsum(g, D, B)" apply (erule Fin_induct) apply (subgoal_tac [2] "cons (x, y) \ D = cons (x, y \ D) ") @@ -106,19 +106,19 @@ lemma msetsum_Un_disjoint: - "[| C \ Fin(A); D \ Fin(A); C \ D = 0; - \x \ A. multiset(g(x)) & mset_of(g(x))\B |] - ==> msetsum(g, C \ D, B) = msetsum(g, C, B) +# msetsum(g,D, B)" + "\C \ Fin(A); D \ Fin(A); C \ D = 0; + \x \ A. multiset(g(x)) & mset_of(g(x))\B\ + \ msetsum(g, C \ D, B) = msetsum(g, C, B) +# msetsum(g,D, B)" apply (subst msetsum_Un_Int [symmetric]) apply (auto simp add: msetsum_multiset) done lemma UN_Fin_lemma [rule_format (no_asm)]: - "I \ Fin(A) ==> (\i \ I. C(i) \ Fin(B)) \ (\i \ I. C(i)):Fin(B)" + "I \ Fin(A) \ (\i \ I. C(i) \ Fin(B)) \ (\i \ I. C(i)):Fin(B)" by (erule Fin_induct, auto) lemma msetsum_UN_disjoint [rule_format (no_asm)]: - "[| I \ Fin(K); \i \ K. C(i) \ Fin(A) |] ==> + "\I \ Fin(K); \i \ K. C(i) \ Fin(A)\ \ (\x \ A. multiset(f(x)) & mset_of(f(x))\B) \ (\i \ I. \j \ I. i\j \ C(i) \ C(j) = 0) \ msetsum(f, \i \ I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)" @@ -138,9 +138,9 @@ done lemma msetsum_addf: - "[| C \ Fin(A); + "\C \ Fin(A); \x \ A. multiset(f(x)) & mset_of(f(x))\B; - \x \ A. multiset(g(x)) & mset_of(g(x))\B |] ==> + \x \ A. multiset(g(x)) & mset_of(g(x))\B\ \ msetsum(%x. f(x) +# g(x), C, B) = msetsum(f, C, B) +# msetsum(g, C, B)" apply (subgoal_tac "\x \ A. multiset (f(x) +# g(x)) & mset_of (f(x) +# g(x)) \ B") apply (erule Fin_induct) @@ -148,16 +148,16 @@ done lemma msetsum_cong: - "[| C=D; !!x. x \ D ==> f(x) = g(x) |] - ==> msetsum(f, C, B) = msetsum(g, D, B)" + "\C=D; \x. x \ D \ f(x) = g(x)\ + \ msetsum(f, C, B) = msetsum(g, D, B)" by (simp add: msetsum_def general_setsum_def cong add: fold_cong) -lemma multiset_union_diff: "[| multiset(M); multiset(N) |] ==> M +# N -# N = M" +lemma multiset_union_diff: "\multiset(M); multiset(N)\ \ M +# N -# N = M" by (simp add: multiset_equality) -lemma msetsum_Un: "[| C \ Fin(A); D \ Fin(A); - \x \ A. multiset(f(x)) & mset_of(f(x)) \ B |] - ==> msetsum(f, C \ D, B) = +lemma msetsum_Un: "\C \ Fin(A); D \ Fin(A); + \x \ A. multiset(f(x)) & mset_of(f(x)) \ B\ + \ msetsum(f, C \ D, B) = msetsum(f, C, B) +# msetsum(f, D, B) -# msetsum(f, C \ D, B)" apply (subgoal_tac "C \ D \ Fin (A) & C \ D \ Fin (A) ") apply clarify @@ -170,7 +170,7 @@ by (simp add: nsetsum_def) lemma nsetsum_cons [simp]: - "[| Finite(C); x\C |] ==> nsetsum(g, cons(x, C))= g(x) #+ nsetsum(g, C)" + "\Finite(C); x\C\ \ nsetsum(g, cons(x, C))= g(x) #+ nsetsum(g, C)" apply (simp add: nsetsum_def general_setsum_def) apply (rule_tac A = "cons (x, C)" in fold_typing.fold_cons) apply (auto intro: Finite_cons_lemma simp add: fold_typing_def) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Mutex.thy --- a/src/ZF/UNITY/Mutex.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Mutex.thy Tue Sep 27 16:51:35 2022 +0100 @@ -21,11 +21,11 @@ the usual ZF typechecking: an ill-tyed expressions reduce to the empty set. \ -abbreviation "p == Var([0])" -abbreviation "m == Var([1])" -abbreviation "n == Var([0,0])" -abbreviation "u == Var([0,1])" -abbreviation "v == Var([1,0])" +abbreviation "p \ Var([0])" +abbreviation "m \ Var([1])" +abbreviation "n \ Var([0,0])" +abbreviation "u \ Var([0,1])" +abbreviation "v \ Var([1,0])" axiomatization where \ \Type declarations\ p_type: "type_of(p)=bool & default_val(p)=0" and @@ -36,56 +36,56 @@ definition (** The program for process U **) - "U0 == {:state*state. t = s(u:=1, m:=#1) & s`m = #0}" + "U0 \ {:state*state. t = s(u:=1, m:=#1) & s`m = #0}" definition - "U1 == {:state*state. t = s(p:= s`v, m:=#2) & s`m = #1}" + "U1 \ {:state*state. t = s(p:= s`v, m:=#2) & s`m = #1}" definition - "U2 == {:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}" + "U2 \ {:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}" definition - "U3 == {:state*state. t=s(u:=0, m:=#4) & s`m = #3}" + "U3 \ {:state*state. t=s(u:=0, m:=#4) & s`m = #3}" definition - "U4 == {:state*state. t = s(p:=1, m:=#0) & s`m = #4}" + "U4 \ {:state*state. t = s(p:=1, m:=#0) & s`m = #4}" (** The program for process V **) definition - "V0 == {:state*state. t = s (v:=1, n:=#1) & s`n = #0}" + "V0 \ {:state*state. t = s (v:=1, n:=#1) & s`n = #0}" definition - "V1 == {:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}" + "V1 \ {:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}" definition - "V2 == {:state*state. t = s(n:=#3) & s`p=1 & s`n = #2}" + "V2 \ {:state*state. t = s(n:=#3) & s`p=1 & s`n = #2}" definition - "V3 == {:state*state. t = s (v:=0, n:=#4) & s`n = #3}" + "V3 \ {:state*state. t = s (v:=0, n:=#4) & s`n = #3}" definition - "V4 == {:state*state. t = s (p:=0, n:=#0) & s`n = #4}" + "V4 \ {:state*state. t = s (p:=0, n:=#0) & s`n = #4}" definition - "Mutex == mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0}, + "Mutex \ mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0}, {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4}, Pow(state*state))" (** The correct invariants **) definition - "IU == {s:state. (s`u = 1\(#1 $\ s`m & s`m $\ #3)) + "IU \ {s:state. (s`u = 1\(#1 $\ s`m & s`m $\ #3)) & (s`m = #3 \ s`p=0)}" definition - "IV == {s:state. (s`v = 1 \ (#1 $\ s`n & s`n $\ #3)) + "IV \ {s:state. (s`v = 1 \ (#1 $\ s`n & s`n $\ #3)) & (s`n = #3 \ s`p=1)}" (** The faulty invariant (for U alone) **) definition - "bad_IU == {s:state. (s`u = 1 \ (#1 $\ s`m & s`m $\ #3))& + "bad_IU \ {s:state. (s`u = 1 \ (#1 $\ s`m & s`m $\ #3))& (#3 $\ s`m & s`m $\ #4 \ s`p=0)}" @@ -93,27 +93,27 @@ declare p_type [simp] u_type [simp] v_type [simp] m_type [simp] n_type [simp] -lemma u_value_type: "s \ state ==>s`u \ bool" +lemma u_value_type: "s \ state \s`u \ bool" apply (unfold state_def) apply (drule_tac a = u in apply_type, auto) done -lemma v_value_type: "s \ state ==> s`v \ bool" +lemma v_value_type: "s \ state \ s`v \ bool" apply (unfold state_def) apply (drule_tac a = v in apply_type, auto) done -lemma p_value_type: "s \ state ==> s`p \ bool" +lemma p_value_type: "s \ state \ s`p \ bool" apply (unfold state_def) apply (drule_tac a = p in apply_type, auto) done -lemma m_value_type: "s \ state ==> s`m \ int" +lemma m_value_type: "s \ state \ s`m \ int" apply (unfold state_def) apply (drule_tac a = m in apply_type, auto) done -lemma n_value_type: "s \ state ==>s`n \ int" +lemma n_value_type: "s \ state \s`n \ int" apply (unfold state_def) apply (drule_tac a = n in apply_type, auto) done @@ -174,14 +174,14 @@ done (*The safety property: mutual exclusion*) -lemma mutual_exclusion: "Mutex \ Always({s \ state. ~(s`m = #3 & s`n = #3)})" +lemma mutual_exclusion: "Mutex \ Always({s \ state. \(s`m = #3 & s`n = #3)})" apply (rule Always_weaken) apply (rule Always_Int_I [OF IU IV], auto) done (*The bad invariant FAILS in V1*) -lemma less_lemma: "[| x$<#1; #3 $\ x |] ==> P" +lemma less_lemma: "\x$<#1; #3 $\ x\ \ P" apply (drule_tac j = "#1" and k = "#3" in zless_zle_trans) apply (drule_tac [2] j = x in zle_zless_trans, auto) done @@ -232,7 +232,7 @@ lemma U_lemma1: "Mutex \ {s \ state. s`m = #1} \w {s \ state. s`p =1}" by (rule LeadsTo_Trans [OF U_F1 [THEN LeadsTo_weaken_R] U_lemma2], blast) -lemma eq_123: "i \ int ==> (#1 $\ i & i $\ #3) \ (i=#1 | i=#2 | i=#3)" +lemma eq_123: "i \ int \ (#1 $\ i & i $\ #3) \ (i=#1 | i=#2 | i=#3)" apply auto apply (auto simp add: neq_iff_zless) apply (drule_tac [4] j = "#3" and i = i in zle_zless_trans) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/State.thy --- a/src/ZF/UNITY/State.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/State.thy Tue Sep 27 16:51:35 2022 +0100 @@ -21,20 +21,20 @@ default_val :: "i=>i" definition - "state == \x \ var. cons(default_val(x), type_of(x))" + "state \ \x \ var. cons(default_val(x), type_of(x))" definition - "st0 == \x \ var. default_val(x)" + "st0 \ \x \ var. default_val(x)" definition st_set :: "i=>o" where (* To prevent typing conditions like `A<=state' from being used in combination with the rules `constrains_weaken', etc. *) - "st_set(A) == A<=state" + "st_set(A) \ A<=state" definition st_compl :: "i=>i" where - "st_compl(A) == state-A" + "st_compl(A) \ state-A" lemma st0_in_state [simp,TC]: "st0 \ state" @@ -59,16 +59,16 @@ (* Intersection *) -lemma st_set_Int [intro!]: "st_set(A) | st_set(B) ==> st_set(A \ B)" +lemma st_set_Int [intro!]: "st_set(A) | st_set(B) \ st_set(A \ B)" by (simp add: st_set_def, auto) lemma st_set_Inter [intro!]: - "(S=0) | (\A \ S. st_set(A)) ==> st_set(\(S))" + "(S=0) | (\A \ S. st_set(A)) \ st_set(\(S))" apply (simp add: st_set_def Inter_def, auto) done (* Diff *) -lemma st_set_DiffI [intro!]: "st_set(A) ==> st_set(A - B)" +lemma st_set_DiffI [intro!]: "st_set(A) \ st_set(A - B)" by (simp add: st_set_def, auto) lemma Collect_Int_state [simp]: "Collect(state,P) \ state = Collect(state,P)" @@ -80,18 +80,18 @@ (* Introduction and destruction rules for st_set *) -lemma st_setI: "A \ state ==> st_set(A)" +lemma st_setI: "A \ state \ st_set(A)" by (simp add: st_set_def) -lemma st_setD: "st_set(A) ==> A<=state" +lemma st_setD: "st_set(A) \ A<=state" by (simp add: st_set_def) -lemma st_set_subset: "[| st_set(A); B<=A |] ==> st_set(B)" +lemma st_set_subset: "\st_set(A); B<=A\ \ st_set(B)" by (simp add: st_set_def, auto) lemma state_update_type: - "[| s \ state; x \ var; y \ type_of(x) |] ==> s(x:=y):state" + "\s \ state; x \ var; y \ type_of(x)\ \ s(x:=y):state" apply (simp add: state_def) apply (blast intro: update_type) done @@ -103,12 +103,12 @@ by (simp add: st_compl_def) lemma st_compl_Collect [simp]: - "st_compl({s \ state. P(s)}) = {s \ state. ~P(s)}" + "st_compl({s \ state. P(s)}) = {s \ state. \P(s)}" by (simp add: st_compl_def, auto) (*For using "disjunction" (union over an index set) to eliminate a variable.*) lemma UN_conj_eq: - "\d\D. f(d) \ A ==> (\k\A. {d\D. P(d) & f(d) = k}) = {d\D. P(d)}" + "\d\D. f(d) \ A \ (\k\A. {d\D. P(d) & f(d) = k}) = {d\D. P(d)}" by blast end diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/SubstAx.thy Tue Sep 27 16:51:35 2022 +0100 @@ -14,18 +14,18 @@ definition (* The definitions below are not `conventional', but yield simpler rules *) Ensures :: "[i,i] => i" (infixl \Ensures\ 60) where - "A Ensures B == {F \ program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" + "A Ensures B \ {F \ program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" definition LeadsTo :: "[i, i] => i" (infixl \\w\ 60) where - "A \w B == {F \ program. F:(reachable(F) \ A) \ (reachable(F) \ B)}" + "A \w B \ {F \ program. F:(reachable(F) \ A) \ (reachable(F) \ B)}" (*Resembles the previous definition of LeadsTo*) (* Equivalence with the HOL-like definition *) lemma LeadsTo_eq: -"st_set(B)==> A \w B = {F \ program. F:(reachable(F) \ A) \ B}" +"st_set(B)\ A \w B = {F \ program. F:(reachable(F) \ A) \ B}" apply (unfold LeadsTo_def) apply (blast dest: psp_stable2 leadsToD2 constrainsD2 intro: leadsTo_weaken) done @@ -36,35 +36,35 @@ (*** Specialized laws for handling invariants ***) (** Conjoining an Always property **) -lemma Always_LeadsTo_pre: "F \ Always(I) ==> (F:(I \ A) \w A') \ (F \ A \w A')" +lemma Always_LeadsTo_pre: "F \ Always(I) \ (F:(I \ A) \w A') \ (F \ A \w A')" by (simp add: LeadsTo_def Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2) -lemma Always_LeadsTo_post: "F \ Always(I) ==> (F \ A \w (I \ A')) \ (F \ A \w A')" +lemma Always_LeadsTo_post: "F \ Always(I) \ (F \ A \w (I \ A')) \ (F \ A \w A')" apply (unfold LeadsTo_def) apply (simp add: Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2) done (* Like 'Always_LeadsTo_pre RS iffD1', but with premises in the good order *) -lemma Always_LeadsToI: "[| F \ Always(C); F \ (C \ A) \w A' |] ==> F \ A \w A'" +lemma Always_LeadsToI: "\F \ Always(C); F \ (C \ A) \w A'\ \ F \ A \w A'" by (blast intro: Always_LeadsTo_pre [THEN iffD1]) (* Like 'Always_LeadsTo_post RS iffD2', but with premises in the good order *) -lemma Always_LeadsToD: "[| F \ Always(C); F \ A \w A' |] ==> F \ A \w (C \ A')" +lemma Always_LeadsToD: "\F \ Always(C); F \ A \w A'\ \ F \ A \w (C \ A')" by (blast intro: Always_LeadsTo_post [THEN iffD2]) (*** Introduction rules \ Basis, Trans, Union ***) -lemma LeadsTo_Basis: "F \ A Ensures B ==> F \ A \w B" +lemma LeadsTo_Basis: "F \ A Ensures B \ F \ A \w B" by (auto simp add: Ensures_def LeadsTo_def) lemma LeadsTo_Trans: - "[| F \ A \w B; F \ B \w C |] ==> F \ A \w C" + "\F \ A \w B; F \ B \w C\ \ F \ A \w C" apply (simp (no_asm_use) add: LeadsTo_def) apply (blast intro: leadsTo_Trans) done lemma LeadsTo_Union: -"[|(!!A. A \ S ==> F \ A \w B); F \ program|]==>F \ \(S) \w B" +"\(\A. A \ S \ F \ A \w B); F \ program\\F \ \(S) \w B" apply (simp add: LeadsTo_def) apply (subst Int_Union_Union2) apply (rule leadsTo_UN, auto) @@ -72,23 +72,23 @@ (*** Derived rules ***) -lemma leadsTo_imp_LeadsTo: "F \ A \ B ==> F \ A \w B" +lemma leadsTo_imp_LeadsTo: "F \ A \ B \ F \ A \w B" apply (frule leadsToD2, clarify) apply (simp (no_asm_simp) add: LeadsTo_eq) apply (blast intro: leadsTo_weaken_L) done (*Useful with cancellation, disjunction*) -lemma LeadsTo_Un_duplicate: "F \ A \w (A' \ A') ==> F \ A \w A'" +lemma LeadsTo_Un_duplicate: "F \ A \w (A' \ A') \ F \ A \w A'" by (simp add: Un_ac) lemma LeadsTo_Un_duplicate2: - "F \ A \w (A' \ C \ C) ==> F \ A \w (A' \ C)" + "F \ A \w (A' \ C \ C) \ F \ A \w (A' \ C)" by (simp add: Un_ac) lemma LeadsTo_UN: - "[|(!!i. i \ I ==> F \ A(i) \w B); F \ program|] - ==>F:(\i \ I. A(i)) \w B" + "\(\i. i \ I \ F \ A(i) \w B); F \ program\ + \F:(\i \ I. A(i)) \w B" apply (simp add: LeadsTo_def) apply (simp (no_asm_simp) del: UN_simps add: Int_UN_distrib) apply (rule leadsTo_UN, auto) @@ -96,7 +96,7 @@ (*Binary union introduction rule*) lemma LeadsTo_Un: - "[| F \ A \w C; F \ B \w C |] ==> F \ (A \ B) \w C" + "\F \ A \w C; F \ B \w C\ \ F \ (A \ B) \w C" apply (subst Un_eq_Union) apply (rule LeadsTo_Union) apply (auto dest: LeadsTo_type [THEN subsetD]) @@ -104,11 +104,11 @@ (*Lets us look at the starting state*) lemma single_LeadsTo_I: - "[|(!!s. s \ A ==> F:{s} \w B); F \ program|]==>F \ A \w B" + "\(\s. s \ A \ F:{s} \w B); F \ program\\F \ A \w B" apply (subst UN_singleton [symmetric], rule LeadsTo_UN, auto) done -lemma subset_imp_LeadsTo: "[| A \ B; F \ program |] ==> F \ A \w B" +lemma subset_imp_LeadsTo: "\A \ B; F \ program\ \ F \ A \w B" apply (simp (no_asm_simp) add: LeadsTo_def) apply (blast intro: subset_imp_leadsTo) done @@ -122,33 +122,33 @@ by (auto dest: LeadsTo_type [THEN subsetD] simp add: LeadsTo_eq) declare LeadsTo_state [iff] -lemma LeadsTo_weaken_R: "[| F \ A \w A'; A'<=B'|] ==> F \ A \w B'" +lemma LeadsTo_weaken_R: "\F \ A \w A'; A'<=B'\ \ F \ A \w B'" apply (unfold LeadsTo_def) apply (auto intro: leadsTo_weaken_R) done -lemma LeadsTo_weaken_L: "[| F \ A \w A'; B \ A |] ==> F \ B \w A'" +lemma LeadsTo_weaken_L: "\F \ A \w A'; B \ A\ \ F \ B \w A'" apply (unfold LeadsTo_def) apply (auto intro: leadsTo_weaken_L) done -lemma LeadsTo_weaken: "[| F \ A \w A'; B<=A; A'<=B' |] ==> F \ B \w B'" +lemma LeadsTo_weaken: "\F \ A \w A'; B<=A; A'<=B'\ \ F \ B \w B'" by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans) lemma Always_LeadsTo_weaken: -"[| F \ Always(C); F \ A \w A'; C \ B \ A; C \ A' \ B' |] - ==> F \ B \w B'" +"\F \ Always(C); F \ A \w A'; C \ B \ A; C \ A' \ B'\ + \ F \ B \w B'" apply (blast dest: Always_LeadsToI intro: LeadsTo_weaken Always_LeadsToD) done (** Two theorems for "proof lattices" **) -lemma LeadsTo_Un_post: "F \ A \w B ==> F:(A \ B) \w B" +lemma LeadsTo_Un_post: "F \ A \w B \ F:(A \ B) \w B" by (blast dest: LeadsTo_type [THEN subsetD] intro: LeadsTo_Un subset_imp_LeadsTo) -lemma LeadsTo_Trans_Un: "[| F \ A \w B; F \ B \w C |] - ==> F \ (A \ B) \w C" +lemma LeadsTo_Trans_Un: "\F \ A \w B; F \ B \w C\ + \ F \ (A \ B) \w C" apply (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans dest: LeadsTo_type [THEN subsetD]) done @@ -166,14 +166,14 @@ (** More rules using the premise "Always(I)" **) -lemma EnsuresI: "[| F:(A-B) Co (A \ B); F \ transient (A-B) |] ==> F \ A Ensures B" +lemma EnsuresI: "\F:(A-B) Co (A \ B); F \ transient (A-B)\ \ F \ A Ensures B" apply (simp add: Ensures_def Constrains_eq_constrains) apply (blast intro: ensuresI constrains_weaken transient_strengthen dest: constrainsD2) done -lemma Always_LeadsTo_Basis: "[| F \ Always(I); F \ (I \ (A-A')) Co (A \ A'); - F \ transient (I \ (A-A')) |] - ==> F \ A \w A'" +lemma Always_LeadsTo_Basis: "\F \ Always(I); F \ (I \ (A-A')) Co (A \ A'); + F \ transient (I \ (A-A'))\ + \ F \ A \w A'" apply (rule Always_LeadsToI, assumption) apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen) done @@ -181,36 +181,36 @@ (*Set difference: maybe combine with leadsTo_weaken_L?? This is the most useful form of the "disjunction" rule*) lemma LeadsTo_Diff: - "[| F \ (A-B) \w C; F \ (A \ B) \w C |] ==> F \ A \w C" + "\F \ (A-B) \w C; F \ (A \ B) \w C\ \ F \ A \w C" by (blast intro: LeadsTo_Un LeadsTo_weaken) lemma LeadsTo_UN_UN: - "[|(!!i. i \ I ==> F \ A(i) \w A'(i)); F \ program |] - ==> F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" + "\(\i. i \ I \ F \ A(i) \w A'(i)); F \ program\ + \ F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" apply (rule LeadsTo_Union, auto) apply (blast intro: LeadsTo_weaken_R) done (*Binary union version*) lemma LeadsTo_Un_Un: - "[| F \ A \w A'; F \ B \w B' |] ==> F:(A \ B) \w (A' \ B')" + "\F \ A \w A'; F \ B \w B'\ \ F:(A \ B) \w (A' \ B')" by (blast intro: LeadsTo_Un LeadsTo_weaken_R) (** The cancellation law **) -lemma LeadsTo_cancel2: "[| F \ A \w(A' \ B); F \ B \w B' |] ==> F \ A \w (A' \ B')" +lemma LeadsTo_cancel2: "\F \ A \w(A' \ B); F \ B \w B'\ \ F \ A \w (A' \ B')" by (blast intro: LeadsTo_Un_Un subset_imp_LeadsTo LeadsTo_Trans dest: LeadsTo_type [THEN subsetD]) lemma Un_Diff: "A \ (B - A) = A \ B" by auto -lemma LeadsTo_cancel_Diff2: "[| F \ A \w (A' \ B); F \ (B-A') \w B' |] ==> F \ A \w (A' \ B')" +lemma LeadsTo_cancel_Diff2: "\F \ A \w (A' \ B); F \ (B-A') \w B'\ \ F \ A \w (A' \ B')" apply (rule LeadsTo_cancel2) prefer 2 apply assumption apply (simp (no_asm_simp) add: Un_Diff) done -lemma LeadsTo_cancel1: "[| F \ A \w (B \ A'); F \ B \w B' |] ==> F \ A \w (B' \ A')" +lemma LeadsTo_cancel1: "\F \ A \w (B \ A'); F \ B \w B'\ \ F \ A \w (B' \ A')" apply (simp add: Un_commute) apply (blast intro!: LeadsTo_cancel2) done @@ -218,7 +218,7 @@ lemma Diff_Un2: "(B - A) \ A = B \ A" by auto -lemma LeadsTo_cancel_Diff1: "[| F \ A \w (B \ A'); F \ (B-A') \w B' |] ==> F \ A \w (B' \ A')" +lemma LeadsTo_cancel_Diff1: "\F \ A \w (B \ A'); F \ (B-A') \w B'\ \ F \ A \w (B' \ A')" apply (rule LeadsTo_cancel1) prefer 2 apply assumption apply (simp (no_asm_simp) add: Diff_Un2) @@ -227,7 +227,7 @@ (** The impossibility law **) (*The set "A" may be non-empty, but it contains no reachable states*) -lemma LeadsTo_empty: "F \ A \w 0 ==> F \ Always (state -A)" +lemma LeadsTo_empty: "F \ A \w 0 \ F \ Always (state -A)" apply (simp (no_asm_use) add: LeadsTo_def Always_eq_includes_reachable) apply (cut_tac reachable_type) apply (auto dest!: leadsTo_empty) @@ -236,26 +236,26 @@ (** PSP \ Progress-Safety-Progress **) (*Special case of PSP \ Misra's "stable conjunction"*) -lemma PSP_Stable: "[| F \ A \w A'; F \ Stable(B) |]==> F:(A \ B) \w (A' \ B)" +lemma PSP_Stable: "\F \ A \w A'; F \ Stable(B)\\ F:(A \ B) \w (A' \ B)" apply (simp add: LeadsTo_def Stable_eq_stable, clarify) apply (drule psp_stable, assumption) apply (simp add: Int_ac) done -lemma PSP_Stable2: "[| F \ A \w A'; F \ Stable(B) |] ==> F \ (B \ A) \w (B \ A')" +lemma PSP_Stable2: "\F \ A \w A'; F \ Stable(B)\ \ F \ (B \ A) \w (B \ A')" apply (simp (no_asm_simp) add: PSP_Stable Int_ac) done -lemma PSP: "[| F \ A \w A'; F \ B Co B'|]==> F \ (A \ B') \w ((A' \ B) \ (B' - B))" +lemma PSP: "\F \ A \w A'; F \ B Co B'\\ F \ (A \ B') \w ((A' \ B) \ (B' - B))" apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains) apply (blast dest: psp intro: leadsTo_weaken) done -lemma PSP2: "[| F \ A \w A'; F \ B Co B' |]==> F:(B' \ A) \w ((B \ A') \ (B' - B))" +lemma PSP2: "\F \ A \w A'; F \ B Co B'\\ F:(B' \ A) \w ((B \ A') \ (B' - B))" by (simp (no_asm_simp) add: PSP Int_ac) lemma PSP_Unless: -"[| F \ A \w A'; F \ B Unless B'|]==> F:(A \ B) \w ((A' \ B) \ B')" +"\F \ A \w A'; F \ B Unless B'\\ F:(A \ B) \w ((A' \ B) \ B')" apply (unfold op_Unless_def) apply (drule PSP, assumption) apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo) @@ -264,11 +264,11 @@ (*** Induction rules ***) (** Meta or object quantifier ????? **) -lemma LeadsTo_wf_induct: "[| wf(r); +lemma LeadsTo_wf_induct: "\wf(r); \m \ I. F \ (A \ f-``{m}) \w ((A \ f-``(converse(r) `` {m})) \ B); - field(r)<=I; A<=f-``I; F \ program |] - ==> F \ A \w B" + field(r)<=I; A<=f-``I; F \ program\ + \ F \ A \w B" apply (simp (no_asm_use) add: LeadsTo_def) apply auto apply (erule_tac I = I and f = f in leadsTo_wf_induct, safe) @@ -278,8 +278,8 @@ done -lemma LessThan_induct: "[| \m \ nat. F:(A \ f-``{m}) \w ((A \ f-``m) \ B); - A<=f-``nat; F \ program |] ==> F \ A \w B" +lemma LessThan_induct: "\\m \ nat. F:(A \ f-``{m}) \w ((A \ f-``m) \ B); + A<=f-``nat; F \ program\ \ F \ A \w 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]) @@ -297,16 +297,16 @@ (*** Completion \ Binary and General Finite versions ***) -lemma Completion: "[| F \ A \w (A' \ C); F \ A' Co (A' \ C); - F \ B \w (B' \ C); F \ B' Co (B' \ C) |] - ==> F \ (A \ B) \w ((A' \ B') \ C)" +lemma Completion: "\F \ A \w (A' \ C); F \ A' Co (A' \ C); + F \ B \w (B' \ C); F \ B' Co (B' \ C)\ + \ F \ (A \ B) \w ((A' \ B') \ C)" apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains Int_Un_distrib) apply (blast intro: completion leadsTo_weaken) done lemma Finite_completion_aux: - "[| I \ Fin(X);F \ program |] - ==> (\i \ I. F \ (A(i)) \w (A'(i) \ C)) \ + "\I \ Fin(X);F \ program\ + \ (\i \ I. F \ (A(i)) \w (A'(i) \ C)) \ (\i \ I. F \ (A'(i)) Co (A'(i) \ C)) \ F \ (\i \ I. A(i)) \w ((\i \ I. A'(i)) \ C)" apply (erule Fin_induct) @@ -317,16 +317,16 @@ done lemma Finite_completion: - "[| I \ Fin(X); !!i. i \ I ==> F \ A(i) \w (A'(i) \ C); - !!i. i \ I ==> F \ A'(i) Co (A'(i) \ C); - F \ program |] - ==> F \ (\i \ I. A(i)) \w ((\i \ I. A'(i)) \ C)" + "\I \ Fin(X); \i. i \ I \ F \ A(i) \w (A'(i) \ C); + \i. i \ I \ F \ A'(i) Co (A'(i) \ C); + F \ program\ + \ F \ (\i \ I. A(i)) \w ((\i \ I. A'(i)) \ C)" by (blast intro: Finite_completion_aux [THEN mp, THEN mp]) lemma Stable_completion: - "[| F \ A \w A'; F \ Stable(A'); - F \ B \w B'; F \ Stable(B') |] - ==> F \ (A \ B) \w (A' \ B')" + "\F \ A \w A'; F \ Stable(A'); + F \ B \w B'; F \ Stable(B')\ + \ F \ (A \ B) \w (A' \ B')" apply (unfold Stable_def) apply (rule_tac C1 = 0 in Completion [THEN LeadsTo_weaken_R]) prefer 5 @@ -335,10 +335,10 @@ done lemma Finite_stable_completion: - "[| I \ Fin(X); - (!!i. i \ I ==> F \ A(i) \w A'(i)); - (!!i. i \ I ==>F \ Stable(A'(i))); F \ program |] - ==> F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" + "\I \ Fin(X); + (\i. i \ I \ F \ A(i) \w A'(i)); + (\i. i \ I \F \ Stable(A'(i))); F \ program\ + \ F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" apply (unfold Stable_def) apply (rule_tac C1 = 0 in Finite_completion [THEN LeadsTo_weaken_R], simp_all) apply (rule_tac [3] subset_refl, auto) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/UNITY.thy --- a/src/ZF/UNITY/UNITY.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/UNITY.thy Tue Sep 27 16:51:35 2022 +0100 @@ -14,7 +14,7 @@ definition program :: i where - "program == {: + "program \ {: Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)). id(state) \ acts & id(state) \ allowed}" @@ -22,80 +22,80 @@ 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) == + "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))" + "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" + "programify(F) \ if F \ program then F else SKIP" definition RawInit :: "i=>i" where - "RawInit(F) == fst(F)" + "RawInit(F) \ fst(F)" definition Init :: "i=>i" where - "Init(F) == RawInit(programify(F))" + "Init(F) \ RawInit(programify(F))" definition RawActs :: "i=>i" where - "RawActs(F) == cons(id(state), fst(snd(F)))" + "RawActs(F) \ cons(id(state), fst(snd(F)))" definition Acts :: "i=>i" where - "Acts(F) == RawActs(programify(F))" + "Acts(F) \ RawActs(programify(F))" definition RawAllowedActs :: "i=>i" where - "RawAllowedActs(F) == cons(id(state), snd(snd(F)))" + "RawAllowedActs(F) \ cons(id(state), snd(snd(F)))" definition AllowedActs :: "i=>i" where - "AllowedActs(F) == RawAllowedActs(programify(F))" + "AllowedActs(F) \ RawAllowedActs(programify(F))" definition Allowed :: "i =>i" where - "Allowed(F) == {G \ program. Acts(G) \ AllowedActs(F)}" + "Allowed(F) \ {G \ program. Acts(G) \ AllowedActs(F)}" definition initially :: "i=>i" where - "initially(A) == {F \ program. Init(F)\A}" + "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)}" + "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)" + "A unless B \ (A - B) co (A \ B)" definition stable :: "i=>i" where - "stable(A) == A co A" + "stable(A) \ A co A" definition strongest_rhs :: "[i, i] => i" where - "strongest_rhs(F, A) == \({B \ Pow(state). F \ A co B})" + "strongest_rhs(F, A) \ \({B \ Pow(state). F \ A co B})" definition invariant :: "i => i" where - "invariant(A) == initially(A) \ stable(A)" + "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))" + "f comp g \ %x. f(g(x))" definition pg_compl :: "i=>i" where - "pg_compl(X)== program - X" + "pg_compl(X)\ program - X" text\SKIP\ lemma SKIP_in_program [iff,TC]: "SKIP \ program" @@ -105,7 +105,7 @@ subsection\The function \<^term>\programify\, the coercion from anything to program\ -lemma programify_program [simp]: "F \ program ==> programify(F)=F" +lemma programify_program [simp]: "F \ program \ programify(F)=F" by (force simp add: programify_def) lemma programify_in_program [iff,TC]: "programify(F) \ program" @@ -127,13 +127,13 @@ subsection\The Inspectors for Programs\ -lemma id_in_RawActs: "F \ program ==>id(state) \ RawActs(F)" +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)" +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)" @@ -149,14 +149,14 @@ subsection\Types of the Inspectors\ -lemma RawInit_type: "F \ program ==> RawInit(F)\state" +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)" +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)" + "F \ program \ RawAllowedActs(F)\Pow(state*state)" by (auto simp add: program_def RawAllowedActs_def) lemma Init_type: "Init(F)\state" @@ -176,11 +176,11 @@ by (simp add: RawAllowedActs_type AllowedActs_def) text\Needed in Behaviors\ -lemma ActsD: "[| act \ Acts(F); \ act |] ==> s \ state & s' \ state" +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" + "\act \ AllowedActs(F); \ act\ \ s \ state & s' \ state" by (blast dest: AllowedActs_type [THEN subsetD]) subsection\Simplification rules involving \<^term>\state\, \<^term>\Init\, @@ -280,7 +280,7 @@ text\Equality of UNITY programs\ lemma raw_surjective_mk_program: - "F \ program ==> mk_program(RawInit(F), RawActs(F), RawAllowedActs(F))=F" + "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 @@ -290,45 +290,45 @@ 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" + "\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" + "\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) \ + "\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" + "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))" + "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))" + "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 & + "\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 @@ -336,12 +336,12 @@ 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'))" + "\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)" +lemma def_set_simp: "A \ B \ (x \ A) \ (x \ B)" by auto @@ -351,15 +351,15 @@ 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'" + "\(\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" + "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)" +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" @@ -380,18 +380,18 @@ text\monotonic in 2nd argument\ lemma constrains_weaken_R: - "[| F \ A co A'; A'\B' |] ==> F \ A co B'" + "\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'" + "\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'" + "\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 @@ -400,12 +400,12 @@ subsection\Constrains and Union\ lemma constrains_Un: - "[| F \ A co A'; F \ B co B' |] ==> F \ (A \ B) co (A' \ B')" + "\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))" + "\\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: @@ -413,7 +413,7 @@ 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)" + "i \ I \ (\i \ I. A(i)) co B = (\i \ I. A(i) co B)" by (force simp add: constrains_def st_set_def) @@ -423,16 +423,16 @@ 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))" + "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')" + "\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))" + "\\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) @@ -442,22 +442,22 @@ (* 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 |]==> +"\\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'" + "\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" +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')" +"\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 @@ -468,12 +468,12 @@ 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" +lemma unlessI: "\F \ (A-B) co (A \ B)\ \ F \ A unless B" apply (unfold unless_def) apply (blast dest: constrainsD2) done -lemma unlessD: "F :A unless B ==> F \ (A-B) co (A \ B)" +lemma unlessD: "F :A unless B \ F \ (A-B) co (A \ B)" by (unfold unless_def, auto) @@ -482,10 +482,10 @@ lemma initially_type: "initially(A) \ program" by (unfold initially_def, blast) -lemma initiallyI: "[| F \ program; Init(F)\A |] ==> F \ initially(A)" +lemma initiallyI: "\F \ program; Init(F)\A\ \ F \ initially(A)" by (unfold initially_def, blast) -lemma initiallyD: "F \ initially(A) ==> Init(F)\A" +lemma initiallyD: "F \ initially(A) \ Init(F)\A" by (unfold initially_def, blast) @@ -494,13 +494,13 @@ lemma stable_type: "stable(A)\program" by (unfold stable_def constrains_def, blast) -lemma stableI: "F \ A co A ==> F \ stable(A)" +lemma stableI: "F \ A co A \ F \ stable(A)" by (unfold stable_def, assumption) -lemma stableD: "F \ stable(A) ==> F \ A co A" +lemma stableD: "F \ stable(A) \ F \ A co A" by (unfold stable_def, assumption) -lemma stableD2: "F \ stable(A) ==> F \ program & st_set(A)" +lemma stableD2: "F \ stable(A) \ F \ program & st_set(A)" by (unfold stable_def constrains_def, auto) lemma stable_state [simp]: "stable(state) = program" @@ -514,49 +514,49 @@ subsection\Union and Intersection with \<^term>\stable\\ lemma stable_Un: - "[| F \ stable(A); F \ stable(A') |] ==> F \ stable(A \ A')" + "\F \ stable(A); F \ stable(A')\ \ F \ stable(A \ A')" apply (unfold 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))" + "\\i. i\I \ F \ stable(A(i)); F \ program\ + \ F \ stable (\i \ I. A(i))" apply (unfold stable_def) apply (blast intro: constrains_UN) done lemma stable_Int: - "[| F \ stable(A); F \ stable(A') |] ==> F \ stable (A \ A')" + "\F \ stable(A); F \ stable(A')\ \ F \ stable (A \ A')" apply (unfold 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))" + "\\i. i \ I \ F \ stable(A(i)); F \ program\ + \ F \ stable (\i \ I. A(i))" apply (unfold 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)})" + "\\z. F \ stable({s \ state. P(s, z)}); F \ program\ + \ F \ stable({s \ state. \z. P(s, z)})" apply (unfold 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')" + "\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')" + "\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) *) +(* \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\\ @@ -566,15 +566,15 @@ apply (blast dest: stable_type [THEN subsetD]) done -lemma invariantI: "[| Init(F)\A; F \ stable(A) |] ==> F \ invariant(A)" +lemma invariantI: "\Init(F)\A; F \ stable(A)\ \ F \ invariant(A)" apply (unfold invariant_def initially_def) apply (frule stable_type [THEN subsetD], auto) done -lemma invariantD: "F \ invariant(A) ==> Init(F)\A & F \ stable(A)" +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)" +lemma invariantD2: "F \ invariant(A) \ F \ program & st_set(A)" apply (unfold invariant_def) apply (blast dest: stableD2) done @@ -582,7 +582,7 @@ 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)" + "\F \ invariant(A); F \ invariant(B)\ \ F \ invariant(A \ B)" apply (unfold invariant_def initially_def) apply (simp add: stable_Int, blast) done @@ -591,30 +591,30 @@ subsection\The Elimination Theorem\ (** The "free" m has become universally quantified! - Should the premise be !!m instead of \m ? Would make it harder + 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))" + "\\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))" + "\\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))" + "\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" + "\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 \ diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/Union.thy --- a/src/ZF/UNITY/Union.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/Union.thy Tue Sep 27 16:51:35 2022 +0100 @@ -16,28 +16,28 @@ definition (*FIXME: conjoin Init(F) \ Init(G) \ 0 *) ok :: "[i, i] => o" (infixl \ok\ 65) where - "F ok G == Acts(F) \ AllowedActs(G) & + "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)))" + "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 + "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), + "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 & + "safety_prop(X) \ X\program & SKIP \ X & (\G \ program. Acts(G) \ (\F \ X. Acts(F)) \ G \ X)" syntax @@ -45,7 +45,7 @@ "_JOIN" :: "[pttrn, i, i] => i" (\(3\_ \ _./ _)\ 10) translations - "\x \ A. B" == "CONST JOIN(A, (\x. B))" + "\x \ A. B" == "CONST JOIN(A, (\x. B))" "\x y. B" == "\x. \y. B" "\x. B" == "CONST JOIN(CONST state, (\x. B))" @@ -177,14 +177,14 @@ by (rule program_equalityI, auto) lemma JOIN_cong [cong]: - "[| I=J; !!i. i \ J ==> F(i) = G(i) |] ==> + "\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))" +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 @@ -208,7 +208,7 @@ 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)" +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 @@ -221,7 +221,7 @@ 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)" + "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 @@ -245,8 +245,8 @@ *) lemma Join_constrains_weaken: - "[| F \ A co A'; G \ B co B' |] - ==> F \ G \ (A \ B) co (A' \ B')" + "\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) @@ -254,7 +254,7 @@ (*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))" + 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) @@ -274,7 +274,7 @@ done lemma initially_JOIN_I: - assumes major: "(!!i. i \ I ==>F(i) \ initially(A))" + assumes major: "(\i. i \ I \F(i) \ initially(A))" and minor: "i \ I" shows "(\i \ I. F(i)) \ initially(A)" apply (cut_tac minor) @@ -284,7 +284,7 @@ done lemma invariant_JOIN_I: - assumes major: "(!!i. i \ I ==> F(i) \ invariant(A))" + assumes major: "(\i. i \ I \ F(i) \ invariant(A))" and minor: "i \ I" shows "(\i \ I. F(i)) \ invariant(A)" apply (cut_tac minor) @@ -302,12 +302,12 @@ by (simp add: stable_def) lemma initially_JoinI [intro!]: - "[| F \ initially(A); G \ initially(A) |] ==> F \ G \ initially(A)" + "\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)" + "\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) @@ -316,13 +316,13 @@ (* 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))))" +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 \ (\i \ I. F(i)) \ transient(A) \ (\i \ I. programify(F(i)) \ transient(A))" apply (auto simp add: transient_def JOIN_def) apply (unfold st_set_def) @@ -336,16 +336,16 @@ apply (auto simp add: transient_def Join_def Int_Un_distrib2) done -lemma Join_transient_I1: "F \ transient(A) ==> F \ G \ transient(A)" +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)" +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) *) +(*If I=0 it degenerates to (SKIP \ A ensures B) = False, i.e. to \(A\B) *) lemma JOIN_ensures: - "i \ I ==> + "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))" @@ -362,8 +362,8 @@ done lemma stable_Join_constrains: - "[| F \ stable(A); G \ A co A' |] - ==> F \ G \ A co A'" + "\F \ stable(A); G \ A co A'\ + \ F \ G \ A co A'" apply (unfold 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) @@ -372,7 +372,7 @@ (*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)" + "\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) @@ -381,7 +381,7 @@ (*As above, but exchanging the roles of F and G*) lemma stable_Join_Always2: - "[| F \ invariant(A); G \ stable(A) |] ==> F \ G \ Always(A)" + "\F \ invariant(A); G \ stable(A)\ \ F \ G \ Always(A)" apply (subst Join_commute) apply (blast intro: stable_Join_Always1) done @@ -389,7 +389,7 @@ lemma stable_Join_ensures1: - "[| F \ stable(A); G \ A ensures B |] ==> F \ G \ A ensures B" + "\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) @@ -400,7 +400,7 @@ (*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" + "\F \ A ensures B; G \ stable(A)\ \ F \ G \ A ensures B" apply (subst Join_commute) apply (blast intro: stable_Join_ensures1) done @@ -433,7 +433,7 @@ 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)" +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))" @@ -447,7 +447,7 @@ 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)" +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) @@ -474,7 +474,7 @@ done lemma Allowed_JOIN [simp]: - "i \ I ==> + "i \ I \ Allowed(JOIN(I,F)) = (\i \ I. Allowed(programify(F(i))))" apply (auto simp add: Allowed_def, blast) done @@ -494,7 +494,7 @@ 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)" + "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) @@ -503,7 +503,7 @@ done lemma safety_prop_AllowedActs_iff_Allowed: - "safety_prop(X) ==> + "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) @@ -511,7 +511,7 @@ lemma Allowed_eq: - "safety_prop(X) ==> Allowed(mk_program(init, acts, \F \ X. Acts(F))) = X" + "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) @@ -519,8 +519,8 @@ done lemma def_prg_Allowed: - "[| F == mk_program (init, acts, \F \ X. Acts(F)); safety_prop(X) |] - ==> Allowed(F) = X" + "\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!*) @@ -530,17 +530,17 @@ (* To be used with resolution *) lemma safety_prop_constrainsI [iff]: - "[| A\B; st_set(A) |] ==>safety_prop(A co B)" + "\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))" +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)" + "\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) @@ -549,7 +549,7 @@ (* 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)))" + 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) @@ -565,8 +565,8 @@ 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))" +"\F \ mk_program(init,acts, \G \ X. Acts(G)); safety_prop(X)\ + \ F ok G \ (programify(G) \ X & acts \ Pow(state*state) \ AllowedActs(G))" apply (unfold ok_def) apply (drule_tac G = G in safety_prop_Acts_iff) apply (cut_tac F = G in AllowedActs_type) diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/WFair.thy --- a/src/ZF/UNITY/WFair.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/WFair.thy Tue Sep 27 16:51:35 2022 +0100 @@ -17,12 +17,12 @@ (* 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) & + "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)" + "A ensures B \ ((A-B) co (A \ B)) \ transient(A-B)" consts @@ -33,9 +33,9 @@ domains "leads(D, F)" \ "Pow(D)*Pow(D)" intros - Basis: "[| F \ A ensures B; A \ Pow(D); B \ Pow(D) |] ==> :leads(D, F)" - Trans: "[| \ leads(D, F); \ leads(D, F) |] ==> :leads(D, F)" - Union: "[| S \ Pow({A \ S. :leads(D, F)}); B \ Pow(D); S \ Pow(Pow(D)) |] ==> + Basis: "\F \ A ensures B; A \ Pow(D); B \ Pow(D)\ \ :leads(D, F)" + Trans: "\ \ leads(D, F); \ leads(D, F)\ \ :leads(D, F)" + Union: "\S \ Pow({A \ S. :leads(D, F)}); B \ Pow(D); S \ Pow(Pow(D))\ \ <\(S),B>:leads(D, F)" monos Pow_mono @@ -44,12 +44,12 @@ definition (* The Visible version of the LEADS-TO relation*) leadsTo :: "[i, i] => i" (infixl \\\ 60) where - "A \ B == {F \ program. :leads(state, F)}" + "A \ B \ {F \ program. :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})" + "wlt(F, B) \ \({A \ Pow(state). F \ A \ B})" (** Ad-hoc set-theory rules **) @@ -65,27 +65,27 @@ by (unfold transient_def, auto) lemma transientD2: -"F \ transient(A) ==> F \ program & st_set(A)" +"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" +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)" +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)" +"\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" + "\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" @@ -96,7 +96,7 @@ apply (auto intro: st0_in_state) done -lemma transient_state2: "state<=B ==> transient(B) = 0" +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) @@ -115,36 +115,36 @@ by (simp add: ensures_def constrains_def, auto) lemma ensuresI: -"[|F:(A-B) co (A \ B); F \ transient(A-B)|]==>F \ A ensures B" +"\F:(A-B) co (A \ B); F \ transient(A-B)\\F \ A ensures B" apply (unfold 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" +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)" +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'" +lemma ensures_weaken_R: "\F \ A ensures A'; A'<=B'\ \ F \ A ensures B'" apply (unfold 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)" + "\F \ stable(C); F \ A ensures B\ \ F:(C \ A) ensures (C \ B)" apply (unfold 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" +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) @@ -153,7 +153,7 @@ 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" +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 ***) @@ -164,13 +164,13 @@ by (unfold leadsTo_def, auto) lemma leadsToD2: -"F \ A \ B ==> F \ program & st_set(A) & st_set(B)" +"F \ A \ B \ F \ program & st_set(A) & st_set(B)" apply (unfold 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" + "\F \ A ensures B; st_set(A); st_set(B)\ \ F \ A \ B" apply (unfold leadsTo_def st_set_def) apply (cut_tac ensures_type) apply (auto intro: leads.Basis) @@ -178,66 +178,66 @@ 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 *) +(* \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" +lemma leadsTo_Trans: "\F \ A \ B; F \ B \ C\\F \ A \ C" apply (unfold 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)" +lemma transient_imp_leadsTo: "F \ transient(A) \ F \ A \ (state-A)" apply (unfold 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'" +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)" + "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" + "\\A. A \ S \ F \ A \ B; F \ program; st_set(B)\ + \ F \ \(S) \ B" apply (unfold 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" + "\\A. A \ S \F \ (A \ C) \ B; F \ program; st_set(B)\ + \ F \ (\(S)Int C)\ B" apply (unfold 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" + "\\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" + "\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" + "\\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" +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)" @@ -251,21 +251,21 @@ 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'" +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'" +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'" +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" +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) @@ -285,45 +285,45 @@ 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" + "\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))" + "\\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')" +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')" +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')" +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')" +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')" + "\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) @@ -332,11 +332,11 @@ (*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)" + 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) @@ -350,13 +350,13 @@ (* 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)" + 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) @@ -379,11 +379,11 @@ (** 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; + "\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)" + \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\ @@ -394,11 +394,11 @@ lemma leadsTo_induct_pre: - "[| F \ za \ zb; + "\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)" + \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) @@ -410,7 +410,7 @@ (** The impossibility law **) lemma leadsTo_empty: - "F \ A \ 0 ==> A=0" + "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)+ @@ -423,7 +423,7 @@ text\Special case of PSP: Misra's "stable conjunction"\ lemma psp_stable: - "[| F \ A \ A'; F \ stable(B) |] ==> F:(A \ B) \ (A' \ B)" + "\F \ A \ A'; F \ stable(B)\ \ F:(A \ B) \ (A' \ B)" apply (unfold stable_def) apply (frule leadsToD2) apply (erule leadsTo_induct) @@ -435,12 +435,12 @@ done -lemma psp_stable2: "[|F \ A \ A'; F \ stable(B) |]==>F: (B \ A) \ (B \ A')" +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))" +"\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) (*speeds up the proof*) apply clarify @@ -448,7 +448,7 @@ done lemma psp: -"[|F \ A \ A'; F \ B co B'; st_set(B')|]==> F:(A \ B') \ ((A' \ B) \ (B' - B))" +"\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) @@ -463,13 +463,13 @@ done -lemma psp2: "[| F \ A \ A'; F \ B co B'; st_set(B') |] - ==> F \ (B' \ A) \ ((B \ A') \ (B' - B))" +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')" + "\F \ A \ A'; F \ B unless B'; st_set(B); st_set(B')\ + \ F \ (A \ B) \ ((A' \ B) \ B')" apply (unfold unless_def) apply (subgoal_tac "st_set (A) &st_set (A') ") prefer 2 apply (blast dest: leadsToD2) @@ -481,13 +481,13 @@ 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); +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" + ((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) @@ -497,13 +497,13 @@ done (** Meta or object quantifier ? **) -lemma leadsTo_wf_induct: "[| wf(r); +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" + ((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) @@ -522,7 +522,7 @@ done -lemma Image_inverse_lessThan: "k measure(A, %x. x) -`` {k} = k" +lemma Image_inverse_lessThan: "k measure(A, %x. x) -`` {k} = k" apply (rule equalityI) apply (auto simp add: measure_def) apply (blast intro: ltD) @@ -534,10 +534,10 @@ (*Alternative proof is via the lemma F \ (A \ f-`(lessThan m)) \ B*) lemma lessThan_induct: - "[| A<=f-``nat; + "\A<=f-``nat; F \ program; st_set(A); st_set(B); - \m \ nat. F:(A \ f-``{m}) \ ((A \ f -`` m) \ B) |] - ==> F \ A \ 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]) @@ -561,10 +561,10 @@ apply (blast dest: leadsToD2 intro!: leadsTo_Union) done -(* [| F \ program; st_set(B) |] ==> F \ wlt(F, B) \ B *) +(* \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)" +lemma leadsTo_subset: "F \ A \ B \ A \ wlt(F, B)" apply (unfold wlt_def) apply (frule leadsToD2) apply (auto simp add: st_set_def) @@ -577,24 +577,24 @@ done (*Misra's property W4*) -lemma wlt_increasing: "[| F \ program; st_set(B) |] ==> B \ wlt(F,B)" +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; + "\B \ A2; F \ (A1 - B) co (A1 \ B); - F \ (A2 - C) co (A2 \ C) |] - ==> F \ (A1 \ A2 - C) co (A1 \ A2 \ C)" + 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')" + \ \B \ Pow(state). A<=B & F \ B \ A' & F \ (B-A') co (B \ A')" apply (frule leadsToD2) apply (erule leadsTo_induct) txt\Basis\ @@ -619,7 +619,7 @@ (*Misra's property W5*) -lemma wlt_constrains_wlt: "[| F \ program; st_set(B) |] ==>F \ (wlt(F, B) - B) co (wlt(F,B))" +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) ") @@ -630,10 +630,10 @@ subsection\Completion: Binary and General Finite versions\ -lemma completion_aux: "[| W = wlt(F, (B' \ C)); +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)" + 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 @@ -665,7 +665,7 @@ lemmas completion = refl [THEN completion_aux] lemma finite_completion_aux: - "[| I \ Fin(X); F \ program; st_set(C) |] ==> + "\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)" @@ -677,16 +677,16 @@ 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)" + "\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')" + "\F \ A \ A'; F \ stable(A'); + F \ B \ B'; F \ stable(B')\ + \ F \ (A \ B) \ (A' \ B')" apply (unfold stable_def) apply (rule_tac C1 = 0 in completion [THEN leadsTo_weaken_R], simp+) apply (blast dest: leadsToD2) @@ -694,10 +694,10 @@ 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))" + "\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))" apply (unfold stable_def) apply (subgoal_tac "st_set (\i \ I. A' (i))") prefer 2 apply (blast dest: leadsToD2) diff -r f2094906e491 -r e44d86131648 src/ZF/Univ.thy --- a/src/ZF/Univ.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Univ.thy Tue Sep 27 16:51:35 2022 +0100 @@ -15,26 +15,26 @@ definition Vfrom :: "[i,i]=>i" where - "Vfrom(A,i) == transrec(i, %x f. A \ (\y\x. Pow(f`y)))" + "Vfrom(A,i) \ transrec(i, %x f. A \ (\y\x. Pow(f`y)))" abbreviation Vset :: "i=>i" where - "Vset(x) == Vfrom(0,x)" + "Vset(x) \ Vfrom(0,x)" definition Vrec :: "[i, [i,i]=>i] =>i" where - "Vrec(a,H) == transrec(rank(a), %x g. \z\Vset(succ(x)). + "Vrec(a,H) \ transrec(rank(a), %x g. \z\Vset(succ(x)). H(z, \w\Vset(x). g`rank(w)`w)) ` a" definition Vrecursor :: "[[i,i]=>i, i] =>i" where - "Vrecursor(H,a) == transrec(rank(a), %x g. \z\Vset(succ(x)). + "Vrecursor(H,a) \ transrec(rank(a), %x g. \z\Vset(succ(x)). H(\w\Vset(x). g`rank(w)`w, z)) ` a" definition univ :: "i=>i" where - "univ(A) == Vfrom(A,nat)" + "univ(A) \ Vfrom(A,nat)" subsection\Immediate Consequences of the Definition of \<^term>\Vfrom(A,i)\\ @@ -46,7 +46,7 @@ subsubsection\Monotonicity\ lemma Vfrom_mono [rule_format]: - "A<=B ==> \j. i<=j \ Vfrom(A,i) \ Vfrom(B,j)" + "A<=B \ \j. i<=j \ Vfrom(A,i) \ Vfrom(B,j)" apply (rule_tac a=i in eps_induct) apply (rule impI [THEN allI]) apply (subst Vfrom [of A]) @@ -55,7 +55,7 @@ apply (erule UN_mono, blast) done -lemma VfromI: "[| a \ Vfrom(A,j); j a \ Vfrom(A,i)" +lemma VfromI: "\a \ Vfrom(A,j); j \ a \ Vfrom(A,i)" by (blast dest: Vfrom_mono [OF subset_refl le_imp_subset [OF leI]]) @@ -96,7 +96,7 @@ subsection\Basic Closure Properties\ -lemma zero_in_Vfrom: "y:x ==> 0 \ Vfrom(A,x)" +lemma zero_in_Vfrom: "y:x \ 0 \ Vfrom(A,x)" by (subst Vfrom, blast) lemma i_subset_Vfrom: "i \ Vfrom(A,i)" @@ -111,25 +111,25 @@ lemmas A_into_Vfrom = A_subset_Vfrom [THEN subsetD] -lemma subset_mem_Vfrom: "a \ Vfrom(A,i) ==> a \ Vfrom(A,succ(i))" +lemma subset_mem_Vfrom: "a \ Vfrom(A,i) \ a \ Vfrom(A,succ(i))" by (subst Vfrom, blast) subsubsection\Finite sets and ordered pairs\ -lemma singleton_in_Vfrom: "a \ Vfrom(A,i) ==> {a} \ Vfrom(A,succ(i))" +lemma singleton_in_Vfrom: "a \ Vfrom(A,i) \ {a} \ Vfrom(A,succ(i))" by (rule subset_mem_Vfrom, safe) lemma doubleton_in_Vfrom: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i) |] ==> {a,b} \ Vfrom(A,succ(i))" + "\a \ Vfrom(A,i); b \ Vfrom(A,i)\ \ {a,b} \ Vfrom(A,succ(i))" by (rule subset_mem_Vfrom, safe) lemma Pair_in_Vfrom: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i) |] ==> \ Vfrom(A,succ(succ(i)))" + "\a \ Vfrom(A,i); b \ Vfrom(A,i)\ \ \ Vfrom(A,succ(succ(i)))" apply (unfold Pair_def) apply (blast intro: doubleton_in_Vfrom) done -lemma succ_in_Vfrom: "a \ Vfrom(A,i) ==> succ(a) \ Vfrom(A,succ(succ(i)))" +lemma succ_in_Vfrom: "a \ Vfrom(A,i) \ succ(a) \ Vfrom(A,succ(succ(i)))" apply (intro subset_mem_Vfrom succ_subsetI, assumption) apply (erule subset_trans) apply (rule Vfrom_mono [OF subset_refl subset_succI]) @@ -140,7 +140,7 @@ lemma Vfrom_0: "Vfrom(A,0) = A" by (subst Vfrom, blast) -lemma Vfrom_succ_lemma: "Ord(i) ==> Vfrom(A,succ(i)) = A \ Pow(Vfrom(A,i))" +lemma Vfrom_succ_lemma: "Ord(i) \ Vfrom(A,succ(i)) = A \ Pow(Vfrom(A,i))" apply (rule Vfrom [THEN trans]) apply (rule equalityI [THEN subst_context, OF _ succI1 [THEN RepFunI, THEN Union_upper]]) @@ -159,7 +159,7 @@ (*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces the conclusion to be Vfrom(A,\(X)) = A \ (\y\X. Vfrom(A,y)) *) -lemma Vfrom_Union: "y:X ==> Vfrom(A,\(X)) = (\y\X. Vfrom(A,y))" +lemma Vfrom_Union: "y:X \ Vfrom(A,\(X)) = (\y\X. Vfrom(A,y))" apply (subst Vfrom) apply (rule equalityI) txt\first inclusion\ @@ -181,15 +181,15 @@ (*NB. limit ordinals are non-empty: Vfrom(A,0) = A = A \ (\y\0. Vfrom(A,y)) *) lemma Limit_Vfrom_eq: - "Limit(i) ==> Vfrom(A,i) = (\y\i. Vfrom(A,y))" + "Limit(i) \ Vfrom(A,i) = (\y\i. Vfrom(A,y))" apply (rule Limit_has_0 [THEN ltD, THEN Vfrom_Union, THEN subst], assumption) apply (simp add: Limit_Union_eq) done lemma Limit_VfromE: - "[| a \ Vfrom(A,i); ~R ==> Limit(i); - !!x. [| x Vfrom(A,x) |] ==> R - |] ==> R" + "\a \ Vfrom(A,i); \R \ Limit(i); + \x. \x Vfrom(A,x)\ \ R +\ \ R" apply (rule classical) apply (rule Limit_Vfrom_eq [THEN equalityD1, THEN subsetD, THEN UN_E]) prefer 2 apply assumption @@ -198,7 +198,7 @@ done lemma singleton_in_VLimit: - "[| a \ Vfrom(A,i); Limit(i) |] ==> {a} \ Vfrom(A,i)" + "\a \ Vfrom(A,i); Limit(i)\ \ {a} \ Vfrom(A,i)" apply (erule Limit_VfromE, assumption) apply (erule singleton_in_Vfrom [THEN VfromI]) apply (blast intro: Limit_has_succ) @@ -211,7 +211,7 @@ text\Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)\ lemma doubleton_in_VLimit: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i) |] ==> {a,b} \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i)\ \ {a,b} \ Vfrom(A,i)" apply (erule Limit_VfromE, assumption) apply (erule Limit_VfromE, assumption) apply (blast intro: VfromI [OF doubleton_in_Vfrom] @@ -219,7 +219,7 @@ done lemma Pair_in_VLimit: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i) |] ==> \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i)\ \ \ Vfrom(A,i)" txt\Infer that a, b occur at ordinals x,xa < i.\ apply (erule Limit_VfromE, assumption) apply (erule Limit_VfromE, assumption) @@ -228,7 +228,7 @@ Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt) done -lemma product_VLimit: "Limit(i) ==> Vfrom(A,i) * Vfrom(A,i) \ Vfrom(A,i)" +lemma product_VLimit: "Limit(i) \ Vfrom(A,i) * Vfrom(A,i) \ Vfrom(A,i)" by (blast intro: Pair_in_VLimit) lemmas Sigma_subset_VLimit = @@ -237,29 +237,29 @@ lemmas nat_subset_VLimit = subset_trans [OF nat_le_Limit [THEN le_imp_subset] i_subset_Vfrom] -lemma nat_into_VLimit: "[| n: nat; Limit(i) |] ==> n \ Vfrom(A,i)" +lemma nat_into_VLimit: "\n: nat; Limit(i)\ \ n \ Vfrom(A,i)" by (blast intro: nat_subset_VLimit [THEN subsetD]) subsubsection\Closure under Disjoint Union\ lemmas zero_in_VLimit = Limit_has_0 [THEN ltD, THEN zero_in_Vfrom] -lemma one_in_VLimit: "Limit(i) ==> 1 \ Vfrom(A,i)" +lemma one_in_VLimit: "Limit(i) \ 1 \ Vfrom(A,i)" by (blast intro: nat_into_VLimit) lemma Inl_in_VLimit: - "[| a \ Vfrom(A,i); Limit(i) |] ==> Inl(a) \ Vfrom(A,i)" + "\a \ Vfrom(A,i); Limit(i)\ \ Inl(a) \ Vfrom(A,i)" apply (unfold Inl_def) apply (blast intro: zero_in_VLimit Pair_in_VLimit) done lemma Inr_in_VLimit: - "[| b \ Vfrom(A,i); Limit(i) |] ==> Inr(b) \ Vfrom(A,i)" + "\b \ Vfrom(A,i); Limit(i)\ \ Inr(b) \ Vfrom(A,i)" apply (unfold Inr_def) apply (blast intro: one_in_VLimit Pair_in_VLimit) done -lemma sum_VLimit: "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) \ Vfrom(C,i)" +lemma sum_VLimit: "Limit(i) \ Vfrom(C,i)+Vfrom(C,i) \ Vfrom(C,i)" by (blast intro!: Inl_in_VLimit Inr_in_VLimit) lemmas sum_subset_VLimit = subset_trans [OF sum_mono sum_VLimit] @@ -268,14 +268,14 @@ subsection\Properties assuming \<^term>\Transset(A)\\ -lemma Transset_Vfrom: "Transset(A) ==> Transset(Vfrom(A,i))" +lemma Transset_Vfrom: "Transset(A) \ Transset(Vfrom(A,i))" apply (rule_tac a=i in eps_induct) apply (subst Vfrom) apply (blast intro!: Transset_Union_family Transset_Un Transset_Pow) done lemma Transset_Vfrom_succ: - "Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))" + "Transset(A) \ Vfrom(A, succ(i)) = Pow(Vfrom(A,i))" apply (rule Vfrom_succ [THEN trans]) apply (rule equalityI [OF _ Un_upper2]) apply (rule Un_least [OF _ subset_refl]) @@ -283,26 +283,26 @@ apply (erule Transset_Vfrom [THEN Transset_iff_Pow [THEN iffD1]]) done -lemma Transset_Pair_subset: "[| \ C; Transset(C) |] ==> a: C & b: C" +lemma Transset_Pair_subset: "\ \ C; Transset(C)\ \ a: C & b: C" by (unfold Pair_def Transset_def, blast) lemma Transset_Pair_subset_VLimit: - "[| \ Vfrom(A,i); Transset(A); Limit(i) |] - ==> \ Vfrom(A,i)" + "\ \ Vfrom(A,i); Transset(A); Limit(i)\ + \ \ Vfrom(A,i)" apply (erule Transset_Pair_subset [THEN conjE]) apply (erule Transset_Vfrom) apply (blast intro: Pair_in_VLimit) done lemma Union_in_Vfrom: - "[| X \ Vfrom(A,j); Transset(A) |] ==> \(X) \ Vfrom(A, succ(j))" + "\X \ Vfrom(A,j); Transset(A)\ \ \(X) \ Vfrom(A, succ(j))" apply (drule Transset_Vfrom) apply (rule subset_mem_Vfrom) apply (unfold Transset_def, blast) done lemma Union_in_VLimit: - "[| X \ Vfrom(A,i); Limit(i); Transset(A) |] ==> \(X) \ Vfrom(A,i)" + "\X \ Vfrom(A,i); Limit(i); Transset(A)\ \ \(X) \ Vfrom(A,i)" apply (rule Limit_VfromE, assumption+) apply (blast intro: Limit_has_succ VfromI Union_in_Vfrom) done @@ -315,10 +315,10 @@ text\General theorem for membership in Vfrom(A,i) when i is a limit ordinal\ lemma in_VLimit: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); - !!x y j. [| j Vfrom(A,j); y \ Vfrom(A,j) |] - ==> \k. h(x,y) \ Vfrom(A,k) & k h(a,b) \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); + \x y j. \j Vfrom(A,j); y \ Vfrom(A,j)\ + \ \k. h(x,y) \ Vfrom(A,k) & k + \ h(a,b) \ Vfrom(A,i)" txt\Infer that a, b occur at ordinals x,xa < i.\ apply (erule Limit_VfromE, assumption) apply (erule Limit_VfromE, assumption, atomize) @@ -332,8 +332,8 @@ subsubsection\Products\ lemma prod_in_Vfrom: - "[| a \ Vfrom(A,j); b \ Vfrom(A,j); Transset(A) |] - ==> a*b \ Vfrom(A, succ(succ(succ(j))))" + "\a \ Vfrom(A,j); b \ Vfrom(A,j); Transset(A)\ + \ a*b \ Vfrom(A, succ(succ(succ(j))))" apply (drule Transset_Vfrom) apply (rule subset_mem_Vfrom) apply (unfold Transset_def) @@ -341,8 +341,8 @@ done lemma prod_in_VLimit: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); Transset(A) |] - ==> a*b \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); Transset(A)\ + \ a*b \ Vfrom(A,i)" apply (erule in_VLimit, assumption+) apply (blast intro: prod_in_Vfrom Limit_has_succ) done @@ -350,8 +350,8 @@ subsubsection\Disjoint Sums, or Quine Ordered Pairs\ lemma sum_in_Vfrom: - "[| a \ Vfrom(A,j); b \ Vfrom(A,j); Transset(A); 1:j |] - ==> a+b \ Vfrom(A, succ(succ(succ(j))))" + "\a \ Vfrom(A,j); b \ Vfrom(A,j); Transset(A); 1:j\ + \ a+b \ Vfrom(A, succ(succ(succ(j))))" apply (unfold sum_def) apply (drule Transset_Vfrom) apply (rule subset_mem_Vfrom) @@ -360,8 +360,8 @@ done lemma sum_in_VLimit: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); Transset(A) |] - ==> a+b \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); Transset(A)\ + \ a+b \ Vfrom(A,i)" apply (erule in_VLimit, assumption+) apply (blast intro: sum_in_Vfrom Limit_has_succ) done @@ -369,7 +369,7 @@ subsubsection\Function Space!\ lemma fun_in_Vfrom: - "[| a \ Vfrom(A,j); b \ Vfrom(A,j); Transset(A) |] ==> + "\a \ Vfrom(A,j); b \ Vfrom(A,j); Transset(A)\ \ a->b \ Vfrom(A, succ(succ(succ(succ(j)))))" apply (unfold Pi_def) apply (drule Transset_Vfrom) @@ -385,14 +385,14 @@ done lemma fun_in_VLimit: - "[| a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); Transset(A) |] - ==> a->b \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i); Transset(A)\ + \ a->b \ Vfrom(A,i)" apply (erule in_VLimit, assumption+) apply (blast intro: fun_in_Vfrom Limit_has_succ) done lemma Pow_in_Vfrom: - "[| a \ Vfrom(A,j); Transset(A) |] ==> Pow(a) \ Vfrom(A, succ(succ(j)))" + "\a \ Vfrom(A,j); Transset(A)\ \ Pow(a) \ Vfrom(A, succ(succ(j)))" apply (drule Transset_Vfrom) apply (rule subset_mem_Vfrom) apply (unfold Transset_def) @@ -400,7 +400,7 @@ done lemma Pow_in_VLimit: - "[| a \ Vfrom(A,i); Limit(i); Transset(A) |] ==> Pow(a) \ Vfrom(A,i)" + "\a \ Vfrom(A,i); Limit(i); Transset(A)\ \ Pow(a) \ Vfrom(A,i)" by (blast elim: Limit_VfromE intro: Limit_has_succ Pow_in_Vfrom VfromI) @@ -414,7 +414,7 @@ subsubsection\Characterisation of the elements of \<^term>\Vset(i)\\ -lemma VsetD [rule_format]: "Ord(i) ==> \b. b \ Vset(i) \ rank(b) < i" +lemma VsetD [rule_format]: "Ord(i) \ \b. b \ Vset(i) \ rank(b) < i" apply (erule trans_induct) apply (subst Vset, safe) apply (subst rank) @@ -422,18 +422,18 @@ done lemma VsetI_lemma [rule_format]: - "Ord(i) ==> \b. rank(b) \ i \ b \ Vset(i)" + "Ord(i) \ \b. rank(b) \ i \ b \ Vset(i)" apply (erule trans_induct) apply (rule allI) apply (subst Vset) apply (blast intro!: rank_lt [THEN ltD]) done -lemma VsetI: "rank(x) x \ Vset(i)" +lemma VsetI: "rank(x) x \ Vset(i)" by (blast intro: VsetI_lemma elim: ltE) text\Merely a lemma for the next result\ -lemma Vset_Ord_rank_iff: "Ord(i) ==> b \ Vset(i) \ rank(b) < i" +lemma Vset_Ord_rank_iff: "Ord(i) \ b \ Vset(i) \ rank(b) < i" by (blast intro: VsetD VsetI) lemma Vset_rank_iff [simp]: "b \ Vset(a) \ rank(b) < rank(a)" @@ -444,7 +444,7 @@ text\This is rank(rank(a)) = rank(a)\ declare Ord_rank [THEN rank_of_Ord, simp] -lemma rank_Vset: "Ord(i) ==> rank(Vset(i)) = i" +lemma rank_Vset: "Ord(i) \ rank(Vset(i)) = i" apply (subst rank) apply (rule equalityI, safe) apply (blast intro: VsetD [THEN ltD]) @@ -453,7 +453,7 @@ Ord_in_Ord [THEN rank_of_Ord, THEN ssubst]) done -lemma Finite_Vset: "i \ nat ==> Finite(Vset(i))" +lemma Finite_Vset: "i \ nat \ Finite(Vset(i))" apply (erule nat_induct) apply (simp add: Vfrom_0) apply (simp add: Vset_succ) @@ -467,7 +467,7 @@ done lemma Int_Vset_subset: - "[| !!i. Ord(i) ==> a \ Vset(i) \ b |] ==> a \ b" + "\\i. Ord(i) \ a \ Vset(i) \ b\ \ a \ b" apply (rule subset_trans) apply (rule Int_greatest [OF subset_refl arg_subset_Vset_rank]) apply (blast intro: Ord_rank) @@ -496,9 +496,9 @@ apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def) done -text\This form avoids giant explosions in proofs. NOTE USE OF ==\ +text\This form avoids giant explosions in proofs. NOTE USE OF \\ lemma def_Vrec: - "[| !!x. h(x)==Vrec(x,H) |] ==> + "\\x. h(x)\Vrec(x,H)\ \ h(a) = H(a, \x\Vset(rank(a)). h(x))" apply simp apply (rule Vrec) @@ -512,9 +512,9 @@ apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def) done -text\This form avoids giant explosions in proofs. NOTE USE OF ==\ +text\This form avoids giant explosions in proofs. NOTE USE OF \\ lemma def_Vrecursor: - "h == Vrecursor(H) ==> h(a) = H(\x\Vset(rank(a)). h(x), a)" + "h \ Vrecursor(H) \ h(a) = H(\x\Vset(rank(a)). h(x), a)" apply simp apply (rule Vrecursor) done @@ -522,13 +522,13 @@ subsection\The Datatype Universe: \<^term>\univ(A)\\ -lemma univ_mono: "A<=B ==> univ(A) \ univ(B)" +lemma univ_mono: "A<=B \ univ(A) \ univ(B)" apply (unfold univ_def) apply (erule Vfrom_mono) apply (rule subset_refl) done -lemma Transset_univ: "Transset(A) ==> Transset(univ(A))" +lemma Transset_univ: "Transset(A) \ Transset(univ(A))" apply (unfold univ_def) apply (erule Transset_Vfrom) done @@ -540,23 +540,23 @@ apply (rule Limit_nat [THEN Limit_Vfrom_eq]) done -lemma subset_univ_eq_Int: "c \ univ(A) ==> c = (\i\nat. c \ Vfrom(A,i))" +lemma subset_univ_eq_Int: "c \ univ(A) \ c = (\i\nat. c \ Vfrom(A,i))" apply (rule subset_UN_iff_eq [THEN iffD1]) apply (erule univ_eq_UN [THEN subst]) done lemma univ_Int_Vfrom_subset: - "[| a \ univ(X); - !!i. i:nat ==> a \ Vfrom(X,i) \ b |] - ==> a \ b" + "\a \ univ(X); + \i. i:nat \ a \ Vfrom(X,i) \ b\ + \ a \ b" apply (subst subset_univ_eq_Int, assumption) apply (rule UN_least, simp) done lemma univ_Int_Vfrom_eq: - "[| a \ univ(X); b \ univ(X); - !!i. i:nat ==> a \ Vfrom(X,i) = b \ Vfrom(X,i) - |] ==> a = b" + "\a \ univ(X); b \ univ(X); + \i. i:nat \ a \ Vfrom(X,i) = b \ Vfrom(X,i) +\ \ a = b" apply (rule equalityI) apply (rule univ_Int_Vfrom_subset, assumption) apply (blast elim: equalityCE) @@ -583,25 +583,25 @@ subsubsection\Closure under Unordered and Ordered Pairs\ -lemma singleton_in_univ: "a: univ(A) ==> {a} \ univ(A)" +lemma singleton_in_univ: "a: univ(A) \ {a} \ univ(A)" apply (unfold univ_def) apply (blast intro: singleton_in_VLimit Limit_nat) done lemma doubleton_in_univ: - "[| a: univ(A); b: univ(A) |] ==> {a,b} \ univ(A)" + "\a: univ(A); b: univ(A)\ \ {a,b} \ univ(A)" apply (unfold univ_def) apply (blast intro: doubleton_in_VLimit Limit_nat) done lemma Pair_in_univ: - "[| a: univ(A); b: univ(A) |] ==> \ univ(A)" + "\a: univ(A); b: univ(A)\ \ \ univ(A)" apply (unfold univ_def) apply (blast intro: Pair_in_VLimit Limit_nat) done lemma Union_in_univ: - "[| X: univ(A); Transset(A) |] ==> \(X) \ univ(A)" + "\X: univ(A); Transset(A)\ \ \(X) \ univ(A)" apply (unfold univ_def) apply (blast intro: Union_in_VLimit Limit_nat) done @@ -619,7 +619,7 @@ apply (rule i_subset_Vfrom) done -text\n:nat ==> n:univ(A)\ +text\n:nat \ n:univ(A)\ lemmas nat_into_univ = nat_subset_univ [THEN subsetD] subsubsection\Instances for 1 and 2\ @@ -643,12 +643,12 @@ subsubsection\Closure under Disjoint Union\ -lemma Inl_in_univ: "a: univ(A) ==> Inl(a) \ univ(A)" +lemma Inl_in_univ: "a: univ(A) \ Inl(a) \ univ(A)" apply (unfold univ_def) apply (erule Inl_in_VLimit [OF _ Limit_nat]) done -lemma Inr_in_univ: "b: univ(A) ==> Inr(b) \ univ(A)" +lemma Inr_in_univ: "b: univ(A) \ Inr(b) \ univ(A)" apply (unfold univ_def) apply (erule Inr_in_VLimit [OF _ Limit_nat]) done @@ -661,7 +661,7 @@ lemmas sum_subset_univ = subset_trans [OF sum_mono sum_univ] lemma Sigma_subset_univ: - "[|A \ univ(D); \x. x \ A \ B(x) \ univ(D)|] ==> Sigma(A,B) \ univ(D)" + "\A \ univ(D); \x. x \ A \ B(x) \ univ(D)\ \ Sigma(A,B) \ univ(D)" apply (simp add: univ_def) apply (blast intro: Sigma_subset_VLimit del: subsetI) done @@ -677,14 +677,14 @@ subsubsection\Closure under Finite Powerset\ lemma Fin_Vfrom_lemma: - "[| b: Fin(Vfrom(A,i)); Limit(i) |] ==> \j. b \ Vfrom(A,j) & jb: Fin(Vfrom(A,i)); Limit(i)\ \ \j. b \ Vfrom(A,j) & j Fin(Vfrom(A,i)) \ Vfrom(A,i)" +lemma Fin_VLimit: "Limit(i) \ Fin(Vfrom(A,i)) \ Vfrom(A,i)" apply (rule subsetI) apply (drule Fin_Vfrom_lemma, safe) apply (rule Vfrom [THEN ssubst]) @@ -701,7 +701,7 @@ subsubsection\Closure under Finite Powers: Functions from a Natural Number\ lemma nat_fun_VLimit: - "[| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) \ Vfrom(A,i)" + "\n: nat; Limit(i)\ \ n -> Vfrom(A,i) \ Vfrom(A,i)" apply (erule nat_fun_subset_Fin [THEN subset_trans]) apply (blast del: subsetI intro: subset_refl Fin_subset_VLimit Sigma_subset_VLimit nat_subset_VLimit) @@ -709,7 +709,7 @@ lemmas nat_fun_subset_VLimit = subset_trans [OF Pi_mono nat_fun_VLimit] -lemma nat_fun_univ: "n: nat ==> n -> univ(A) \ univ(A)" +lemma nat_fun_univ: "n: nat \ n -> univ(A) \ univ(A)" apply (unfold univ_def) apply (erule nat_fun_VLimit [OF _ Limit_nat]) done @@ -719,7 +719,7 @@ text\General but seldom-used version; normally the domain is fixed\ lemma FiniteFun_VLimit1: - "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) \ Vfrom(A,i)" + "Limit(i) \ Vfrom(A,i) -||> Vfrom(A,i) \ Vfrom(A,i)" apply (rule FiniteFun.dom_subset [THEN subset_trans]) apply (blast del: subsetI intro: Fin_subset_VLimit Sigma_subset_VLimit subset_refl) @@ -732,20 +732,20 @@ text\Version for a fixed domain\ lemma FiniteFun_VLimit: - "[| W \ Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) \ Vfrom(A,i)" + "\W \ Vfrom(A,i); Limit(i)\ \ W -||> Vfrom(A,i) \ Vfrom(A,i)" apply (rule subset_trans) apply (erule FiniteFun_mono [OF _ subset_refl]) apply (erule FiniteFun_VLimit1) done lemma FiniteFun_univ: - "W \ univ(A) ==> W -||> univ(A) \ univ(A)" + "W \ univ(A) \ W -||> univ(A) \ univ(A)" apply (unfold univ_def) apply (erule FiniteFun_VLimit [OF _ Limit_nat]) done lemma FiniteFun_in_univ: - "[| f: W -||> univ(A); W \ univ(A) |] ==> f \ univ(A)" + "\f: W -||> univ(A); W \ univ(A)\ \ f \ univ(A)" by (erule FiniteFun_univ [THEN subsetD], assumption) text\Remove \\\ from the rule above\ @@ -758,8 +758,8 @@ text\This version says a, b exist one level down, in the smaller set Vfrom(X,i)\ lemma doubleton_in_Vfrom_D: - "[| {a,b} \ Vfrom(X,succ(i)); Transset(X) |] - ==> a \ Vfrom(X,i) & b \ Vfrom(X,i)" + "\{a,b} \ Vfrom(X,succ(i)); Transset(X)\ + \ a \ Vfrom(X,i) & b \ Vfrom(X,i)" by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD], assumption, fast) @@ -775,14 +775,14 @@ **) lemma Pair_in_Vfrom_D: - "[| \ Vfrom(X,succ(i)); Transset(X) |] - ==> a \ Vfrom(X,i) & b \ Vfrom(X,i)" + "\ \ Vfrom(X,succ(i)); Transset(X)\ + \ a \ Vfrom(X,i) & b \ Vfrom(X,i)" apply (unfold Pair_def) apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D) done lemma product_Int_Vfrom_subset: - "Transset(X) ==> + "Transset(X) \ (a*b) \ Vfrom(X, succ(i)) \ (a \ Vfrom(X,i)) * (b \ Vfrom(X,i))" by (blast dest!: Pair_in_Vfrom_D) diff -r f2094906e491 -r e44d86131648 src/ZF/WF.thy --- a/src/ZF/WF.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/WF.thy Tue Sep 27 16:51:35 2022 +0100 @@ -21,58 +21,58 @@ definition wf :: "i=>o" where (*r is a well-founded relation*) - "wf(r) == \Z. Z=0 | (\x\Z. \y. :r \ ~ y \ Z)" + "wf(r) \ \Z. Z=0 | (\x\Z. \y. :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)" + "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}))))" + "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))" + "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))" + "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})))" + "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)" + "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)" +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)" +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)" +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)" +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)" +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)" +lemma wf_subset: "\wf(s); r<=s\ \ wf(r)" by (simp add: wf_def, fast) subsubsection\Introduction Rules for \<^term>\wf_on\\ @@ -80,7 +80,7 @@ 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. :r |] ==> False" + assumes prem: "\Z u. \Z<=A; u \ Z; \x\Z. \y\Z. :r\ \ False" shows "wf[A](r)" apply (unfold wf_on_def wf_def) apply (rule equals0I [THEN disjCI, THEN allI]) @@ -89,10 +89,10 @@ 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. : r \ y \ B) \ x \ B ==> A<=B\\ + \<^prop>\\B. \x\A. (\y. : r \ y \ B) \ x \ B \ A<=B\\ lemma wf_onI2: - assumes prem: "!!y B. [| \x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A |] - ==> y \ B" + assumes prem: "\y B. \\x\A. (\y\A. :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]) @@ -106,11 +106,11 @@ 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. : r \ P(y) |] ==> P(x) |] - ==> P(a)" + "\wf(r); + \x.\\y. : r \ P(y)\ \ P(x)\ + \ P(a)" apply (unfold wf_def) -apply (erule_tac x = "{z \ domain(r). ~ P(z)}" in allE) +apply (erule_tac x = "{z \ domain(r). \ P(z)}" in allE) apply blast done @@ -118,9 +118,9 @@ 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. : r \ P(y) |] ==> P(x) |] - ==> P(a)" + "\wf(r); a \ A; field(r)<=A; + \x.\x \ A; \y. : 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 @@ -129,9 +129,9 @@ by blast lemma wf_on_induct_raw [consumes 2, induct set: wf_on]: - "[| wf[A](r); a \ A; - !!x.[| x \ A; \y\A. : r \ P(y) |] ==> P(x) - |] ==> P(a)" + "\wf[A](r); a \ A; + \x.\x \ A; \y\A. : r \ P(y)\ \ P(x) +\ \ P(a)" apply (unfold wf_on_def) apply (erule wf_induct2, assumption) apply (rule field_Int_square, blast) @@ -144,10 +144,10 @@ 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. :r \ y \ B) \ x \ B; y \ A|] - ==> y \ B |] - ==> wf(r)" + "\field(r)<=A; + \y B. \\x\A. (\y\A. :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 @@ -157,34 +157,34 @@ subsection\Basic Properties of Well-Founded Relations\ -lemma wf_not_refl: "wf(r) ==> \ r" +lemma wf_not_refl: "wf(r) \ \ r" by (erule_tac a=a in wf_induct, blast) -lemma wf_not_sym [rule_format]: "wf(r) ==> \x. :r \ \ r" +lemma wf_not_sym [rule_format]: "wf(r) \ \x. :r \ \ r" by (erule_tac a=a in wf_induct, blast) -(* @{term"[| wf(r); \ r; ~P ==> \ r |] ==> P"} *) +(* @{term"\wf(r); \ r; \P \ \ r\ \ P"} *) lemmas wf_asym = wf_not_sym [THEN swap] -lemma wf_on_not_refl: "[| wf[A](r); a \ A |] ==> \ r" +lemma wf_on_not_refl: "\wf[A](r); 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 \ :r \ \r)" + "\wf[A](r); a \ A\ \ (\b. b\A \ :r \ \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 ==> \ r; - \ r ==> Z; ~Z ==> a \ A; ~Z ==> b \ A |] ==> Z" + "\wf[A](r); \Z \ \ r; + \ 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); :r; :r; :r; a \ A; b \ A; c \ A |] ==> P" + "\wf[A](r); :r; :r; :r; a \ A; b \ A; c \ A\ \ P" apply (subgoal_tac "\y\A. \z\A. :r \ :r \ :r \ P", blast) apply (erule_tac a=a in wf_on_induct, assumption, blast) @@ -195,14 +195,14 @@ 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^+)" + "\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^+)" +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) @@ -219,7 +219,7 @@ subsection\The Predicate \<^term>\is_recfun\\ -lemma is_recfun_type: "is_recfun(r,a,H,f) ==> f \ r-``{a} -> range(f)" +lemma is_recfun_type: "is_recfun(r,a,H,f) \ f \ r-``{a} -> range(f)" apply (unfold is_recfun_def) apply (erule ssubst) apply (rule lamI [THEN rangeI, THEN lam_type], assumption) @@ -228,7 +228,7 @@ lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function] lemma apply_recfun: - "[| is_recfun(r,a,H,f); :r |] ==> f`x = H(x, restrict(f,r-``{x}))" + "\is_recfun(r,a,H,f); :r\ \ f`x = H(x, restrict(f,r-``{x}))" apply (unfold 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) @@ -236,8 +236,8 @@ done lemma is_recfun_equal [rule_format]: - "[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] - ==> :r \ :r \ f`x=g`x" + "\wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g)\ + \ :r \ :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) @@ -253,9 +253,9 @@ done lemma is_recfun_cut: - "[| wf(r); trans(r); - is_recfun(r,a,H,f); is_recfun(r,b,H,g); :r |] - ==> restrict(f, r-``{b}) = g" + "\wf(r); trans(r); + is_recfun(r,a,H,f); is_recfun(r,b,H,g); :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) @@ -266,23 +266,23 @@ 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" + "\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" + "\is_recfun(r,a,H,f); wf(r); trans(r)\ \ the_recfun(r,a,H) = f" apply (unfold 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))" + "\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))" + "\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) @@ -311,13 +311,13 @@ subsection\Unfolding \<^term>\wftrec(r,a,H)\\ lemma the_recfun_cut: - "[| wf(r); trans(r); :r |] - ==> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)" + "\wf(r); trans(r); :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) |] ==> + "\wf(r); trans(r)\ \ wftrec(r,a,H) = H(a, \x\r-``{a}. wftrec(r,x,H))" apply (unfold wftrec_def) apply (subst unfold_the_recfun [unfolded is_recfun_def]) @@ -329,7 +329,7 @@ (*NOT SUITABLE FOR REWRITING: it is recursive!*) lemma wfrec: - "wf(r) ==> wfrec(r,a,H) = H(a, \x\r-``{a}. wfrec(r,x,H))" + "wf(r) \ wfrec(r,a,H) = H(a, \x\r-``{a}. wfrec(r,x,H))" apply (unfold wfrec_def) apply (erule wf_trancl [THEN wftrec, THEN ssubst]) apply (rule trans_trancl) @@ -338,18 +338,18 @@ apply (rule subset_refl) done -(*This form avoids giant explosions in proofs. NOTE USE OF == *) +(*This form avoids giant explosions in proofs. NOTE USE OF \ *) lemma def_wfrec: - "[| !!x. h(x)==wfrec(r,x,H); wf(r) |] ==> + "\\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)" + "\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) @@ -357,7 +357,7 @@ lemma wfrec_on: - "[| wf[A](r); a \ A |] ==> + "\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) apply (erule wfrec [THEN trans]) diff -r f2094906e491 -r e44d86131648 src/ZF/ZF.thy --- a/src/ZF/ZF.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ZF.thy Tue Sep 27 16:51:35 2022 +0100 @@ -15,27 +15,27 @@ definition iterates_omega :: "[i=>i,i] => i" (\(_^\ '(_'))\ [60,1000] 60) where - "F^\ (x) == \n\nat. F^n (x)" + "F^\ (x) \ \n\nat. F^n (x)" lemma iterates_triv: - "[| n\nat; F(x) = x |] ==> F^n (x) = x" + "\n\nat; F(x) = x\ \ F^n (x) = x" by (induct n rule: nat_induct, simp_all) lemma iterates_type [TC]: - "[| n \ nat; a \ A; !!x. x \ A ==> F(x) \ A |] - ==> F^n (a) \ A" + "\n \ nat; a \ A; \x. x \ A \ F(x) \ A\ + \ F^n (a) \ A" by (induct n rule: nat_induct, simp_all) lemma iterates_omega_triv: - "F(x) = x ==> F^\ (x) = x" + "F(x) = x \ F^\ (x) = x" by (simp add: iterates_omega_def iterates_triv) lemma Ord_iterates [simp]: - "[| n\nat; !!i. Ord(i) ==> Ord(F(i)); Ord(x) |] - ==> Ord(F^n (x))" + "\n\nat; \i. Ord(i) \ Ord(F(i)); Ord(x)\ + \ Ord(F^n (x))" by (induct n rule: nat_induct, simp_all) -lemma iterates_commute: "n \ nat ==> F(F^n (x)) = F^n (F(x))" +lemma iterates_commute: "n \ nat \ F(F^n (x)) = F^n (F(x))" by (induct_tac n, simp_all) @@ -46,7 +46,7 @@ definition transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where - "transrec3(k, a, b, c) == + "transrec3(k, a, b, c) \ transrec(k, \x r. if x=0 then a else if Limit(x) then c(x, \y\x. r`y) @@ -60,7 +60,7 @@ by (rule transrec3_def [THEN def_transrec, THEN trans], simp) lemma transrec3_Limit: - "Limit(i) ==> + "Limit(i) \ transrec3(i,a,b,c) = c(i, \j\i. transrec3(j,a,b,c))" by (rule transrec3_def [THEN def_transrec, THEN trans], force) diff -r f2094906e491 -r e44d86131648 src/ZF/ZF_Base.thy --- a/src/ZF/ZF_Base.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ZF_Base.thy Tue Sep 27 16:51:35 2022 +0100 @@ -49,7 +49,7 @@ 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))" + where "Replace(A,P) \ PrimReplace(A, %x y. (\!z. P(x,z)) & P(x,y))" syntax "_Replace" :: "[pttrn, pttrn, i, o] => i" (\(1{_ ./ _ \ _, _})\) @@ -60,7 +60,7 @@ (* 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)}" + where "RepFun(A,f) \ {y . x\A, y=f(x)}" syntax "_RepFun" :: "[i, pttrn, i] => i" (\(1{_ ./ _ \ _})\ [51,0,51]) @@ -71,7 +71,7 @@ (* 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)}" + where "Collect(A,P) \ {y . x\A, x=y & P(x)}" syntax "_Collect" :: "[pttrn, i, o] \ i" (\(1{_ \ _ ./ _})\) @@ -82,7 +82,7 @@ subsection \General union and intersection\ definition Inter :: "i => i" (\\_\ [90] 90) - where "\(A) == { x\\(A) . \y\A. x\y}" + where "\(A) \ { x\\(A) . \y\A. x\y}" syntax "_UNION" :: "[pttrn, i, i] => i" (\(3\_\_./ _)\ 10) @@ -98,25 +98,25 @@ 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)}" + 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) }" + where "A - B \ { x\A . \(x\B) }" definition Un :: "[i, i] \ i" (infixl \\\ 65) \ \binary union\ - where "A \ B == \(Upair(A,B))" + where "A \ B \ \(Upair(A,B))" definition Int :: "[i, i] \ i" (infixl \\\ 70) \ \binary intersection\ - where "A \ B == \(Upair(A,B))" + where "A \ B \ \(Upair(A,B))" definition cons :: "[i, i] => i" - where "cons(a,A) == Upair(a,a) \ A" + where "cons(a,A) \ Upair(a,a) \ A" definition succ :: "i => i" - where "succ(i) == cons(i, i)" + where "succ(i) \ cons(i, i)" nonterminal "is" syntax @@ -154,30 +154,30 @@ 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)})" + 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" + 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)" + 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}}" + where "Pair(a,b) \ {{a,a}, {a,b}}" definition fst :: "i \ i" - where "fst(p) == THE a. \b. p = Pair(a, b)" + where "fst(p) \ THE a. \b. p = Pair(a, b)" definition snd :: "i \ i" - where "snd(p) == THE b. \a. p = Pair(a, b)" + 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))" + where "split(c) \ \p. c(fst(p), snd(p))" (* Patterns -- extends pre-defined type "pttrn" used in abstractions *) nonterminal patterns @@ -193,7 +193,7 @@ "\\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\}" + 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)" @@ -203,44 +203,44 @@ (*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\}" + 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\}" + where "domain(r) \ {x. w\r, \y. w=\x,y\}" definition range :: "i \ i" - where "range(r) == domain(converse(r))" + where "range(r) \ domain(converse(r))" definition field :: "i \ i" - where "field(r) == domain(r) \ range(r)" + 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\" + 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')" + 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}" + 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" + 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\}" + 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}" + 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})" + 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)}" + 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)" @@ -267,7 +267,7 @@ function_space (infixr \->\ 60) and Subset (infixl \<=\ 50) and mem (infixl \:\ 50) and - not_mem (infixl \~:\ 50) + not_mem (infixl \\:\ 50) syntax (ASCII) "_Ball" :: "[pttrn, i, o] => o" (\(3ALL _:_./ _)\ 10) @@ -287,30 +287,30 @@ subsection \Substitution\ (*Useful examples: singletonI RS subst_elem, subst_elem RSN (2,IntI) *) -lemma subst_elem: "[| b\A; a=b |] ==> a\A" +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)" +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)" +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" + "\\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" +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)" +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!*) @@ -319,11 +319,11 @@ (*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))" + "\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))" + "(\x. x \ A \ P(x)) \ Trueprop (\x\A. P(x))" by (simp only: Ball_def atomize_all atomize_imp) lemmas [symmetric, rulify] = atomize_ball @@ -332,27 +332,27 @@ subsection\Bounded existential quantifier\ -lemma bexI [intro]: "[| P(x); x: A |] ==> \x\A. P(x)" +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)" +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)" +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" +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!!*) +(*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))" + "\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) @@ -360,34 +360,34 @@ subsection\Rules for subsets\ lemma subsetI [intro!]: - "(!!x. x\A ==> x\B) ==> A \ B" + "(\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" +lemma subsetD [elim]: "\A \ B; c\A\ \ c\B" apply (unfold subset_def) apply (erule bspec, assumption) done (*Classical elimination rule*) lemma subsetCE [elim]: - "[| A \ B; c\A ==> P; c\B ==> P |] ==> P" + "\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" +lemma rev_subsetD: "\c\A; A<=B\ \ c\B" by blast -lemma contra_subsetD: "[| A \ B; c \ B |] ==> c \ A" +lemma contra_subsetD: "\A \ B; c \ B\ \ c \ A" by blast -lemma rev_contra_subsetD: "[| c \ B; A \ B |] ==> c \ A" +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" +lemma subset_trans: "\A<=B; B<=C\ \ A<=C" by blast (*Useful for proving A<=B by rewriting in some cases*) @@ -404,25 +404,25 @@ subsection\Rules for equality\ (*Anti-symmetry of the subset relation*) -lemma equalityI [intro]: "[| A \ B; B \ A |] ==> A = B" +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" +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" +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" + "\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)" + "A = B \ (\x. x \ A <-> x \ B)" by auto @@ -436,26 +436,26 @@ (*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 |] ==> + "\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" + "\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" + "\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) |] ==> + "\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) @@ -464,23 +464,23 @@ subsection\Rules for RepFun\ -lemma RepFunI: "a \ A ==> f(a) \ {f(x). x\A}" +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}" +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 |] ==> + "\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)" + "\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))" @@ -496,21 +496,21 @@ 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)}" +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" +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" +lemma CollectD1: "a \ {x\A. P(x)} \ a\A" by (erule CollectE, assumption) -lemma CollectD2: "a \ {x\A. P(x)} ==> P(a)" +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))" + "\A=B; \x. x\B \ P(x) <-> Q(x)\ + \ Collect(A, %x. P(x)) = Collect(B, %x. Q(x))" by (simp add: Collect_def) @@ -519,10 +519,10 @@ 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)" +lemma UnionI [intro]: "\B: C; A: B\ \ A: \(C)" by (simp, blast) -lemma UnionE [elim!]: "[| A \ \(C); !!B.[| A: B; B: C |] ==> R |] ==> R" +lemma UnionE [elim!]: "\A \ \(C); \B.\A: B; B: C\ \ R\ \ R" by (simp, blast) @@ -533,16 +533,16 @@ 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))" +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" + "\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))" + "\A=B; \x. x\B \ C(x)=D(x)\ \ (\x\A. C(x)) = (\x\B. D(x))" by simp @@ -568,18 +568,18 @@ lemma empty_subsetI [simp]: "0 \ A" by blast -lemma equals0I: "[| !!y. y\A ==> False |] ==> A=0" +lemma equals0I: "\\y. y\A \ False\ \ A=0" by blast -lemma equals0D [dest]: "A=0 ==> a \ A" +lemma equals0D [dest]: "A=0 \ a \ A" by blast declare sym [THEN equals0D, dest] -lemma not_emptyI: "a\A ==> A \ 0" +lemma not_emptyI: "a\A \ A \ 0" by blast -lemma not_emptyE: "[| A \ 0; !!x. x\A ==> R |] ==> R" +lemma not_emptyE: "\A \ 0; \x. x\A \ R\ \ R" by blast @@ -591,17 +591,17 @@ (* Intersection is well-behaved only if the family is non-empty! *) lemma InterI [intro!]: - "[| !!x. x: C ==> A: x; C\0 |] ==> A \ \(C)" + "\\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" +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" + "\A \ \(C); B\C \ R; A\B \ R\ \ R" by (simp add: Inter_def, blast) @@ -612,14 +612,14 @@ 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))" +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)" +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))" + "\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*) @@ -627,10 +627,10 @@ subsection\Rules for Powersets\ -lemma PowI: "A \ B ==> A \ Pow(B)" +lemma PowI: "A \ B \ A \ Pow(B)" by (erule Pow_iff [THEN iffD2]) -lemma PowD: "A \ Pow(B) ==> A<=B" +lemma PowD: "A \ Pow(B) \ A<=B" by (erule Pow_iff [THEN iffD1]) declare Pow_iff [iff] diff -r f2094906e491 -r e44d86131648 src/ZF/Zorn.thy --- a/src/ZF/Zorn.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/Zorn.thy Tue Sep 27 16:51:35 2022 +0100 @@ -12,27 +12,27 @@ definition Subset_rel :: "i=>i" where - "Subset_rel(A) == {z \ A*A . \x y. z= & x<=y & x\y}" + "Subset_rel(A) \ {z \ A*A . \x y. z= & x<=y & x\y}" definition chain :: "i=>i" where - "chain(A) == {F \ Pow(A). \X\F. \Y\F. X<=Y | Y<=X}" + "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}" + "super(A,c) \ {d \ chain(A). c<=d & c\d}" definition maxchain :: "i=>i" where - "maxchain(A) == {c \ chain(A). super(A,c)=0}" + "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}" + "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)" +lemma Union_in_Pow: "Y \ Pow(Pow(A)) \ \(Y) \ Pow(A)" by blast @@ -47,10 +47,10 @@ inductive domains "TFin(S,next)" \ "Pow(S)" intros - nextI: "[| x \ TFin(S,next); next \ increasing(S) |] - ==> next`x \ TFin(S,next)" + nextI: "\x \ TFin(S,next); next \ increasing(S)\ + \ next`x \ TFin(S,next)" - Pow_UnionI: "Y \ Pow(TFin(S,next)) ==> \(Y) \ TFin(S,next)" + Pow_UnionI: "Y \ Pow(TFin(S,next)) \ \(Y) \ TFin(S,next)" monos Pow_mono con_defs increasing_def @@ -59,22 +59,22 @@ subsection\Mathematical Preamble\ -lemma Union_lemma0: "(\x\C. x<=A | B<=x) ==> \(C)<=A | B<=\(C)" +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" + "\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)" +lemma increasingD1: "f \ increasing(A) \ f \ Pow(A)->Pow(A)" apply (unfold increasing_def) apply (erule CollectD1) done -lemma increasingD2: "[| f \ increasing(A); x<=A |] ==> x \ f`x" +lemma increasingD2: "\f \ increasing(A); x<=A\ \ x \ f`x" by (unfold increasing_def, blast) lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI] @@ -84,10 +84,10 @@ 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)" + "\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+) @@ -98,9 +98,9 @@ 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" + "\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*) @@ -110,8 +110,8 @@ 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" + "\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\\ @@ -135,14 +135,14 @@ 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" + "\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" + "\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) @@ -153,7 +153,7 @@ text\Lemma 3 of section 3.3\ lemma equal_next_upper: - "[| n \ TFin(S,next); m \ TFin(S,next); m = next`m |] ==> n \ m" + "\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) @@ -161,8 +161,8 @@ 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))" + "\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]) @@ -197,15 +197,15 @@ done lemma choice_super: - "[| ch \ (\X \ Pow(chain(S)) - {0}. X); X \ chain(S); X \ maxchain(S) |] - ==> ch ` super(S,X) \ super(S,X)" + "\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" + "\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) @@ -213,7 +213,7 @@ text\This justifies Definition 4.4\ lemma Hausdorff_next_exists: - "ch \ (\X \ Pow(chain(S))-{0}. X) ==> + "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). @@ -236,12 +236,12 @@ text\Lemma 4\ lemma TFin_chain_lemma4: - "[| c \ TFin(S,next); + "\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)" + 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]]) @@ -277,10 +277,10 @@ 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)" + "\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" +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) @@ -299,7 +299,7 @@ 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" + "\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) @@ -321,8 +321,8 @@ text\Lemma 5\ lemma TFin_well_lemma5: - "[| n \ TFin(S,next); Z \ TFin(S,next); z:Z; ~ \(Z) \ Z |] - ==> \m \ Z. n \ m" + "\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) @@ -332,7 +332,7 @@ done text\Well-ordering of \<^term>\TFin(S,next)\\ -lemma well_ord_TFin_lemma: "[| Z \ TFin(S,next); z \ Z |] ==> \(Z) \ Z" +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) @@ -344,7 +344,7 @@ 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)))" + \ well_ord(TFin(S,next), Subset_rel(TFin(S,next)))" apply (rule well_ordI) apply (unfold Subset_rel_def linear_def) txt\Prove the well-foundedness goal\ @@ -364,14 +364,14 @@ 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" + "\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) ==> + "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)" @@ -391,10 +391,10 @@ text\The construction of the injection\ lemma choice_imp_injection: - "[| ch \ (\X \ Pow(S)-{0}. X); + "\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})) + \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) @@ -406,7 +406,7 @@ 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}) +apply (subgoal_tac "\ next ` Union({y \ TFin (S,next) . x \ y}) \ \({y \ TFin (S,next) . x \ y}) ") prefer 2 apply (simp del: Union_iff @@ -439,7 +439,7 @@ "Chain(r) = {A \ Pow(field(r)). \a\A. \b\A. \ r | \ r}" lemma mono_Chain: - "r \ s ==> Chain(r) \ Chain(s)" + "r \ s \ Chain(r) \ Chain(s)" unfolding Chain_def by blast diff -r f2094906e491 -r e44d86131648 src/ZF/equalities.thy --- a/src/ZF/equalities.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/equalities.thy Tue Sep 27 16:51:35 2022 +0100 @@ -10,7 +10,7 @@ text\These cover union, intersection, converse, domain, range, etc. Philippe de Groote proved many of the inclusions.\ -lemma in_mono: "A\B ==> x\A \ x\B" +lemma in_mono: "A\B \ x\A \ x\B" by blast lemma the_eq_0 [simp]: "(THE x. False) = 0" @@ -39,22 +39,22 @@ lemma converse_iff [simp]: "\ converse(r) \ \r" by (unfold converse_def, blast) -lemma converseI [intro!]: "\r ==> \converse(r)" +lemma converseI [intro!]: "\r \ \converse(r)" by (unfold converse_def, blast) -lemma converseD: " \ converse(r) ==> \ r" +lemma converseD: " \ converse(r) \ \ r" by (unfold converse_def, blast) lemma converseE [elim!]: - "[| yx \ converse(r); - !!x y. [| yx=; \r |] ==> P |] - ==> P" + "\yx \ converse(r); + \x y. \yx=; \r\ \ P\ + \ P" by (unfold converse_def, blast) -lemma converse_converse: "r\Sigma(A,B) ==> converse(converse(r)) = r" +lemma converse_converse: "r\Sigma(A,B) \ converse(converse(r)) = r" by blast -lemma converse_type: "r\A*B ==> converse(r)\B*A" +lemma converse_type: "r\A*B \ converse(r)\B*A" by blast lemma converse_prod [simp]: "converse(A*B) = B*A" @@ -64,13 +64,13 @@ by blast lemma converse_subset_iff: - "A \ Sigma(X,Y) ==> converse(A) \ converse(B) \ A \ B" + "A \ Sigma(X,Y) \ converse(A) \ converse(B) \ A \ B" by blast subsection\Finite Set Constructions Using \<^term>\cons\\ -lemma cons_subsetI: "[| a\C; B\C |] ==> cons(a,B) \ C" +lemma cons_subsetI: "\a\C; B\C\ \ cons(a,B) \ C" by blast lemma subset_consI: "B \ cons(a,B)" @@ -80,7 +80,7 @@ by blast (*A safe special case of subset elimination, adding no new variables - [| cons(a,B) \ C; [| a \ C; B \ C |] ==> R |] ==> R *) + \cons(a,B) \ C; \a \ C; B \ C\ \ R\ \ R *) lemmas cons_subsetE = cons_subset_iff [THEN iffD1, THEN conjE] lemma subset_empty_iff: "A\0 \ A=0" @@ -96,16 +96,16 @@ lemma cons_commute: "cons(a, cons(b, C)) = cons(b, cons(a, C))" by blast -lemma cons_absorb: "a: B ==> cons(a,B) = B" +lemma cons_absorb: "a: B \ cons(a,B) = B" by blast -lemma cons_Diff: "a: B ==> cons(a, B-{a}) = B" +lemma cons_Diff: "a: B \ cons(a, B-{a}) = B" by blast lemma Diff_cons_eq: "cons(a,B) - C = (if a\C then B-C else cons(a,B-C))" by auto -lemma equal_singleton: "[| a: C; \y. y \C \ y=b |] ==> C = {b}" +lemma equal_singleton: "\a: C; \y. y \C \ y=b\ \ C = {b}" by blast lemma [simp]: "cons(a,cons(a,B)) = cons(a,B)" @@ -113,10 +113,10 @@ (** singletons **) -lemma singleton_subsetI: "a\C ==> {a} \ C" +lemma singleton_subsetI: "a\C \ {a} \ C" by blast -lemma singleton_subsetD: "{a} \ C ==> a\C" +lemma singleton_subsetD: "{a} \ C \ a\C" by blast @@ -127,11 +127,11 @@ (*But if j is an ordinal or is transitive, then @{term"i\j"} implies @{term"i\j"}! See @{text"Ord_succ_subsetI}*) -lemma succ_subsetI: "[| i\j; i\j |] ==> succ(i)\j" +lemma succ_subsetI: "\i\j; i\j\ \ succ(i)\j" by (unfold succ_def, blast) lemma succ_subsetE: - "[| succ(i) \ j; [| i\j; i\j |] ==> P |] ==> P" + "\succ(i) \ j; \i\j; i\j\ \ P\ \ P" by (unfold succ_def, blast) lemma succ_subset_iff: "succ(a) \ B \ (a \ B & a \ B)" @@ -151,7 +151,7 @@ lemma Int_lower2: "A \ B \ B" by blast -lemma Int_greatest: "[| C\A; C\B |] ==> C \ A \ B" +lemma Int_greatest: "\C\A; C\B\ \ C \ A \ B" by blast lemma Int_cons: "cons(a,B) \ C \ cons(a, B \ C)" @@ -175,10 +175,10 @@ (*Intersection is an AC-operator*) lemmas Int_ac= Int_assoc Int_left_absorb Int_commute Int_left_commute -lemma Int_absorb1: "B \ A ==> A \ B = B" +lemma Int_absorb1: "B \ A \ A \ B = B" by blast -lemma Int_absorb2: "A \ B ==> A \ B = A" +lemma Int_absorb2: "A \ B \ A \ B = A" by blast lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" @@ -193,7 +193,7 @@ lemma subset_Int_iff2: "A\B \ B \ A = A" by (blast elim!: equalityE) -lemma Int_Diff_eq: "C\A ==> (A-B) \ C = C-B" +lemma Int_Diff_eq: "C\A \ (A-B) \ C = C-B" by blast lemma Int_cons_left: @@ -220,7 +220,7 @@ lemma Un_upper2: "B \ A \ B" by blast -lemma Un_least: "[| A\C; B\C |] ==> A \ B \ C" +lemma Un_least: "\A\C; B\C\ \ A \ B \ C" by blast lemma Un_cons: "cons(a,B) \ C = cons(a, B \ C)" @@ -244,10 +244,10 @@ (*Union is an AC-operator*) lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute -lemma Un_absorb1: "A \ B ==> A \ B = B" +lemma Un_absorb1: "A \ B \ A \ B = B" by blast -lemma Un_absorb2: "B \ A ==> A \ B = A" +lemma Un_absorb2: "B \ A \ A \ B = A" by blast lemma Un_Int_distrib: "(A \ B) \ C = (A \ C) \ (B \ C)" @@ -270,7 +270,7 @@ lemma Diff_subset: "A-B \ A" by blast -lemma Diff_contains: "[| C\A; C \ B = 0 |] ==> C \ A-B" +lemma Diff_contains: "\C\A; C \ B = 0\ \ C \ A-B" by blast lemma subset_Diff_cons_iff: "B \ A - cons(c,C) \ B\A-C & c \ B" @@ -279,7 +279,7 @@ lemma Diff_cancel: "A - A = 0" by blast -lemma Diff_triv: "A \ B = 0 ==> A - B = A" +lemma Diff_triv: "A \ B = 0 \ A - B = A" by blast lemma empty_Diff [simp]: "0 - A = 0" @@ -291,24 +291,24 @@ lemma Diff_eq_0_iff: "A - B = 0 \ A \ B" by (blast elim: equalityE) -(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) +(*NOT SUITABLE FOR REWRITING since {a} \ cons(a,0)*) lemma Diff_cons: "A - cons(a,B) = A - B - {a}" by blast -(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) +(*NOT SUITABLE FOR REWRITING since {a} \ cons(a,0)*) lemma Diff_cons2: "A - cons(a,B) = A - {a} - B" by blast lemma Diff_disjoint: "A \ (B-A) = 0" by blast -lemma Diff_partition: "A\B ==> A \ (B-A) = B" +lemma Diff_partition: "A\B \ A \ (B-A) = B" by blast lemma subset_Un_Diff: "A \ B \ (A - B)" by blast -lemma double_complement: "[| A\B; B\C |] ==> B-(C-A) = A" +lemma double_complement: "\A\B; B\C\ \ B-(C-A) = A" by blast lemma double_complement_Un: "(A \ B) - (B-A) = A" @@ -349,10 +349,10 @@ lemma Union_subset_iff: "\(A) \ C \ (\x\A. x \ C)" by blast -lemma Union_upper: "B\A ==> B \ \(A)" +lemma Union_upper: "B\A \ B \ \(A)" by blast -lemma Union_least: "[| !!x. x\A ==> x\C |] ==> \(A) \ C" +lemma Union_least: "\\x. x\A \ x\C\ \ \(A) \ C" by blast lemma Union_cons [simp]: "\(cons(a,B)) = a \ \(B)" @@ -375,33 +375,33 @@ (** Big Intersection is the greatest lower bound of a nonempty set **) -lemma Inter_subset_iff: "A\0 ==> C \ \(A) \ (\x\A. C \ x)" +lemma Inter_subset_iff: "A\0 \ C \ \(A) \ (\x\A. C \ x)" by blast -lemma Inter_lower: "B\A ==> \(A) \ B" +lemma Inter_lower: "B\A \ \(A) \ B" by blast -lemma Inter_greatest: "[| A\0; !!x. x\A ==> C\x |] ==> C \ \(A)" +lemma Inter_greatest: "\A\0; \x. x\A \ C\x\ \ C \ \(A)" by blast (** Intersection of a family of sets **) -lemma INT_lower: "x\A ==> (\x\A. B(x)) \ B(x)" +lemma INT_lower: "x\A \ (\x\A. B(x)) \ B(x)" by blast -lemma INT_greatest: "[| A\0; !!x. x\A ==> C\B(x) |] ==> C \ (\x\A. B(x))" +lemma INT_greatest: "\A\0; \x. x\A \ C\B(x)\ \ C \ (\x\A. B(x))" by force lemma Inter_0 [simp]: "\(0) = 0" by (unfold Inter_def, blast) lemma Inter_Un_subset: - "[| z\A; z\B |] ==> \(A) \ \(B) \ \(A \ B)" + "\z\A; z\B\ \ \(A) \ \(B) \ \(A \ B)" by blast (* A good challenge: Inter is ill-behaved on the empty set *) lemma Inter_Un_distrib: - "[| A\0; B\0 |] ==> \(A \ B) = \(A) \ \(B)" + "\A\0; B\0\ \ \(A \ B) = \(A) \ \(B)" by blast lemma Union_singleton: "\({b}) = b" @@ -422,10 +422,10 @@ lemma UN_subset_iff: "(\x\A. B(x)) \ C \ (\x\A. B(x) \ C)" by blast -lemma UN_upper: "x\A ==> B(x) \ (\x\A. B(x))" +lemma UN_upper: "x\A \ B(x) \ (\x\A. B(x))" by (erule RepFunI [THEN Union_upper]) -lemma UN_least: "[| !!x. x\A ==> B(x)\C |] ==> (\x\A. B(x)) \ C" +lemma UN_least: "\\x. x\A \ B(x)\C\ \ (\x\A. B(x)) \ C" by blast lemma Union_eq_UN: "\(A) = (\x\A. x)" @@ -456,14 +456,14 @@ lemma Int_UN_distrib: "B \ (\i\I. A(i)) = (\i\I. B \ A(i))" by blast -lemma Un_INT_distrib: "I\0 ==> B \ (\i\I. A(i)) = (\i\I. B \ A(i))" +lemma Un_INT_distrib: "I\0 \ B \ (\i\I. A(i)) = (\i\I. B \ A(i))" by auto lemma Int_UN_distrib2: "(\i\I. A(i)) \ (\j\J. B(j)) = (\i\I. \j\J. A(i) \ B(j))" by blast -lemma Un_INT_distrib2: "[| I\0; J\0 |] ==> +lemma Un_INT_distrib2: "\I\0; J\0\ \ (\i\I. A(i)) \ (\j\J. B(j)) = (\i\I. \j\J. A(i) \ B(j))" by auto @@ -480,7 +480,7 @@ by (auto simp add: Inter_def) lemma INT_Union_eq: - "0 \ A ==> (\x\ \(A). B(x)) = (\y\A. \x\y. B(x))" + "0 \ A \ (\x\ \(A). B(x)) = (\y\A. \x\y. B(x))" apply (subgoal_tac "\x\A. x\0") prefer 2 apply blast apply (force simp add: Inter_def ball_conj_distrib) @@ -488,7 +488,7 @@ lemma INT_UN_eq: "(\x\A. B(x) \ 0) - ==> (\z\ (\x\A. B(x)). C(z)) = (\x\A. \z\ B(x). C(z))" + \ (\z\ (\x\A. B(x)). C(z)) = (\x\A. \z\ B(x). C(z))" apply (subst INT_Union_eq, blast) apply (simp add: Inter_def) done @@ -502,7 +502,7 @@ by blast lemma INT_Int_distrib: - "I\0 ==> (\i\I. A(i) \ B(i)) = (\i\I. A(i)) \ (\i\I. B(i))" + "I\0 \ (\i\I. A(i) \ B(i)) = (\i\I. A(i)) \ (\i\I. B(i))" by (blast elim!: not_emptyE) lemma UN_Int_subset: @@ -511,10 +511,10 @@ (** Devlin, page 12, exercise 5: Complements **) -lemma Diff_UN: "I\0 ==> B - (\i\I. A(i)) = (\i\I. B - A(i))" +lemma Diff_UN: "I\0 \ B - (\i\I. A(i)) = (\i\I. B - A(i))" by (blast elim!: not_emptyE) -lemma Diff_INT: "I\0 ==> B - (\i\I. A(i)) = (\i\I. B - A(i))" +lemma Diff_INT: "I\0 \ B - (\i\I. A(i)) = (\i\I. B - A(i))" by (blast elim!: not_emptyE) @@ -583,17 +583,17 @@ lemma domain_iff: "a: domain(r) \ (\y. \ r)" by (unfold domain_def, blast) -lemma domainI [intro]: "\ r ==> a: domain(r)" +lemma domainI [intro]: "\ r \ a: domain(r)" by (unfold domain_def, blast) lemma domainE [elim!]: - "[| a \ domain(r); !!y. \ r ==> P |] ==> P" + "\a \ domain(r); \y. \ r \ P\ \ P" by (unfold domain_def, blast) lemma domain_subset: "domain(Sigma(A,B)) \ A" by blast -lemma domain_of_prod: "b\B ==> domain(A*B) = A" +lemma domain_of_prod: "b\B \ domain(A*B) = A" by blast lemma domain_0 [simp]: "domain(0) = 0" @@ -620,12 +620,12 @@ (** Range **) -lemma rangeI [intro]: "\ r ==> b \ range(r)" +lemma rangeI [intro]: "\ r \ b \ range(r)" apply (unfold range_def) apply (erule converseI [THEN domainI]) done -lemma rangeE [elim!]: "[| b \ range(r); !!x. \ r ==> P |] ==> P" +lemma rangeE [elim!]: "\b \ range(r); \x. \ r \ P\ \ P" by (unfold range_def, blast) lemma range_subset: "range(A*B) \ B" @@ -634,7 +634,7 @@ apply (rule domain_subset) done -lemma range_of_prod: "a\A ==> range(A*B) = B" +lemma range_of_prod: "a\A \ range(A*B) = B" by blast lemma range_0 [simp]: "range(0) = 0" @@ -661,21 +661,21 @@ (** Field **) -lemma fieldI1: "\ r ==> a \ field(r)" +lemma fieldI1: "\ r \ a \ field(r)" by (unfold field_def, blast) -lemma fieldI2: "\ r ==> b \ field(r)" +lemma fieldI2: "\ r \ b \ field(r)" by (unfold field_def, blast) lemma fieldCI [intro]: - "(~ \r ==> \ r) ==> a \ field(r)" + "(\ \r \ \ r) \ a \ field(r)" apply (unfold field_def, blast) done lemma fieldE [elim!]: - "[| a \ field(r); - !!x. \ r ==> P; - !!x. \ r ==> P |] ==> P" + "\a \ field(r); + \x. \ r \ P; + \x. \ r \ P\ \ P" by (unfold field_def, blast) lemma field_subset: "field(A*B) \ A \ B" @@ -691,13 +691,13 @@ apply (rule Un_upper2) done -lemma domain_times_range: "r \ Sigma(A,B) ==> r \ domain(r)*range(r)" +lemma domain_times_range: "r \ Sigma(A,B) \ r \ domain(r)*range(r)" by blast -lemma field_times_field: "r \ Sigma(A,B) ==> r \ field(r)*field(r)" +lemma field_times_field: "r \ Sigma(A,B) \ r \ field(r)*field(r)" by blast -lemma relation_field_times_field: "relation(r) ==> r \ field(r)*field(r)" +lemma relation_field_times_field: "relation(r) \ r \ field(r)*field(r)" by (simp add: relation_def, blast) lemma field_of_prod: "field(A*A) = A" @@ -722,18 +722,18 @@ by blast (** The Union of a set of relations is a relation -- Lemma for fun_Union **) -lemma rel_Union: "(\x\S. \A B. x \ A*B) ==> +lemma rel_Union: "(\x\S. \A B. x \ A*B) \ \(S) \ domain(\(S)) * range(\(S))" by blast (** The Union of 2 relations is a relation (Lemma for fun_Un) **) -lemma rel_Un: "[| r \ A*B; s \ C*D |] ==> (r \ s) \ (A \ C) * (B \ D)" +lemma rel_Un: "\r \ A*B; s \ C*D\ \ (r \ s) \ (A \ C) * (B \ D)" by blast -lemma domain_Diff_eq: "[| \ r; c\b |] ==> domain(r-{}) = domain(r)" +lemma domain_Diff_eq: "\ \ r; c\b\ \ domain(r-{}) = domain(r)" by blast -lemma range_Diff_eq: "[| \ r; c\a |] ==> range(r-{}) = range(r)" +lemma range_Diff_eq: "\ \ r; c\a\ \ range(r-{}) = range(r)" by blast @@ -745,14 +745,14 @@ lemma image_singleton_iff: "b \ r``{a} \ \r" by (rule image_iff [THEN iff_trans], blast) -lemma imageI [intro]: "[| \ r; a\A |] ==> b \ r``A" +lemma imageI [intro]: "\\ r; a\A\ \ b \ r``A" by (unfold image_def, blast) lemma imageE [elim!]: - "[| b: r``A; !!x.[| \ r; x\A |] ==> P |] ==> P" + "\b: r``A; \x.\\ r; x\A\ \ P\ \ P" by (unfold image_def, blast) -lemma image_subset: "r \ A*B ==> r``C \ B" +lemma image_subset: "r \ A*B \ r``C \ B" by blast lemma image_0 [simp]: "r``0 = 0" @@ -774,7 +774,7 @@ lemma image_Int_square_subset: "(r \ A*A)``B \ (r``B) \ A" by blast -lemma image_Int_square: "B\A ==> (r \ A*A)``B = (r``B) \ A" +lemma image_Int_square: "B\A \ (r \ A*A)``B = (r``B) \ A" by blast @@ -798,15 +798,15 @@ lemma vimage_singleton_iff: "a \ r-``{b} \ \r" by (rule vimage_iff [THEN iff_trans], blast) -lemma vimageI [intro]: "[| \ r; b\B |] ==> a \ r-``B" +lemma vimageI [intro]: "\\ r; b\B\ \ a \ r-``B" by (unfold vimage_def, blast) lemma vimageE [elim!]: - "[| a: r-``B; !!x.[| \ r; x\B |] ==> P |] ==> P" + "\a: r-``B; \x.\\ r; x\B\ \ P\ \ P" apply (unfold vimage_def, blast) done -lemma vimage_subset: "r \ A*B ==> r-``C \ A" +lemma vimage_subset: "r \ A*B \ r-``C \ A" apply (unfold vimage_def) apply (erule converse_type [THEN image_subset]) done @@ -825,19 +825,19 @@ by blast lemma function_vimage_Int: - "function(f) ==> f-``(A \ B) = (f-``A) \ (f-``B)" + "function(f) \ f-``(A \ B) = (f-``A) \ (f-``B)" by (unfold function_def, blast) -lemma function_vimage_Diff: "function(f) ==> f-``(A-B) = (f-``A) - (f-``B)" +lemma function_vimage_Diff: "function(f) \ f-``(A-B) = (f-``A) - (f-``B)" by (unfold function_def, blast) -lemma function_image_vimage: "function(f) ==> f `` (f-`` A) \ A" +lemma function_image_vimage: "function(f) \ f `` (f-`` A) \ A" by (unfold function_def, blast) lemma vimage_Int_square_subset: "(r \ A*A)-``B \ (r-``B) \ A" by blast -lemma vimage_Int_square: "B\A ==> (r \ A*A)-``B = (r-``B) \ A" +lemma vimage_Int_square: "B\A \ (r \ A*A)-``B = (r-``B) \ A" by blast @@ -903,13 +903,13 @@ lemma Pow_Int_eq [simp]: "Pow(A \ B) = Pow(A) \ Pow(B)" by blast -lemma Pow_INT_eq: "A\0 ==> Pow(\x\A. B(x)) = (\x\A. Pow(B(x)))" +lemma Pow_INT_eq: "A\0 \ Pow(\x\A. B(x)) = (\x\A. Pow(B(x)))" by (blast elim!: not_emptyE) subsection\RepFun\ -lemma RepFun_subset: "[| !!x. x\A ==> f(x) \ B |] ==> {f(x). x\A} \ B" +lemma RepFun_subset: "\\x. x\A \ f(x) \ B\ \ {f(x). x\A} \ B" by blast lemma RepFun_eq_0_iff [simp]: "{f(x).x\A}=0 \ A=0" diff -r f2094906e491 -r e44d86131648 src/ZF/ex/BinEx.thy --- a/src/ZF/ex/BinEx.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/BinEx.thy Tue Sep 27 16:51:35 2022 +0100 @@ -23,7 +23,7 @@ lemma "$- #65745 = #-65745" by simp (*80 msec*) -(* negation of ~54321 *) +(* negation of \54321 *) lemma "$- #-54321 = #54321" by simp (*90 msec*) @@ -62,7 +62,7 @@ by simp -(*** Quotient and remainder!! [they could be faster] ***) +(*** Quotient and remainder\[they could be faster] ***) lemma "#23 zdiv #3 = #7" by simp diff -r f2094906e491 -r e44d86131648 src/ZF/ex/CoUnit.thy --- a/src/ZF/ex/CoUnit.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/CoUnit.thy Tue Sep 27 16:51:35 2022 +0100 @@ -25,7 +25,7 @@ "counit" = Con ("x \ counit") inductive_cases ConE: "Con(x) \ counit" - \ \USELESS because folding on \<^term>\Con(xa) == xa\ fails.\ + \ \USELESS because folding on \<^term>\Con(xa) \ xa\ fails.\ lemma Con_iff: "Con(x) = Con(y) \ x = y" \ \Proving freeness results.\ @@ -73,7 +73,7 @@ done lemma counit2_Int_Vset_subset [rule_format]: - "Ord(i) ==> \x y. x \ counit2 \ y \ counit2 \ x \ Vset(i) \ y" + "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>)") @@ -85,7 +85,7 @@ addSEs [@{thm Ord_in_Ord}, @{thm Pair_inject}]) 1\) done -lemma counit2_implies_equal: "[| x \ counit2; y \ counit2 |] ==> x = y" +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 diff -r f2094906e491 -r e44d86131648 src/ZF/ex/Commutation.thy --- a/src/ZF/ex/Commutation.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/Commutation.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,40 +9,40 @@ definition square :: "[i, i, i, i] => o" where - "square(r,s,t,u) == + "square(r,s,t,u) \ (\a b. \ r \ (\c. \ s \ (\x. \ t & \ u)))" definition commute :: "[i, i] => o" where - "commute(r,s) == square(r,s,s,r)" + "commute(r,s) \ square(r,s,s,r)" definition diamond :: "i=>o" where - "diamond(r) == commute(r, r)" + "diamond(r) \ commute(r, r)" definition strip :: "i=>o" where - "strip(r) == commute(r^*, r)" + "strip(r) \ commute(r^*, r)" definition Church_Rosser :: "i => o" where - "Church_Rosser(r) == (\x y. \ (r \ converse(r))^* \ + "Church_Rosser(r) \ (\x y. \ (r \ converse(r))^* \ (\z. \ r^* & \ r^*))" definition confluent :: "i=>o" where - "confluent(r) == diamond(r^*)" + "confluent(r) \ diamond(r^*)" -lemma square_sym: "square(r,s,t,u) ==> square(s,r,u,t)" +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)" +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^*)" + "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) @@ -51,18 +51,18 @@ (* A special case of square_rtrancl_on *) lemma diamond_strip: - "diamond(r) ==> strip(r)" + "diamond(r) \ strip(r)" apply (unfold diamond_def commute_def strip_def) apply (rule square_rtrancl, simp_all) done (*** commute ***) -lemma commute_sym: "commute(r,s) ==> commute(s,r)" +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^*)" + "commute(r,s) \ field(r)=field(s) \ commute(r^*,s^*)" apply (unfold commute_def) apply (rule square_rtrancl) apply (rule square_sym [THEN square_rtrancl, THEN square_sym]) @@ -70,31 +70,31 @@ done -lemma confluentD: "confluent(r) ==> diamond(r^*)" +lemma confluentD: "confluent(r) \ diamond(r^*)" by (simp add: confluent_def) -lemma strip_confluent: "strip(r) ==> confluent(r)" +lemma strip_confluent: "strip(r) \ confluent(r)" apply (unfold 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)" +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)" + "\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)" + "diamond(r) \ confluent(r)" apply (unfold 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)" + "\confluent(r); confluent(s); commute(r^*, s^*); + relation(r); relation(s)\ \ confluent(r \ s)" apply (unfold confluent_def) apply (rule rtrancl_Un_rtrancl [THEN subst], auto) apply (blast dest: diamond_Un intro: diamond_confluent [THEN confluentD]) @@ -102,7 +102,7 @@ lemma diamond_to_confluence: - "[| diamond(r); s \ r; r<= s^* |] ==> confluent(s)" + "\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]) @@ -112,7 +112,7 @@ (*** Church_Rosser ***) lemma Church_Rosser1: - "Church_Rosser(r) ==> confluent(r)" + "Church_Rosser(r) \ confluent(r)" apply (unfold confluent_def Church_Rosser_def square_def commute_def diamond_def, auto) apply (drule converseI) @@ -125,7 +125,7 @@ lemma Church_Rosser2: - "confluent(r) ==> Church_Rosser(r)" + "confluent(r) \ Church_Rosser(r)" apply (unfold confluent_def Church_Rosser_def square_def commute_def diamond_def, auto) apply (frule fieldI1) diff -r f2094906e491 -r e44d86131648 src/ZF/ex/Group.thy --- a/src/ZF/ex/Group.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/Group.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,23 +19,23 @@ definition carrier :: "i => i" where - "carrier(M) == fst(M)" + "carrier(M) \ fst(M)" definition mmult :: "[i, i, i] => i" (infixl \\\\ 70) where - "mmult(M,x,y) == fst(snd(M)) ` " + "mmult(M,x,y) \ fst(snd(M)) ` " definition one :: "i => i" (\\\\) where - "one(M) == fst(snd(snd(M)))" + "one(M) \ fst(snd(snd(M)))" definition update_carrier :: "[i,i] => i" where - "update_carrier(M,A) == " + "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>)" + "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]: @@ -133,7 +133,7 @@ with x xG show "x \ \ = x" by simp qed have inv_ex: - "!!x. x \ carrier(G) ==> \y \ carrier(G). y \ x = \ & x \ y = \" + "\x. x \ carrier(G) \ \y \ carrier(G). y \ x = \ & x \ y = \" proof - fix x assume x: "x \ carrier(G)" @@ -295,7 +295,7 @@ monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp] lemma subgroup_nonempty: - "~ subgroup(0,G)" + "\ subgroup(0,G)" by (blast dest: subgroup.one_closed) @@ -303,7 +303,7 @@ definition DirProdGroup :: "[i,i] => i" (infixr \\\ 80) where - "G \ H == carrier(H), + "G \ H \ carrier(H), (\<, > \ (carrier(G) \ carrier(H)) \ (carrier(G) \ carrier(H)). \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>), @@ -328,8 +328,8 @@ by (simp add: DirProdGroup_def) lemma mult_DirProdGroup [simp]: - "[|g \ carrier(G); h \ carrier(H); g' \ carrier(G); h' \ carrier(H)|] - ==> \\<^bsub>G \ H\<^esub> = \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>" + "\g \ carrier(G); h \ carrier(H); g' \ carrier(G); h' \ carrier(H)\ + \ \\<^bsub>G \ H\<^esub> = \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>" by (simp add: DirProdGroup_def) lemma inv_DirProdGroup [simp]: @@ -345,7 +345,7 @@ definition hom :: "[i,i] => i" where - "hom(G,H) == + "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))}" @@ -371,7 +371,7 @@ definition iso :: "[i,i] => i" (infixr \\\ 60) where - "G \ H == hom(G,H) \ bij(carrier(G), carrier(H))" + "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) @@ -490,8 +490,8 @@ 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" + 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" @@ -503,7 +503,7 @@ definition BijGroup :: "i=>i" where - "BijGroup(S) == + "BijGroup(S) \ \ bij(S,S) \ bij(S,S). g O f, id(S), 0>" @@ -532,17 +532,17 @@ left_comp_inverse [OF bij_is_inj]) done -lemma iso_is_bij: "h \ G \ H ==> h \ bij(carrier(G), carrier(H))" +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)" + "auto(G) \ iso(G,G)" definition AutoGroup :: "i=>i" where - "AutoGroup(G) == update_carrier(BijGroup(carrier(G)), auto(G))" + "AutoGroup(G) \ update_carrier(BijGroup(carrier(G)), auto(G))" lemma (in group) id_in_auto: "id(carrier(G)) \ auto(G)" @@ -577,23 +577,23 @@ definition r_coset :: "[i,i,i] => i" (infixl \#>\\ 60) where - "H #>\<^bsub>G\<^esub> a == \h\H. {h \\<^bsub>G\<^esub> a}" + "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}" + "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}" + "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}" + "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}" + "set_inv\<^bsub>G\<^esub> H \ \h\H. {inv\<^bsub>G\<^esub> h}" locale normal = subgroup + group + @@ -654,7 +654,7 @@ subsection \Normal subgroups\ -lemma normal_imp_subgroup: "H \ G ==> subgroup(H,G)" +lemma normal_imp_subgroup: "H \ G \ subgroup(H,G)" by (simp add: normal_def subgroup_def) lemma (in group) normalI: @@ -861,7 +861,7 @@ definition r_congruent :: "[i,i] => i" (\rcong\ _\ [60] 60) where - "rcong\<^bsub>G\<^esub> H == { \ carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \\<^bsub>G\<^esub> y \ H}" + "rcong\<^bsub>G\<^esub> H \ { \ carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \\<^bsub>G\<^esub> y \ H}" lemma (in subgroup) equiv_rcong: @@ -941,7 +941,7 @@ definition order :: "i => i" where - "order(S) == |carrier(S)|" + "order(S) \ |carrier(S)|" lemma (in group) rcos_self: assumes "subgroup(H, G)" @@ -1022,7 +1022,7 @@ definition FactGroup :: "[i,i] => i" (infixl \Mod\ 65) where \ \Actually defined for groups rather than monoids\ - "G Mod H == + "G Mod H \ G\<^esub> H, \ \ (rcosets\<^bsub>G\<^esub> H) \ (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>" lemma (in normal) setmult_closed: @@ -1086,7 +1086,7 @@ definition kernel :: "[i,i,i] => i" where \ \the kernel of a homomorphism\ - "kernel(G,H,h) == {x \ carrier(G). h ` x = \\<^bsub>H\<^esub>}" + "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) @@ -1127,13 +1127,13 @@ 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'" + "\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)" + "\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: diff -r f2094906e491 -r e44d86131648 src/ZF/ex/LList.thy --- a/src/ZF/ex/LList.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/LList.thy Tue Sep 27 16:51:35 2022 +0100 @@ -33,21 +33,21 @@ domains "lleq(A)" \ "llist(A) * llist(A)" intros LNil: " \ lleq(A)" - LCons: "[| a \ A; \ lleq(A) |] - ==> \ 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))" + "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))" + 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 @@ -77,7 +77,7 @@ lemma LCons_iff: "LCons(a,l)=LCons(a',l') \ a=a' & l=l'" by auto -lemma LNil_LCons_iff: "~ LNil=LCons(a,l)" +lemma LNil_LCons_iff: "\ LNil=LCons(a,l)" by auto (* @@ -90,7 +90,7 @@ (*** Lemmas to justify using "llist" in other recursive type definitions ***) -lemma llist_mono: "A \ B ==> llist(A) \ llist(B)" +lemma llist_mono: "A \ B \ llist(A) \ llist(B)" apply (unfold llist.defs ) apply (rule gfp_mono) apply (rule llist.bnd_mono) @@ -107,7 +107,7 @@ declare Ord_in_Ord [elim!] lemma llist_quniv_lemma: - "Ord(i) ==> l \ llist(quniv(A)) \ l \ Vset(i) \ univ(eclose(A))" + "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))\ @@ -142,7 +142,7 @@ (*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*) lemma lleq_Int_Vset_subset: - "Ord(i) ==> \ lleq(A) \ l \ Vset(i) \ l'" + "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)\ @@ -156,20 +156,20 @@ qed (*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*) -lemma lleq_symmetric: " \ lleq(A) ==> \ lleq(A)" +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'" +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)" + "\l=l'; l \ llist(A)\ \ \ lleq(A)" apply (rule_tac X = "{. l \ llist (A) }" in lleq.coinduct) apply blast apply safe @@ -195,12 +195,12 @@ 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)" +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)" +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) @@ -212,7 +212,7 @@ flip_LCons [simp] not_type [simp] -lemma bool_Int_subset_univ: "b \ bool ==> b \ X \ univ(eclose(A))" +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!] @@ -221,7 +221,7 @@ (*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))" + "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)\ @@ -237,10 +237,10 @@ qed qed -lemma flip_in_quniv: "l \ llist(bool) ==> flip(l) \ quniv(bool)" +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)" +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) @@ -248,7 +248,7 @@ apply (erule_tac a=la in llist.cases, auto) done -lemma flip_flip: "l \ llist(bool) ==> flip(flip(l)) = l" +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 diff -r f2094906e491 -r e44d86131648 src/ZF/ex/Limit.thy --- a/src/ZF/ex/Limit.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/Limit.thy Tue Sep 27 16:51:35 2022 +0100 @@ -21,15 +21,15 @@ definition rel :: "[i,i,i]=>o" where - "rel(D,x,y) == :snd(D)" + "rel(D,x,y) \ :snd(D)" definition set :: "i=>i" where - "set(D) == fst(D)" + "set(D) \ fst(D)" definition po :: "i=>o" where - "po(D) == + "po(D) \ (\x \ set(D). rel(D,x,x)) & (\x \ set(D). \y \ set(D). \z \ set(D). rel(D,x,y) \ rel(D,y,z) \ rel(D,x,z)) & @@ -38,65 +38,65 @@ definition chain :: "[i,i]=>o" where (* Chains are object level functions nat->set(D) *) - "chain(D,X) == X \ nat->set(D) & (\n \ nat. rel(D,X`n,X`(succ(n))))" + "chain(D,X) \ X \ nat->set(D) & (\n \ nat. rel(D,X`n,X`(succ(n))))" definition isub :: "[i,i,i]=>o" where - "isub(D,X,x) == x \ set(D) & (\n \ nat. rel(D,X`n,x))" + "isub(D,X,x) \ x \ set(D) & (\n \ nat. rel(D,X`n,x))" definition islub :: "[i,i,i]=>o" where - "islub(D,X,x) == isub(D,X,x) & (\y. isub(D,X,y) \ rel(D,x,y))" + "islub(D,X,x) \ isub(D,X,x) & (\y. isub(D,X,y) \ rel(D,x,y))" definition lub :: "[i,i]=>i" where - "lub(D,X) == THE x. islub(D,X,x)" + "lub(D,X) \ THE x. islub(D,X,x)" definition cpo :: "i=>o" where - "cpo(D) == po(D) & (\X. chain(D,X) \ (\x. islub(D,X,x)))" + "cpo(D) \ po(D) & (\X. chain(D,X) \ (\x. islub(D,X,x)))" definition pcpo :: "i=>o" where - "pcpo(D) == cpo(D) & (\x \ set(D). \y \ set(D). rel(D,x,y))" + "pcpo(D) \ cpo(D) & (\x \ set(D). \y \ set(D). rel(D,x,y))" definition bot :: "i=>i" where - "bot(D) == THE x. x \ set(D) & (\y \ set(D). rel(D,x,y))" + "bot(D) \ THE x. x \ set(D) & (\y \ set(D). rel(D,x,y))" definition mono :: "[i,i]=>i" where - "mono(D,E) == + "mono(D,E) \ {f \ set(D)->set(E). \x \ set(D). \y \ set(D). rel(D,x,y) \ rel(E,f`x,f`y)}" definition cont :: "[i,i]=>i" where - "cont(D,E) == + "cont(D,E) \ {f \ mono(D,E). \X. chain(D,X) \ f`(lub(D,X)) = lub(E,\n \ nat. f`(X`n))}" definition cf :: "[i,i]=>i" where - "cf(D,E) == + "cf(D,E) \ cont(D,E)*cont(D,E). \x \ set(D). rel(E,(fst(y))`x,(snd(y))`x)}>" definition suffix :: "[i,i]=>i" where - "suffix(X,n) == \m \ nat. X`(n #+ m)" + "suffix(X,n) \ \m \ nat. X`(n #+ m)" definition subchain :: "[i,i]=>o" where - "subchain(X,Y) == \m \ nat. \n \ nat. X`m = Y`(m #+ n)" + "subchain(X,Y) \ \m \ nat. \n \ nat. X`m = Y`(m #+ n)" definition dominate :: "[i,i,i]=>o" where - "dominate(D,X,Y) == \m \ nat. \n \ nat. rel(D,X`m,Y`n)" + "dominate(D,X,Y) \ \m \ nat. \n \ nat. rel(D,X`m,Y`n)" definition matrix :: "[i,i]=>o" where - "matrix(D,M) == + "matrix(D,M) \ M \ nat -> (nat -> set(D)) & (\n \ nat. \m \ nat. rel(D,M`n`m,M`succ(n)`m)) & (\n \ nat. \m \ nat. rel(D,M`n`m,M`n`succ(m))) & @@ -104,22 +104,22 @@ definition projpair :: "[i,i,i,i]=>o" where - "projpair(D,E,e,p) == + "projpair(D,E,e,p) \ e \ cont(D,E) & p \ cont(E,D) & p O e = id(set(D)) & rel(cf(E,E),e O p,id(set(E)))" definition emb :: "[i,i,i]=>o" where - "emb(D,E,e) == \p. projpair(D,E,e,p)" + "emb(D,E,e) \ \p. projpair(D,E,e,p)" definition Rp :: "[i,i,i]=>i" where - "Rp(D,E,e) == THE p. projpair(D,E,e,p)" + "Rp(D,E,e) \ THE p. projpair(D,E,e,p)" definition (* Twice, constructions on cpos are more difficult. *) iprod :: "i=>i" where - "iprod(DD) == + "iprod(DD) \ <(\n \ nat. set(DD`n)), {x:(\n \ nat. set(DD`n))*(\n \ nat. set(DD`n)). \n \ nat. rel(DD`n,fst(x)`n,snd(x)`n)}>" @@ -127,66 +127,66 @@ definition mkcpo :: "[i,i=>o]=>i" where (* Cannot use rel(D), is meta fun, need two more args *) - "mkcpo(D,P) == + "mkcpo(D,P) \ <{x \ set(D). P(x)},{x \ set(D)*set(D). rel(D,fst(x),snd(x))}>" definition subcpo :: "[i,i]=>o" where - "subcpo(D,E) == + "subcpo(D,E) \ set(D) \ set(E) & (\x \ set(D). \y \ set(D). rel(D,x,y) \ rel(E,x,y)) & (\X. chain(D,X) \ lub(E,X):set(D))" definition subpcpo :: "[i,i]=>o" where - "subpcpo(D,E) == subcpo(D,E) & bot(E):set(D)" + "subpcpo(D,E) \ subcpo(D,E) & bot(E):set(D)" definition emb_chain :: "[i,i]=>o" where - "emb_chain(DD,ee) == + "emb_chain(DD,ee) \ (\n \ nat. cpo(DD`n)) & (\n \ nat. emb(DD`n,DD`succ(n),ee`n))" definition Dinf :: "[i,i]=>i" where - "Dinf(DD,ee) == + "Dinf(DD,ee) \ mkcpo(iprod(DD)) (%x. \n \ nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)" definition e_less :: "[i,i,i,i]=>i" where (* Valid for m \ n only. *) - "e_less(DD,ee,m,n) == rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)" + "e_less(DD,ee,m,n) \ rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)" definition e_gr :: "[i,i,i,i]=>i" where (* Valid for n \ m only. *) - "e_gr(DD,ee,m,n) == + "e_gr(DD,ee,m,n) \ rec(m#-n,id(set(DD`n)), %x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))" definition eps :: "[i,i,i,i]=>i" where - "eps(DD,ee,m,n) == if(m \ n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))" + "eps(DD,ee,m,n) \ if(m \ n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))" definition rho_emb :: "[i,i,i]=>i" where - "rho_emb(DD,ee,n) == \x \ set(DD`n). \m \ nat. eps(DD,ee,n,m)`x" + "rho_emb(DD,ee,n) \ \x \ set(DD`n). \m \ nat. eps(DD,ee,n,m)`x" definition rho_proj :: "[i,i,i]=>i" where - "rho_proj(DD,ee,n) == \x \ set(Dinf(DD,ee)). x`n" + "rho_proj(DD,ee,n) \ \x \ set(Dinf(DD,ee)). x`n" definition commute :: "[i,i,i,i=>i]=>o" where - "commute(DD,ee,E,r) == + "commute(DD,ee,E,r) \ (\n \ nat. emb(DD`n,E,r(n))) & (\m \ nat. \n \ nat. m \ n \ r(n) O eps(DD,ee,m,n) = r(m))" definition mediating :: "[i,i,i=>i,i=>i,i]=>o" where - "mediating(E,G,r,f,t) == emb(E,G,t) & (\n \ nat. f(n) = t O r(n))" + "mediating(E,G,r,f,t) \ emb(E,G,t) & (\n \ nat. f(n) = t O r(n))" lemmas nat_linear_le = Ord_linear_le [OF nat_into_Ord nat_into_Ord] @@ -196,102 +196,102 @@ (* Basic results. *) (*----------------------------------------------------------------------*) -lemma set_I: "x \ fst(D) ==> x \ set(D)" +lemma set_I: "x \ fst(D) \ x \ set(D)" by (simp add: set_def) -lemma rel_I: ":snd(D) ==> rel(D,x,y)" +lemma rel_I: ":snd(D) \ rel(D,x,y)" by (simp add: rel_def) -lemma rel_E: "rel(D,x,y) ==> :snd(D)" +lemma rel_E: "rel(D,x,y) \ :snd(D)" by (simp add: rel_def) (*----------------------------------------------------------------------*) (* I/E/D rules for po and cpo. *) (*----------------------------------------------------------------------*) -lemma po_refl: "[|po(D); x \ set(D)|] ==> rel(D,x,x)" +lemma po_refl: "\po(D); x \ set(D)\ \ rel(D,x,x)" by (unfold po_def, blast) -lemma po_trans: "[|po(D); rel(D,x,y); rel(D,y,z); x \ set(D); - y \ set(D); z \ set(D)|] ==> rel(D,x,z)" +lemma po_trans: "\po(D); rel(D,x,y); rel(D,y,z); x \ set(D); + y \ set(D); z \ set(D)\ \ rel(D,x,z)" by (unfold po_def, blast) lemma po_antisym: - "[|po(D); rel(D,x,y); rel(D,y,x); x \ set(D); y \ set(D)|] ==> x = y" + "\po(D); rel(D,x,y); rel(D,y,x); x \ set(D); y \ set(D)\ \ x = y" by (unfold po_def, blast) lemma poI: - "[| !!x. x \ set(D) ==> rel(D,x,x); - !!x y z. [| rel(D,x,y); rel(D,y,z); x \ set(D); y \ set(D); z \ set(D)|] - ==> rel(D,x,z); - !!x y. [| rel(D,x,y); rel(D,y,x); x \ set(D); y \ set(D)|] ==> x=y |] - ==> po(D)" + "\\x. x \ set(D) \ rel(D,x,x); + \x y z. \rel(D,x,y); rel(D,y,z); x \ set(D); y \ set(D); z \ set(D)\ + \ rel(D,x,z); + \x y. \rel(D,x,y); rel(D,y,x); x \ set(D); y \ set(D)\ \ x=y\ + \ po(D)" by (unfold po_def, blast) -lemma cpoI: "[| po(D); !!X. chain(D,X) ==> islub(D,X,x(D,X))|] ==> cpo(D)" +lemma cpoI: "\po(D); \X. chain(D,X) \ islub(D,X,x(D,X))\ \ cpo(D)" by (simp add: cpo_def, blast) -lemma cpo_po: "cpo(D) ==> po(D)" +lemma cpo_po: "cpo(D) \ po(D)" by (simp add: cpo_def) -lemma cpo_refl [simp,intro!,TC]: "[|cpo(D); x \ set(D)|] ==> rel(D,x,x)" +lemma cpo_refl [simp,intro!,TC]: "\cpo(D); x \ set(D)\ \ rel(D,x,x)" by (blast intro: po_refl cpo_po) lemma cpo_trans: - "[|cpo(D); rel(D,x,y); rel(D,y,z); x \ set(D); - y \ set(D); z \ set(D)|] ==> rel(D,x,z)" + "\cpo(D); rel(D,x,y); rel(D,y,z); x \ set(D); + y \ set(D); z \ set(D)\ \ rel(D,x,z)" by (blast intro: cpo_po po_trans) lemma cpo_antisym: - "[|cpo(D); rel(D,x,y); rel(D,y,x); x \ set(D); y \ set(D)|] ==> x = y" + "\cpo(D); rel(D,x,y); rel(D,y,x); x \ set(D); y \ set(D)\ \ x = y" by (blast intro: cpo_po po_antisym) -lemma cpo_islub: "[|cpo(D); chain(D,X); !!x. islub(D,X,x) ==> R|] ==> R" +lemma cpo_islub: "\cpo(D); chain(D,X); \x. islub(D,X,x) \ R\ \ R" by (simp add: cpo_def, blast) (*----------------------------------------------------------------------*) (* Theorems about isub and islub. *) (*----------------------------------------------------------------------*) -lemma islub_isub: "islub(D,X,x) ==> isub(D,X,x)" +lemma islub_isub: "islub(D,X,x) \ isub(D,X,x)" by (simp add: islub_def) -lemma islub_in: "islub(D,X,x) ==> x \ set(D)" +lemma islub_in: "islub(D,X,x) \ x \ set(D)" by (simp add: islub_def isub_def) -lemma islub_ub: "[|islub(D,X,x); n \ nat|] ==> rel(D,X`n,x)" +lemma islub_ub: "\islub(D,X,x); n \ nat\ \ rel(D,X`n,x)" by (simp add: islub_def isub_def) -lemma islub_least: "[|islub(D,X,x); isub(D,X,y)|] ==> rel(D,x,y)" +lemma islub_least: "\islub(D,X,x); isub(D,X,y)\ \ rel(D,x,y)" by (simp add: islub_def) lemma islubI: - "[|isub(D,X,x); !!y. isub(D,X,y) ==> rel(D,x,y)|] ==> islub(D,X,x)" + "\isub(D,X,x); \y. isub(D,X,y) \ rel(D,x,y)\ \ islub(D,X,x)" by (simp add: islub_def) lemma isubI: - "[|x \ set(D); !!n. n \ nat ==> rel(D,X`n,x)|] ==> isub(D,X,x)" + "\x \ set(D); \n. n \ nat \ rel(D,X`n,x)\ \ isub(D,X,x)" by (simp add: isub_def) lemma isubE: - "[|isub(D,X,x); [|x \ set(D); !!n. n \ nat==>rel(D,X`n,x)|] ==> P - |] ==> P" + "\isub(D,X,x); \x \ set(D); \n. n \ nat\rel(D,X`n,x)\ \ P +\ \ P" by (simp add: isub_def) -lemma isubD1: "isub(D,X,x) ==> x \ set(D)" +lemma isubD1: "isub(D,X,x) \ x \ set(D)" by (simp add: isub_def) -lemma isubD2: "[|isub(D,X,x); n \ nat|]==>rel(D,X`n,x)" +lemma isubD2: "\isub(D,X,x); n \ nat\\rel(D,X`n,x)" by (simp add: isub_def) -lemma islub_unique: "[|islub(D,X,x); islub(D,X,y); cpo(D)|] ==> x = y" +lemma islub_unique: "\islub(D,X,x); islub(D,X,y); cpo(D)\ \ x = y" by (blast intro: cpo_antisym islub_least islub_isub islub_in) (*----------------------------------------------------------------------*) (* lub gives the least upper bound of chains. *) (*----------------------------------------------------------------------*) -lemma cpo_lub: "[|chain(D,X); cpo(D)|] ==> islub(D,X,lub(D,X))" +lemma cpo_lub: "\chain(D,X); cpo(D)\ \ islub(D,X,lub(D,X))" apply (simp add: lub_def) apply (best elim: cpo_islub intro: theI islub_unique) done @@ -301,29 +301,29 @@ (*----------------------------------------------------------------------*) lemma chainI: - "[|X \ nat->set(D); !!n. n \ nat ==> rel(D,X`n,X`succ(n))|] ==> chain(D,X)" + "\X \ nat->set(D); \n. n \ nat \ rel(D,X`n,X`succ(n))\ \ chain(D,X)" by (simp add: chain_def) -lemma chain_fun: "chain(D,X) ==> X \ nat -> set(D)" +lemma chain_fun: "chain(D,X) \ X \ nat -> set(D)" by (simp add: chain_def) -lemma chain_in [simp,TC]: "[|chain(D,X); n \ nat|] ==> X`n \ set(D)" +lemma chain_in [simp,TC]: "\chain(D,X); n \ nat\ \ X`n \ set(D)" apply (simp add: chain_def) apply (blast dest: apply_type) done lemma chain_rel [simp,TC]: - "[|chain(D,X); n \ nat|] ==> rel(D, X ` n, X ` succ(n))" + "\chain(D,X); n \ nat\ \ rel(D, X ` n, X ` succ(n))" by (simp add: chain_def) lemma chain_rel_gen_add: - "[|chain(D,X); cpo(D); n \ nat; m \ nat|] ==> rel(D,X`n,(X`(m #+ n)))" + "\chain(D,X); cpo(D); n \ nat; m \ nat\ \ rel(D,X`n,(X`(m #+ n)))" apply (induct_tac m) apply (auto intro: cpo_trans) done lemma chain_rel_gen: - "[|n \ m; chain(D,X); cpo(D); m \ nat|] ==> rel(D,X`n,X`m)" + "\n \ m; chain(D,X); cpo(D); m \ nat\ \ rel(D,X`n,X`m)" apply (frule lt_nat_in_nat, erule nat_succI) apply (erule rev_mp) (*prepare the induction*) apply (induct_tac m) @@ -335,46 +335,46 @@ (*----------------------------------------------------------------------*) lemma pcpoI: - "[|!!y. y \ set(D)==>rel(D,x,y); x \ set(D); cpo(D)|]==>pcpo(D)" + "\\y. y \ set(D)\rel(D,x,y); x \ set(D); cpo(D)\\pcpo(D)" by (simp add: pcpo_def, auto) -lemma pcpo_cpo [TC]: "pcpo(D) ==> cpo(D)" +lemma pcpo_cpo [TC]: "pcpo(D) \ cpo(D)" by (simp add: pcpo_def) lemma pcpo_bot_ex1: - "pcpo(D) ==> \! x. x \ set(D) & (\y \ set(D). rel(D,x,y))" + "pcpo(D) \ \! x. x \ set(D) & (\y \ set(D). rel(D,x,y))" apply (simp add: pcpo_def) apply (blast intro: cpo_antisym) done lemma bot_least [TC]: - "[| pcpo(D); y \ set(D)|] ==> rel(D,bot(D),y)" + "\pcpo(D); y \ set(D)\ \ rel(D,bot(D),y)" apply (simp add: bot_def) apply (best intro: pcpo_bot_ex1 [THEN theI2]) done lemma bot_in [TC]: - "pcpo(D) ==> bot(D):set(D)" + "pcpo(D) \ bot(D):set(D)" apply (simp add: bot_def) apply (best intro: pcpo_bot_ex1 [THEN theI2]) done lemma bot_unique: - "[| pcpo(D); x \ set(D); !!y. y \ set(D) ==> rel(D,x,y)|] ==> x = bot(D)" + "\pcpo(D); x \ set(D); \y. y \ set(D) \ rel(D,x,y)\ \ x = bot(D)" by (blast intro: cpo_antisym pcpo_cpo bot_in bot_least) (*----------------------------------------------------------------------*) (* Constant chains and lubs and cpos. *) (*----------------------------------------------------------------------*) -lemma chain_const: "[|x \ set(D); cpo(D)|] ==> chain(D,(\n \ nat. x))" +lemma chain_const: "\x \ set(D); cpo(D)\ \ chain(D,(\n \ nat. x))" by (simp add: chain_def) lemma islub_const: - "[|x \ set(D); cpo(D)|] ==> islub(D,(\n \ nat. x),x)" + "\x \ set(D); cpo(D)\ \ islub(D,(\n \ nat. x),x)" by (simp add: islub_def isub_def, blast) -lemma lub_const: "[|x \ set(D); cpo(D)|] ==> lub(D,\n \ nat. x) = x" +lemma lub_const: "\x \ set(D); cpo(D)\ \ lub(D,\n \ nat. x) = x" by (blast intro: islub_unique cpo_lub chain_const islub_const) (*----------------------------------------------------------------------*) @@ -382,18 +382,18 @@ (*----------------------------------------------------------------------*) lemma isub_suffix: - "[| chain(D,X); cpo(D) |] ==> isub(D,suffix(X,n),x) \ isub(D,X,x)" + "\chain(D,X); cpo(D)\ \ isub(D,suffix(X,n),x) \ isub(D,X,x)" apply (simp add: isub_def suffix_def, safe) apply (drule_tac x = na in bspec) apply (auto intro: cpo_trans chain_rel_gen_add) done lemma islub_suffix: - "[|chain(D,X); cpo(D)|] ==> islub(D,suffix(X,n),x) \ islub(D,X,x)" + "\chain(D,X); cpo(D)\ \ islub(D,suffix(X,n),x) \ islub(D,X,x)" by (simp add: islub_def isub_suffix) lemma lub_suffix: - "[|chain(D,X); cpo(D)|] ==> lub(D,suffix(X,n)) = lub(D,X)" + "\chain(D,X); cpo(D)\ \ lub(D,suffix(X,n)) = lub(D,X)" by (simp add: lub_def islub_suffix) (*----------------------------------------------------------------------*) @@ -401,31 +401,31 @@ (*----------------------------------------------------------------------*) lemma dominateI: - "[| !!m. m \ nat ==> n(m):nat; !!m. m \ nat ==> rel(D,X`m,Y`n(m))|] - ==> dominate(D,X,Y)" + "\\m. m \ nat \ n(m):nat; \m. m \ nat \ rel(D,X`m,Y`n(m))\ + \ dominate(D,X,Y)" by (simp add: dominate_def, blast) lemma dominate_isub: - "[|dominate(D,X,Y); isub(D,Y,x); cpo(D); - X \ nat->set(D); Y \ nat->set(D)|] ==> isub(D,X,x)" + "\dominate(D,X,Y); isub(D,Y,x); cpo(D); + X \ nat->set(D); Y \ nat->set(D)\ \ isub(D,X,x)" apply (simp add: isub_def dominate_def) apply (blast intro: cpo_trans intro!: apply_funtype) done lemma dominate_islub: - "[|dominate(D,X,Y); islub(D,X,x); islub(D,Y,y); cpo(D); - X \ nat->set(D); Y \ nat->set(D)|] ==> rel(D,x,y)" + "\dominate(D,X,Y); islub(D,X,x); islub(D,Y,y); cpo(D); + X \ nat->set(D); Y \ nat->set(D)\ \ rel(D,x,y)" apply (simp add: islub_def) apply (blast intro: dominate_isub) done lemma subchain_isub: - "[|subchain(Y,X); isub(D,X,x)|] ==> isub(D,Y,x)" + "\subchain(Y,X); isub(D,X,x)\ \ isub(D,Y,x)" by (simp add: isub_def subchain_def, force) lemma dominate_islub_eq: - "[|dominate(D,X,Y); subchain(Y,X); islub(D,X,x); islub(D,Y,y); cpo(D); - X \ nat->set(D); Y \ nat->set(D)|] ==> x = y" + "\dominate(D,X,Y); subchain(Y,X); islub(D,X,x); islub(D,Y,y); cpo(D); + X \ nat->set(D); Y \ nat->set(D)\ \ x = y" by (blast intro: cpo_antisym dominate_islub islub_least subchain_isub islub_isub islub_in) @@ -434,55 +434,55 @@ (* Matrix. *) (*----------------------------------------------------------------------*) -lemma matrix_fun: "matrix(D,M) ==> M \ nat -> (nat -> set(D))" +lemma matrix_fun: "matrix(D,M) \ M \ nat -> (nat -> set(D))" by (simp add: matrix_def) -lemma matrix_in_fun: "[|matrix(D,M); n \ nat|] ==> M`n \ nat -> set(D)" +lemma matrix_in_fun: "\matrix(D,M); n \ nat\ \ M`n \ nat -> set(D)" by (blast intro: apply_funtype matrix_fun) -lemma matrix_in: "[|matrix(D,M); n \ nat; m \ nat|] ==> M`n`m \ set(D)" +lemma matrix_in: "\matrix(D,M); n \ nat; m \ nat\ \ M`n`m \ set(D)" by (blast intro: apply_funtype matrix_in_fun) lemma matrix_rel_1_0: - "[|matrix(D,M); n \ nat; m \ nat|] ==> rel(D,M`n`m,M`succ(n)`m)" + "\matrix(D,M); n \ nat; m \ nat\ \ rel(D,M`n`m,M`succ(n)`m)" by (simp add: matrix_def) lemma matrix_rel_0_1: - "[|matrix(D,M); n \ nat; m \ nat|] ==> rel(D,M`n`m,M`n`succ(m))" + "\matrix(D,M); n \ nat; m \ nat\ \ rel(D,M`n`m,M`n`succ(m))" by (simp add: matrix_def) lemma matrix_rel_1_1: - "[|matrix(D,M); n \ nat; m \ nat|] ==> rel(D,M`n`m,M`succ(n)`succ(m))" + "\matrix(D,M); n \ nat; m \ nat\ \ rel(D,M`n`m,M`succ(n)`succ(m))" by (simp add: matrix_def) -lemma fun_swap: "f \ X->Y->Z ==> (\y \ Y. \x \ X. f`x`y):Y->X->Z" +lemma fun_swap: "f \ X->Y->Z \ (\y \ Y. \x \ X. f`x`y):Y->X->Z" by (blast intro: lam_type apply_funtype) lemma matrix_sym_axis: - "matrix(D,M) ==> matrix(D,\m \ nat. \n \ nat. M`n`m)" + "matrix(D,M) \ matrix(D,\m \ nat. \n \ nat. M`n`m)" by (simp add: matrix_def fun_swap) lemma matrix_chain_diag: - "matrix(D,M) ==> chain(D,\n \ nat. M`n`n)" + "matrix(D,M) \ chain(D,\n \ nat. M`n`n)" apply (simp add: chain_def) apply (auto intro: lam_type matrix_in matrix_rel_1_1) done lemma matrix_chain_left: - "[|matrix(D,M); n \ nat|] ==> chain(D,M`n)" + "\matrix(D,M); n \ nat\ \ chain(D,M`n)" apply (unfold chain_def) apply (auto intro: matrix_fun [THEN apply_type] matrix_in matrix_rel_0_1) done lemma matrix_chain_right: - "[|matrix(D,M); m \ nat|] ==> chain(D,\n \ nat. M`n`m)" + "\matrix(D,M); m \ nat\ \ chain(D,\n \ nat. M`n`m)" apply (simp add: chain_def) apply (auto intro: lam_type matrix_in matrix_rel_1_0) done lemma matrix_chainI: - assumes xprem: "!!x. x \ nat==>chain(D,M`x)" - and yprem: "!!y. y \ nat==>chain(D,\x \ nat. M`x`y)" + assumes xprem: "\x. x \ nat\chain(D,M`x)" + and yprem: "\y. y \ nat\chain(D,\x \ nat. M`x`y)" and Mfun: "M \ nat->nat->set(D)" and cpoD: "cpo(D)" shows "matrix(D,M)" @@ -506,13 +506,13 @@ qed lemma lemma2: - "[|x \ nat; m \ nat; rel(D,(\n \ nat. M`n`m1)`x,(\n \ nat. M`n`m1)`m)|] - ==> rel(D,M`x`m1,M`m`m1)" + "\x \ nat; m \ nat; rel(D,(\n \ nat. M`n`m1)`x,(\n \ nat. M`n`m1)`m)\ + \ rel(D,M`x`m1,M`m`m1)" by simp lemma isub_lemma: - "[|isub(D, \n \ nat. M`n`n, y); matrix(D,M); cpo(D)|] - ==> isub(D, \n \ nat. lub(D,\m \ nat. M`n`m), y)" + "\isub(D, \n \ nat. M`n`n, y); matrix(D,M); cpo(D)\ + \ isub(D, \n \ nat. lub(D,\m \ nat. M`n`m), y)" proof (simp add: isub_def, safe) fix n assume DM: "matrix(D, M)" and D: "cpo(D)" and n: "n \ nat" and y: "y \ set(D)" @@ -546,7 +546,7 @@ qed lemma matrix_chain_lub: - "[|matrix(D,M); cpo(D)|] ==> chain(D,\n \ nat. lub(D,\m \ nat. M`n`m))" + "\matrix(D,M); cpo(D)\ \ chain(D,\n \ nat. lub(D,\m \ nat. M`n`m))" proof (simp add: chain_def, intro conjI ballI) assume "matrix(D, M)" "cpo(D)" thus "(\x\nat. lub(D, Lambda(nat, (`)(M ` x)))) \ nat \ set(D)" @@ -591,16 +591,16 @@ by (simp add: lub_def) lemma lub_matrix_diag: - "[|matrix(D,M); cpo(D)|] - ==> lub(D,(\n \ nat. lub(D,\m \ nat. M`n`m))) = + "\matrix(D,M); cpo(D)\ + \ lub(D,(\n \ nat. lub(D,\m \ nat. M`n`m))) = lub(D,(\n \ nat. M`n`n))" apply (simp (no_asm) add: lub_matrix_diag_aux1 lub_matrix_diag_aux2) apply (simp add: islub_def isub_eq) done lemma lub_matrix_diag_sym: - "[|matrix(D,M); cpo(D)|] - ==> lub(D,(\m \ nat. lub(D,\n \ nat. M`n`m))) = + "\matrix(D,M); cpo(D)\ + \ lub(D,(\m \ nat. lub(D,\n \ nat. M`n`m))) = lub(D,(\n \ nat. M`n`n))" by (drule matrix_sym_axis [THEN lub_matrix_diag], auto) @@ -609,49 +609,49 @@ (*----------------------------------------------------------------------*) lemma monoI: - "[|f \ set(D)->set(E); - !!x y. [|rel(D,x,y); x \ set(D); y \ set(D)|] ==> rel(E,f`x,f`y)|] - ==> f \ mono(D,E)" + "\f \ set(D)->set(E); + \x y. \rel(D,x,y); x \ set(D); y \ set(D)\ \ rel(E,f`x,f`y)\ + \ f \ mono(D,E)" by (simp add: mono_def) -lemma mono_fun: "f \ mono(D,E) ==> f \ set(D)->set(E)" +lemma mono_fun: "f \ mono(D,E) \ f \ set(D)->set(E)" by (simp add: mono_def) -lemma mono_map: "[|f \ mono(D,E); x \ set(D)|] ==> f`x \ set(E)" +lemma mono_map: "\f \ mono(D,E); x \ set(D)\ \ f`x \ set(E)" by (blast intro!: mono_fun [THEN apply_type]) lemma mono_mono: - "[|f \ mono(D,E); rel(D,x,y); x \ set(D); y \ set(D)|] ==> rel(E,f`x,f`y)" + "\f \ mono(D,E); rel(D,x,y); x \ set(D); y \ set(D)\ \ rel(E,f`x,f`y)" by (simp add: mono_def) lemma contI: - "[|f \ set(D)->set(E); - !!x y. [|rel(D,x,y); x \ set(D); y \ set(D)|] ==> rel(E,f`x,f`y); - !!X. chain(D,X) ==> f`lub(D,X) = lub(E,\n \ nat. f`(X`n))|] - ==> f \ cont(D,E)" + "\f \ set(D)->set(E); + \x y. \rel(D,x,y); x \ set(D); y \ set(D)\ \ rel(E,f`x,f`y); + \X. chain(D,X) \ f`lub(D,X) = lub(E,\n \ nat. f`(X`n))\ + \ f \ cont(D,E)" by (simp add: cont_def mono_def) -lemma cont2mono: "f \ cont(D,E) ==> f \ mono(D,E)" +lemma cont2mono: "f \ cont(D,E) \ f \ mono(D,E)" by (simp add: cont_def) -lemma cont_fun [TC]: "f \ cont(D,E) ==> f \ set(D)->set(E)" +lemma cont_fun [TC]: "f \ cont(D,E) \ f \ set(D)->set(E)" apply (simp add: cont_def) apply (rule mono_fun, blast) done -lemma cont_map [TC]: "[|f \ cont(D,E); x \ set(D)|] ==> f`x \ set(E)" +lemma cont_map [TC]: "\f \ cont(D,E); x \ set(D)\ \ f`x \ set(E)" by (blast intro!: cont_fun [THEN apply_type]) declare comp_fun [TC] lemma cont_mono: - "[|f \ cont(D,E); rel(D,x,y); x \ set(D); y \ set(D)|] ==> rel(E,f`x,f`y)" + "\f \ cont(D,E); rel(D,x,y); x \ set(D); y \ set(D)\ \ rel(E,f`x,f`y)" apply (simp add: cont_def) apply (blast intro!: mono_mono) done lemma cont_lub: - "[|f \ cont(D,E); chain(D,X)|] ==> f`(lub(D,X)) = lub(E,\n \ nat. f`(X`n))" + "\f \ cont(D,E); chain(D,X)\ \ f`(lub(D,X)) = lub(E,\n \ nat. f`(X`n))" by (simp add: cont_def) (*----------------------------------------------------------------------*) @@ -659,13 +659,13 @@ (*----------------------------------------------------------------------*) lemma mono_chain: - "[|f \ mono(D,E); chain(D,X)|] ==> chain(E,\n \ nat. f`(X`n))" + "\f \ mono(D,E); chain(D,X)\ \ chain(E,\n \ nat. f`(X`n))" apply (simp (no_asm) add: chain_def) apply (blast intro: lam_type mono_map chain_in mono_mono chain_rel) done lemma cont_chain: - "[|f \ cont(D,E); chain(D,X)|] ==> chain(E,\n \ nat. f`(X`n))" + "\f \ cont(D,E); chain(D,X)\ \ chain(E,\n \ nat. f`(X`n))" by (blast intro: mono_chain cont2mono) (*----------------------------------------------------------------------*) @@ -674,21 +674,21 @@ (* The following development more difficult with cpo-as-relation approach. *) -lemma cf_cont: "f \ set(cf(D,E)) ==> f \ cont(D,E)" +lemma cf_cont: "f \ set(cf(D,E)) \ f \ cont(D,E)" by (simp add: set_def cf_def) lemma cont_cf: (* Non-trivial with relation *) - "f \ cont(D,E) ==> f \ set(cf(D,E))" + "f \ cont(D,E) \ f \ set(cf(D,E))" by (simp add: set_def cf_def) (* rel_cf originally an equality. Now stated as two rules. Seemed easiest. *) lemma rel_cfI: - "[|!!x. x \ set(D) ==> rel(E,f`x,g`x); f \ cont(D,E); g \ cont(D,E)|] - ==> rel(cf(D,E),f,g)" + "\\x. x \ set(D) \ rel(E,f`x,g`x); f \ cont(D,E); g \ cont(D,E)\ + \ rel(cf(D,E),f,g)" by (simp add: rel_I cf_def) -lemma rel_cf: "[|rel(cf(D,E),f,g); x \ set(D)|] ==> rel(E,f`x,g`x)" +lemma rel_cf: "\rel(cf(D,E),f,g); x \ set(D)\ \ rel(E,f`x,g`x)" by (simp add: rel_def cf_def) (*----------------------------------------------------------------------*) @@ -696,15 +696,15 @@ (*----------------------------------------------------------------------*) lemma chain_cf: - "[| chain(cf(D,E),X); x \ set(D)|] ==> chain(E,\n \ nat. X`n`x)" + "\chain(cf(D,E),X); x \ set(D)\ \ chain(E,\n \ nat. X`n`x)" apply (rule chainI) apply (blast intro: lam_type apply_funtype cont_fun cf_cont chain_in, simp) apply (blast intro: rel_cf chain_rel) done lemma matrix_lemma: - "[|chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E) |] - ==> matrix(E,\x \ nat. \xa \ nat. X`x`(Xa`xa))" + "\chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E)\ + \ matrix(E,\x \ nat. \xa \ nat. X`x`(Xa`xa))" apply (rule matrix_chainI, auto) apply (force intro: chainI lam_type apply_funtype cont_fun cf_cont cont_mono) apply (force intro: chainI lam_type apply_funtype cont_fun cf_cont rel_cf) @@ -745,8 +745,8 @@ qed lemma islub_cf: - "[| chain(cf(D,E),X); cpo(D); cpo(E)|] - ==> islub(cf(D,E), X, \x \ set(D). lub(E,\n \ nat. X`n`x))" + "\chain(cf(D,E),X); cpo(D); cpo(E)\ + \ islub(cf(D,E), X, \x \ set(D). lub(E,\n \ nat. X`n`x))" apply (rule islubI) apply (rule isubI) apply (rule chain_cf_lub_cont [THEN cont_cf], assumption+) @@ -762,7 +762,7 @@ done lemma cpo_cf [TC]: - "[| cpo(D); cpo(E)|] ==> cpo(cf(D,E))" + "\cpo(D); cpo(E)\ \ cpo(cf(D,E))" apply (rule poI [THEN cpoI]) apply (rule rel_cfI) apply (assumption | rule cpo_refl cf_cont [THEN cont_fun, THEN apply_type] @@ -780,13 +780,13 @@ done lemma lub_cf: - "[| chain(cf(D,E),X); cpo(D); cpo(E)|] - ==> lub(cf(D,E), X) = (\x \ set(D). lub(E,\n \ nat. X`n`x))" + "\chain(cf(D,E),X); cpo(D); cpo(E)\ + \ lub(cf(D,E), X) = (\x \ set(D). lub(E,\n \ nat. X`n`x))" by (blast intro: islub_unique cpo_lub islub_cf cpo_cf) lemma const_cont [TC]: - "[|y \ set(E); cpo(D); cpo(E)|] ==> (\x \ set(D).y) \ cont(D,E)" + "\y \ set(E); cpo(D); cpo(E)\ \ (\x \ set(D).y) \ cont(D,E)" apply (rule contI) prefer 2 apply simp apply (blast intro: lam_type) @@ -794,19 +794,19 @@ done lemma cf_least: - "[|cpo(D); pcpo(E); y \ cont(D,E)|]==>rel(cf(D,E),(\x \ set(D).bot(E)),y)" + "\cpo(D); pcpo(E); y \ cont(D,E)\\rel(cf(D,E),(\x \ set(D).bot(E)),y)" apply (rule rel_cfI, simp, typecheck) done lemma pcpo_cf: - "[|cpo(D); pcpo(E)|] ==> pcpo(cf(D,E))" + "\cpo(D); pcpo(E)\ \ pcpo(cf(D,E))" apply (rule pcpoI) apply (assumption | rule cf_least bot_in const_cont [THEN cont_cf] cf_cont cpo_cf pcpo_cpo)+ done lemma bot_cf: - "[|cpo(D); pcpo(E)|] ==> bot(cf(D,E)) = (\x \ set(D).bot(E))" + "\cpo(D); pcpo(E)\ \ bot(cf(D,E)) = (\x \ set(D).bot(E))" by (blast intro: bot_unique [symmetric] pcpo_cf cf_least bot_in [THEN const_cont, THEN cont_cf] cf_cont pcpo_cpo) @@ -815,13 +815,13 @@ (*----------------------------------------------------------------------*) lemma id_cont [TC,intro!]: - "cpo(D) ==> id(set(D)) \ cont(D,D)" + "cpo(D) \ id(set(D)) \ cont(D,D)" by (simp add: id_type contI cpo_lub [THEN islub_in] chain_fun [THEN eta]) lemmas comp_cont_apply = cont_fun [THEN comp_fun_apply] lemma comp_pres_cont [TC]: - "[| f \ cont(D',E); g \ cont(D,D'); cpo(D)|] ==> f O g \ cont(D,E)" + "\f \ cont(D',E); g \ cont(D,D'); cpo(D)\ \ f O g \ cont(D,E)" apply (rule contI) apply (rule_tac [2] comp_cont_apply [THEN ssubst]) apply (rule_tac [4] comp_cont_apply [THEN ssubst]) @@ -837,9 +837,9 @@ lemma comp_mono: - "[| f \ cont(D',E); g \ cont(D,D'); f':cont(D',E); g':cont(D,D'); - rel(cf(D',E),f,f'); rel(cf(D,D'),g,g'); cpo(D); cpo(E) |] - ==> rel(cf(D,E),f O g,f' O g')" + "\f \ cont(D',E); g \ cont(D,D'); f':cont(D',E); g':cont(D,D'); + rel(cf(D',E),f,f'); rel(cf(D,D'),g,g'); cpo(D); cpo(E)\ + \ rel(cf(D,E),f O g,f' O g')" apply (rule rel_cfI) apply (subst comp_cont_apply) apply (rule_tac [3] comp_cont_apply [THEN ssubst]) @@ -848,8 +848,8 @@ done lemma chain_cf_comp: - "[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(E)|] - ==> chain(cf(D,E),\n \ nat. X`n O Y`n)" + "\chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(E)\ + \ chain(cf(D,E),\n \ nat. X`n O Y`n)" apply (rule chainI) defer 1 apply simp @@ -864,8 +864,8 @@ done lemma comp_lubs: - "[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(D'); cpo(E)|] - ==> lub(cf(D',E),X) O lub(cf(D,D'),Y) = lub(cf(D,E),\n \ nat. X`n O Y`n)" + "\chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(D'); cpo(E)\ + \ lub(cf(D',E),X) O lub(cf(D,D'),Y) = lub(cf(D,E),\n \ nat. X`n O Y`n)" apply (rule fun_extension) apply (rule_tac [3] lub_cf [THEN ssubst]) apply (assumption | @@ -891,23 +891,23 @@ (*----------------------------------------------------------------------*) lemma projpairI: - "[| e \ cont(D,E); p \ cont(E,D); p O e = id(set(D)); - rel(cf(E,E))(e O p)(id(set(E)))|] ==> projpair(D,E,e,p)" + "\e \ cont(D,E); p \ cont(E,D); p O e = id(set(D)); + rel(cf(E,E))(e O p)(id(set(E)))\ \ projpair(D,E,e,p)" by (simp add: projpair_def) -lemma projpair_e_cont: "projpair(D,E,e,p) ==> e \ cont(D,E)" +lemma projpair_e_cont: "projpair(D,E,e,p) \ e \ cont(D,E)" by (simp add: projpair_def) -lemma projpair_p_cont: "projpair(D,E,e,p) ==> p \ cont(E,D)" +lemma projpair_p_cont: "projpair(D,E,e,p) \ p \ cont(E,D)" by (simp add: projpair_def) -lemma projpair_ep_cont: "projpair(D,E,e,p) ==> e \ cont(D,E) & p \ cont(E,D)" +lemma projpair_ep_cont: "projpair(D,E,e,p) \ e \ cont(D,E) & p \ cont(E,D)" by (simp add: projpair_def) -lemma projpair_eq: "projpair(D,E,e,p) ==> p O e = id(set(D))" +lemma projpair_eq: "projpair(D,E,e,p) \ p O e = id(set(D))" by (simp add: projpair_def) -lemma projpair_rel: "projpair(D,E,e,p) ==> rel(cf(E,E))(e O p)(id(set(E)))" +lemma projpair_rel: "projpair(D,E,e,p) \ rel(cf(E,E))(e O p)(id(set(E)))" by (simp add: projpair_def) @@ -924,8 +924,8 @@ and comp_id = fun_is_rel [THEN right_comp_id] lemma projpair_unique_aux1: - "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p'); - rel(cf(D,E),e,e')|] ==> rel(cf(E,D),p',p)" + "\cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p'); + rel(cf(D,E),e,e')\ \ rel(cf(E,D),p',p)" apply (rule_tac b=p' in projpair_p_cont [THEN cont_fun, THEN id_comp, THEN subst], assumption) apply (rule projpair_eq [THEN subst], assumption) @@ -947,8 +947,8 @@ text\Proof's very like the previous one. Is there a pattern that could be exploited?\ lemma projpair_unique_aux2: - "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p'); - rel(cf(E,D),p',p)|] ==> rel(cf(D,E),e,e')" + "\cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p'); + rel(cf(E,D),p',p)\ \ rel(cf(D,E),e,e')" apply (rule_tac b=e in projpair_e_cont [THEN cont_fun, THEN comp_id, THEN subst], assumption) @@ -968,26 +968,26 @@ lemma projpair_unique: - "[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p')|] - ==> (e=e')\(p=p')" + "\cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p')\ + \ (e=e')\(p=p')" by (blast intro: cpo_antisym projpair_unique_aux1 projpair_unique_aux2 cpo_cf cont_cf dest: projpair_ep_cont) (* Slightly different, more asms, since THE chooses the unique element. *) lemma embRp: - "[|emb(D,E,e); cpo(D); cpo(E)|] ==> projpair(D,E,e,Rp(D,E,e))" + "\emb(D,E,e); cpo(D); cpo(E)\ \ projpair(D,E,e,Rp(D,E,e))" apply (simp add: emb_def Rp_def) apply (blast intro: theI2 projpair_unique [THEN iffD1]) done -lemma embI: "projpair(D,E,e,p) ==> emb(D,E,e)" +lemma embI: "projpair(D,E,e,p) \ emb(D,E,e)" by (simp add: emb_def, auto) -lemma Rp_unique: "[|projpair(D,E,e,p); cpo(D); cpo(E)|] ==> Rp(D,E,e) = p" +lemma Rp_unique: "\projpair(D,E,e,p); cpo(D); cpo(E)\ \ Rp(D,E,e) = p" by (blast intro: embRp embI projpair_unique [THEN iffD1]) -lemma emb_cont [TC]: "emb(D,E,e) ==> e \ cont(D,E)" +lemma emb_cont [TC]: "emb(D,E,e) \ e \ cont(D,E)" apply (simp add: emb_def) apply (blast intro: projpair_e_cont) done @@ -1000,7 +1000,7 @@ lemma embRp_eq_thm: - "[|emb(D,E,e); x \ set(D); cpo(D); cpo(E)|] ==> Rp(D,E,e)`(e`x) = x" + "\emb(D,E,e); x \ set(D); cpo(D); cpo(E)\ \ Rp(D,E,e)`(e`x) = x" apply (rule comp_fun_apply [THEN subst]) apply (assumption | rule Rp_cont emb_cont cont_fun)+ apply (subst embRp_eq) @@ -1013,17 +1013,17 @@ (*----------------------------------------------------------------------*) lemma projpair_id: - "cpo(D) ==> projpair(D,D,id(set(D)),id(set(D)))" + "cpo(D) \ projpair(D,D,id(set(D)),id(set(D)))" apply (simp add: projpair_def) apply (blast intro: cpo_cf cont_cf) done lemma emb_id: - "cpo(D) ==> emb(D,D,id(set(D)))" + "cpo(D) \ emb(D,D,id(set(D)))" by (auto intro: embI projpair_id) lemma Rp_id: - "cpo(D) ==> Rp(D,D,id(set(D))) = id(set(D))" + "cpo(D) \ Rp(D,D,id(set(D))) = id(set(D))" by (auto intro: Rp_unique projpair_id) @@ -1036,8 +1036,8 @@ (* Proof in Isa/ZF: 23 lines (compared to 56: 60% reduction). *) lemma comp_lemma: - "[|emb(D,D',e); emb(D',E,e'); cpo(D); cpo(D'); cpo(E)|] - ==> projpair(D,E,e' O e,(Rp(D,D',e)) O (Rp(D',E,e')))" + "\emb(D,D',e); emb(D',E,e'); cpo(D); cpo(D'); cpo(E)\ + \ projpair(D,E,e' O e,(Rp(D,D',e)) O (Rp(D',E,e')))" apply (simp add: projpair_def, safe) apply (assumption | rule comp_pres_cont Rp_cont emb_cont)+ apply (rule comp_assoc [THEN subst]) @@ -1071,27 +1071,27 @@ (*----------------------------------------------------------------------*) lemma iprodI: - "x:(\n \ nat. set(DD`n)) ==> x \ set(iprod(DD))" + "x:(\n \ nat. set(DD`n)) \ x \ set(iprod(DD))" by (simp add: set_def iprod_def) lemma iprodE: - "x \ set(iprod(DD)) ==> x:(\n \ nat. set(DD`n))" + "x \ set(iprod(DD)) \ x:(\n \ nat. set(DD`n))" by (simp add: set_def iprod_def) (* Contains typing conditions in contrast to HOL-ST *) lemma rel_iprodI: - "[|!!n. n \ nat ==> rel(DD`n,f`n,g`n); f:(\n \ nat. set(DD`n)); - g:(\n \ nat. set(DD`n))|] ==> rel(iprod(DD),f,g)" + "\\n. n \ nat \ rel(DD`n,f`n,g`n); f:(\n \ nat. set(DD`n)); + g:(\n \ nat. set(DD`n))\ \ rel(iprod(DD),f,g)" by (simp add: iprod_def rel_I) lemma rel_iprodE: - "[|rel(iprod(DD),f,g); n \ nat|] ==> rel(DD`n,f`n,g`n)" + "\rel(iprod(DD),f,g); n \ nat\ \ rel(DD`n,f`n,g`n)" by (simp add: iprod_def rel_def) lemma chain_iprod: - "[|chain(iprod(DD),X); !!n. n \ nat ==> cpo(DD`n); n \ nat|] - ==> chain(DD`n,\m \ nat. X`m`n)" + "\chain(iprod(DD),X); \n. n \ nat \ cpo(DD`n); n \ nat\ + \ chain(DD`n,\m \ nat. X`m`n)" apply (unfold chain_def, safe) apply (rule lam_type) apply (rule apply_type) @@ -1101,8 +1101,8 @@ done lemma islub_iprod: - "[|chain(iprod(DD),X); !!n. n \ nat ==> cpo(DD`n)|] - ==> islub(iprod(DD),X,\n \ nat. lub(DD`n,\m \ nat. X`m`n))" + "\chain(iprod(DD),X); \n. n \ nat \ cpo(DD`n)\ + \ islub(iprod(DD),X,\n \ nat. lub(DD`n,\m \ nat. X`m`n))" apply (simp add: islub_def isub_def, safe) apply (rule iprodI) apply (blast intro: lam_type chain_iprod [THEN cpo_lub, THEN islub_in]) @@ -1123,7 +1123,7 @@ done lemma cpo_iprod [TC]: - "(!!n. n \ nat ==> cpo(DD`n)) ==> cpo(iprod(DD))" + "(\n. n \ nat \ cpo(DD`n)) \ cpo(iprod(DD))" apply (assumption | rule cpoI poI)+ apply (rule rel_iprodI) (*not repeated: want to solve 1, leave 2 unchanged *) apply (simp | rule cpo_refl iprodE [THEN apply_type] iprodE)+ @@ -1140,8 +1140,8 @@ lemma lub_iprod: - "[|chain(iprod(DD),X); !!n. n \ nat ==> cpo(DD`n)|] - ==> lub(iprod(DD),X) = (\n \ nat. lub(DD`n,\m \ nat. X`m`n))" + "\chain(iprod(DD),X); \n. n \ nat \ cpo(DD`n)\ + \ lub(iprod(DD),X) = (\n \ nat. lub(DD`n,\m \ nat. X`m`n))" by (blast intro: cpo_lub [THEN islub_unique] islub_iprod cpo_iprod) @@ -1150,39 +1150,39 @@ (*----------------------------------------------------------------------*) lemma subcpoI: - "[|set(D)<=set(E); - !!x y. [|x \ set(D); y \ set(D)|] ==> rel(D,x,y)\rel(E,x,y); - !!X. chain(D,X) ==> lub(E,X) \ set(D)|] ==> subcpo(D,E)" + "\set(D)<=set(E); + \x y. \x \ set(D); y \ set(D)\ \ rel(D,x,y)\rel(E,x,y); + \X. chain(D,X) \ lub(E,X) \ set(D)\ \ subcpo(D,E)" by (simp add: subcpo_def) -lemma subcpo_subset: "subcpo(D,E) ==> set(D)<=set(E)" +lemma subcpo_subset: "subcpo(D,E) \ set(D)<=set(E)" by (simp add: subcpo_def) lemma subcpo_rel_eq: - "[|subcpo(D,E); x \ set(D); y \ set(D)|] ==> rel(D,x,y)\rel(E,x,y)" + "\subcpo(D,E); x \ set(D); y \ set(D)\ \ rel(D,x,y)\rel(E,x,y)" by (simp add: subcpo_def) lemmas subcpo_relD1 = subcpo_rel_eq [THEN iffD1] lemmas subcpo_relD2 = subcpo_rel_eq [THEN iffD2] -lemma subcpo_lub: "[|subcpo(D,E); chain(D,X)|] ==> lub(E,X) \ set(D)" +lemma subcpo_lub: "\subcpo(D,E); chain(D,X)\ \ lub(E,X) \ set(D)" by (simp add: subcpo_def) -lemma chain_subcpo: "[|subcpo(D,E); chain(D,X)|] ==> chain(E,X)" +lemma chain_subcpo: "\subcpo(D,E); chain(D,X)\ \ chain(E,X)" by (blast intro: Pi_type [THEN chainI] chain_fun subcpo_relD1 subcpo_subset [THEN subsetD] chain_in chain_rel) -lemma ub_subcpo: "[|subcpo(D,E); chain(D,X); isub(D,X,x)|] ==> isub(E,X,x)" +lemma ub_subcpo: "\subcpo(D,E); chain(D,X); isub(D,X,x)\ \ isub(E,X,x)" by (blast intro: isubI subcpo_relD1 subcpo_relD1 chain_in isubD1 isubD2 subcpo_subset [THEN subsetD] chain_in chain_rel) lemma islub_subcpo: - "[|subcpo(D,E); cpo(E); chain(D,X)|] ==> islub(D,X,lub(E,X))" + "\subcpo(D,E); cpo(E); chain(D,X)\ \ islub(D,X,lub(E,X))" by (blast intro: islubI isubI subcpo_lub subcpo_relD2 chain_in islub_ub islub_least cpo_lub chain_subcpo isubD1 ub_subcpo) -lemma subcpo_cpo: "[|subcpo(D,E); cpo(E)|] ==> cpo(D)" +lemma subcpo_cpo: "\subcpo(D,E); cpo(E)\ \ cpo(D)" apply (assumption | rule cpoI poI)+ apply (simp add: subcpo_rel_eq) apply (assumption | rule cpo_refl subcpo_subset [THEN subsetD])+ @@ -1193,39 +1193,39 @@ apply (fast intro: islub_subcpo) done -lemma lub_subcpo: "[|subcpo(D,E); cpo(E); chain(D,X)|] ==> lub(D,X) = lub(E,X)" +lemma lub_subcpo: "\subcpo(D,E); cpo(E); chain(D,X)\ \ lub(D,X) = lub(E,X)" by (blast intro: cpo_lub [THEN islub_unique] islub_subcpo subcpo_cpo) (*----------------------------------------------------------------------*) (* Making subcpos using mkcpo. *) (*----------------------------------------------------------------------*) -lemma mkcpoI: "[|x \ set(D); P(x)|] ==> x \ set(mkcpo(D,P))" +lemma mkcpoI: "\x \ set(D); P(x)\ \ x \ set(mkcpo(D,P))" by (simp add: set_def mkcpo_def) -lemma mkcpoD1: "x \ set(mkcpo(D,P))==> x \ set(D)" +lemma mkcpoD1: "x \ set(mkcpo(D,P))\ x \ set(D)" by (simp add: set_def mkcpo_def) -lemma mkcpoD2: "x \ set(mkcpo(D,P))==> P(x)" +lemma mkcpoD2: "x \ set(mkcpo(D,P))\ P(x)" by (simp add: set_def mkcpo_def) -lemma rel_mkcpoE: "rel(mkcpo(D,P),x,y) ==> rel(D,x,y)" +lemma rel_mkcpoE: "rel(mkcpo(D,P),x,y) \ rel(D,x,y)" by (simp add: rel_def mkcpo_def) lemma rel_mkcpo: - "[|x \ set(D); y \ set(D)|] ==> rel(mkcpo(D,P),x,y) \ rel(D,x,y)" + "\x \ set(D); y \ set(D)\ \ rel(mkcpo(D,P),x,y) \ rel(D,x,y)" by (simp add: mkcpo_def rel_def set_def) lemma chain_mkcpo: - "chain(mkcpo(D,P),X) ==> chain(D,X)" + "chain(mkcpo(D,P),X) \ chain(D,X)" apply (rule chainI) apply (blast intro: Pi_type chain_fun chain_in [THEN mkcpoD1]) apply (blast intro: rel_mkcpo [THEN iffD1] chain_rel mkcpoD1 chain_in) done lemma subcpo_mkcpo: - "[|!!X. chain(mkcpo(D,P),X) ==> P(lub(D,X)); cpo(D)|] - ==> subcpo(mkcpo(D,P),D)" + "\\X. chain(mkcpo(D,P),X) \ P(lub(D,X)); cpo(D)\ + \ subcpo(mkcpo(D,P),D)" apply (intro subcpoI subsetI rel_mkcpo) apply (erule mkcpoD1)+ apply (blast intro: mkcpoI cpo_lub [THEN islub_in] chain_mkcpo) @@ -1236,15 +1236,15 @@ (*----------------------------------------------------------------------*) lemma emb_chainI: - "[|!!n. n \ nat ==> cpo(DD`n); - !!n. n \ nat ==> emb(DD`n,DD`succ(n),ee`n)|] ==> emb_chain(DD,ee)" + "\\n. n \ nat \ cpo(DD`n); + \n. n \ nat \ emb(DD`n,DD`succ(n),ee`n)\ \ emb_chain(DD,ee)" by (simp add: emb_chain_def) -lemma emb_chain_cpo [TC]: "[|emb_chain(DD,ee); n \ nat|] ==> cpo(DD`n)" +lemma emb_chain_cpo [TC]: "\emb_chain(DD,ee); n \ nat\ \ cpo(DD`n)" by (simp add: emb_chain_def) lemma emb_chain_emb: - "[|emb_chain(DD,ee); n \ nat|] ==> emb(DD`n,DD`succ(n),ee`n)" + "\emb_chain(DD,ee); n \ nat\ \ emb(DD`n,DD`succ(n),ee`n)" by (simp add: emb_chain_def) (*----------------------------------------------------------------------*) @@ -1252,44 +1252,44 @@ (*----------------------------------------------------------------------*) lemma DinfI: - "[|x:(\n \ nat. set(DD`n)); - !!n. n \ nat ==> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n|] - ==> x \ set(Dinf(DD,ee))" + "\x:(\n \ nat. set(DD`n)); + \n. n \ nat \ Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n\ + \ x \ set(Dinf(DD,ee))" apply (simp add: Dinf_def) apply (blast intro: mkcpoI iprodI) done -lemma Dinf_prod: "x \ set(Dinf(DD,ee)) ==> x:(\n \ nat. set(DD`n))" +lemma Dinf_prod: "x \ set(Dinf(DD,ee)) \ x:(\n \ nat. set(DD`n))" apply (simp add: Dinf_def) apply (erule mkcpoD1 [THEN iprodE]) done lemma Dinf_eq: - "[|x \ set(Dinf(DD,ee)); n \ nat|] - ==> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n" + "\x \ set(Dinf(DD,ee)); n \ nat\ + \ Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n" apply (simp add: Dinf_def) apply (blast dest: mkcpoD2) done lemma rel_DinfI: - "[|!!n. n \ nat ==> rel(DD`n,x`n,y`n); - x:(\n \ nat. set(DD`n)); y:(\n \ nat. set(DD`n))|] - ==> rel(Dinf(DD,ee),x,y)" + "\\n. n \ nat \ rel(DD`n,x`n,y`n); + x:(\n \ nat. set(DD`n)); y:(\n \ nat. set(DD`n))\ + \ rel(Dinf(DD,ee),x,y)" apply (simp add: Dinf_def) apply (blast intro: rel_mkcpo [THEN iffD2] rel_iprodI iprodI) done -lemma rel_Dinf: "[|rel(Dinf(DD,ee),x,y); n \ nat|] ==> rel(DD`n,x`n,y`n)" +lemma rel_Dinf: "\rel(Dinf(DD,ee),x,y); n \ nat\ \ rel(DD`n,x`n,y`n)" apply (simp add: Dinf_def) apply (erule rel_mkcpoE [THEN rel_iprodE], assumption) done -lemma chain_Dinf: "chain(Dinf(DD,ee),X) ==> chain(iprod(DD),X)" +lemma chain_Dinf: "chain(Dinf(DD,ee),X) \ chain(iprod(DD),X)" apply (simp add: Dinf_def) apply (erule chain_mkcpo) done -lemma subcpo_Dinf: "emb_chain(DD,ee) ==> subcpo(Dinf(DD,ee),iprod(DD))" +lemma subcpo_Dinf: "emb_chain(DD,ee) \ subcpo(Dinf(DD,ee),iprod(DD))" apply (simp add: Dinf_def) apply (rule subcpo_mkcpo) apply (simp add: Dinf_def [symmetric]) @@ -1308,7 +1308,7 @@ (* Simple example of existential reasoning in Isabelle versus HOL. *) -lemma cpo_Dinf: "emb_chain(DD,ee) ==> cpo(Dinf(DD,ee))" +lemma cpo_Dinf: "emb_chain(DD,ee) \ cpo(Dinf(DD,ee))" apply (rule subcpo_cpo) apply (erule subcpo_Dinf) apply (auto intro: cpo_iprod emb_chain_cpo) @@ -1318,8 +1318,8 @@ the proof steps are essentially the same (I think). *) lemma lub_Dinf: - "[|chain(Dinf(DD,ee),X); emb_chain(DD,ee)|] - ==> lub(Dinf(DD,ee),X) = (\n \ nat. lub(DD`n,\m \ nat. X`m`n))" + "\chain(Dinf(DD,ee),X); emb_chain(DD,ee)\ + \ lub(Dinf(DD,ee),X) = (\n \ nat. lub(DD`n,\m \ nat. X`m`n))" apply (subst subcpo_Dinf [THEN lub_subcpo]) apply (auto intro: cpo_iprod emb_chain_cpo lub_iprod chain_Dinf) done @@ -1330,7 +1330,7 @@ (*----------------------------------------------------------------------*) lemma e_less_eq: - "m \ nat ==> e_less(DD,ee,m,m) = id(set(DD`m))" + "m \ nat \ e_less(DD,ee,m,m) = id(set(DD`m))" by (simp add: e_less_def) lemma lemma_succ_sub: "succ(m#+n)#-m = succ(natify(n))" @@ -1341,12 +1341,12 @@ by (simp add: e_less_def) lemma le_exists: - "[| m \ n; !!x. [|n=m#+x; x \ nat|] ==> Q; n \ nat |] ==> Q" + "\m \ n; \x. \n=m#+x; x \ nat\ \ Q; n \ nat\ \ Q" apply (drule less_imp_succ_add, auto) done -lemma e_less_le: "[| m \ n; n \ nat |] - ==> e_less(DD,ee,m,succ(n)) = ee`n O e_less(DD,ee,m,n)" +lemma e_less_le: "\m \ n; n \ nat\ + \ e_less(DD,ee,m,succ(n)) = ee`n O e_less(DD,ee,m,n)" apply (rule le_exists, assumption) apply (simp add: e_less_add, assumption) done @@ -1354,12 +1354,12 @@ (* All theorems assume variables m and n are natural numbers. *) lemma e_less_succ: - "m \ nat ==> e_less(DD,ee,m,succ(m)) = ee`m O id(set(DD`m))" + "m \ nat \ e_less(DD,ee,m,succ(m)) = ee`m O id(set(DD`m))" by (simp add: e_less_le e_less_eq) lemma e_less_succ_emb: - "[|!!n. n \ nat ==> emb(DD`n,DD`succ(n),ee`n); m \ nat|] - ==> e_less(DD,ee,m,succ(m)) = ee`m" + "\\n. n \ nat \ emb(DD`n,DD`succ(n),ee`n); m \ nat\ + \ e_less(DD,ee,m,succ(m)) = ee`m" apply (simp add: e_less_succ) apply (blast intro: emb_cont cont_fun comp_id) done @@ -1368,8 +1368,8 @@ (* In any case the one below was very easy to write. *) lemma emb_e_less_add: - "[| emb_chain(DD,ee); m \ nat |] - ==> emb(DD`m, DD`(m#+k), e_less(DD,ee,m,m#+k))" + "\emb_chain(DD,ee); m \ nat\ + \ emb(DD`m, DD`(m#+k), e_less(DD,ee,m,m#+k))" apply (subgoal_tac "emb (DD`m, DD` (m#+natify (k)), e_less (DD,ee,m,m#+natify (k))) ") apply (rule_tac [2] n = "natify (k) " in nat_induct) @@ -1380,8 +1380,8 @@ done lemma emb_e_less: - "[| m \ n; emb_chain(DD,ee); n \ nat |] - ==> emb(DD`m, DD`n, e_less(DD,ee,m,n))" + "\m \ n; emb_chain(DD,ee); n \ nat\ + \ emb(DD`m, DD`n, e_less(DD,ee,m,n))" apply (frule lt_nat_in_nat) apply (erule nat_succI) (* same proof as e_less_le *) @@ -1389,7 +1389,7 @@ apply (simp add: emb_e_less_add, assumption) done -lemma comp_mono_eq: "[|f=f'; g=g'|] ==> f O g = f' O g'" +lemma comp_mono_eq: "\f=f'; g=g'\ \ f O g = f' O g'" by simp (* Note the object-level implication for induction on k. This @@ -1397,8 +1397,8 @@ Therefore this theorem is only a lemma. *) lemma e_less_split_add_lemma [rule_format]: - "[| emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> n \ k \ + "\emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ n \ k \ e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)" apply (induct_tac k) apply (simp add: e_less_eq id_type [THEN id_comp]) @@ -1419,41 +1419,41 @@ done lemma e_less_split_add: - "[| n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)" + "\n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)" by (blast intro: e_less_split_add_lemma) lemma e_gr_eq: - "m \ nat ==> e_gr(DD,ee,m,m) = id(set(DD`m))" + "m \ nat \ e_gr(DD,ee,m,m) = id(set(DD`m))" by (simp add: e_gr_def) lemma e_gr_add: - "[|n \ nat; k \ nat|] - ==> e_gr(DD,ee,succ(n#+k),n) = + "\n \ nat; k \ nat\ + \ e_gr(DD,ee,succ(n#+k),n) = e_gr(DD,ee,n#+k,n) O Rp(DD`(n#+k),DD`succ(n#+k),ee`(n#+k))" by (simp add: e_gr_def) lemma e_gr_le: - "[|n \ m; m \ nat; n \ nat|] - ==> e_gr(DD,ee,succ(m),n) = e_gr(DD,ee,m,n) O Rp(DD`m,DD`succ(m),ee`m)" + "\n \ m; m \ nat; n \ nat\ + \ e_gr(DD,ee,succ(m),n) = e_gr(DD,ee,m,n) O Rp(DD`m,DD`succ(m),ee`m)" apply (erule le_exists) apply (simp add: e_gr_add, assumption+) done lemma e_gr_succ: - "m \ nat ==> e_gr(DD,ee,succ(m),m) = id(set(DD`m)) O Rp(DD`m,DD`succ(m),ee`m)" + "m \ nat \ e_gr(DD,ee,succ(m),m) = id(set(DD`m)) O Rp(DD`m,DD`succ(m),ee`m)" by (simp add: e_gr_le e_gr_eq) (* Cpo asm's due to THE uniqueness. *) -lemma e_gr_succ_emb: "[|emb_chain(DD,ee); m \ nat|] - ==> e_gr(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)" +lemma e_gr_succ_emb: "\emb_chain(DD,ee); m \ nat\ + \ e_gr(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)" apply (simp add: e_gr_succ) apply (blast intro: id_comp Rp_cont cont_fun emb_chain_cpo emb_chain_emb) done lemma e_gr_fun_add: - "[|emb_chain(DD,ee); n \ nat; k \ nat|] - ==> e_gr(DD,ee,n#+k,n): set(DD`(n#+k))->set(DD`n)" + "\emb_chain(DD,ee); n \ nat; k \ nat\ + \ e_gr(DD,ee,n#+k,n): set(DD`(n#+k))->set(DD`n)" apply (induct_tac k) apply (simp add: e_gr_eq id_type) apply (simp add: e_gr_add) @@ -1461,15 +1461,15 @@ done lemma e_gr_fun: - "[|n \ m; emb_chain(DD,ee); m \ nat; n \ nat|] - ==> e_gr(DD,ee,m,n): set(DD`m)->set(DD`n)" + "\n \ m; emb_chain(DD,ee); m \ nat; n \ nat\ + \ e_gr(DD,ee,m,n): set(DD`m)->set(DD`n)" apply (rule le_exists, assumption) apply (simp add: e_gr_fun_add, assumption+) done lemma e_gr_split_add_lemma: - "[| emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> m \ k \ + "\emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ m \ k \ e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)" apply (induct_tac k) apply (rule impI) @@ -1491,20 +1491,20 @@ done lemma e_gr_split_add: - "[| m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)" + "\m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)" apply (blast intro: e_gr_split_add_lemma [THEN mp]) done lemma e_less_cont: - "[|m \ n; emb_chain(DD,ee); m \ nat; n \ nat|] - ==> e_less(DD,ee,m,n):cont(DD`m,DD`n)" + "\m \ n; emb_chain(DD,ee); m \ nat; n \ nat\ + \ e_less(DD,ee,m,n):cont(DD`m,DD`n)" apply (blast intro: emb_cont emb_e_less) done lemma e_gr_cont: - "[|n \ m; emb_chain(DD,ee); m \ nat; n \ nat|] - ==> e_gr(DD,ee,m,n):cont(DD`m,DD`n)" + "\n \ m; emb_chain(DD,ee); m \ nat; n \ nat\ + \ e_gr(DD,ee,m,n):cont(DD`m,DD`n)" apply (erule rev_mp) apply (induct_tac m) apply (simp add: le0_iff e_gr_eq nat_0I) @@ -1520,8 +1520,8 @@ (* Considerably shorter.... 57 against 26 *) lemma e_less_e_gr_split_add: - "[|n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> e_less(DD,ee,m,m#+n) = e_gr(DD,ee,m#+k,m#+n) O e_less(DD,ee,m,m#+k)" + "\n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ e_less(DD,ee,m,m#+n) = e_gr(DD,ee,m#+k,m#+n) O e_less(DD,ee,m,m#+k)" (* Use mp to prepare for induction. *) apply (erule rev_mp) apply (induct_tac k) @@ -1547,8 +1547,8 @@ (* Again considerably shorter, and easy to obtain from the previous thm. *) lemma e_gr_e_less_split_add: - "[|m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> e_gr(DD,ee,n#+m,n) = e_gr(DD,ee,n#+k,n) O e_less(DD,ee,n#+m,n#+k)" + "\m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ e_gr(DD,ee,n#+m,n) = e_gr(DD,ee,n#+k,n) O e_less(DD,ee,n#+m,n#+k)" (* Use mp to prepare for induction. *) apply (erule rev_mp) apply (induct_tac k) @@ -1572,54 +1572,54 @@ lemma emb_eps: - "[|m \ n; emb_chain(DD,ee); m \ nat; n \ nat|] - ==> emb(DD`m,DD`n,eps(DD,ee,m,n))" + "\m \ n; emb_chain(DD,ee); m \ nat; n \ nat\ + \ emb(DD`m,DD`n,eps(DD,ee,m,n))" apply (simp add: eps_def) apply (blast intro: emb_e_less) done lemma eps_fun: - "[|emb_chain(DD,ee); m \ nat; n \ nat|] - ==> eps(DD,ee,m,n): set(DD`m)->set(DD`n)" + "\emb_chain(DD,ee); m \ nat; n \ nat\ + \ eps(DD,ee,m,n): set(DD`m)->set(DD`n)" apply (simp add: eps_def) apply (auto intro: e_less_cont [THEN cont_fun] not_le_iff_lt [THEN iffD1, THEN leI] e_gr_fun nat_into_Ord) done -lemma eps_id: "n \ nat ==> eps(DD,ee,n,n) = id(set(DD`n))" +lemma eps_id: "n \ nat \ eps(DD,ee,n,n) = id(set(DD`n))" by (simp add: eps_def e_less_eq) lemma eps_e_less_add: - "[|m \ nat; n \ nat|] ==> eps(DD,ee,m,m#+n) = e_less(DD,ee,m,m#+n)" + "\m \ nat; n \ nat\ \ eps(DD,ee,m,m#+n) = e_less(DD,ee,m,m#+n)" by (simp add: eps_def add_le_self) lemma eps_e_less: - "[|m \ n; m \ nat; n \ nat|] ==> eps(DD,ee,m,n) = e_less(DD,ee,m,n)" + "\m \ n; m \ nat; n \ nat\ \ eps(DD,ee,m,n) = e_less(DD,ee,m,n)" by (simp add: eps_def) lemma eps_e_gr_add: - "[|n \ nat; k \ nat|] ==> eps(DD,ee,n#+k,n) = e_gr(DD,ee,n#+k,n)" + "\n \ nat; k \ nat\ \ eps(DD,ee,n#+k,n) = e_gr(DD,ee,n#+k,n)" by (simp add: eps_def e_less_eq e_gr_eq) lemma eps_e_gr: - "[|n \ m; m \ nat; n \ nat|] ==> eps(DD,ee,m,n) = e_gr(DD,ee,m,n)" + "\n \ m; m \ nat; n \ nat\ \ eps(DD,ee,m,n) = e_gr(DD,ee,m,n)" apply (erule le_exists) apply (simp_all add: eps_e_gr_add) done lemma eps_succ_ee: - "[|!!n. n \ nat ==> emb(DD`n,DD`succ(n),ee`n); m \ nat|] - ==> eps(DD,ee,m,succ(m)) = ee`m" + "\\n. n \ nat \ emb(DD`n,DD`succ(n),ee`n); m \ nat\ + \ eps(DD,ee,m,succ(m)) = ee`m" by (simp add: eps_e_less le_succ_iff e_less_succ_emb) lemma eps_succ_Rp: - "[|emb_chain(DD,ee); m \ nat|] - ==> eps(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)" + "\emb_chain(DD,ee); m \ nat\ + \ eps(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)" by (simp add: eps_e_gr le_succ_iff e_gr_succ_emb) lemma eps_cont: - "[|emb_chain(DD,ee); m \ nat; n \ nat|] ==> eps(DD,ee,m,n): cont(DD`m,DD`n)" + "\emb_chain(DD,ee); m \ nat; n \ nat\ \ eps(DD,ee,m,n): cont(DD`m,DD`n)" apply (rule_tac i = m and j = n in nat_linear_le) apply (simp_all add: eps_e_less e_less_cont eps_e_gr e_gr_cont) done @@ -1627,29 +1627,29 @@ (* Theorems about splitting. *) lemma eps_split_add_left: - "[|n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,m#+k) = eps(DD,ee,m#+n,m#+k) O eps(DD,ee,m,m#+n)" + "\n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,m#+k) = eps(DD,ee,m#+n,m#+k) O eps(DD,ee,m,m#+n)" apply (simp add: eps_e_less add_le_self add_le_mono) apply (auto intro: e_less_split_add) done lemma eps_split_add_left_rev: - "[|n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,m#+n) = eps(DD,ee,m#+k,m#+n) O eps(DD,ee,m,m#+k)" + "\n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,m#+n) = eps(DD,ee,m#+k,m#+n) O eps(DD,ee,m,m#+k)" apply (simp add: eps_e_less_add eps_e_gr add_le_self add_le_mono) apply (auto intro: e_less_e_gr_split_add) done lemma eps_split_add_right: - "[|m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,n#+k,n) = eps(DD,ee,n#+m,n) O eps(DD,ee,n#+k,n#+m)" + "\m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,n#+k,n) = eps(DD,ee,n#+m,n) O eps(DD,ee,n#+k,n#+m)" apply (simp add: eps_e_gr add_le_self add_le_mono) apply (auto intro: e_gr_split_add) done lemma eps_split_add_right_rev: - "[|m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,n#+m,n) = eps(DD,ee,n#+k,n) O eps(DD,ee,n#+m,n#+k)" + "\m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,n#+m,n) = eps(DD,ee,n#+k,n) O eps(DD,ee,n#+m,n#+k)" apply (simp add: eps_e_gr_add eps_e_less add_le_self add_le_mono) apply (auto intro: e_gr_e_less_split_add) done @@ -1657,38 +1657,38 @@ (* Arithmetic *) lemma le_exists_lemma: - "[| n \ k; k \ m; - !!p q. [|p \ q; k=n#+p; m=n#+q; p \ nat; q \ nat|] ==> R; - m \ nat |]==>R" + "\n \ k; k \ m; + \p q. \p \ q; k=n#+p; m=n#+q; p \ nat; q \ nat\ \ R; + m \ nat\\R" apply (rule le_exists, assumption) prefer 2 apply (simp add: lt_nat_in_nat) apply (rule le_trans [THEN le_exists], assumption+, force+) done lemma eps_split_left_le: - "[|m \ k; k \ n; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" + "\m \ k; k \ n; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" apply (rule le_exists_lemma, assumption+) apply (auto intro: eps_split_add_left) done lemma eps_split_left_le_rev: - "[|m \ n; n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" + "\m \ n; n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" apply (rule le_exists_lemma, assumption+) apply (auto intro: eps_split_add_left_rev) done lemma eps_split_right_le: - "[|n \ k; k \ m; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" + "\n \ k; k \ m; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" apply (rule le_exists_lemma, assumption+) apply (auto intro: eps_split_add_right) done lemma eps_split_right_le_rev: - "[|n \ m; m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" + "\n \ m; m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" apply (rule le_exists_lemma, assumption+) apply (auto intro: eps_split_add_right_rev) done @@ -1696,8 +1696,8 @@ (* The desired two theorems about `splitting'. *) lemma eps_split_left: - "[|m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" + "\m \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" apply (rule nat_linear_le) apply (rule_tac [4] eps_split_right_le_rev) prefer 4 apply assumption @@ -1708,8 +1708,8 @@ done lemma eps_split_right: - "[|n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat|] - ==> eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" + "\n \ k; emb_chain(DD,ee); m \ nat; n \ nat; k \ nat\ + \ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)" apply (rule nat_linear_le) apply (rule_tac [3] eps_split_left_le_rev) prefer 3 apply assumption @@ -1726,8 +1726,8 @@ (* Considerably shorter. *) lemma rho_emb_fun: - "[|emb_chain(DD,ee); n \ nat|] - ==> rho_emb(DD,ee,n): set(DD`n) -> set(Dinf(DD,ee))" + "\emb_chain(DD,ee); n \ nat\ + \ rho_emb(DD,ee,n): set(DD`n) -> set(Dinf(DD,ee))" apply (simp add: rho_emb_def) apply (assumption | rule lam_type DinfI eps_cont [THEN cont_fun, THEN apply_type])+ @@ -1760,21 +1760,21 @@ done lemma rho_emb_apply1: - "x \ set(DD`n) ==> rho_emb(DD,ee,n)`x = (\m \ nat. eps(DD,ee,n,m)`x)" + "x \ set(DD`n) \ rho_emb(DD,ee,n)`x = (\m \ nat. eps(DD,ee,n,m)`x)" by (simp add: rho_emb_def) lemma rho_emb_apply2: - "[|x \ set(DD`n); m \ nat|] ==> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x" + "\x \ set(DD`n); m \ nat\ \ rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x" by (simp add: rho_emb_def) -lemma rho_emb_id: "[| x \ set(DD`n); n \ nat|] ==> rho_emb(DD,ee,n)`x`n = x" +lemma rho_emb_id: "\x \ set(DD`n); n \ nat\ \ rho_emb(DD,ee,n)`x`n = x" by (simp add: rho_emb_apply2 eps_id) (* Shorter proof, 23 against 62. *) lemma rho_emb_cont: - "[|emb_chain(DD,ee); n \ nat|] - ==> rho_emb(DD,ee,n): cont(DD`n,Dinf(DD,ee))" + "\emb_chain(DD,ee); n \ nat\ + \ rho_emb(DD,ee,n): cont(DD`n,Dinf(DD,ee))" apply (rule contI) apply (assumption | rule rho_emb_fun)+ apply (rule rel_DinfI) @@ -1807,8 +1807,8 @@ (* 32 vs 61, using safe_tac with imp in asm would be unfortunate (5steps) *) lemma eps1_aux1: - "[|m \ n; emb_chain(DD,ee); x \ set(Dinf(DD,ee)); m \ nat; n \ nat|] - ==> rel(DD`n,e_less(DD,ee,m,n)`(x`m),x`n)" + "\m \ n; emb_chain(DD,ee); x \ set(Dinf(DD,ee)); m \ nat; n \ nat\ + \ rel(DD`n,e_less(DD,ee,m,n)`(x`m),x`n)" apply (erule rev_mp) (* For induction proof *) apply (induct_tac n) apply (rule impI) @@ -1851,8 +1851,8 @@ (* 18 vs 40 *) lemma eps1_aux2: - "[|n \ m; emb_chain(DD,ee); x \ set(Dinf(DD,ee)); m \ nat; n \ nat|] - ==> rel(DD`n,e_gr(DD,ee,m,n)`(x`m),x`n)" + "\n \ m; emb_chain(DD,ee); x \ set(Dinf(DD,ee)); m \ nat; n \ nat\ + \ rel(DD`n,e_gr(DD,ee,m,n)`(x`m),x`n)" apply (erule rev_mp) (* For induction proof *) apply (induct_tac m) apply (rule impI) @@ -1875,8 +1875,8 @@ done lemma eps1: - "[|emb_chain(DD,ee); x \ set(Dinf(DD,ee)); m \ nat; n \ nat|] - ==> rel(DD`n,eps(DD,ee,m,n)`(x`m),x`n)" + "\emb_chain(DD,ee); x \ set(Dinf(DD,ee)); m \ nat; n \ nat\ + \ rel(DD`n,eps(DD,ee,m,n)`(x`m),x`n)" apply (simp add: eps_def) apply (blast intro: eps1_aux1 not_le_iff_lt [THEN iffD1, THEN leI, THEN eps1_aux2] nat_into_Ord) @@ -1887,8 +1887,8 @@ Look for occurrences of rel_cfI, rel_DinfI, etc to evaluate the problem. *) lemma lam_Dinf_cont: - "[| emb_chain(DD,ee); n \ nat |] - ==> (\x \ set(Dinf(DD,ee)). x`n) \ cont(Dinf(DD,ee),DD`n)" + "\emb_chain(DD,ee); n \ nat\ + \ (\x \ set(Dinf(DD,ee)). x`n) \ cont(Dinf(DD,ee),DD`n)" apply (rule contI) apply (assumption | rule lam_type apply_type Dinf_prod)+ apply simp @@ -1899,8 +1899,8 @@ done lemma rho_projpair: - "[| emb_chain(DD,ee); n \ nat |] - ==> projpair(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n),rho_proj(DD,ee,n))" + "\emb_chain(DD,ee); n \ nat\ + \ projpair(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n),rho_proj(DD,ee,n))" apply (simp add: rho_proj_def) apply (rule projpairI) apply (assumption | rule rho_emb_cont)+ @@ -1934,11 +1934,11 @@ done lemma emb_rho_emb: - "[| emb_chain(DD,ee); n \ nat |] ==> emb(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))" + "\emb_chain(DD,ee); n \ nat\ \ emb(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))" by (auto simp add: emb_def intro: exI rho_projpair) -lemma rho_proj_cont: "[| emb_chain(DD,ee); n \ nat |] - ==> rho_proj(DD,ee,n) \ cont(Dinf(DD,ee),DD`n)" +lemma rho_proj_cont: "\emb_chain(DD,ee); n \ nat\ + \ rho_proj(DD,ee,n) \ cont(Dinf(DD,ee),DD`n)" by (auto intro: rho_projpair projpair_p_cont) (*----------------------------------------------------------------------*) @@ -1946,24 +1946,24 @@ (*----------------------------------------------------------------------*) lemma commuteI: - "[| !!n. n \ nat ==> emb(DD`n,E,r(n)); - !!m n. [|m \ n; m \ nat; n \ nat|] ==> r(n) O eps(DD,ee,m,n) = r(m) |] - ==> commute(DD,ee,E,r)" + "\\n. n \ nat \ emb(DD`n,E,r(n)); + \m n. \m \ n; m \ nat; n \ nat\ \ r(n) O eps(DD,ee,m,n) = r(m)\ + \ commute(DD,ee,E,r)" by (simp add: commute_def) lemma commute_emb [TC]: - "[| commute(DD,ee,E,r); n \ nat |] ==> emb(DD`n,E,r(n))" + "\commute(DD,ee,E,r); n \ nat\ \ emb(DD`n,E,r(n))" by (simp add: commute_def) lemma commute_eq: - "[| commute(DD,ee,E,r); m \ n; m \ nat; n \ nat |] - ==> r(n) O eps(DD,ee,m,n) = r(m) " + "\commute(DD,ee,E,r); m \ n; m \ nat; n \ nat\ + \ r(n) O eps(DD,ee,m,n) = r(m) " by (simp add: commute_def) (* Shorter proof: 11 vs 46 lines. *) lemma rho_emb_commute: - "emb_chain(DD,ee) ==> commute(DD,ee,Dinf(DD,ee),rho_emb(DD,ee))" + "emb_chain(DD,ee) \ commute(DD,ee,Dinf(DD,ee),rho_emb(DD,ee))" apply (rule commuteI) apply (assumption | rule emb_rho_emb)+ apply (rule fun_extension) (* Manual instantiation in HOL. *) @@ -1976,14 +1976,14 @@ apply (auto intro: eps_fun) done -lemma le_succ: "n \ nat ==> n \ succ(n)" +lemma le_succ: "n \ nat \ n \ succ(n)" by (simp add: le_succ_iff) (* Shorter proof: 21 vs 83 (106 - 23, due to OAssoc complication) *) lemma commute_chain: - "[| commute(DD,ee,E,r); emb_chain(DD,ee); cpo(E) |] - ==> chain(cf(E,E),\n \ nat. r(n) O Rp(DD`n,E,r(n)))" + "\commute(DD,ee,E,r); emb_chain(DD,ee); cpo(E)\ + \ chain(cf(E,E),\n \ nat. r(n) O Rp(DD`n,E,r(n)))" apply (rule chainI) apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont emb_cont emb_chain_cpo, @@ -2007,25 +2007,25 @@ done lemma rho_emb_chain: - "emb_chain(DD,ee) ==> + "emb_chain(DD,ee) \ chain(cf(Dinf(DD,ee),Dinf(DD,ee)), \n \ nat. rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))" by (auto intro: commute_chain rho_emb_commute cpo_Dinf) lemma rho_emb_chain_apply1: - "[| emb_chain(DD,ee); x \ set(Dinf(DD,ee)) |] - ==> chain(Dinf(DD,ee), + "\emb_chain(DD,ee); x \ set(Dinf(DD,ee))\ + \ chain(Dinf(DD,ee), \n \ nat. (rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))`x)" by (drule rho_emb_chain [THEN chain_cf], assumption, simp) lemma chain_iprod_emb_chain: - "[| chain(iprod(DD),X); emb_chain(DD,ee); n \ nat |] - ==> chain(DD`n,\m \ nat. X `m `n)" + "\chain(iprod(DD),X); emb_chain(DD,ee); n \ nat\ + \ chain(DD`n,\m \ nat. X `m `n)" by (auto intro: chain_iprod emb_chain_cpo) lemma rho_emb_chain_apply2: - "[| emb_chain(DD,ee); x \ set(Dinf(DD,ee)); n \ nat |] ==> + "\emb_chain(DD,ee); x \ set(Dinf(DD,ee)); n \ nat\ \ chain (DD`n, \xa \ nat. @@ -2037,7 +2037,7 @@ (* Shorter proof: 32 vs 72 (roughly), Isabelle proof has lemmas. *) lemma rho_emb_lub: - "emb_chain(DD,ee) ==> + "emb_chain(DD,ee) \ lub(cf(Dinf(DD,ee),Dinf(DD,ee)), \n \ nat. rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))) = id(set(Dinf(DD,ee)))" @@ -2078,9 +2078,9 @@ done lemma theta_chain: (* almost same proof as commute_chain *) - "[| commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G) |] - ==> chain(cf(E,G),\n \ nat. f(n) O Rp(DD`n,E,r(n)))" + "\commute(DD,ee,E,r); commute(DD,ee,G,f); + emb_chain(DD,ee); cpo(E); cpo(G)\ + \ chain(cf(E,G),\n \ nat. f(n) O Rp(DD`n,E,r(n)))" apply (rule chainI) apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont emb_cont emb_chain_cpo, @@ -2105,9 +2105,9 @@ done lemma theta_proj_chain: (* similar proof to theta_chain *) - "[| commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G) |] - ==> chain(cf(G,E),\n \ nat. r(n) O Rp(DD`n,G,f(n)))" + "\commute(DD,ee,E,r); commute(DD,ee,G,f); + emb_chain(DD,ee); cpo(E); cpo(G)\ + \ chain(cf(G,E),\n \ nat. r(n) O Rp(DD`n,G,f(n)))" apply (rule chainI) apply (blast intro: lam_type cont_cf comp_pres_cont commute_emb Rp_cont emb_cont emb_chain_cpo, simp) @@ -2138,9 +2138,9 @@ (* Controlled simplification inside lambda: introduce lemmas *) lemma commute_O_lemma: - "[| commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G); x \ nat |] - ==> r(x) O Rp(DD ` x, G, f(x)) O f(x) O Rp(DD ` x, E, r(x)) = + "\commute(DD,ee,E,r); commute(DD,ee,G,f); + emb_chain(DD,ee); cpo(E); cpo(G); x \ nat\ + \ r(x) O Rp(DD ` x, G, f(x)) O f(x) O Rp(DD ` x, E, r(x)) = r(x) O Rp(DD ` x, E, r(x))" apply (rule_tac s1 = "f (x) " in comp_assoc [THEN subst]) apply (subst embRp_eq) @@ -2152,10 +2152,10 @@ (* Shorter proof (but lemmas): 19 vs 79 (103 - 24, due to OAssoc) *) lemma theta_projpair: - "[| lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); + "\lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G) |] - ==> projpair + emb_chain(DD,ee); cpo(E); cpo(G)\ + \ projpair (E,G, lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n))), lub(cf(G,E), \n \ nat. r(n) O Rp(DD`n,G,f(n))))" @@ -2182,17 +2182,17 @@ done lemma emb_theta: - "[| lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); + "\lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G) |] - ==> emb(E,G,lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n))))" + emb_chain(DD,ee); cpo(E); cpo(G)\ + \ emb(E,G,lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n))))" apply (simp add: emb_def) apply (blast intro: theta_projpair) done lemma mono_lemma: - "[| g \ cont(D,D'); cpo(D); cpo(D'); cpo(E) |] - ==> (\f \ cont(D',E). f O g) \ mono(cf(D',E),cf(D,E))" + "\g \ cont(D,D'); cpo(D); cpo(D'); cpo(E)\ + \ (\f \ cont(D',E). f O g) \ mono(cf(D',E),cf(D,E))" apply (rule monoI) apply (simp add: set_def cf_def) apply (drule cf_cont)+ @@ -2201,9 +2201,9 @@ done lemma commute_lam_lemma: - "[| commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G); n \ nat |] - ==> (\na \ nat. (\f \ cont(E, G). f O r(n)) ` + "\commute(DD,ee,E,r); commute(DD,ee,G,f); + emb_chain(DD,ee); cpo(E); cpo(G); n \ nat\ + \ (\na \ nat. (\f \ cont(E, G). f O r(n)) ` ((\n \ nat. f(n) O Rp(DD ` n, E, r(n))) ` na)) = (\na \ nat. (f(na) O Rp(DD ` na, E, r(na))) O r(n))" apply (rule fun_extension) @@ -2212,18 +2212,18 @@ done lemma chain_lemma: - "[| commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G); n \ nat |] - ==> chain(cf(DD`n,G),\x \ nat. (f(x) O Rp(DD ` x, E, r(x))) O r(n))" + "\commute(DD,ee,E,r); commute(DD,ee,G,f); + emb_chain(DD,ee); cpo(E); cpo(G); n \ nat\ + \ chain(cf(DD`n,G),\x \ nat. (f(x) O Rp(DD ` x, E, r(x))) O r(n))" apply (rule commute_lam_lemma [THEN subst]) apply (blast intro: theta_chain emb_chain_cpo commute_emb [THEN emb_cont, THEN mono_lemma, THEN mono_chain])+ done lemma suffix_lemma: - "[| commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G); cpo(DD`x); x \ nat |] - ==> suffix(\n \ nat. (f(n) O Rp(DD`n,E,r(n))) O r(x),x) = + "\commute(DD,ee,E,r); commute(DD,ee,G,f); + emb_chain(DD,ee); cpo(E); cpo(G); cpo(DD`x); x \ nat\ + \ suffix(\n \ nat. (f(n) O Rp(DD`n,E,r(n))) O r(x),x) = (\n \ nat. f(x))" apply (simp add: suffix_def) apply (rule lam_type [THEN fun_extension]) @@ -2241,20 +2241,20 @@ lemma mediatingI: - "[|emb(E,G,t); !!n. n \ nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)" + "\emb(E,G,t); \n. n \ nat \ f(n) = t O r(n)\\mediating(E,G,r,f,t)" by (simp add: mediating_def) -lemma mediating_emb: "mediating(E,G,r,f,t) ==> emb(E,G,t)" +lemma mediating_emb: "mediating(E,G,r,f,t) \ emb(E,G,t)" by (simp add: mediating_def) -lemma mediating_eq: "[| mediating(E,G,r,f,t); n \ nat |] ==> f(n) = t O r(n)" +lemma mediating_eq: "\mediating(E,G,r,f,t); n \ nat\ \ f(n) = t O r(n)" by (simp add: mediating_def) lemma lub_universal_mediating: - "[| lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); + "\lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G) |] - ==> mediating(E,G,r,f,lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n))))" + emb_chain(DD,ee); cpo(E); cpo(G)\ + \ mediating(E,G,r,f,lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n))))" apply (assumption | rule mediatingI emb_theta)+ apply (rule_tac b = "r (n) " in lub_const [THEN subst]) apply (rule_tac [3] comp_lubs [THEN ssubst]) @@ -2268,11 +2268,11 @@ done lemma lub_universal_unique: - "[| mediating(E,G,r,f,t); + "\mediating(E,G,r,f,t); lub(cf(E,E), \n \ nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); commute(DD,ee,E,r); commute(DD,ee,G,f); - emb_chain(DD,ee); cpo(E); cpo(G) |] - ==> t = lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n)))" + emb_chain(DD,ee); cpo(E); cpo(G)\ + \ t = lub(cf(E,G), \n \ nat. f(n) O Rp(DD`n,E,r(n)))" apply (rule_tac b = t in comp_id [THEN subst]) apply (erule_tac [2] subst) apply (rule_tac [2] b = t in lub_const [THEN subst]) @@ -2289,7 +2289,7 @@ (*---------------------------------------------------------------------*) theorem Dinf_universal: - "[| commute(DD,ee,G,f); emb_chain(DD,ee); cpo(G) |] ==> + "\commute(DD,ee,G,f); emb_chain(DD,ee); cpo(G)\ \ mediating (Dinf(DD,ee),G,rho_emb(DD,ee),f, lub(cf(Dinf(DD,ee),G), diff -r f2094906e491 -r e44d86131648 src/ZF/ex/NatSum.thy --- a/src/ZF/ex/NatSum.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/NatSum.thy Tue Sep 27 16:51:35 2022 +0100 @@ -12,7 +12,7 @@ thanks to new simprocs. Thanks to Sloane's On-Line Encyclopedia of Integer Sequences, - http://www.research.att.com/~njas/sequences/ + http://www.research.att.com/\njas/sequences/ *) @@ -27,41 +27,41 @@ declare zdiff_zmult_distrib [simp] zdiff_zmult_distrib2 [simp] (*The sum of the first n odd numbers equals n squared.*) -lemma sum_of_odds: "n \ nat ==> sum (%i. i $+ i $+ #1, n) = $#n $* $#n" +lemma sum_of_odds: "n \ nat \ sum (%i. i $+ i $+ #1, n) = $#n $* $#n" by (induct_tac "n", auto) (*The sum of the first n odd squares*) lemma sum_of_odd_squares: - "n \ nat ==> #3 $* sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) = + "n \ nat \ #3 $* sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) = $#n $* (#4 $* $#n $* $#n $- #1)" by (induct_tac "n", auto) (*The sum of the first n odd cubes*) lemma sum_of_odd_cubes: "n \ nat - ==> sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) = + \ sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) = $#n $* $#n $* (#2 $* $#n $* $#n $- #1)" by (induct_tac "n", auto) (*The sum of the first n positive integers equals n(n+1)/2.*) lemma sum_of_naturals: - "n \ nat ==> #2 $* sum(%i. i, succ(n)) = $#n $* $#succ(n)" + "n \ nat \ #2 $* sum(%i. i, succ(n)) = $#n $* $#succ(n)" by (induct_tac "n", auto) lemma sum_of_squares: - "n \ nat ==> #6 $* sum (%i. i$*i, succ(n)) = + "n \ nat \ #6 $* sum (%i. i$*i, succ(n)) = $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1)" by (induct_tac "n", auto) lemma sum_of_cubes: - "n \ nat ==> #4 $* sum (%i. i$*i$*i, succ(n)) = + "n \ nat \ #4 $* sum (%i. i$*i$*i, succ(n)) = $#n $* $#n $* ($#n $+ #1) $* ($#n $+ #1)" by (induct_tac "n", auto) (** Sum of fourth powers **) lemma sum_of_fourth_powers: - "n \ nat ==> #30 $* sum (%i. i$*i$*i$*i, succ(n)) = + "n \ nat \ #30 $* sum (%i. i$*i$*i$*i, succ(n)) = $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1) $* (#3 $* $#n $* $#n $+ #3 $* $#n $- #1)" by (induct_tac "n", auto) diff -r f2094906e491 -r e44d86131648 src/ZF/ex/Primes.thy --- a/src/ZF/ex/Primes.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/Primes.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,67 +9,67 @@ definition divides :: "[i,i]=>o" (infixl \dvd\ 50) where - "m dvd n == m \ nat & n \ nat & (\k \ nat. n = m#*k)" + "m dvd n \ m \ nat & n \ nat & (\k \ nat. n = m#*k)" definition is_gcd :: "[i,i,i]=>o" \ \definition of great common divisor\ where - "is_gcd(p,m,n) == ((p dvd m) & (p dvd n)) & + "is_gcd(p,m,n) \ ((p dvd m) & (p dvd n)) & (\d\nat. (d dvd m) & (d dvd n) \ d dvd p)" definition gcd :: "[i,i]=>i" \ \Euclid's algorithm for the gcd\ where - "gcd(m,n) == transrec(natify(n), + "gcd(m,n) \ transrec(natify(n), %n f. \m \ nat. if n=0 then m else f`(m mod n)`n) ` natify(m)" definition coprime :: "[i,i]=>o" \ \the coprime relation\ where - "coprime(m,n) == gcd(m,n) = 1" + "coprime(m,n) \ gcd(m,n) = 1" definition prime :: i \ \the set of prime numbers\ where - "prime == {p \ nat. 1

m \ nat. m dvd p \ m=1 | m=p)}" + "prime \ {p \ nat. 1

m \ nat. m dvd p \ m=1 | m=p)}" subsection\The Divides Relation\ -lemma dvdD: "m dvd n ==> m \ nat & n \ nat & (\k \ nat. n = m#*k)" +lemma dvdD: "m dvd n \ m \ nat & n \ nat & (\k \ nat. n = m#*k)" by (unfold divides_def, assumption) lemma dvdE: - "[|m dvd n; !!k. [|m \ nat; n \ nat; k \ nat; n = m#*k|] ==> P|] ==> P" + "\m dvd n; \k. \m \ nat; n \ nat; k \ nat; n = m#*k\ \ P\ \ P" by (blast dest!: dvdD) lemmas dvd_imp_nat1 = dvdD [THEN conjunct1] lemmas dvd_imp_nat2 = dvdD [THEN conjunct2, THEN conjunct1] -lemma dvd_0_right [simp]: "m \ nat ==> m dvd 0" +lemma dvd_0_right [simp]: "m \ nat \ m dvd 0" apply (simp add: divides_def) apply (fast intro: nat_0I mult_0_right [symmetric]) done -lemma dvd_0_left: "0 dvd m ==> m = 0" +lemma dvd_0_left: "0 dvd m \ m = 0" by (simp add: divides_def) -lemma dvd_refl [simp]: "m \ nat ==> m dvd m" +lemma dvd_refl [simp]: "m \ nat \ m dvd m" apply (simp add: divides_def) apply (fast intro: nat_1I mult_1_right [symmetric]) done -lemma dvd_trans: "[| m dvd n; n dvd p |] ==> m dvd p" +lemma dvd_trans: "\m dvd n; n dvd p\ \ m dvd p" by (auto simp add: divides_def intro: mult_assoc mult_type) -lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m=n" +lemma dvd_anti_sym: "\m dvd n; n dvd m\ \ m=n" apply (simp add: divides_def) apply (force dest: mult_eq_self_implies_10 simp add: mult_assoc mult_eq_1_iff) done -lemma dvd_mult_left: "[|(i#*j) dvd k; i \ nat|] ==> i dvd k" +lemma dvd_mult_left: "\(i#*j) dvd k; i \ nat\ \ i dvd k" by (auto simp add: divides_def mult_assoc) -lemma dvd_mult_right: "[|(i#*j) dvd k; j \ nat|] ==> j dvd k" +lemma dvd_mult_right: "\(i#*j) dvd k; j \ nat\ \ j dvd k" apply (simp add: divides_def, clarify) apply (rule_tac x = "i#*ka" in bexI) apply (simp add: mult_ac) @@ -91,14 +91,14 @@ by (simp add: gcd_def) lemma gcd_non_0_raw: - "[| 0 nat |] ==> gcd(m,n) = gcd(n, m mod n)" + "\0 nat\ \ gcd(m,n) = gcd(n, m mod n)" apply (simp add: gcd_def) apply (rule_tac P = "%z. left (z) = right" for left right in transrec [THEN ssubst]) apply (simp add: ltD [THEN mem_imp_not_eq, THEN not_sym] mod_less_divisor [THEN ltD]) done -lemma gcd_non_0: "0 < natify(n) ==> gcd(m,n) = gcd(n, m mod n)" +lemma gcd_non_0: "0 < natify(n) \ gcd(m,n) = gcd(n, m mod n)" apply (cut_tac m = m and n = "natify (n) " in gcd_non_0_raw) apply auto done @@ -106,17 +106,17 @@ lemma gcd_1 [simp]: "gcd(m,1) = 1" by (simp (no_asm_simp) add: gcd_non_0) -lemma dvd_add: "[| k dvd a; k dvd b |] ==> k dvd (a #+ b)" +lemma dvd_add: "\k dvd a; k dvd b\ \ k dvd (a #+ b)" apply (simp add: divides_def) apply (fast intro: add_mult_distrib_left [symmetric] add_type) done -lemma dvd_mult: "k dvd n ==> k dvd (m #* n)" +lemma dvd_mult: "k dvd n \ k dvd (m #* n)" apply (simp add: divides_def) apply (fast intro: mult_left_commute mult_type) done -lemma dvd_mult2: "k dvd m ==> k dvd (m #* n)" +lemma dvd_mult2: "k dvd m \ k dvd (m #* n)" apply (subst mult_commute) apply (blast intro: dvd_mult) done @@ -126,7 +126,7 @@ lemmas dvdI2 [simp] = dvd_refl [THEN dvd_mult2] lemma dvd_mod_imp_dvd_raw: - "[| a \ nat; b \ nat; k dvd b; k dvd (a mod b) |] ==> k dvd a" + "\a \ nat; b \ nat; k dvd b; k dvd (a mod b)\ \ k dvd a" apply (case_tac "b=0") apply (simp add: DIVISION_BY_ZERO_MOD) apply (blast intro: mod_div_equality [THEN subst] @@ -134,17 +134,17 @@ intro!: dvd_add dvd_mult mult_type mod_type div_type) done -lemma dvd_mod_imp_dvd: "[| k dvd (a mod b); k dvd b; a \ nat |] ==> k dvd a" +lemma dvd_mod_imp_dvd: "\k dvd (a mod b); k dvd b; a \ nat\ \ k dvd a" apply (cut_tac b = "natify (b)" in dvd_mod_imp_dvd_raw) apply auto apply (simp add: divides_def) done (*Imitating TFL*) -lemma gcd_induct_lemma [rule_format (no_asm)]: "[| n \ nat; +lemma gcd_induct_lemma [rule_format (no_asm)]: "\n \ nat; \m \ nat. P(m,0); - \m \ nat. \n \ nat. 0 P(n, m mod n) \ P(m,n) |] - ==> \m \ nat. P (m,n)" + \m \ nat. \n \ nat. 0 P(n, m mod n) \ P(m,n)\ + \ \m \ nat. P (m,n)" apply (erule_tac i = n in complete_induct) apply (case_tac "x=0") apply (simp (no_asm_simp)) @@ -154,10 +154,10 @@ apply (blast intro: mod_less_divisor [THEN ltD]) done -lemma gcd_induct: "!!P. [| m \ nat; n \ nat; - !!m. m \ nat ==> P(m,0); - !!m n. [|m \ nat; n \ nat; 0 P(m,n) |] - ==> P (m,n)" +lemma gcd_induct: "\P. \m \ nat; n \ nat; + \m. m \ nat \ P(m,0); + \m n. \m \ nat; n \ nat; 0 \ P(m,n)\ + \ P (m,n)" by (blast intro: gcd_induct_lemma) @@ -176,25 +176,25 @@ text\Property 1: gcd(a,b) divides a and b\ lemma gcd_dvd_both: - "[| m \ nat; n \ nat |] ==> gcd (m, n) dvd m & gcd (m, n) dvd n" + "\m \ nat; n \ nat\ \ gcd (m, n) dvd m & gcd (m, n) dvd n" apply (rule_tac m = m and n = n in gcd_induct) apply (simp_all add: gcd_non_0) apply (blast intro: dvd_mod_imp_dvd_raw nat_into_Ord [THEN Ord_0_lt]) done -lemma gcd_dvd1 [simp]: "m \ nat ==> gcd(m,n) dvd m" +lemma gcd_dvd1 [simp]: "m \ nat \ gcd(m,n) dvd m" apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both) apply auto done -lemma gcd_dvd2 [simp]: "n \ nat ==> gcd(m,n) dvd n" +lemma gcd_dvd2 [simp]: "n \ nat \ gcd(m,n) dvd n" apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both) apply auto done text\if f divides a and b then f divides gcd(a,b)\ -lemma dvd_mod: "[| f dvd a; f dvd b |] ==> f dvd (a mod b)" +lemma dvd_mod: "\f dvd a; f dvd b\ \ f dvd (a mod b)" apply (simp add: divides_def) apply (case_tac "b=0") apply (simp add: DIVISION_BY_ZERO_MOD, auto) @@ -205,19 +205,19 @@ if f divides a and f divides b then f divides gcd(a,b)\ lemma gcd_greatest_raw [rule_format]: - "[| m \ nat; n \ nat; f \ nat |] - ==> (f dvd m) \ (f dvd n) \ f dvd gcd(m,n)" + "\m \ nat; n \ nat; f \ nat\ + \ (f dvd m) \ (f dvd n) \ f dvd gcd(m,n)" apply (rule_tac m = m and n = n in gcd_induct) apply (simp_all add: gcd_non_0 dvd_mod) done -lemma gcd_greatest: "[| f dvd m; f dvd n; f \ nat |] ==> f dvd gcd(m,n)" +lemma gcd_greatest: "\f dvd m; f dvd n; f \ nat\ \ f dvd gcd(m,n)" apply (rule gcd_greatest_raw) apply (auto simp add: divides_def) done -lemma gcd_greatest_iff [simp]: "[| k \ nat; m \ nat; n \ nat |] - ==> (k dvd gcd (m, n)) \ (k dvd m & k dvd n)" +lemma gcd_greatest_iff [simp]: "\k \ nat; m \ nat; n \ nat\ + \ (k dvd gcd (m, n)) \ (k dvd m & k dvd n)" by (blast intro!: gcd_greatest gcd_dvd1 gcd_dvd2 intro: dvd_trans) @@ -225,12 +225,12 @@ text\The GCD exists and function gcd computes it.\ -lemma is_gcd: "[| m \ nat; n \ nat |] ==> is_gcd(gcd(m,n), m, n)" +lemma is_gcd: "\m \ nat; n \ nat\ \ is_gcd(gcd(m,n), m, n)" by (simp add: is_gcd_def) text\The GCD is unique\ -lemma is_gcd_unique: "[|is_gcd(m,a,b); is_gcd(n,a,b); m\nat; n\nat|] ==> m=n" +lemma is_gcd_unique: "\is_gcd(m,a,b); is_gcd(n,a,b); m\nat; n\nat\ \ m=n" apply (simp add: is_gcd_def) apply (blast intro: dvd_anti_sym) done @@ -238,7 +238,7 @@ lemma is_gcd_commute: "is_gcd(k,m,n) \ is_gcd(k,n,m)" by (simp add: is_gcd_def, blast) -lemma gcd_commute_raw: "[| m \ nat; n \ nat |] ==> gcd(m,n) = gcd(n,m)" +lemma gcd_commute_raw: "\m \ nat; n \ nat\ \ gcd(m,n) = gcd(n,m)" apply (rule is_gcd_unique) apply (rule is_gcd) apply (rule_tac [3] is_gcd_commute [THEN iffD1]) @@ -250,8 +250,8 @@ apply auto done -lemma gcd_assoc_raw: "[| k \ nat; m \ nat; n \ nat |] - ==> gcd (gcd (k, m), n) = gcd (k, gcd (m, n))" +lemma gcd_assoc_raw: "\k \ nat; m \ nat; n \ nat\ + \ gcd (gcd (k, m), n) = gcd (k, gcd (m, n))" apply (rule is_gcd_unique) apply (rule is_gcd) apply (simp_all add: is_gcd_def) @@ -289,7 +289,7 @@ lemma gcd_add2' [simp]: "gcd (m, n #+ m) = gcd (m, n)" by (subst add_commute, rule gcd_add2) -lemma gcd_add_mult_raw: "k \ nat ==> gcd (m, k #* m #+ n) = gcd (m, n)" +lemma gcd_add_mult_raw: "k \ nat \ gcd (m, k #* m #+ n) = gcd (m, n)" apply (erule nat_induct) apply (auto simp add: gcd_add2 add_assoc) done @@ -303,8 +303,8 @@ subsection\Multiplication Laws\ lemma gcd_mult_distrib2_raw: - "[| k \ nat; m \ nat; n \ nat |] - ==> k #* gcd (m, n) = gcd (k #* m, k #* n)" + "\k \ nat; m \ nat; n \ nat\ + \ k #* gcd (m, n) = gcd (k #* m, k #* n)" apply (erule_tac m = m and n = n in gcd_induct, assumption) apply simp apply (case_tac "k = 0", simp) @@ -324,25 +324,25 @@ by (cut_tac k = k and n = 1 in gcd_mult, auto) lemma relprime_dvd_mult: - "[| gcd (k,n) = 1; k dvd (m #* n); m \ nat |] ==> k dvd m" + "\gcd (k,n) = 1; k dvd (m #* n); m \ nat\ \ k dvd m" apply (cut_tac k = m and m = k and n = n in gcd_mult_distrib2, auto) apply (erule_tac b = m in ssubst) apply (simp add: dvd_imp_nat1) done lemma relprime_dvd_mult_iff: - "[| gcd (k,n) = 1; m \ nat |] ==> k dvd (m #* n) \ k dvd m" + "\gcd (k,n) = 1; m \ nat\ \ k dvd (m #* n) \ k dvd m" by (blast intro: dvdI2 relprime_dvd_mult dvd_trans) lemma prime_imp_relprime: - "[| p \ prime; ~ (p dvd n); n \ nat |] ==> gcd (p, n) = 1" + "\p \ prime; \ (p dvd n); n \ nat\ \ gcd (p, n) = 1" apply (simp add: prime_def, clarify) apply (drule_tac x = "gcd (p,n)" in bspec) apply auto apply (cut_tac m = p and n = n in gcd_dvd2, auto) done -lemma prime_into_nat: "p \ prime ==> p \ nat" +lemma prime_into_nat: "p \ prime \ p \ nat" by (simp add: prime_def) lemma prime_nonzero: "p \ prime \ p\0" @@ -354,12 +354,12 @@ one of those primes.\ lemma prime_dvd_mult: - "[|p dvd m #* n; p \ prime; m \ nat; n \ nat |] ==> p dvd m \ p dvd n" + "\p dvd m #* n; p \ prime; m \ nat; n \ nat\ \ p dvd m \ p dvd n" by (blast intro: relprime_dvd_mult prime_imp_relprime prime_into_nat) lemma gcd_mult_cancel_raw: - "[|gcd (k,n) = 1; m \ nat; n \ nat|] ==> gcd (k #* m, n) = gcd (m, n)" + "\gcd (k,n) = 1; m \ nat; n \ nat\ \ gcd (k #* m, n) = gcd (m, n)" apply (rule dvd_anti_sym) apply (rule gcd_greatest) apply (rule relprime_dvd_mult [of _ k]) @@ -369,7 +369,7 @@ apply (blast intro: dvdI1 gcd_dvd1 dvd_trans) done -lemma gcd_mult_cancel: "gcd (k,n) = 1 ==> gcd (k #* m, n) = gcd (m, n)" +lemma gcd_mult_cancel: "gcd (k,n) = 1 \ gcd (k #* m, n) = gcd (m, n)" apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_mult_cancel_raw) apply auto done diff -r f2094906e491 -r e44d86131648 src/ZF/ex/Ramsey.thy --- a/src/ZF/ex/Ramsey.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/Ramsey.thy Tue Sep 27 16:51:35 2022 +0100 @@ -28,23 +28,23 @@ definition Symmetric :: "i=>o" where - "Symmetric(E) == (\x y. :E \ :E)" + "Symmetric(E) \ (\x y. :E \ :E)" definition Atleast :: "[i,i]=>o" where \ \not really necessary: ZF defines cardinality\ - "Atleast(n,S) == (\f. f \ inj(n,S))" + "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 \ \ E)" + "Clique(C,V,E) \ (C \ V) & (\x \ C. \y \ C. x\y \ \ E)" definition Indept :: "[i,i,i]=>o" where - "Indept(I,V,E) == (I \ V) & (\x \ I. \y \ I. x\y \ \ E)" + "Indept(I,V,E) \ (I \ V) & (\x \ I. \y \ I. x\y \ \ E)" definition Ramsey :: "[i,i,i]=>o" where - "Ramsey(n,i,j) == \V E. Symmetric(E) & Atleast(n,V) \ + "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))" @@ -53,13 +53,13 @@ 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)" +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)" +lemma Indept_superset: "\Indept(I,V',E); V'<=V\ \ Indept(I,V,E)" by (unfold Indept_def, blast) (*** Atleast ***) @@ -68,31 +68,31 @@ by (unfold Atleast_def inj_def Pi_def function_def, blast) lemma Atleast_succD: - "Atleast(succ(m),A) ==> \x \ A. Atleast(m, A-{x})" + "Atleast(succ(m),A) \ \x \ A. Atleast(m, A-{x})" apply (unfold 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)" + "\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))" + "\Atleast(m,B); b\ B\ \ Atleast(succ(m), cons(b,B))" apply (unfold 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)" + "\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!!*) + the same proof could be used\*) lemma pigeon2 [rule_format]: - "m \ nat ==> + "m \ nat \ \n \ nat. \A B. Atleast((m#+n) #- succ(0), A \ B) \ Atleast(m,A) | Atleast(n,B)" apply (induct_tac "m") @@ -137,16 +137,16 @@ (*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)})" +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!*) +(*For the Atleast part, proves \(a \ I) from the second premise!*) lemma Indept_succ: - "[| Indept(I, {z \ V-{a}. \ E}, E); Symmetric(E); a \ V; - Atleast(j,I) |] ==> + "\Indept(I, {z \ V-{a}. \ 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) apply (blast intro!: Atleast_succI) @@ -154,8 +154,8 @@ lemma Clique_succ: - "[| Clique(C, {z \ V-{a}. :E}, E); Symmetric(E); a \ V; - Atleast(j,C) |] ==> + "\Clique(C, {z \ V-{a}. :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) apply (blast intro!: Atleast_succI) @@ -165,8 +165,8 @@ (*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))" + "\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.:E" in Atleast_partition [THEN disjE], @@ -181,7 +181,7 @@ (** 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)" +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) @@ -191,7 +191,7 @@ done (*Final statement in a tidy form, without succ(...) *) -lemma ramsey: "[| i \ nat; j \ nat |] ==> \n \ nat. Ramsey(n,i,j)" +lemma ramsey: "\i \ nat; j \ nat\ \ \n \ nat. Ramsey(n,i,j)" by (blast dest: ramsey_lemma) end diff -r f2094906e491 -r e44d86131648 src/ZF/ex/Ring.thy --- a/src/ZF/ex/Ring.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/Ring.thy Tue Sep 27 16:51:35 2022 +0100 @@ -43,7 +43,7 @@ definition a_inv :: "[i,i] => i" (\\\ _\ [81] 80) where - "a_inv(R) == m_inv ()" + "a_inv(R) \ m_inv ()" definition minus :: "[i,i,i] => i" (\(_ \\ _)\ [65,66] 65) where @@ -105,7 +105,7 @@ by (simp add: a_inv_def group.inv_closed [OF a_group, simplified]) lemma (in abelian_monoid) a_closed [intro, simp]: - "[| x \ carrier(G); y \ carrier(G) |] ==> x \ y \ carrier(G)" + "\x \ carrier(G); y \ carrier(G)\ \ x \ y \ carrier(G)" by (rule monoid.m_closed [OF a_monoid, simplified, simplified ring_add_def [symmetric]]) @@ -173,7 +173,7 @@ lemmas (in abelian_monoid) a_ac = a_assoc a_comm a_lcomm text \ - The following proofs are from Jacobson, Basic Algebra I, pp.~88--89 + The following proofs are from Jacobson, Basic Algebra I, pp.\88--89 \ context ring @@ -232,7 +232,7 @@ definition ring_hom :: "[i,i] => i" where - "ring_hom(R,S) == + "ring_hom(R,S) \ {h \ carrier(R) -> carrier(S). (\x y. x \ carrier(R) & y \ carrier(R) \ h ` (x \\<^bsub>R\<^esub> y) = (h ` x) \\<^bsub>S\<^esub> (h ` y) & diff -r f2094906e491 -r e44d86131648 src/ZF/ex/misc.thy --- a/src/ZF/ex/misc.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/ex/misc.thy Tue Sep 27 16:51:35 2022 +0100 @@ -31,7 +31,7 @@ text\A weird property of ordered pairs.\ -lemma "b\c ==> \ = " +lemma "b\c \ \ = " by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast) text\These two are cited in Benzmueller and Kohlhase's system description of @@ -44,15 +44,15 @@ by (blast intro!: equalityI) text\trivial example of term synthesis: apparently hard for some provers!\ -schematic_goal "a \ b ==> a:?X & b \ ?X" +schematic_goal "a \ b \ a:?X & b \ ?X" by blast text\Nice blast benchmark. Proved in 0.3s; old tactics can't manage it!\ -lemma "\x \ S. \y \ S. x \ y ==> \z. S \ {z}" +lemma "\x \ S. \y \ S. x \ y \ \z. S \ {z}" by blast text\variant of the benchmark above\ -lemma "\x \ S. \(S) \ x ==> \z. S \ {z}" +lemma "\x \ S. \(S) \ x \ \z. S \ {z}" by blast (*Example 12 (credited to Peter Andrews) from @@ -87,10 +87,10 @@ by force text\Another version, with meta-level rewriting\ -lemma "(!! A f B g. hom(A,f,B,g) == +lemma "(\A f B g. hom(A,f,B,g) \ {H \ A->B. f \ A*A->A & g \ B*B->B & (\x \ A. \y \ A. H`(f`) = g`)}) - ==> J \ hom(A,f,B,g) & K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" + \ J \ hom(A,f,B,g) & K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" by force @@ -105,38 +105,38 @@ comp_mem_injD2 comp_mem_surjD2 lemma pastre1: - "[| (h O g O f) \ inj(A,A); + "\(h O g O f) \ inj(A,A); (f O h O g) \ surj(B,B); (g O f O h) \ surj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre3: - "[| (h O g O f) \ surj(A,A); + "\(h O g O f) \ surj(A,A); (f O h O g) \ surj(B,B); (g O f O h) \ inj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre4: - "[| (h O g O f) \ surj(A,A); + "\(h O g O f) \ surj(A,A); (f O h O g) \ inj(B,B); (g O f O h) \ inj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre5: - "[| (h O g O f) \ inj(A,A); + "\(h O g O f) \ inj(A,A); (f O h O g) \ surj(B,B); (g O f O h) \ inj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) lemma pastre6: - "[| (h O g O f) \ inj(A,A); + "\(h O g O f) \ inj(A,A); (f O h O g) \ inj(B,B); (g O f O h) \ surj(C,C); - f \ A->B; g \ B->C; h \ C->A |] ==> h \ bij(C,A)" + f \ A->B; g \ B->C; h \ C->A\ \ h \ bij(C,A)" by (unfold bij_def, blast) diff -r f2094906e491 -r e44d86131648 src/ZF/func.thy --- a/src/ZF/func.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/func.thy Tue Sep 27 16:51:35 2022 +0100 @@ -9,11 +9,11 @@ subsection\The Pi Operator: Dependent Function Space\ -lemma subset_Sigma_imp_relation: "r \ Sigma(A,B) ==> relation(r)" +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" + "relation(r) \ converse(converse(r)) = r" by (simp add: relation_def, blast) lemma relation_restrict [simp]: "relation(restrict(r,A))" @@ -28,23 +28,23 @@ "f \ Pi(A,B) \ f<=Sigma(A,B) & (\x\A. \!y. : f)" by (unfold Pi_def function_def, blast) -lemma fun_is_function: "f \ Pi(A,B) ==> function(f)" +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)" + "\function(f); relation(f)\ \ f \ domain(f) -> range(f)" by (simp add: Pi_iff relation_def, blast) lemma functionI: - "[| !!x y y'. [| :r; :r |] ==> y=y' |] ==> function(r)" + "\\x y 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)" +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')" + "\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 @@ -52,57 +52,57 @@ 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" +lemma fun_weaken_type: "\f \ A->B; B<=D\ \ f \ A->D" by (unfold Pi_def, best) subsection\Function Application\ -lemma apply_equality2: "[| : f; : f; f \ Pi(A,B) |] ==> b=c" +lemma apply_equality2: "\: f; : f; f \ Pi(A,B)\ \ b=c" by (unfold Pi_def function_def, blast) -lemma function_apply_equality: "[| : f; function(f) |] ==> f`a = b" +lemma function_apply_equality: "\: f; function(f)\ \ f`a = b" by (unfold apply_def function_def, blast) -lemma apply_equality: "[| : f; f \ Pi(A,B) |] ==> f`a = b" +lemma apply_equality: "\: f; f \ Pi(A,B)\ \ f`a = b" apply (unfold 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" +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 = " +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" +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" +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)" +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" +lemma apply_funtype: "\f \ A->B; a \ A\ \ f`a \ B" by (blast dest: apply_type) -lemma apply_iff: "f \ Pi(A,B) ==> : f \ a \ A & f`a = b" +lemma apply_iff: "f \ Pi(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)" +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 @@ -114,38 +114,38 @@ 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)" + "\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: "[| \ f; f \ Pi(A,B) |] ==> a \ A" +lemma domain_type: "\ \ f; f \ Pi(A,B)\ \ a \ A" by (blast dest: fun_is_rel) -lemma range_type: "[| \ f; f \ Pi(A,B) |] ==> b \ B(a)" +lemma range_type: "\ \ f; f \ Pi(A,B)\ \ b \ B(a)" by (blast dest: fun_is_rel) -lemma Pair_mem_PiD: "[| : f; f \ Pi(A,B) |] ==> a \ A & b \ B(a) & f`a = b" +lemma Pair_mem_PiD: "\: 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))" +lemma lamI: "a \ A \ \ (\x\A. b(x))" apply (unfold lam_def) apply (erule RepFunI) done lemma lamE: - "[| p: (\x\A. b(x)); !!x.[| x \ A; p= |] ==> P - |] ==> P" + "\p: (\x\A. b(x)); \x.\x \ A; p=\ \ P +\ \ P" by (simp add: lam_def, blast) -lemma lamD: "[| : (\x\A. b(x)) |] ==> c = b(a)" +lemma lamD: "\: (\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)" + "\\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}" @@ -160,7 +160,7 @@ 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)" +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" @@ -171,17 +171,17 @@ (*congruence rule for lambda abstraction*) lemma lam_cong [cong]: - "[| A=A'; !!x. x \ A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')" + "\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)" + "(\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)" +lemma lam_eqE: "\(\x\A. f(x)) = (\x\A. g(x)); a \ A\ \ f(a)=g(a)" by (fast intro!: lamI elim: equalityE lamE) @@ -207,26 +207,26 @@ (*Semi-extensionality!*) lemma fun_subset: - "[| f \ Pi(A,B); g \ Pi(C,D); A<=C; - !!x. x \ A ==> f`x = g`x |] ==> f<=g" + "\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" + "\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" +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" + "\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)" +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]) @@ -234,7 +234,7 @@ (*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" + 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]) @@ -244,12 +244,12 @@ subsection\Images of Functions\ -lemma image_lam: "C \ A ==> (\x\A. b(x)) `` C = {b(x). x \ C}" +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))" + \ {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) @@ -261,10 +261,10 @@ (*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}" + "\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}" +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 @@ -274,7 +274,7 @@ 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)" + "\f \ Pi(A,B); x \ A\ \ f `` cons(x,y) = cons(f`x, f``y)" by (blast dest: apply_equality apply_Pair) @@ -284,10 +284,10 @@ by (unfold restrict_def, blast) lemma function_restrictI: - "function(f) ==> function(restrict(f,A))" + "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)" +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)" @@ -308,13 +308,13 @@ apply (auto simp add: domain_def) done -lemma restrict_idem: "f \ Sigma(A,B) ==> restrict(f,A) = f" +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" + "\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" @@ -327,11 +327,11 @@ by (simp add: restrict apply_0) lemma restrict_lam_eq: - "A<=C ==> restrict(\x\C. b(x), A) = (\x\A. b(x))" + "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))" + "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) @@ -343,14 +343,14 @@ (** 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))" + "\\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 |] ==> + "\\f\S. \C D. f \ C->D; + \f\S. \y\S. f<=y | y<=f\ \ \(S) \ domain(\(S)) -> range(\(S))" apply (unfold Pi_def) apply (blast intro!: rel_Union function_Union) @@ -368,48 +368,48 @@ subset_trans [OF _ Un_upper2] lemma fun_disjoint_Un: - "[| f \ A->B; g \ C->D; A \ C = 0 |] - ==> (f \ g) \ (A \ C) -> (B \ D)" + "\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" +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" +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" +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)" +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)" +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(,f) \ cons(c,A) -> cons(b,B)" + "\f \ A->B; c\A\ \ cons(,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(,f) \ cons(c,A) -> B" + "\f \ A->B; c\A; b \ B\ \ cons(,f) \ cons(c,A) -> B" by (blast intro: fun_extend [THEN fun_weaken_type]) lemma extend_apply: - "c \ domain(f) ==> cons(,f)`a = (if a=c then b else f`a)" + "c \ domain(f) \ cons(,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(,f)`a = (if a=c then b else f`a)" + "\f \ A->B; c\A\ \ cons(,f)`a = (if a=c then b else f`a)" apply (rule extend_apply) apply (simp add: Pi_def, blast) done @@ -418,7 +418,7 @@ (*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(, f)})" + "c \ A \ cons(c,A) -> B = (\f \ A->B. \b\B. {cons(, f)})" apply (rule equalityI) apply (safe elim!: fun_extend3) (*Inclusion of left into right*) @@ -443,7 +443,7 @@ definition update :: "[i,i,i] => i" where - "update(f,a,b) == \x\cons(a, domain(f)). if(x=a, b, f`x)" + "update(f,a,b) \ \x\cons(a, domain(f)). if(x=a, b, f`x)" nonterminal updbinds and updbind @@ -467,7 +467,7 @@ apply (simp_all add: apply_0) done -lemma update_idem: "[| f`x = y; f \ Pi(A,B); x \ A |] ==> f(x:=y) = f" +lemma update_idem: "\f`x = y; f \ Pi(A,B); x \ A\ \ f(x:=y) = f" apply (unfold update_def) apply (simp add: domain_of_fun cons_absorb) apply (rule fun_extension) @@ -475,13 +475,13 @@ done -(* [| f \ Pi(A, B); x \ A |] ==> f(x := f`x) = f *) +(* \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)" +lemma update_type: "\f \ Pi(A,B); x \ A; y \ B(x)\ \ f(x:=y) \ Pi(A, B)" apply (unfold update_def) apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type) done @@ -493,96 +493,96 @@ (*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)" +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}" +lemma RepFun_mono: "A<=B \ {f(x). x \ A} \ {f(x). x \ B}" by blast -lemma Pow_mono: "A<=B ==> Pow(A) \ Pow(B)" +lemma Pow_mono: "A<=B \ Pow(A) \ Pow(B)" by blast -lemma Union_mono: "A<=B ==> \(A) \ \(B)" +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))" + "\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)" +lemma Inter_anti_mono: "\A<=B; A\0\ \ \(B) \ \(A)" by blast -lemma cons_mono: "C<=D ==> cons(a,C) \ cons(a,D)" +lemma cons_mono: "C<=D \ cons(a,C) \ cons(a,D)" by blast -lemma Un_mono: "[| A<=C; B<=D |] ==> A \ B \ C \ D" +lemma Un_mono: "\A<=C; B<=D\ \ A \ B \ C \ D" by blast -lemma Int_mono: "[| A<=C; B<=D |] ==> A \ B \ C \ D" +lemma Int_mono: "\A<=C; B<=D\ \ A \ B \ C \ D" by blast -lemma Diff_mono: "[| A<=C; D<=B |] ==> A-B \ C-D" +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)" + "\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" +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" +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)" +lemma lam_mono: "A<=B \ Lambda(A,c) \ Lambda(B,c)" apply (unfold lam_def) apply (erule RepFun_mono) done subsubsection\Converse, Domain, Range, Field\ -lemma converse_mono: "r<=s ==> converse(r) \ converse(s)" +lemma converse_mono: "r<=s \ converse(r) \ converse(s)" by blast -lemma domain_mono: "r<=s ==> domain(r)<=domain(s)" +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)" +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)" +lemma field_mono: "r<=s \ field(r)<=field(s)" by blast -lemma field_rel_subset: "r \ A*A ==> field(r) \ A" +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. :r ==> :s; A<=B |] ==> r``A \ s``B" + "\\x y. :r \ :s; A<=B\ \ r``A \ s``B" by blast lemma vimage_pair_mono: - "[| !! x y. :r ==> :s; A<=B |] ==> r-``A \ s-``B" + "\\x y. :r \ :s; A<=B\ \ r-``A \ s-``B" by blast -lemma image_mono: "[| r<=s; A<=B |] ==> r``A \ s``B" +lemma image_mono: "\r<=s; A<=B\ \ r``A \ s``B" by blast -lemma vimage_mono: "[| r<=s; A<=B |] ==> r-``A \ s-``B" +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)" + "\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*) @@ -592,7 +592,7 @@ (* 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))" + "\f \ Pi(X, Y); A \ X\ \ (\x\f``A. P(x)) \ (\x\A. P(f`x))" apply safe apply rule prefer 2 apply assumption @@ -601,7 +601,7 @@ done lemma ball_image_simp: - "[| f \ Pi(X, Y); A \ X |] ==> (\x\f``A. P(x)) \ (\x\A. P(f`x))" + "\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 diff -r f2094906e491 -r e44d86131648 src/ZF/pair.thy --- a/src/ZF/pair.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/pair.thy Tue Sep 27 16:51:35 2022 +0100 @@ -54,7 +54,7 @@ declare sym [THEN Pair_neq_0, elim!] -lemma Pair_neq_fst: "=a ==> P" +lemma Pair_neq_fst: "=a \ P" proof (unfold Pair_def) assume eq: "{{a, a}, {a, b}} = a" have "{a, a} \ {{a, a}, {a, b}}" by (rule consI1) @@ -63,7 +63,7 @@ ultimately show "P" by (rule mem_asym) qed -lemma Pair_neq_snd: "=b ==> P" +lemma Pair_neq_snd: "=b \ P" proof (unfold Pair_def) assume eq: "{{a, a}, {a, b}} = b" have "{a, b} \ {{a, a}, {a, b}}" by blast @@ -80,7 +80,7 @@ lemma Sigma_iff [simp]: ": Sigma(A,B) \ a \ A & b \ B(a)" by (simp add: Sigma_def) -lemma SigmaI [TC,intro!]: "[| a \ A; b \ B(a) |] ==> \ Sigma(A,B)" +lemma SigmaI [TC,intro!]: "\a \ A; b \ B(a)\ \ \ Sigma(A,B)" by simp lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1] @@ -88,19 +88,19 @@ (*The general elimination rule*) lemma SigmaE [elim!]: - "[| c \ Sigma(A,B); - !!x y.[| x \ A; y \ B(x); c= |] ==> P - |] ==> P" + "\c \ Sigma(A,B); + \x y.\x \ A; y \ B(x); c=\ \ P +\ \ P" by (unfold Sigma_def, blast) lemma SigmaE2 [elim!]: - "[| \ Sigma(A,B); - [| a \ A; b \ B(a) |] ==> P - |] ==> P" + "\ \ Sigma(A,B); + \a \ A; b \ B(a)\ \ P +\ \ P" by (unfold Sigma_def, blast) lemma Sigma_cong: - "[| A=A'; !!x. x \ A' ==> B(x)=B'(x) |] ==> + "\A=A'; \x. x \ A' \ B(x)=B'(x)\ \ Sigma(A,B) = Sigma(A',B')" by (simp add: Sigma_def) @@ -126,46 +126,46 @@ lemma snd_conv [simp]: "snd() = b" by (simp add: snd_def) -lemma fst_type [TC]: "p \ Sigma(A,B) ==> fst(p) \ A" +lemma fst_type [TC]: "p \ Sigma(A,B) \ fst(p) \ A" by auto -lemma snd_type [TC]: "p \ Sigma(A,B) ==> snd(p) \ B(fst(p))" +lemma snd_type [TC]: "p \ Sigma(A,B) \ snd(p) \ B(fst(p))" by auto -lemma Pair_fst_snd_eq: "a \ Sigma(A,B) ==> = a" +lemma Pair_fst_snd_eq: "a \ Sigma(A,B) \ = a" by auto subsection\The Eliminator, \<^term>\split\\ (*A META-equality, so that it applies to higher types as well...*) -lemma split [simp]: "split(%x y. c(x,y), ) == c(a,b)" +lemma split [simp]: "split(%x y. c(x,y), ) \ c(a,b)" by (simp add: split_def) lemma split_type [TC]: - "[| p \ Sigma(A,B); - !!x y.[| x \ A; y \ B(x) |] ==> c(x,y):C() - |] ==> split(%x y. c(x,y), p) \ C(p)" + "\p \ Sigma(A,B); + \x y.\x \ A; y \ B(x)\ \ c(x,y):C() +\ \ split(%x y. c(x,y), p) \ C(p)" by (erule SigmaE, auto) lemma expand_split: - "u \ A*B ==> + "u \ A*B \ R(split(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" by (auto simp add: split_def) subsection\A version of \<^term>\split\ for Formulae: Result Type \<^typ>\o\\ -lemma splitI: "R(a,b) ==> split(R, )" +lemma splitI: "R(a,b) \ split(R, )" by (simp add: split_def) lemma splitE: - "[| split(R,z); z \ Sigma(A,B); - !!x y. [| z = ; R(x,y) |] ==> P - |] ==> P" + "\split(R,z); z \ Sigma(A,B); + \x y. \z = ; R(x,y)\ \ P +\ \ P" by (auto simp add: split_def) -lemma splitD: "split(R,) ==> R(a,b)" +lemma splitD: "split(R,) \ R(a,b)" by (simp add: split_def) text \ diff -r f2094906e491 -r e44d86131648 src/ZF/upair.thy --- a/src/ZF/upair.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/upair.thy Tue Sep 27 16:51:35 2022 +0100 @@ -19,7 +19,7 @@ ML_file \Tools/typechk.ML\ lemma atomize_ball [symmetric, rulify]: - "(!!x. x \ A ==> P(x)) == Trueprop (\x\A. P(x))" + "(\x. x \ A \ P(x)) \ Trueprop (\x\A. P(x))" by (simp add: Ball_def atomize_all atomize_imp) @@ -34,7 +34,7 @@ lemma UpairI2: "b \ Upair(a,b)" by simp -lemma UpairE: "[| a \ Upair(b,c); a=b ==> P; a=c ==> P |] ==> P" +lemma UpairE: "\a \ Upair(b,c); a=b \ P; a=c \ P\ \ P" by (simp, blast) subsection\Rules for Binary Union, Defined via \<^term>\Upair\\ @@ -44,23 +44,23 @@ apply (blast intro: UpairI1 UpairI2 elim: UpairE) done -lemma UnI1: "c \ A ==> c \ A \ B" +lemma UnI1: "c \ A \ c \ A \ B" by simp -lemma UnI2: "c \ B ==> c \ A \ B" +lemma UnI2: "c \ B \ c \ A \ B" by simp declare UnI1 [elim?] UnI2 [elim?] -lemma UnE [elim!]: "[| c \ A \ B; c \ A ==> P; c \ B ==> P |] ==> P" +lemma UnE [elim!]: "\c \ A \ B; c \ A \ P; c \ B \ P\ \ P" by (simp, blast) (*Stronger version of the rule above*) -lemma UnE': "[| c \ A \ B; c \ A ==> P; [| c \ B; c\A |] ==> P |] ==> P" +lemma UnE': "\c \ A \ B; c \ A \ P; \c \ B; c\A\ \ P\ \ P" by (simp, blast) (*Classical introduction rule: no commitment to A vs B*) -lemma UnCI [intro!]: "(c \ B ==> c \ A) ==> c \ A \ B" +lemma UnCI [intro!]: "(c \ B \ c \ A) \ c \ A \ B" by (simp, blast) subsection\Rules for Binary Intersection, Defined via \<^term>\Upair\\ @@ -70,16 +70,16 @@ apply (blast intro: UpairI1 UpairI2 elim: UpairE) done -lemma IntI [intro!]: "[| c \ A; c \ B |] ==> c \ A \ B" +lemma IntI [intro!]: "\c \ A; c \ B\ \ c \ A \ B" by simp -lemma IntD1: "c \ A \ B ==> c \ A" +lemma IntD1: "c \ A \ B \ c \ A" by simp -lemma IntD2: "c \ A \ B ==> c \ B" +lemma IntD2: "c \ A \ B \ c \ B" by simp -lemma IntE [elim!]: "[| c \ A \ B; [| c \ A; c \ B |] ==> P |] ==> P" +lemma IntE [elim!]: "\c \ A \ B; \c \ A; c \ B\ \ P\ \ P" by simp @@ -88,16 +88,16 @@ lemma Diff_iff [simp]: "c \ A-B \ (c \ A & c\B)" by (unfold Diff_def, blast) -lemma DiffI [intro!]: "[| c \ A; c \ B |] ==> c \ A - B" +lemma DiffI [intro!]: "\c \ A; c \ B\ \ c \ A - B" by simp -lemma DiffD1: "c \ A - B ==> c \ A" +lemma DiffD1: "c \ A - B \ c \ A" by simp -lemma DiffD2: "c \ A - B ==> c \ B" +lemma DiffD2: "c \ A - B \ c \ B" by simp -lemma DiffE [elim!]: "[| c \ A - B; [| c \ A; c\B |] ==> P |] ==> P" +lemma DiffE [elim!]: "\c \ A - B; \c \ A; c\B\ \ P\ \ P" by simp @@ -114,19 +114,19 @@ by simp -lemma consI2: "a \ B ==> a \ cons(b,B)" +lemma consI2: "a \ B \ a \ cons(b,B)" by simp -lemma consE [elim!]: "[| a \ cons(b,A); a=b ==> P; a \ A ==> P |] ==> P" +lemma consE [elim!]: "\a \ cons(b,A); a=b \ P; a \ A \ P\ \ P" by (simp, blast) (*Stronger version of the rule above*) lemma consE': - "[| a \ cons(b,A); a=b ==> P; [| a \ A; a\b |] ==> P |] ==> P" + "\a \ cons(b,A); a=b \ P; \a \ A; a\b\ \ P\ \ P" by (simp, blast) (*Classical introduction rule*) -lemma consCI [intro!]: "(a\B ==> a=b) ==> a \ cons(b,B)" +lemma consCI [intro!]: "(a\B \ a=b) \ a \ cons(b,B)" by (simp, blast) lemma cons_not_0 [simp]: "cons(a,B) \ 0" @@ -151,16 +151,16 @@ subsection\Descriptions\ lemma the_equality [intro]: - "[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a" + "\P(a); \x. P(x) \ x=a\ \ (THE x. P(x)) = a" apply (unfold the_def) apply (fast dest: subst) done (* Only use this if you already know \!x. P(x) *) -lemma the_equality2: "[| \!x. P(x); P(a) |] ==> (THE x. P(x)) = a" +lemma the_equality2: "\\!x. P(x); P(a)\ \ (THE x. P(x)) = a" by blast -lemma theI: "\!x. P(x) ==> P(THE x. P(x))" +lemma theI: "\!x. P(x) \ P(THE x. P(x))" apply (erule ex1E) apply (subst the_equality) apply (blast+) @@ -170,15 +170,15 @@ @{term "THE x.P(x)"} rewrites to @{term "THE x.Q(x)"} *) (*If it's "undefined", it's zero!*) -lemma the_0: "~ (\!x. P(x)) ==> (THE x. P(x))=0" +lemma the_0: "\ (\!x. P(x)) \ (THE x. P(x))=0" apply (unfold the_def) apply (blast elim!: ReplaceE) done (*Easier to apply than theI: conclusion has only one occurrence of P*) lemma theI2: - assumes p1: "~ Q(0) ==> \!x. P(x)" - and p2: "!!x. P(x) ==> Q(x)" + assumes p1: "\ Q(0) \ \!x. P(x)" + and p2: "\x. P(x) \ Q(x)" shows "Q(THE x. P(x))" apply (rule classical) apply (rule p2) @@ -205,25 +205,25 @@ (*Never use with case splitting, or if P is known to be true or false*) lemma if_cong: - "[| P\Q; Q ==> a=c; ~Q ==> b=d |] - ==> (if P then a else b) = (if Q then c else d)" + "\P\Q; Q \ a=c; \Q \ b=d\ + \ (if P then a else b) = (if Q then c else d)" by (simp add: if_def cong add: conj_cong) (*Prevents simplification of x and y \ faster and allows the execution of functional programs. NOW THE DEFAULT.*) -lemma if_weak_cong: "P\Q ==> (if P then x else y) = (if Q then x else y)" +lemma if_weak_cong: "P\Q \ (if P then x else y) = (if Q then x else y)" by simp (*Not needed for rewriting, since P would rewrite to True anyway*) -lemma if_P: "P ==> (if P then a else b) = a" +lemma if_P: "P \ (if P then a else b) = a" by (unfold if_def, blast) (*Not needed for rewriting, since P would rewrite to False anyway*) -lemma if_not_P: "~P ==> (if P then a else b) = b" +lemma if_not_P: "\P \ (if P then a else b) = b" by (unfold if_def, blast) lemma split_if [split]: - "P(if Q then x else y) \ ((Q \ P(x)) & (~Q \ P(y)))" + "P(if Q then x else y) \ ((Q \ P(x)) & (\Q \ P(y)))" by (case_tac Q, simp_all) (** Rewrite rules for boolean case-splitting: faster than split_if [split] @@ -238,16 +238,16 @@ lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 (*Logically equivalent to split_if_mem2*) -lemma if_iff: "a: (if P then x else y) \ P & a \ x | ~P & a \ y" +lemma if_iff: "a: (if P then x else y) \ P & a \ x | \P & a \ y" by simp lemma if_type [TC]: - "[| P ==> a \ A; ~P ==> b \ A |] ==> (if P then a else b): A" + "\P \ a \ A; \P \ b \ A\ \ (if P then a else b): A" by simp (** Splitting IFs in the assumptions **) -lemma split_if_asm: "P(if Q then x else y) \ (~((Q & ~P(x)) | (~Q & ~P(y))))" +lemma split_if_asm: "P(if Q then x else y) \ (\((Q & \P(x)) | (\Q & \P(y))))" by simp lemmas if_splits = split_if split_if_asm @@ -256,14 +256,14 @@ subsection\Consequences of Foundation\ (*was called mem_anti_sym*) -lemma mem_asym: "[| a \ b; ~P ==> b \ a |] ==> P" +lemma mem_asym: "\a \ b; \P \ b \ a\ \ P" apply (rule classical) apply (rule_tac A1 = "{a,b}" in foundation [THEN disjE]) apply (blast elim!: equalityE)+ done (*was called mem_anti_refl*) -lemma mem_irrefl: "a \ a ==> P" +lemma mem_irrefl: "a \ a \ P" by (blast intro: mem_asym) (*mem_irrefl should NOT be added to default databases: @@ -275,10 +275,10 @@ done (*Good for proving inequalities by rewriting*) -lemma mem_imp_not_eq: "a \ A ==> a \ A" +lemma mem_imp_not_eq: "a \ A \ a \ A" by (blast elim!: mem_irrefl) -lemma eq_imp_not_mem: "a=A ==> a \ A" +lemma eq_imp_not_mem: "a=A \ a \ A" by (blast intro: elim: mem_irrefl) subsection\Rules for Successor\ @@ -289,16 +289,16 @@ lemma succI1 [simp]: "i \ succ(i)" by (simp add: succ_iff) -lemma succI2: "i \ j ==> i \ succ(j)" +lemma succI2: "i \ j \ i \ succ(j)" by (simp add: succ_iff) lemma succE [elim!]: - "[| i \ succ(j); i=j ==> P; i \ j ==> P |] ==> P" + "\i \ succ(j); i=j \ P; i \ j \ P\ \ P" apply (simp add: succ_iff, blast) done (*Classical introduction rule*) -lemma succCI [intro!]: "(i\j ==> i=j) ==> i \ succ(j)" +lemma succCI [intro!]: "(i\j \ i=j) \ i \ succ(j)" by (simp add: succ_iff, blast) lemma succ_not_0 [simp]: "succ(n) \ 0" @@ -309,7 +309,7 @@ declare succ_not_0 [THEN not_sym, simp] declare sym [THEN succ_neq_0, elim!] -(* @{term"succ(c) \ B ==> c \ B"} *) +(* @{term"succ(c) \ B \ c \ B"} *) lemmas succ_subsetD = succI1 [THEN [2] subsetD] (* @{term"succ(b) \ b"} *) @@ -327,7 +327,7 @@ "(\x\A. P(x) & Q) \ (\x\A. P(x)) & (A=0 | Q)" "(\x\A. P(x) | Q) \ ((\x\A. P(x)) | Q)" "(\x\A. P(x) \ Q) \ ((\x\A. P(x)) \ Q)" - "(~(\x\A. P(x))) \ (\x\A. ~P(x))" + "(\(\x\A. P(x))) \ (\x\A. \P(x))" "(\x\0.P(x)) \ True" "(\x\succ(i).P(x)) \ P(i) & (\x\i. P(x))" "(\x\cons(a,B).P(x)) \ P(a) & (\x\B. P(x))" @@ -363,7 +363,7 @@ "(\x\cons(a,B).P(x)) \ P(a) | (\x\B. P(x))" "(\x\RepFun(A,f). P(x)) \ (\y\A. P(f(y)))" "(\x\\(A).P(x)) \ (\y\A. \x\y. P(x))" - "(~(\x\A. P(x))) \ (\x\A. ~P(x))" + "(\(\x\A. P(x))) \ (\x\A. \P(x))" by blast+ lemma bex_simps2: