# HG changeset patch # User paulson # Date 1664294603 -3600 # Node ID 0c18df79b1c86dc0b85150ef68421b0b4b6f76c0 # Parent e44d861316485dd381d9b2afe07a9000b2b38781 more modernisation of syntax diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC15_WO6.thy --- a/src/ZF/AC/AC15_WO6.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC15_WO6.thy Tue Sep 27 17:03:23 2022 +0100 @@ -16,7 +16,7 @@ So we don't have to prove all implications of both cases. Moreover we don't need to prove AC13(1) \ AC1 and AC11 \ AC14 as -Rubin & Rubin do. +Rubin \ Rubin do. *) theory AC15_WO6 @@ -46,7 +46,7 @@ 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: @@ -54,10 +54,10 @@ by (unfold pairwise_disjoint_def, blast) lemma lemma3: - "\B \ {cons(0, x*nat). x \ A}. pairwise_disjoint(f`B) & - sets_of_size_between(f`B, 2, n) & \(f`B)=B - \ \B \ A. \! u. u \ f`cons(0, B*nat) & u \ cons(0, B*nat) & - 0 \ u & 2 \ u & u \ n" + "\B \ {cons(0, x*nat). x \ A}. pairwise_disjoint(f`B) \ + sets_of_size_between(f`B, 2, n) \ \(f`B)=B + \ \B \ A. \! u. u \ f`cons(0, B*nat) \ u \ cons(0, B*nat) \ + 0 \ u \ 2 \ u \ u \ n" apply (unfold sets_of_size_between_def) apply (rule ballI) apply (erule_tac x="cons(0, B*nat)" in ballE) @@ -67,7 +67,7 @@ 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" +apply (rule_tac x = "\x \ RepFun(A,P). \ j. \a\A. x=P(a) \ f`a=j" in exI) apply (rule_tac d = "%y. P (converse (f) `y) " in lam_injective) apply (erule RepFunE) @@ -103,10 +103,10 @@ lemma ex_fun_AC13_AC15: "\\B \ {cons(0, x*nat). x \ A}. - pairwise_disjoint(f`B) & - sets_of_size_between(f`B, 2, succ(n)) & \(f`B)=B; + 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" + \ \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) @@ -147,7 +147,7 @@ by (fast intro!: ltI dest!: ltD) lemma AC15_WO6_aux1: - "\x \ Pow(A)-{0}. f`x\0 & f`x \ x & f`x \ m + "\x \ Pow(A)-{0}. f`x\0 \ f`x \ x \ f`x \ m \ (\i<\ x. HH(f,A,x)={A}. HH(f,A,i)) = A" apply (simp add: Ord_Least [THEN OUN_eq_UN]) apply (rule equalityI) @@ -157,7 +157,7 @@ done lemma AC15_WO6_aux2: - "\x \ Pow(A)-{0}. f`x\0 & f`x \ x & f`x \ m + "\x \ Pow(A)-{0}. f`x\0 \ f`x \ x \ f`x \ m \ \x < (\ x. HH(f,A,x)={A}). HH(f,A,x) \ m" apply (rule oallI) apply (drule ltD [THEN less_Least_subset_x]) @@ -251,7 +251,7 @@ by (unfold AC13_def AC14_def AC15_def, fast) (* ********************************************************************** *) -(* The redundant proofs; however cited by Rubin & Rubin *) +(* The redundant proofs; however cited by Rubin \ Rubin *) (* ********************************************************************** *) (* ********************************************************************** *) @@ -262,7 +262,7 @@ by (fast elim!: not_emptyE lepoll_1_is_sing) lemma AC13_AC1_lemma: - "\B \ A. f(B)\0 & f(B)<=B & f(B) \ 1 + "\B \ A. f(B)\0 \ f(B)<=B \ f(B) \ 1 \ (\x \ A. THE y. f(x)={y}) \ (\X \ A. X)" apply (rule lam_type) apply (drule bspec, assumption) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC16_WO4.thy --- a/src/ZF/AC/AC16_WO4.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC16_WO4.thy Tue Sep 27 17:03:23 2022 +0100 @@ -16,8 +16,8 @@ lemma lemma1: "\Finite(A); 0 nat\ - \ \a f. Ord(a) & domain(f) = a & - (\bb m)" + \ \a f. Ord(a) \ domain(f) = a \ + (\b (\b m)" apply (unfold Finite_def) apply (erule bexE) apply (drule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]]) @@ -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]) @@ -133,7 +133,7 @@ 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" + \ \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 @@ -162,7 +162,7 @@ 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" + \ \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) @@ -208,10 +208,10 @@ 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}" + 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 " + \! w. w \ t_n \ z \ w " and disjoint[iff]: "x \ y = 0" and "includes": "t_n \ {v \ Pow(x \ y). v \ succ(k #+ m)}" and WO_R[iff]: "well_ord(y,R)" @@ -265,7 +265,7 @@ lemma ex1_superset_a: "\l \ a; a \ y; b \ y - a; u \ x\ - \ \! c. c \ s(u) & a \ c & b \ c" + \ \! 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) @@ -277,7 +277,7 @@ 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)" + \ (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+) @@ -289,7 +289,7 @@ \ 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" +apply (rule_tac f3 = "\b \ y-a. THE c. c \ s (u) \ a \ c \ b \ c" in exI [THEN lepoll_def [THEN def_imp_iff, THEN iffD2]]) apply (simp add: inj_def) apply (rule conjI) @@ -343,7 +343,7 @@ lemma cons_cons_in: "\z \ xa \ (x - {u}); l \ a; a \ y; u \ x\ - \ \! w. w \ t_n & cons(z, cons(u, a)) \ w" + \ \! 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) @@ -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]) @@ -430,7 +430,7 @@ 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,7 +440,7 @@ 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]) @@ -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]) @@ -543,7 +543,7 @@ done lemma "conclusion": - "\a f. Ord(a) & domain(f) = a & (\bb m)" + "\a f. Ord(a) \ domain(f) = a \ (\b (\b m)" apply (rule well_ord_LL [THEN exE]) apply (rename_tac S) apply (rule_tac x = "ordertype (LL,S)" in exI) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC16_lemmas.thy --- a/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 17:03:23 2022 +0100 @@ -106,7 +106,7 @@ 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) @@ -196,7 +196,7 @@ 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 @@ -220,7 +220,7 @@ 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) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC17_AC1.thy --- a/src/ZF/AC/AC17_AC1.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC17_AC1.thy Tue Sep 27 17:03:23 2022 +0100 @@ -13,7 +13,7 @@ (** AC0 is equivalent to AC1. - AC0 comes from Suppes, AC1 from Rubin & Rubin **) + AC0 comes from Suppes, AC1 from Rubin \ Rubin **) lemma AC0_AC1_lemma: "\f:(\X \ A. X); D \ A\ \ \g. g:(\X \ D. X)" by (fast intro!: lam_type apply_type) @@ -262,7 +262,7 @@ (* ********************************************************************** *) -(* 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" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC18_AC19.thy --- a/src/ZF/AC/AC18_AC19.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC18_AC19.thy Tue Sep 27 17:03:23 2022 +0100 @@ -56,7 +56,7 @@ (* ********************************************************************** *) 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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC7_AC9.thy --- a/src/ZF/AC/AC7_AC9.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC7_AC9.thy Tue Sep 27 17:03:23 2022 +0100 @@ -48,7 +48,7 @@ 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. *) (* ********************************************************************** *) @@ -86,7 +86,7 @@ (* ********************************************************************** *) lemma AC1_AC8_lemma1: - "\B \ A. \B1 B2. B= & B1 \ B2 + "\B \ A. \B1 B2. B= \ B1 \ B2 \ 0 \ { bij(fst(B),snd(B)). B \ A }" apply (unfold eqpoll_def, auto) done @@ -104,13 +104,13 @@ (* ********************************************************************** *) (* AC8 \ AC9 *) -(* - this proof replaces the following two from Rubin & Rubin: *) +(* - this proof replaces the following two from Rubin \ Rubin: *) (* AC8 \ AC1 and AC1 \ AC9 *) (* ********************************************************************** *) lemma AC8_AC9_lemma: "\B1 \ A. \B2 \ A. B1 \ B2 - \ \B \ A*A. \B1 B2. B= & B1 \ B2" + \ \B \ A*A. \B1 B2. B= \ B1 \ B2" by fast lemma AC8_AC9: "AC8 \ AC9" @@ -124,7 +124,7 @@ (* ********************************************************************** *) (* 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 *) +(* by Rubin \ Rubin. But (x * y) is not necessarily equipollent to *) (* (x * y) \ {0} when y is a set of total functions acting from nat to *) (* \(A) -- therefore we have used the set (y * nat) instead of y. *) (* ********************************************************************** *) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/AC_Equiv.thy --- a/src/ZF/AC/AC_Equiv.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/AC_Equiv.thy Tue Sep 27 17:03:23 2022 +0100 @@ -21,21 +21,21 @@ "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 & - (\bb m)" + "WO4(m) \ \A. \a f. Ord(a) \ domain(f)=a \ + (\b (\b 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 - & (\bb m))" + "WO6 \ \A. \m \ nat. 1\m \ (\a f. Ord(a) \ domain(f)=a + \ (\b (\b m))" definition "WO7 \ \A. Finite(A) \ (\R. well_ord(A,R) \ well_ord(A,converse(R)))" @@ -51,7 +51,7 @@ 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 *) @@ -62,7 +62,7 @@ "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)" @@ -77,10 +77,10 @@ "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 @@ -89,45 +89,45 @@ definition "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))" + (\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)) \ - (\n \ nat. 1\n & (\f. \B \ A. (pairwise_disjoint(f`B) & - sets_of_size_between(f`B, 2, succ(n)) & \(f`B)=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 \ - (\m \ nat. 1\m & (\f. \B \ A. f`B\0 & f`B \ B & f`B \ m))" + (\m \ nat. 1\m \ (\f. \B \ A. f`B\0 \ f`B \ B \ f`B \ m))" definition "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))" + (\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}. \f \ Pow(A)-{0} -> A. f`(g`f) \ g`f" locale AC18 = - assumes AC18: "A\0 & (\a \ A. B(a) \ 0) \ + assumes AC18: "A\0 \ (\a \ A. B(a) \ 0) \ ((\a \ A. \b \ B(a). X(a,b)) = (\f \ \a \ A. B(a). \a \ A. X(a, f`a)))" \ \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))" @@ -196,7 +196,7 @@ "\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" +apply (rule_tac x = "\x \ domain(u). \ i. \y. \ u \ f` = i" in exI) apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective) apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/Cardinal_aux.thy --- a/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 17:03:23 2022 +0100 @@ -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 A \ j" by (unfold lesspoll_def, blast dest!: lepoll_imp_ex_le_eqpoll elim!: leE) lemma Un_eqpoll_Inf_Ord: @@ -65,7 +65,7 @@ lemma paired_eqpoll: "{{y,z}. y \ x} \ x" by (unfold eqpoll_def, fast intro!: paired_bij) -lemma ex_eqpoll_disjoint: "\B. B \ A & B \ C = 0" +lemma ex_eqpoll_disjoint: "\B. B \ A \ B \ C = 0" by (fast intro!: paired_eqpoll equals0I elim: mem_asym) (*Finally we reach this result. Surely there's a simpler proof?*) @@ -108,7 +108,7 @@ 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" + \ \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,12 +126,12 @@ done lemma UN_fun_lepoll: - "\\b \ a. f`b \ n & f`b \ T; well_ord(T, R); + "\\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); + "\\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) @@ -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]) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/DC.thy --- a/src/ZF/AC/DC.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/DC.thy Tue Sep 27 17:03:23 2022 +0100 @@ -81,13 +81,13 @@ 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 @@ -103,7 +103,7 @@ 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}" + \ restrict(z2, domain(z1)) = z1}" begin (* ********************************************************************** *) @@ -116,7 +116,7 @@ (* Define XX and RR as follows: *) (* *) (* XX = (\n \ nat. {f \ n->X. \k \ n. \ R}) *) -(* f RR g iff domain(g)=succ(domain(f)) & *) +(* f RR g iff domain(g)=succ(domain(f)) \ *) (* restrict(g, domain(f)) = f *) (* *) (* Then RR satisfies the hypotheses of DC. *) @@ -170,8 +170,8 @@ lemma lemma2: "\\n \ nat. \ RR; f \ nat -> XX; n \ nat\ - \ \k \ nat. f`succ(n) \ k -> X & n \ k - & \ R" + \ \k \ nat. f`succ(n) \ k -> X \ n \ k + \ \ R" apply (induct_tac "n") apply (drule apply_type [OF _ nat_1I]) apply (drule bspec [OF _ nat_0I]) @@ -267,10 +267,10 @@ XX = (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R}) RR = {:Fin(XX)*XX. - (domain(z2)=succ(\f \ z1. domain(f)) & + (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(z2, domain(f)) = f)) | - (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) & - (\f \ z1. restrict(g, domain(f)) = f)) & + (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) \ + (\f \ z1. restrict(g, domain(f)) = f)) \ z2={<0,x>})} Then XX and RR satisfy the hypotheses of DC(omega). @@ -302,20 +302,20 @@ and RR_def: "RR \ {:Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) - & (\f \ z1. restrict(z2, domain(f)) = f)) + \ (\f \ z1. restrict(z2, domain(f)) = f)) | (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) - & (\f \ z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}" + \ (\f \ z1. restrict(g, domain(f)) = f)) \ z2={<0,x>})}" and allRR_def: "allRR \ \b \ {\Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) - & (\f \ z1. domain(f)) = b - & (\f \ z1. restrict(z2,domain(f)) = f))}" + \ (\f \ z1. domain(f)) = b + \ (\f \ z1. restrict(z2,domain(f)) = f))}" begin lemma lemma4: "\range(R) \ domain(R); x \ domain(R)\ - \ RR \ Pow(XX)*XX & + \ 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) @@ -323,7 +323,7 @@ apply (drule Finite_Fin [OF lesspoll_nat_is_Finite PowD], assumption) apply (case_tac "\g \ XX. domain (g) = - succ(\f \ Y. domain(f)) & (\f\Y. restrict(g, domain(f)) = f)") + succ(\f \ Y. domain(f)) \ (\f\Y. restrict(g, domain(f)) = f)") apply (simp add: RR_def, blast) apply (safe del: domainE) apply (unfold XX_def RR_def) @@ -436,7 +436,7 @@ lemma lemma2: "\allRR; f \ nat->XX; range(R) \ domain(R); x \ domain(R); n \ nat\ - \ f`n \ succ(n) -> domain(R) & (\i \ n. :R)" + \ f`n \ succ(n) -> domain(R) \ (\i \ n. :R)" apply (unfold allRR_def) apply (drule ospec) apply (erule ltI [OF _ Ord_nat]) @@ -512,7 +512,7 @@ apply (erule allE impE)+ apply (rule Card_Hartog) apply (erule_tac x = A in allE) -apply (erule_tac x = "{ \ Pow (A) *A . z1 \ Hartog (A) & z2 \ z1}" +apply (erule_tac x = "{ \ Pow (A) *A . z1 \ Hartog (A) \ z2 \ z1}" in allE) apply simp apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE]) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/HH.thy --- a/src/ZF/AC/HH.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/HH.thy Tue Sep 27 17:03:23 2022 +0100 @@ -32,7 +32,7 @@ "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 | b cOrd(a); a \ X\ - \ \Y. Y \ X & (\R. well_ord(Y,R) & ordertype(Y,R)=a)" + \ \Y. Y \ X \ (\R. well_ord(Y,R) \ ordertype(Y,R)=a)" apply (unfold lepoll_def) apply (erule exE) apply (intro exI conjI) @@ -36,7 +36,7 @@ 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) @@ -45,7 +45,7 @@ lemma Ords_lepoll_set_lemma: "(\a. Ord(a) \ a \ X) \ \a. Ord(a) \ - a \ {b. Z \ Pow(X)*Pow(X*X), \Y R. Z= & ordertype(Y,R)=b}" + a \ {b. Z \ Pow(X)*Pow(X*X), \Y R. Z= \ ordertype(Y,R)=b}" apply (intro allI impI) apply (elim allE impE, assumption) apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) @@ -54,7 +54,7 @@ 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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/WO1_AC.thy --- a/src/ZF/AC/WO1_AC.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/WO1_AC.thy Tue Sep 27 17:03:23 2022 +0100 @@ -89,8 +89,8 @@ lemma lemma2: "\WO1; \Finite(B); 1\n\ - \ \C \ Pow(Pow(B)). pairwise_disjoint(C) & - sets_of_size_between(C, 2, succ(n)) & + \ \C \ Pow(Pow(B)). pairwise_disjoint(C) \ + sets_of_size_between(C, 2, succ(n)) \ \(C)=B" apply (drule lemma2_1 [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], assumption) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/WO1_WO7.thy --- a/src/ZF/AC/WO1_WO7.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/WO1_WO7.thy Tue Sep 27 17:03:23 2022 +0100 @@ -2,7 +2,7 @@ Author: Lawrence C Paulson, CU Computer Laboratory Copyright 1998 University of Cambridge -WO7 \ LEMMA \ WO1 (Rubin & Rubin p. 5) +WO7 \ LEMMA \ WO1 (Rubin \ Rubin p. 5) LEMMA is the sentence denoted by (**) Also, WO1 \ WO8 @@ -14,7 +14,7 @@ definition "LEMMA \ - \X. \Finite(X) \ (\R. well_ord(X,R) & \well_ord(X,converse(R)))" + \X. \Finite(X) \ (\R. well_ord(X,R) \ \well_ord(X,converse(R)))" (* ********************************************************************** *) (* It is easy to see that WO7 is equivalent to (**) *) @@ -93,14 +93,14 @@ (* ********************************************************************** *) -(* The proof of WO8 \ WO1 (Rubin & Rubin p. 6) *) +(* The proof of WO8 \ WO1 (Rubin \ Rubin p. 6) *) (* ********************************************************************** *) lemma WO1_WO8: "WO1 \ WO8" by (unfold WO1_def WO8_def, fast) -(* The implication "WO8 \ WO1": a faithful image of Rubin & Rubin's proof*) +(* 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) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/AC/WO2_AC16.thy --- a/src/ZF/AC/WO2_AC16.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/AC/WO2_AC16.thy Tue Sep 27 17:03:23 2022 +0100 @@ -9,7 +9,7 @@ The first instance is trivial, the third not difficult, but the second is very complicated requiring many lemmas. We also need to prove that at any stage gamma the set -(s - \(...) - k_gamma) (Rubin & Rubin page 15) +(s - \(...) - k_gamma) (Rubin \ Rubin page 15) contains m distinct elements (in fact is equipollent to s) *) @@ -22,7 +22,7 @@ "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 & + else r \ {f`(\ i. h`g \ f`i \ (\b f`i \ (\t \ r. \ h`b \ t))))})" (* ********************************************************************** *) @@ -36,7 +36,7 @@ "recfunAC16(f,h,succ(i),a) = (if (\y \ recfunAC16(f,h,i,a). h ` i \ y) then recfunAC16(f,h,i,a) else recfunAC16(f,h,i,a) \ - {f ` (\ j. h ` i \ f ` j & + {f ` (\ j. h ` i \ f ` j \ (\b f`j \ (\t \ recfunAC16(f,h,i,a). \ h`b \ t))))})" apply (simp add: recfunAC16_def) @@ -86,7 +86,7 @@ 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" @@ -96,7 +96,7 @@ 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" @@ -106,25 +106,25 @@ done lemma lemma4: - "\\y X & + "\\y X \ (\xY \ F(y). h(x) \ Y) \ - (\! Y. Y \ F(y) & h(x) \ Y)); + (\! Y. Y \ F(y) \ h(x) \ Y)); x < a\ \ \yzY \ F(y). h(z) \ Y) \ - (\! Y. Y \ F(y) & h(z) \ Y)" + (\! Y. Y \ F(y) \ h(z) \ Y)" apply (intro oallI impI) apply (drule ospec, assumption, clarify) apply (blast elim!: oallE ) done lemma lemma5: - "\\y X & + "\\y X \ (\xY \ F(y). h(x) \ Y) \ - (\! Y. Y \ 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 X \ (\xax \ \x x) - \ (\! Y. Y \ (\x Y))" + \ (\! Y. Y \ (\x h(xa) \ Y))" apply (rule conjI) apply (rule subsetI) apply (erule OUN_E) @@ -148,7 +148,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 gamma the set (s - \(...) - k_gamma) is equipollent to s - (Rubin & Rubin page 15). + (Rubin \ Rubin page 15). *) (* ********************************************************************** *) @@ -177,7 +177,7 @@ done lemma Union_lesspoll: - "\\x \ X. x \ n & x \ T; well_ord(T, R); X \ b; + "\\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)") @@ -238,7 +238,7 @@ lemma Least_eq_imp_ex: "\(\ i. w \ F(i)) = (\ i. z \ F(i)); w \ (\i (\i - \ \b (F(b) - (\c (F(b) - (\c \b (F(b) - (\c z \ (F(b) - (\c\yx Q(x,y)); succ(j) - \ F(j)<=X & (\x Q(x,j))" + "\\y (\x Q(x,y)); succ(j) + \ F(j)<=X \ (\x Q(x,j))" by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]) @@ -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) @@ -406,7 +406,7 @@ 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 & + \ \X \ {Y \ Pow(A). Y\succ(k #+ m)}. h`y \ X \ (\b X \ (\T \ recfunAC16(f, h, y, a). \ h`b \ T))" apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], @@ -433,7 +433,7 @@ h \ 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 & + \ \c f`c \ (\b f`c \ (\T \ recfunAC16(f, h, y, a). \ h`b \ T))" apply (drule ex_next_set, assumption+) @@ -451,11 +451,11 @@ lemma lemma8: "\\xxa \ F(j). P(x, xa)) - \ (\! Y. Y \ F(j) & P(x, Y)); F(j) \ X; - L \ X; P(j, L) & (\x (\xa \ F(j). \P(x, xa)))\ - \ F(j) \ {L} \ X & + \ (\! Y. Y \ F(j) \ P(x, Y)); F(j) \ X; + L \ X; P(j, L) \ (\x (\xa \ F(j). \P(x, xa)))\ + \ F(j) \ {L} \ X \ (\xj | (\xa \ (F(j) \ {L}). P(x, xa)) \ - (\! Y. Y \ (F(j) \ {L}) & P(x, Y)))" + (\! Y. Y \ (F(j) \ {L}) \ P(x, Y)))" apply (rule conjI) apply (fast intro!: singleton_subsetI) apply (rule oallI) @@ -471,9 +471,9 @@ "\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)} & + \ 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))" + (\! Y. Y \ recfunAC16(f, h, b, a) \ h ` x \ Y))" apply (erule lt_induct) apply (frule lt_Ord) apply (erule Ord_cases) @@ -507,12 +507,12 @@ (* ********************************************************************** *) lemma lemma_simp_induct: - "\\b. b F(b) \ S & (\xY \ F(b). f`x \ Y)) - \ (\! Y. Y \ 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 & - (\x \ f``a. \! Y. Y \ (\j Y)" + \ (\j S \ + (\x \ f``a. \! Y. Y \ (\j x \ Y)" apply (rule conjI) apply (rule subsetI) apply (erule OUN_E, blast) @@ -563,7 +563,7 @@ apply (elim exE) apply (rename_tac h) apply (rule_tac x = "\j WO1. Every proof (except WO6 \ WO1 and WO1 \ WO2) are described as "clear" -by Rubin & Rubin (page 2). +by Rubin \ Rubin (page 2). They refer reader to a book by Gödel to see the proof WO1 \ WO2. Fortunately order types made this proof also very easy. *) @@ -17,8 +17,8 @@ (* Auxiliary definitions used in proof *) definition NN :: "i => i" where - "NN(y) \ {m \ nat. \a. \f. Ord(a) & domain(f)=a & - (\bb m)}" + "NN(y) \ {m \ nat. \a. \f. Ord(a) \ domain(f)=a \ + (\b (\b m)}" definition uu :: "[i, i, i, i] => i" where @@ -29,9 +29,9 @@ definition vv1 :: "[i, i, i] => i" where "vv1(f,m,b) \ - let g = \ g. (\d. Ord(d) & (domain(uu(f,b,g,d)) \ 0 & + 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 & + d = \ d. domain(uu(f,b,g,d)) \ 0 \ domain(uu(f,b,g,d)) \ m in if f`b \ 0 then domain(uu(f,b,g,d)) else 0" @@ -127,13 +127,13 @@ (* ********************************************************************** The proof of "WO6 \ WO1". Simplified by L C Paulson. - From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin, + From the book "Equivalents of the Axiom of Choice" by Rubin \ Rubin, pages 2-5 ************************************************************************* *) lemma lt_oadd_odiff_disj: "\k < i++j; Ord(i); Ord(j)\ - \ k < i | (\ k k < i | (\ k k = i ++ (k--i) \ (k--i)bgd m \ (\b 0 \ - (\gd 0 & u(f,b,g,d) \ m)) - | (\b 0 & (\gd 0 \ + (\gd 0 \ u(f,b,g,d) \ m)) + | (\b 0 \ (\gd 0 \ u(f,b,g,d) \ m))" apply (unfold lesspoll_def) apply (blast del: equalityI) @@ -210,7 +210,7 @@ (* ********************************************************************** *) lemma nested_LeastI: "\P(a, b); Ord(a); Ord(b); - Least_a = (\ a. \x. Ord(x) & P(a, x))\ + 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) @@ -218,13 +218,13 @@ done lemmas nested_Least_instance = - nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \ 0 & + nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \ 0 \ domain(uu(f,b,g,d)) \ m"] for f b m lemma gg1_lepoll_m: "\Ord(a); m \ nat; \b0 \ - (\gd 0 & + (\gd 0 \ domain(uu(f,b,g,d)) \ m); \b succ(m); b \ gg1(f,a,m)`b \ m" @@ -439,7 +439,7 @@ done (* ********************************************************************** *) -(* Rubin & Rubin wrote, *) +(* Rubin \ Rubin wrote, *) (* "It follows from (ii) and mathematical induction that if y*y \ y then *) (* y can be well-ordered" *) @@ -471,7 +471,7 @@ 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) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Arith.thy --- a/src/ZF/Arith.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Arith.thy Tue Sep 27 17:03:23 2022 +0100 @@ -393,7 +393,7 @@ 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: @@ -537,7 +537,7 @@ lemma lt_succ_eq_0_disj: "\m\nat; n\nat\ - \ (m < succ(n)) \ (m = 0 | (\j\nat. m = succ(j) & j < n))" + \ (m < succ(n)) \ (m = 0 | (\j\nat. m = succ(j) \ j < n))" by (induct_tac "m", auto) lemma less_diff_conv [rule_format]: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ArithSimp.thy --- a/src/ZF/ArithSimp.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ArithSimp.thy Tue Sep 27 17:03:23 2022 +0100 @@ -303,23 +303,23 @@ apply (simp (no_asm_simp) add: mult_lt_mono2 mult_commute [of _ k]) done -lemma add_eq_0_iff [iff]: "m#+n = 0 \ natify(m)=0 & natify(n)=0" -apply (subgoal_tac "natify (m) #+ natify (n) = 0 \ natify (m) =0 & natify (n) =0") +lemma add_eq_0_iff [iff]: "m#+n = 0 \ natify(m)=0 \ natify(n)=0" +apply (subgoal_tac "natify (m) #+ natify (n) = 0 \ natify (m) =0 \ natify (n) =0") apply (rule_tac [2] n = "natify (m) " in natE) apply (rule_tac [4] n = "natify (n) " in natE) apply auto done -lemma zero_lt_mult_iff [iff]: "0 < m#*n \ 0 < natify(m) & 0 < natify(n)" -apply (subgoal_tac "0 < natify (m) #*natify (n) \ 0 < natify (m) & 0 < natify (n) ") +lemma zero_lt_mult_iff [iff]: "0 < m#*n \ 0 < natify(m) \ 0 < natify(n)" +apply (subgoal_tac "0 < natify (m) #*natify (n) \ 0 < natify (m) \ 0 < natify (n) ") apply (rule_tac [2] n = "natify (m) " in natE) apply (rule_tac [4] n = "natify (n) " in natE) apply (rule_tac [3] n = "natify (n) " in natE) apply auto done -lemma mult_eq_1_iff [iff]: "m#*n = 1 \ natify(m)=1 & natify(n)=1" -apply (subgoal_tac "natify (m) #* natify (n) = 1 \ natify (m) =1 & natify (n) =1") +lemma mult_eq_1_iff [iff]: "m#*n = 1 \ natify(m)=1 \ natify(n)=1" +apply (subgoal_tac "natify (m) #* natify (n) = 1 \ natify (m) =1 \ natify (n) =1") apply (rule_tac [2] n = "natify (m) " in natE) apply (rule_tac [4] n = "natify (n) " in natE) apply 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 m (0 < natify(k) & natify(m) < natify(n))" + "(m#*k < n#*k) \ (0 < natify(k) \ natify(m) < natify(n))" apply (rule iff_trans) apply (rule_tac [2] mult_less_cancel_lemma, auto) done lemma mult_less_cancel1 [simp]: - "(k#*m < k#*n) \ (0 < natify(k) & natify(m) < natify(n))" + "(k#*m < k#*n) \ (0 < natify(k) \ natify(m) < natify(n))" apply (simp (no_asm) add: mult_less_cancel2 mult_commute [of k]) done @@ -374,7 +374,7 @@ 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: @@ -523,7 +523,7 @@ lemma raw_nat_diff_split: "\a:nat; b:nat\ \ - (P(a #- b)) \ ((a < b \P(0)) & (\d\nat. a = b #+ d \ P(d)))" + (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) apply (rule iffI, force, simp) @@ -533,7 +533,7 @@ lemma nat_diff_split: "(P(a #- b)) \ - (natify(a) < natify(b) \P(0)) & (\d\nat. natify(a) = b #+ d \ P(d))" + (natify(a) < natify(b) \P(0)) \ (\d\nat. natify(a) = b #+ d \ P(d))" apply (cut_tac P=P and a="natify(a)" and b="natify(b)" in raw_nat_diff_split) apply simp_all done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Bin.thy --- a/src/ZF/Bin.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Bin.thy Tue Sep 27 17:03:23 2022 +0100 @@ -377,7 +377,7 @@ lemma iszero_integ_of_BIT: "\w \ bin; x \ bool\ - \ iszero (integ_of (w BIT x)) \ (x=0 & iszero (integ_of(w)))" + \ 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 @@ -578,7 +578,7 @@ apply (blast intro: zless_zle_trans) done -lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) \ ($#0 $< z & w $< z)" +lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) \ ($#0 $< z \ w $< z)" apply (case_tac "$#0 $< z") apply (auto simp add: zless_nat_conj_lemma nat_le_int0 not_zless_iff_zle) done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Cardinal.thy Tue Sep 27 17:03:23 2022 +0100 @@ -10,7 +10,7 @@ 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 @@ -22,7 +22,7 @@ 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 @@ -56,9 +56,9 @@ lemma decomposition: "\f \ X->Y; g \ Y->X\ \ - \XA XB YA YB. (XA \ XB = 0) & (XA \ XB = X) & - (YA \ YB = 0) & (YA \ YB = Y) & - f``XA=YA & g``YB=XB" + \XA XB YA YB. (XA \ XB = 0) \ (XA \ XB = X) \ + (YA \ YB = 0) \ (YA \ YB = Y) \ + f``XA=YA \ g``YB=XB" apply (intro exI conjI) apply (rule_tac [6] Banach_last_equation) apply (rule_tac [5] refl) @@ -134,7 +134,7 @@ "\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" +lemma eqpoll_iff: "X \ Y \ X \ Y \ Y \ X" by (blast intro: eqpollI elim!: eqpollE) lemma lepoll_0_is_0: "A \ 0 \ A = 0" @@ -303,13 +303,13 @@ (*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 lemma Ord_Least [intro,simp,TC]: "Ord(\ x. P(x))" -proof (cases "\i. Ord(i) & P(i)") +proof (cases "\i. Ord(i) \ P(i)") case True then obtain i where "P(i)" "Ord(i)" by auto hence " (\ x. P(x)) \ i" by (rule Least_le) @@ -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" @@ -438,7 +438,7 @@ lemma Card_cardinal [iff]: "Card(|A|)" proof (unfold cardinal_def) show "Card(\ i. i \ A)" - proof (cases "\i. Ord (i) & i \ A") + proof (cases "\i. Ord (i) \ i \ A") case False thus ?thesis \ \degenerate case\ by (simp add: Least_0 Card_0) next @@ -620,7 +620,7 @@ 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 @@ -962,7 +962,7 @@ 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]) @@ -1014,7 +1014,7 @@ intro: Un_upper1 [THEN Fin_mono, THEN subsetD] Un_upper2 [THEN Fin_mono, THEN subsetD]) -lemma Finite_Un_iff [simp]: "Finite(A \ B) \ (Finite(A) & Finite(B))" +lemma Finite_Un_iff [simp]: "Finite(A \ B) \ (Finite(A) \ Finite(B))" by (blast intro: subset_Finite Finite_Un) text\The converse must hold too.\ @@ -1040,7 +1040,7 @@ apply (case_tac "a \ A") apply (subgoal_tac [2] "A-{a}=A", auto) apply (rule_tac x = "succ (n) " in bexI) -apply (subgoal_tac "cons (a, A - {a}) = A & cons (n, n) = succ (n) ") +apply (subgoal_tac "cons (a, A - {a}) = A \ cons (n, n) = succ (n) ") apply (drule_tac a = a and b = n in cons_eqpoll_cong) apply (auto dest: mem_irrefl) done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/CardinalArith.thy Tue Sep 27 17:03:23 2022 +0100 @@ -9,7 +9,7 @@ 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 @@ -31,13 +31,13 @@ \ \This definition is more complex than Kunen's but it more easily proved to be a cardinal\ "jump_cardinal(K) \ - \X\Pow(K). {z. r \ Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}" + \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) \ KA" - hence "\c\A. j < c & Card(c)" using A + hence "\c\A. j < c \ Card(c)" using A by (auto simp add: lt_def intro: Card_is_Ord) then obtain c where c: "c\A" "j < c" "Card(c)" by blast @@ -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) @@ -510,17 +510,17 @@ lemma csquare_ltI: "\x \ <, > \ csquare_rel(K)" apply (unfold csquare_rel_def) -apply (subgoal_tac "x y apply *) lemma csquare_or_eqI: - "\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 "x yLimit(K); x y)\ \ ordermap(K*K, csquare_rel(K)) ` < ordermap(K*K, csquare_rel(K)) ` " -apply (subgoal_tac "z well_ord (K*K, csquare_rel (K))") prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ Limit_is_Ord [THEN well_ord_csquare], clarify) apply (rule csquare_ltI [THEN ordermap_mono, THEN ltI]) @@ -745,7 +745,7 @@ (*Allows selective unfolding. Less work than deriving intro/elim rules*) lemma jump_cardinal_iff: "i \ jump_cardinal(K) \ - (\r X. r \ K*K & X \ K & well_ord(X,r) & i = ordertype(X,r))" + (\r X. r \ K*K \ X \ K \ well_ord(X,r) \ i = ordertype(X,r))" apply (unfold jump_cardinal_def) apply (blast del: subsetI) done @@ -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)+ diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Coind/ECR.thy --- a/src/ZF/Coind/ECR.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Coind/ECR.thy Tue Sep 27 17:03:23 2022 +0100 @@ -28,7 +28,7 @@ definition hastyenv :: "[i,i] => o" where "hastyenv(ve,te) \ - ve_dom(ve) = te_dom(te) & + ve_dom(ve) = te_dom(te) \ (\x \ ve_dom(ve). \ HasTyRel)" (* Specialised co-induction rule *) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Coind/Static.thy --- a/src/ZF/Coind/Static.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Coind/Static.thy Tue Sep 27 17:03:23 2022 +0100 @@ -17,9 +17,9 @@ definition isofenv :: "[i,i] => o" where "isofenv(ve,te) \ - ve_dom(ve) = te_dom(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)))" + \c \ Const. ve_app(ve,x) = v_const(c) \ isof(c,te_app(te,x)))" (*** Elaboration ***) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/AC_in_L.thy --- a/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 17:03:23 2022 +0100 @@ -56,7 +56,7 @@ { fix y ys assume "y \ A" and "ys \ list(A)" with Cons - have "\Cons(x,xs),Cons(y,ys)\ \ rlist(A,r) \ x=y & xs = ys \ \Cons(y,ys), Cons(x,xs)\ \ rlist(A,r)" + have "\Cons(x,xs),Cons(y,ys)\ \ rlist(A,r) \ x=y \ xs = ys \ \Cons(y,ys), Cons(x,xs)\ \ rlist(A,r)" apply (rule_tac i = "length(xs)" and j = "length(ys)" in Ord_linear_lt) apply (simp_all add: shorterI) apply (rule linearE [OF r, of x y]) @@ -172,7 +172,7 @@ lemma (in Nat_Times_Nat) fn_iff: "\x \ nat; y \ nat; u \ nat; v \ nat\ - \ (fn` = fn`) \ (x=u & y=v)" + \ (fn` = fn`) \ (x=u \ y=v)" by (blast dest: inj_apply_equality [OF fn_inj]) lemma (in Nat_Times_Nat) enum_type [TC,simp]: @@ -247,8 +247,8 @@ \ \predicate that holds if \<^term>\k\ is a valid index for \<^term>\X\\ "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))} & + arity(p) \ succ(length(env)) \ + X = {x\A. sats(A, p, Cons(x,env))} \ env_form_map(f,r,A,) = k" definition @@ -335,9 +335,9 @@ "rlimit(i,r) \ if Limit(i) then {z: Lset(i) * Lset(i). - \x' x. z = & + \x' x. z = \ (lrank(x') < lrank(x) | - (lrank(x') = lrank(x) & \ r(succ(lrank(x)))))} + (lrank(x') = lrank(x) \ \ r(succ(lrank(x)))))} else 0" definition diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/DPow_absolute.thy --- a/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 17:03:23 2022 +0100 @@ -17,9 +17,9 @@ (* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" "is_formula_rec(M,MH,p,z) \ - \dp[M]. \i[M]. \f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) & + \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)" + successor(M,dp,i) \ fun_apply(M,f,p,z) \ is_transrec(M,MH,i,f)" *) definition @@ -113,7 +113,7 @@ lemma DPow'_eq: "DPow'(A) = {z . ep \ list(A) * formula, \env \ list(A). \p \ formula. - ep = & z = {x\A. sats(A, p, Cons(x,env))}}" + ep = \ z = {x\A. sats(A, p, Cons(x,env))}}" by (simp add: DPow'_def, blast) @@ -192,8 +192,8 @@ and rep: "M(A) \ strong_replacement (M, - \ep z. \env[M]. \p[M]. mem_formula(M,p) & mem_list(M,A,env) & - pair(M,env,p,ep) & + \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': @@ -205,7 +205,7 @@ "M(A) \ strong_replacement (M, \ep z. \env\list(A). \p\formula. - ep = & z = {x \ A . sats(A, p, Cons(x, env))})" + ep = \ z = {x \ A . sats(A, p, Cons(x, env))})" by (insert rep [of A], simp add: Collect_DPow_sats_abs) @@ -223,8 +223,8 @@ is_DPow' :: "[i=>o,i,i] => o" where "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) & + 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: @@ -255,14 +255,14 @@ subsubsection\The Instance of Replacement\ lemma DPow_replacement_Reflects: - "REFLECTS [\x. \u[L]. u \ B & + "REFLECTS [\x. \u[L]. u \ B \ (\env[L]. \p[L]. - mem_formula(L,p) & mem_list(L,A,env) & pair(L,env,p,u) & + mem_formula(L,p) \ mem_list(L,A,env) \ pair(L,env,p,u) \ is_Collect (L, A, is_DPow_sats(L,A,env,p), x)), - \i x. \u \ Lset(i). u \ B & + \i x. \u \ Lset(i). u \ B \ (\env \ Lset(i). \p \ Lset(i). - mem_formula(##Lset(i),p) & mem_list(##Lset(i),A,env) & - pair(##Lset(i),env,p,u) & + mem_formula(##Lset(i),p) \ mem_list(##Lset(i),A,env) \ + pair(##Lset(i),env,p,u) \ is_Collect (##Lset(i), A, is_DPow_sats(##Lset(i),A,env,p), x))]" apply (unfold is_Collect_def) apply (intro FOL_reflections function_reflections mem_formula_reflection @@ -272,8 +272,8 @@ lemma DPow_replacement: "L(A) \ strong_replacement (L, - \ep z. \env[L]. \p[L]. mem_formula(L,p) & mem_list(L,A,env) & - pair(L,env,p,ep) & + \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))" apply (rule strong_replacementI) apply (rule_tac u="{A,B}" @@ -309,7 +309,7 @@ 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 @@ -360,7 +360,7 @@ 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 @@ -412,8 +412,8 @@ (* "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) & + 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))" *) definition @@ -458,11 +458,11 @@ definition transrec_body :: "[i=>o,i,i,i,i] => o" where "transrec_body(M,g,x) \ - \y z. \gy[M]. y \ x & fun_apply(M,g,y,gy) & is_DPow'(M,gy,z)" + \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)" + \ 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 + @@ -470,13 +470,13 @@ "\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. - \r[M]. is_Replace(M, x, transrec_body(M,f,x), r) & + \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))" + \ 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: @@ -535,10 +535,10 @@ subsubsection\The First Instance of Replacement\ lemma strong_rep_Reflects: - "REFLECTS [\u. \v[L]. v \ B & (\gy[L]. - v \ x & fun_apply(L,g,v,gy) & is_DPow'(L,gy,u)), - \i u. \v \ Lset(i). v \ B & (\gy \ Lset(i). - v \ x & fun_apply(##Lset(i),g,v,gy) & is_DPow'(##Lset(i),gy,u))]" + "REFLECTS [\u. \v[L]. v \ B \ (\gy[L]. + v \ x \ fun_apply(L,g,v,gy) \ is_DPow'(L,gy,u)), + \i u. \v \ Lset(i). v \ B \ (\gy \ Lset(i). + v \ x \ fun_apply(##Lset(i),g,v,gy) \ is_DPow'(##Lset(i),gy,u))]" by (intro FOL_reflections function_reflections DPow'_reflection) lemma strong_rep: @@ -555,18 +555,18 @@ subsubsection\The Second Instance of Replacement\ lemma transrec_rep_Reflects: - "REFLECTS [\x. \v[L]. v \ B & - (\y[L]. pair(L,v,y,x) & + "REFLECTS [\x. \v[L]. v \ B \ + (\y[L]. pair(L,v,y,x) \ is_wfrec (L, \x f u. \r[L]. is_Replace (L, x, \y z. - \gy[L]. y \ x & fun_apply(L,f,y,gy) & - is_DPow'(L,gy,z), r) & big_union(L,r,u), mr, v, y)), - \i x. \v \ Lset(i). v \ B & - (\y \ Lset(i). pair(##Lset(i),v,y,x) & + \gy[L]. y \ x \ fun_apply(L,f,y,gy) \ + is_DPow'(L,gy,z), r) \ big_union(L,r,u), mr, v, y)), + \i x. \v \ Lset(i). v \ B \ + (\y \ Lset(i). pair(##Lset(i),v,y,x) \ is_wfrec (##Lset(i), \x f u. \r \ Lset(i). is_Replace (##Lset(i), x, \y z. - \gy \ Lset(i). y \ x & fun_apply(##Lset(i),f,y,gy) & - is_DPow'(##Lset(i),gy,z), r) & + \gy \ Lset(i). y \ x \ fun_apply(##Lset(i),f,y,gy) \ + is_DPow'(##Lset(i),gy,z), r) \ big_union(##Lset(i),r,u), mr, v, y))]" apply (simp only: rex_setclass_is_bex [symmetric]) \ \Convert \\y\Lset(i)\ to \\y[##Lset(i)]\ within the body @@ -579,7 +579,7 @@ lemma transrec_rep: "\L(j)\ \ transrec_replacement(L, \x f u. - \r[L]. is_Replace(L, x, transrec_body(L,f,x), r) & + \r[L]. is_Replace(L, x, transrec_body(L,f,x), r) \ big_union(L, r, u), j)" apply (rule L.transrec_replacementI, assumption) apply (unfold transrec_body_def) @@ -615,7 +615,7 @@ definition constructible :: "[i=>o,i] => o" where "constructible(M,x) \ - \i[M]. \Li[M]. ordinal(M,i) & is_Lset(M,i,Li) & x \ Li" + \i[M]. \Li[M]. ordinal(M,i) \ is_Lset(M,i,Li) \ x \ Li" theorem V_equals_L_in_L: "L(x) \ constructible(L,x)" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Datatype_absolute.thy --- a/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 17:03:23 2022 +0100 @@ -11,7 +11,7 @@ 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 @@ -116,13 +116,13 @@ definition iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where "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), + 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) \ - \sn[M]. \msn[M]. successor(M,n,sn) & membership(M,sn,msn) & + \sn[M]. \msn[M]. successor(M,n,sn) \ membership(M,sn,msn) \ is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)" definition @@ -212,7 +212,7 @@ is_list_functor :: "[i=>o,i,i,i] => o" where "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)" + 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)" @@ -265,9 +265,9 @@ is_formula_functor :: "[i=>o,i,i] => o" where "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) & - cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & + 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)" lemma (in M_trancl) formula_functor_abs [simp]: @@ -287,7 +287,7 @@ by (simp add: list_N_def Nil_def) lemma Cons_in_list_N [simp]: - "Cons(a,l) \ list_N(A,succ(n)) \ a\A & l \ list_N(A,n)" + "Cons(a,l) \ list_N(A,succ(n)) \ a\A \ l \ list_N(A,n)" by (simp add: list_N_def Cons_def) text\These two aren't simprules because they reveal the underlying @@ -343,14 +343,14 @@ definition is_list_N :: "[i=>o,i,i,i] => o" where "is_list_N(M,A,n,Z) \ - \zero[M]. empty(M,zero) & + \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) \ \n[M]. \listn[M]. - finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \ listn" + finite_ordinal(M,n) \ is_list_N(M,A,n,listn) \ l \ listn" definition is_list :: "[i=>o,i,i] => o" where @@ -374,15 +374,15 @@ "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" + "Member(x,y) \ formula_N(succ(n)) \ x \ nat \ y \ nat" by (simp add: formula_N_def Member_def) lemma Equal_in_formula_N [simp]: - "Equal(x,y) \ formula_N(succ(n)) \ x \ nat & y \ nat" + "Equal(x,y) \ formula_N(succ(n)) \ x \ nat \ y \ nat" by (simp add: formula_N_def Equal_def) lemma Nand_in_formula_N [simp]: - "Nand(x,y) \ formula_N(succ(n)) \ x \ formula_N(n) & y \ formula_N(n)" + "Nand(x,y) \ formula_N(succ(n)) \ x \ formula_N(n) \ y \ formula_N(n)" by (simp add: formula_N_def Nand_def) lemma Forall_in_formula_N [simp]: @@ -447,7 +447,7 @@ definition is_formula_N :: "[i=>o,i,i] => o" where "is_formula_N(M,n,Z) \ - \zero[M]. empty(M,zero) & + \zero[M]. empty(M,zero) \ is_iterates(M, is_formula_functor(M), zero, n, Z)" @@ -455,7 +455,7 @@ mem_formula :: "[i=>o,i] => o" where "mem_formula(M,p) \ \n[M]. \formn[M]. - finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \ formn" + finite_ordinal(M,n) \ is_formula_N(M,n,formn) \ p \ formn" definition is_formula :: "[i=>o,i] => o" where @@ -466,12 +466,12 @@ "M(A) \ iterates_replacement(M, is_list_functor(M,A), 0)" and list_replacement2: "M(A) \ strong_replacement(M, - \n y. n\nat & is_iterates(M, is_list_functor(M,A), 0, n, y))" + \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)" and formula_replacement2: "strong_replacement(M, - \n y. n\nat & is_iterates(M, is_formula_functor(M), 0, n, y))" + \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)" @@ -479,7 +479,7 @@ 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)"]]) @@ -527,7 +527,7 @@ subsubsection\Absoluteness of Formulas\ lemma (in M_datatypes) formula_replacement2': - "strong_replacement(M, \n y. n\nat & y = (\X. ((nat*nat) + (nat*nat)) + (X*X + X))^n (0))" + "strong_replacement(M, \n y. n\nat \ y = (\X. ((nat*nat) + (nat*nat)) + (X*X + X))^n (0))" apply (insert formula_replacement2) apply (rule strong_replacement_cong [THEN iffD1]) apply (rule conj_cong [OF iff_refl iterates_abs [of "is_formula_functor(M)"]]) @@ -595,7 +595,7 @@ mem_eclose :: "[i=>o,i,i] => o" where "mem_eclose(M,A,l) \ \n[M]. \eclosen[M]. - finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \ eclosen" + finite_ordinal(M,n) \ is_eclose_n(M,A,n,eclosen) \ l \ eclosen" definition is_eclose :: "[i=>o,i,i] => o" where @@ -607,10 +607,10 @@ "M(A) \ iterates_replacement(M, big_union(M), A)" and eclose_replacement2: "M(A) \ strong_replacement(M, - \n y. n\nat & is_iterates(M, big_union(M), A, n, y))" + \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)"]]) @@ -653,14 +653,14 @@ is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where "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) & + 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) \ \sa[M]. \esa[M]. \mesa[M]. - upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) & + upair(M,a,a,sa) \ is_eclose(M,sa,esa) \ membership(M,esa,mesa) \ wfrec_replacement(M,MH,mesa)" text\The condition \<^term>\Ord(i)\ lets us use the simpler @@ -688,7 +688,7 @@ lemma (in M_eclose) transrec_replacementI: "\M(a); strong_replacement (M, - \x z. \y[M]. pair(M, x, y, z) & + \x z. \y[M]. pair(M, x, y, z) \ is_wfrec(M,MH,Memrel(eclose({a})),x,y))\ \ transrec_replacement(M,MH,a)" by (simp add: transrec_replacement_def wfrec_replacement_def) @@ -701,13 +701,13 @@ is_length :: "[i=>o,i,i,i] => o" where "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" + 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)" -apply (subgoal_tac "M(l) & M(n)") +apply (subgoal_tac "M(l) \ M(n)") prefer 2 apply (blast dest: transM) apply (simp add: is_length_def) apply (blast intro: list_imp_list_N nat_into_Ord list_N_imp_eq_length @@ -748,7 +748,7 @@ definition is_nth :: "[i=>o,i,i,i] => o" where "is_nth(M,n,l,Z) \ - \X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,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)\ @@ -767,46 +767,46 @@ is_Member :: "[i=>o,i,i,i] => o" where \ \because \<^term>\Member(x,y) \ Inl(Inl(\x,y\))\\ "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)" + \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))" by (simp add: is_Member_def Member_def) lemma (in M_trivial) Member_in_M_iff [iff]: - "M(Member(x,y)) \ M(x) & M(y)" + "M(Member(x,y)) \ M(x) \ M(y)" by (simp add: Member_def) 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) \ - \p[M]. \u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,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))" 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)" +lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) \ M(x) \ M(y)" by (simp add: Equal_def) 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) \ - \p[M]. \u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,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))" 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)" +lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) \ M(x) \ M(y)" by (simp add: Nand_def) 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))" @@ -857,13 +857,13 @@ is_depth :: "[i=>o,i,i] => o" where "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" + 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)" -apply (subgoal_tac "M(p) & M(n)") +apply (subgoal_tac "M(p) \ M(n)") prefer 2 apply (blast dest: transM) apply (simp add: is_depth_def) apply (blast intro: formula_imp_formula_N nat_into_Ord formula_N_imp_eq_depth @@ -884,11 +884,11 @@ \ \no constraint on non-formulas\ "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)) & + is_Member(M,x,y,p) \ is_a(x,y,z)) \ (\x[M]. \y[M]. finite_ordinal(M,x) \ finite_ordinal(M,y) \ - is_Equal(M,x,y,p) \ is_b(x,y,z)) & + is_Equal(M,x,y,p) \ is_b(x,y,z)) \ (\x[M]. \y[M]. mem_formula(M,x) \ mem_formula(M,y) \ - is_Nand(M,x,y,p) \ is_c(x,y,z)) & + is_Nand(M,x,y,p) \ is_c(x,y,z)) \ (\x[M]. mem_formula(M,x) \ is_Forall(M,x,p) \ is_d(x,z))" lemma (in M_datatypes) formula_case_abs [simp]: @@ -917,8 +917,8 @@ 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) \ - \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)" + \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)" text\Sufficient conditions to relativize the instance of \<^term>\formula_case\ @@ -968,7 +968,7 @@ and fr_lam_replace: "M(g) \ strong_replacement - (M, \x y. x \ formula & + (M, \x y. x \ formula \ y = \x, formula_rec_case(a,b,c,d,g,x)\)" lemma (in Formula_Rec) formula_rec_case_closed: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Formula.thy --- a/src/ZF/Constructible/Formula.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Formula.thy Tue Sep 27 17:03:23 2022 +0100 @@ -97,7 +97,7 @@ 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]: @@ -116,7 +116,7 @@ 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]: @@ -159,7 +159,7 @@ 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 \ Q) \ sats(A, And(p,q), env)" by (simp) lemma disj_iff_sats: @@ -394,7 +394,7 @@ DPow :: "i => i" where "DPow(A) \ {X \ Pow(A). \env \ list(A). \p \ formula. - arity(p) \ succ(length(env)) & + arity(p) \ succ(length(env)) \ X = {x\A. sats(A, p, Cons(x,env))}}" lemma DPowI: @@ -411,9 +411,9 @@ lemma DPowD: "X \ DPow(A) - \ X \ A & + \ X \ A \ (\env \ list(A). - \p \ formula. arity(p) \ succ(length(env)) & + \p \ formula. arity(p) \ succ(length(env)) \ X = {x\A. sats(A, p, Cons(x,env))})" by (simp add: DPow_def) @@ -593,7 +593,7 @@ 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)))" @@ -841,7 +841,7 @@ 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))" @@ -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) @@ -923,7 +923,7 @@ 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: @@ -940,7 +940,7 @@ 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]) @@ -985,7 +985,7 @@ lemma exists_bigger_env: "\p \ formula; 0 \ A; env \ list(A)\ - \ \env' \ list(A). arity(p) \ succ(length(env')) & + \ \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] diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Internalize.thy --- a/src/ZF/Constructible/Internalize.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Internalize.thy Tue Sep 27 17:03:23 2022 +0100 @@ -8,7 +8,7 @@ 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))))" @@ -38,7 +38,7 @@ 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))))" @@ -68,7 +68,7 @@ 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 @@ -98,7 +98,7 @@ 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) \ @@ -163,8 +163,8 @@ subsubsection\The Formula \<^term>\is_hd\, Internalized\ (* "is_hd(M,xs,H) \ - (is_Nil(M,xs) \ empty(M,H)) & - (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | H=x) & + (is_Nil(M,xs) \ empty(M,H)) \ + (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | H=x) \ (is_quasilist(M,xs) | empty(M,H))" *) definition hd_fm :: "[i,i]=>i" where @@ -200,8 +200,8 @@ subsubsection\The Formula \<^term>\is_tl\, Internalized\ (* "is_tl(M,xs,T) \ - (is_Nil(M,xs) \ T=xs) & - (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | T=l) & + (is_Nil(M,xs) \ T=xs) \ + (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | T=l) \ (is_quasilist(M,xs) | empty(M,T))" *) definition tl_fm :: "[i,i]=>i" where @@ -237,7 +237,7 @@ 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 @@ -282,7 +282,7 @@ (* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" "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))" *) + (\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) \ @@ -322,7 +322,7 @@ subsubsection\The Operator \<^term>\is_Member\, Internalized\ (* "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)" *) + \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) \ @@ -355,7 +355,7 @@ subsubsection\The Operator \<^term>\is_Equal\, Internalized\ (* "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)" *) + \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) \ @@ -388,7 +388,7 @@ subsubsection\The Operator \<^term>\is_Nand\, Internalized\ (* "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)" *) + \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) \ @@ -420,7 +420,7 @@ 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) \ @@ -452,8 +452,8 @@ 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) \ @@ -486,8 +486,8 @@ 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 @@ -522,8 +522,8 @@ 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) \ @@ -568,10 +568,10 @@ "M_is_recfun(M,MH,r,a,f) \ (\z[M]. z \ f \ (\x[M]. \f_r_sx[M]. \y[M]. - MH(x, f_r_sx, y) & pair(M,x,y,z) & + MH(x, f_r_sx, y) \ pair(M,x,y,z) \ (\xa[M]. \sx[M]. \r_sx[M]. - pair(M,x,a,xa) & upair(M,x,x,sx) & - pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) & + pair(M,x,a,xa) \ upair(M,x,x,sx) \ + pre_image(M,r,sx,r_sx) \ restriction(M,f,r_sx,f_r_sx) \ xa \ r)))" apply (simp add: M_is_recfun_def) apply (rule rall_cong, blast) @@ -583,10 +583,10 @@ \z[M]. z \ f \ 2 1 0 new def (\x[M]. \f_r_sx[M]. \y[M]. - MH(x, f_r_sx, y) & pair(M,x,y,z) & + MH(x, f_r_sx, y) \ pair(M,x,y,z) \ (\xa[M]. \sx[M]. \r_sx[M]. - pair(M,x,a,xa) & upair(M,x,x,sx) & - pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) & + pair(M,x,a,xa) \ upair(M,x,x,sx) \ + pre_image(M,r,sx,r_sx) \ restriction(M,f,r_sx,f_r_sx) \ xa \ r)" *) @@ -653,7 +653,7 @@ (* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o" "is_wfrec(M,MH,r,a,z) \ - \f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,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) \ @@ -716,7 +716,7 @@ definition cartprod_fm :: "[i,i,i]=>i" where (* "cartprod(M,A,B,z) \ - \u[M]. u \ z \ (\x[M]. x\A & (\y[M]. y\B & pair(M,x,y,u)))" *) + \u[M]. u \ z \ (\x[M]. x\A \ (\y[M]. y\B \ pair(M,x,y,u)))" *) "cartprod_fm(A,B,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), @@ -752,8 +752,8 @@ (* "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)" *) + 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) \ @@ -824,8 +824,8 @@ (* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" "is_nat_case(M, a, is_b, k, z) \ - (empty(M,k) \ z=a) & - (\m[M]. successor(M,m,k) \ is_b(m,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 @@ -882,7 +882,7 @@ (* iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" "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), + 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 @@ -947,7 +947,7 @@ \<^term>\p\ is enclosed by 9 (??) quantifiers.\ (* "is_iterates(M,isF,v,n,Z) \ - \sn[M]. \msn[M]. successor(M,n,sn) & membership(M,sn,msn) & + \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 @@ -1057,7 +1057,7 @@ (* mem_eclose(M,A,l) \ \n[M]. \eclosen[M]. - finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \ eclosen *) + 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) \ @@ -1125,7 +1125,7 @@ list_functor_fm :: "[i,i,i]=>i" where (* "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)" *) + number1(M,n1) \ cartprod(M,A,X,AX) \ is_sum(M,n1,AX,Z)" *) "list_functor_fm(A,X,Z) \ Exists(Exists( And(number1_fm(1), @@ -1159,7 +1159,7 @@ subsubsection\The Formula \<^term>\is_list_N\, Internalized\ (* "is_list_N(M,A,n,Z) \ - \zero[M]. empty(M,zero) & + \zero[M]. empty(M,zero) \ is_iterates(M, is_list_functor(M,A), zero, n, Z)" *) definition @@ -1202,7 +1202,7 @@ (* mem_list(M,A,l) \ \n[M]. \listn[M]. - finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \ listn *) + 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) \ @@ -1270,9 +1270,9 @@ (* "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) & + 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) \ Exists(Exists(Exists(Exists(Exists( @@ -1310,7 +1310,7 @@ subsubsection\The Formula \<^term>\is_formula_N\, Internalized\ (* "is_formula_N(M,n,Z) \ - \zero[M]. empty(M,zero) & + \zero[M]. empty(M,zero) \ is_iterates(M, is_formula_functor(M), zero, n, Z)" *) definition formula_N_fm :: "[i,i]=>i" where @@ -1352,7 +1352,7 @@ (* mem_formula(M,p) \ \n[M]. \formn[M]. - finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \ formn *) + finite_ordinal(M,n) \ is_formula_N(M,n,formn) \ p \ formn *) definition mem_formula_fm :: "i=>i" where "mem_formula_fm(x) \ @@ -1423,7 +1423,7 @@ "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) & + 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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/L_axioms.thy Tue Sep 27 17:03:23 2022 +0100 @@ -53,8 +53,8 @@ lemma LReplace_in_Lset: "\X \ Lset(i); univalent(L,X,Q); Ord(i)\ - \ \j. Ord(j) & Replace(X, %x y. Q(x,y) & L(y)) \ Lset(j)" -apply (rule_tac x="\y \ Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))" + \ \j. Ord(j) \ Replace(X, %x y. Q(x,y) \ L(y)) \ Lset(j)" +apply (rule_tac x="\y \ Replace(X, %x y. Q(x,y) \ L(y)). succ(lrank(y))" in exI) apply simp apply clarify @@ -65,7 +65,7 @@ lemma LReplace_in_L: "\L(X); univalent(L,X,Q)\ - \ \Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \ Y" + \ \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,7 +79,7 @@ done lemma strong_replacementI [rule_format]: - "\\B[L]. separation(L, %u. \x[L]. x\B & P(x,u))\ + "\\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) @@ -142,7 +142,7 @@ 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))))" @@ -258,7 +258,7 @@ lemma ReflectsD: "\REFLECTS[P,Q]; Ord(i)\ - \ \j. ix \ Lset(j). P(x) \ Q(j,x))" + \ \j. i (\x \ Lset(j). P(x) \ Q(j,x))" apply (unfold L_Reflects_def Closed_Unbounded_def) apply (elim meta_exE, clarify) apply (blast dest!: UnboundedD) @@ -512,7 +512,7 @@ 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))))" @@ -542,7 +542,7 @@ 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) \ @@ -696,7 +696,7 @@ subsubsection\Domain of a Relation, Internalized\ (* "is_domain(M,r,z) \ - \x[M]. (x \ z \ (\w[M]. w\r & (\y[M]. pair(M,x,y,w))))" *) + \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) \ @@ -731,7 +731,7 @@ subsubsection\Range of a Relation, Internalized\ (* "is_range(M,r,z) \ - \y[M]. (y \ z \ (\w[M]. w\r & (\x[M]. pair(M,x,y,w))))" *) + \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) \ @@ -766,8 +766,8 @@ subsubsection\Field of a Relation, Internalized\ (* "is_field(M,r,z) \ - \dr[M]. is_domain(M,r,dr) & - (\rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *) + \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) \ @@ -803,7 +803,7 @@ subsubsection\Image under a Relation, Internalized\ (* "image(M,r,A,z) \ - \y[M]. (y \ z \ (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w))))" *) + \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) \ @@ -839,7 +839,7 @@ subsubsection\Pre-Image under a Relation, Internalized\ (* "pre_image(M,r,A,z) \ - \x[M]. x \ z \ (\w[M]. w\r & (\y[M]. y\A & pair(M,x,y,w)))" *) + \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) \ @@ -876,7 +876,7 @@ (* "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))" *) + 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) \ @@ -981,7 +981,7 @@ subsubsection\Typed Functions, Internalized\ (* "typed_function(M,A,B,r) \ - is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) & + 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 @@ -1042,8 +1042,8 @@ (* "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)" *) + 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) \ @@ -1081,7 +1081,7 @@ subsubsection\Injections, Internalized\ (* "injection(M,A,B,f) \ - typed_function(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 @@ -1123,8 +1123,8 @@ (* surjection :: "[i=>o,i,i,i] => o" "surjection(M,A,B,f) \ - typed_function(M,A,B,f) & - (\y[M]. y\B \ (\x[M]. x\A & fun_apply(M,f,x,y)))" *) + 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) \ @@ -1161,7 +1161,7 @@ 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))" @@ -1194,7 +1194,7 @@ (* "restriction(M,r,A,z) \ - \x[M]. x \ z \ (x \ r & (\u[M]. u\A & (\v[M]. pair(M,u,v,x))))" *) + \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) \ @@ -1230,7 +1230,7 @@ (* order_isomorphism :: "[i=>o,i,i,i,i,i] => o" "order_isomorphism(M,A,r,B,s,f) \ - bijection(M,A,B,f) & + bijection(M,A,B,f) \ (\x[M]. x\A \ (\y[M]. y\A \ (\p[M]. \fx[M]. \fy[M]. \q[M]. pair(M,x,y,p) \ fun_apply(M,f,x,fx) \ fun_apply(M,f,y,fy) \ @@ -1282,8 +1282,8 @@ text\A limit ordinal is a non-empty, successor-closed ordinal\ (* "limit_ordinal(M,a) \ - ordinal(M,a) & \ empty(M,a) & - (\x[M]. x\a \ (\y[M]. y\a & successor(M,x,y)))" *) + 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 @@ -1320,7 +1320,7 @@ subsubsection\Finite Ordinals: The Predicate ``Is A Natural Number''\ (* "finite_ordinal(M,a) \ - ordinal(M,a) & \ limit_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 @@ -1355,7 +1355,7 @@ 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) \ diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Rank.thy --- a/src/ZF/Constructible/Rank.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Rank.thy Tue Sep 27 17:03:23 2022 +0100 @@ -13,25 +13,25 @@ assumes well_ord_iso_separation: "\M(A); M(f); M(r)\ \ separation (M, \x. x\A \ (\y[M]. (\p[M]. - fun_apply(M,f,x,y) & pair(M,y,x,p) & p \ r)))" + fun_apply(M,f,x,y) \ pair(M,y,x,p) \ p \ r)))" and obase_separation: \ \part of the order type formalization\ "\M(A); M(r)\ \ separation(M, \a. \x[M]. \g[M]. \mx[M]. \par[M]. - ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) & + ordinal(M,x) \ membership(M,x,mx) \ pred_set(M,A,a,r,par) \ order_isomorphism(M,par,r,x,mx,g))" and obase_equals_separation: "\M(A); M(r)\ \ separation (M, \x. x\A \ \(\y[M]. \g[M]. - ordinal(M,y) & (\my[M]. \pxr[M]. - membership(M,y,my) & pred_set(M,A,x,r,pxr) & + ordinal(M,y) \ (\my[M]. \pxr[M]. + membership(M,y,my) \ pred_set(M,A,x,r,pxr) \ order_isomorphism(M,pxr,r,y,my,g))))" and omap_replacement: "\M(A); M(r)\ \ strong_replacement(M, \a z. \x[M]. \g[M]. \mx[M]. \par[M]. - ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) & - pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))" + ordinal(M,x) \ pair(M,a,x,z) \ membership(M,x,mx) \ + pred_set(M,A,a,r,par) \ order_isomorphism(M,par,r,x,mx,g))" text\Inductive argument for Kunen's Lemma I 6.1, etc. @@ -138,7 +138,7 @@ 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 @@ -146,12 +146,12 @@ \ \the function that maps wosets to order types\ "omap(M,A,r,f) \ \z[M]. - z \ f \ (\a\A. \x[M]. \g[M]. z = & Ord(x) & + 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 @@ -160,7 +160,7 @@ lemma (in M_ordertype) omap_iff: "\omap(M,A,r,f); M(A); M(f)\ \ z \ f \ - (\a\A. \x[M]. \g[M]. z = & Ord(x) & + (\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) apply (rule iffI) @@ -182,7 +182,7 @@ lemma (in M_ordertype) otype_iff: "\otype(M,A,r,i); M(A); M(r); M(i)\ \ x \ i \ - (M(x) & Ord(x) & + (M(x) \ Ord(x) \ (\a\A. \g[M]. g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x))))" apply (auto simp add: omap_iff otype_def) apply (blast intro: transM) @@ -372,7 +372,7 @@ 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)))" + \ \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) @@ -421,57 +421,57 @@ (\sj msj. M(sj) \ M(msj) \ successor(M,j,sj) \ membership(M,sj,msj) \ M_is_recfun(M, - %x g y. \gx[M]. image(M,g,x,gx) & union(M,i,gx,y), + %x g y. \gx[M]. image(M,g,x,gx) \ union(M,i,gx,y), msj, x, f))" definition is_oadd :: "[i=>o,i,i,i] => o" where "is_oadd(M,i,j,k) \ - (\ ordinal(M,i) & \ ordinal(M,j) & k=0) | - (\ ordinal(M,i) & ordinal(M,j) & k=j) | - (ordinal(M,i) & \ ordinal(M,j) & k=i) | - (ordinal(M,i) & ordinal(M,j) & - (\f fj sj. M(f) & M(fj) & M(sj) & - successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) & - fun_apply(M,f,j,fj) & fj = k))" + (\ ordinal(M,i) \ \ ordinal(M,j) \ k=0) | + (\ ordinal(M,i) \ ordinal(M,j) \ k=j) | + (ordinal(M,i) \ \ ordinal(M,j) \ k=i) | + (ordinal(M,i) \ ordinal(M,j) \ + (\f fj sj. M(f) \ M(fj) \ M(sj) \ + successor(M,j,sj) \ is_oadd_fun(M,i,sj,sj,f) \ + fun_apply(M,f,j,fj) \ fj = k))" definition (*NEEDS RELATIVIZATION*) omult_eqns :: "[i,i,i,i] => o" where "omult_eqns(i,x,g,z) \ - Ord(x) & - (x=0 \ z=0) & - (\j. x = succ(j) \ z = g`j ++ i) & + Ord(x) \ + (x=0 \ z=0) \ + (\j. x = succ(j) \ z = g`j ++ i) \ (Limit(x) \ z = \(g``x))" definition is_omult_fun :: "[i=>o,i,i,i] => o" where "is_omult_fun(M,i,j,f) \ - (\df. M(df) & is_function(M,f) & - is_domain(M,f,df) & subset(M, j, df)) & + (\df. M(df) \ is_function(M,f) \ + is_domain(M,f,df) \ subset(M, j, df)) \ (\x\j. omult_eqns(i,x,f,f`x))" definition is_omult :: "[i=>o,i,i,i] => o" where "is_omult(M,i,j,k) \ - \f fj sj. M(f) & M(fj) & M(sj) & - successor(M,j,sj) & is_omult_fun(M,i,sj,f) & - fun_apply(M,f,j,fj) & fj = k" + \f fj sj. M(f) \ M(fj) \ M(sj) \ + successor(M,j,sj) \ is_omult_fun(M,i,sj,f) \ + fun_apply(M,f,j,fj) \ fj = k" locale M_ord_arith = M_ordertype + assumes oadd_strong_replacement: "\M(i); M(j)\ \ strong_replacement(M, - \x z. \y[M]. pair(M,x,y,z) & - (\f[M]. \fx[M]. is_oadd_fun(M,i,j,x,f) & - image(M,f,x,fx) & y = i \ fx))" + \x z. \y[M]. pair(M,x,y,z) \ + (\f[M]. \fx[M]. is_oadd_fun(M,i,j,x,f) \ + image(M,f,x,fx) \ y = i \ fx))" and omult_strong_replacement': "\M(i); M(j)\ \ strong_replacement(M, - \x z. \y[M]. z = & - (\g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & + \x z. \y[M]. z = \ + (\g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) \ y = (THE z. omult_eqns(i, x, g, z))))" @@ -480,7 +480,7 @@ lemma (in M_ord_arith) is_oadd_fun_iff: "\a\j; M(i); M(j); M(a); M(f)\ \ is_oadd_fun(M,i,j,a,f) \ - f \ a \ range(f) & (\x. M(x) \ x < a \ f`x = i \ f``x)" + f \ a \ range(f) \ (\x. M(x) \ x < a \ f`x = i \ f``x)" apply (frule lt_Ord) apply (simp add: is_oadd_fun_def relation2_def is_recfun_abs [of "%x g. i \ g``x"] @@ -494,8 +494,8 @@ lemma (in M_ord_arith) oadd_strong_replacement': "\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) & + \x z. \y[M]. z = \ + (\g[M]. is_recfun(Memrel(succ(j)),x,%x g. i \ g``x,g) \ y = i \ g``x))" apply (insert oadd_strong_replacement [of i j]) apply (simp add: is_oadd_fun_def relation2_def @@ -547,7 +547,7 @@ lemma (in M_ord_arith) oadd_abs: "\M(i); M(j); M(k)\ \ is_oadd(M,i,j,k) \ k = i++j" -apply (case_tac "Ord(i) & Ord(j)") +apply (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 @@ -575,7 +575,7 @@ lemma the_omult_eqns_0: "(THE z. omult_eqns(i,0,g,z)) = 0" by (simp add: omult_eqns_0) -lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) \ Ord(j) & z = g`j ++ i" +lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) \ Ord(j) \ z = g`j ++ i" by (simp add: omult_eqns_def) lemma the_omult_eqns_succ: @@ -683,8 +683,8 @@ "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) & + (\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) \ @@ -710,7 +710,7 @@ lemma (in M_wfrank) wfrank_strong_replacement': "M(r) \ strong_replacement(M, \x z. \y[M]. \f[M]. - pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) & + pair(M,x,y,z) \ is_recfun(r^+, x, %x f. range(f), f) \ y = range(f))" apply (insert wfrank_strong_replacement [of r]) apply (simp add: relation2_def is_recfun_abs [of "%x. range"]) @@ -730,7 +730,7 @@ wellfoundedrank :: "[i=>o,i,i] => i" where "wellfoundedrank(M,r,A) \ {p. x\A, \y[M]. \f[M]. - p = & is_recfun(r^+, x, %x f. range(f), f) & + p = \ is_recfun(r^+, x, %x f. range(f), f) \ y = range(f)}" lemma (in M_wfrank) exists_wfrank: @@ -803,7 +803,7 @@ apply (simp add: Transset_def, clarify, simp) apply (rename_tac x i f u) apply (frule is_recfun_imp_in_r, assumption) -apply (subgoal_tac "M(u) & M(i) & M(x)") +apply (subgoal_tac "M(u) \ M(i) \ M(x)") prefer 2 apply (blast dest: transM, clarify) apply (rule_tac a=u in rangeI) apply (rule_tac x=u in ReplaceI) @@ -886,7 +886,7 @@ 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") +apply (subgoal_tac "a\A \ b\A") prefer 2 apply blast apply (simp add: lt_def Ord_wellfoundedrank, clarify) apply (frule exists_wfrank [of concl: _ b], erule (1) transM, assumption) @@ -909,7 +909,7 @@ 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))" + \ \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) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Rank_Separation.thy --- a/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 17:03:23 2022 +0100 @@ -18,15 +18,15 @@ lemma well_ord_iso_Reflects: "REFLECTS[\x. x\A \ - (\y[L]. \p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \ r), + (\y[L]. \p[L]. fun_apply(L,f,x,y) \ pair(L,y,x,p) \ p \ r), \i x. x\A \ (\y \ Lset(i). \p \ Lset(i). - fun_apply(##Lset(i),f,x,y) & pair(##Lset(i),y,x,p) & p \ r)]" + fun_apply(##Lset(i),f,x,y) \ pair(##Lset(i),y,x,p) \ p \ r)]" 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]. - fun_apply(L,f,x,y) & pair(L,y,x,p) & p \ r)))" + 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) apply (rule_tac env="[A,f,r]" in DPow_LsetI) @@ -38,10 +38,10 @@ lemma obase_reflects: "REFLECTS[\a. \x[L]. \g[L]. \mx[L]. \par[L]. - ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) & + ordinal(L,x) \ membership(L,x,mx) \ pred_set(L,A,a,r,par) \ order_isomorphism(L,par,r,x,mx,g), \i a. \x \ Lset(i). \g \ Lset(i). \mx \ Lset(i). \par \ Lset(i). - ordinal(##Lset(i),x) & membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) & + ordinal(##Lset(i),x) \ membership(##Lset(i),x,mx) \ pred_set(##Lset(i),A,a,r,par) \ order_isomorphism(##Lset(i),par,r,x,mx,g)]" by (intro FOL_reflections function_reflections fun_plus_reflections) @@ -49,7 +49,7 @@ \ \part of the order type formalization\ "\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) & + 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) apply (rule_tac env="[A,r]" in DPow_LsetI) @@ -61,20 +61,20 @@ lemma obase_equals_reflects: "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) & + 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). - ordinal(##Lset(i),y) & (\my \ Lset(i). \pxr \ Lset(i). - membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) & + 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]. - ordinal(L,y) & (\my[L]. \pxr[L]. - membership(L,y,my) & pred_set(L,A,x,r,pxr) & + 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))))" apply (rule gen_separation_multi [OF obase_equals_reflects, of "{A,r}"], auto) apply (rule_tac env="[A,r]" in DPow_LsetI) @@ -85,13 +85,13 @@ subsubsection\Replacement for \<^term>\omap\\ lemma omap_reflects: - "REFLECTS[\z. \a[L]. a\B & (\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)), - \i z. \a \ Lset(i). a\B & (\x \ Lset(i). \g \ Lset(i). \mx \ Lset(i). + "REFLECTS[\z. \a[L]. a\B \ (\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)), + \i z. \a \ Lset(i). a\B \ (\x \ Lset(i). \g \ Lset(i). \mx \ Lset(i). \par \ Lset(i). - ordinal(##Lset(i),x) & pair(##Lset(i),a,x,z) & - membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) & + ordinal(##Lset(i),x) \ pair(##Lset(i),a,x,z) \ + membership(##Lset(i),x,mx) \ pred_set(##Lset(i),A,a,r,par) \ order_isomorphism(##Lset(i),par,r,x,mx,g))]" by (intro FOL_reflections function_reflections fun_plus_reflections) @@ -99,8 +99,8 @@ "\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))" + 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))" apply (rule strong_replacementI) apply (rule_tac u="{A,r,B}" in gen_separation_multi [OF omap_reflects], auto) apply (rule_tac env="[A,B,r]" in DPow_LsetI) @@ -153,15 +153,15 @@ subsubsection\Replacement for \<^term>\wfrank\\ lemma wfrank_replacement_Reflects: - "REFLECTS[\z. \x[L]. x \ A & + "REFLECTS[\z. \x[L]. x \ A \ (\rplus[L]. tran_closure(L,r,rplus) \ - (\y[L]. \f[L]. pair(L,x,y,z) & - M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) & + (\y[L]. \f[L]. pair(L,x,y,z) \ + M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) \ is_range(L,f,y))), - \i z. \x \ Lset(i). x \ A & + \i z. \x \ Lset(i). x \ A \ (\rplus \ Lset(i). tran_closure(##Lset(i),r,rplus) \ - (\y \ Lset(i). \f \ Lset(i). pair(##Lset(i),x,y,z) & - M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) & + (\y \ Lset(i). \f \ Lset(i). pair(##Lset(i),x,y,z) \ + M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) \ is_range(##Lset(i),f,y)))]" by (intro FOL_reflections function_reflections fun_plus_reflections is_recfun_reflection tran_closure_reflection) @@ -170,8 +170,8 @@ "L(r) \ strong_replacement(L, \x z. \rplus[L]. tran_closure(L,r,rplus) \ - (\y[L]. \f[L]. pair(L,x,y,z) & - M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) & + (\y[L]. \f[L]. pair(L,x,y,z) \ + M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) \ is_range(L,f,y)))" apply (rule strong_replacementI) apply (rule_tac u="{r,B}" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 17:03:23 2022 +0100 @@ -20,14 +20,14 @@ (* "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) & - (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) & empty(M,zero) & - fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & + omega(M,nnat) \ n\nnat \ successor(M,n,n') \ + (\f[M]. typed_function(M,n',A,f) \ + (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) \ empty(M,zero) \ + fun_apply(M,f,zero,x) \ fun_apply(M,f,n,y)) \ (\j[M]. j\n \ (\fj[M]. \sj[M]. \fsj[M]. \ffp[M]. - fun_apply(M,f,j,fj) & successor(M,j,sj) & - fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \ r)))"*) + fun_apply(M,f,j,fj) \ successor(M,j,sj) \ + 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) \ @@ -120,7 +120,7 @@ subsubsection\Transitive Closure of a Relation, Internalized\ (* "tran_closure(M,r,t) \ - \s[M]. rtran_closure(M,r,s) & composition(M,r,s,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) \ @@ -155,9 +155,9 @@ lemma wellfounded_trancl_reflects: "REFLECTS[\x. \w[L]. \wx[L]. \rp[L]. - w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp, + w \ Z \ pair(L,w,x,wx) \ tran_closure(L,r,rp) \ wx \ rp, \i x. \w \ Lset(i). \wx \ Lset(i). \rp \ Lset(i). - w \ Z & pair(##Lset(i),w,x,wx) & tran_closure(##Lset(i),r,rp) & + w \ Z \ pair(##Lset(i),w,x,wx) \ tran_closure(##Lset(i),r,rp) \ wx \ rp]" by (intro FOL_reflections function_reflections fun_plus_reflections tran_closure_reflection) @@ -166,7 +166,7 @@ "\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)" + w \ Z \ pair(L,w,x,wx) \ tran_closure(L,r,rp) \ wx \ rp)" apply (rule gen_separation_multi [OF wellfounded_trancl_reflects, of "{r,Z}"], auto) apply (rule_tac env="[r,Z]" in DPow_LsetI) @@ -218,16 +218,16 @@ lemma list_replacement2_Reflects: "REFLECTS - [\x. \u[L]. u \ B & u \ nat & + [\x. \u[L]. u \ B \ u \ nat \ is_iterates(L, is_list_functor(L, A), 0, u, x), - \i x. \u \ Lset(i). u \ B & u \ nat & + \i x. \u \ Lset(i). u \ B \ u \ nat \ is_iterates(##Lset(i), is_list_functor(##Lset(i), A), 0, u, x)]" by (intro FOL_reflections is_iterates_reflection list_functor_reflection) lemma list_replacement2: "L(A) \ strong_replacement(L, - \n y. n\nat & is_iterates(L, is_list_functor(L,A), 0, n, y))" + \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}" in gen_separation_multi [OF list_replacement2_Reflects], @@ -245,9 +245,9 @@ need to expand iterates_replacement and wfrec_replacement*) lemma formula_replacement1_Reflects: "REFLECTS - [\x. \u[L]. u \ B & (\y[L]. pair(L,u,y,x) & + [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)), - \i x. \u \ Lset(i). u \ B & (\y \ Lset(i). pair(##Lset(i), u, y, x) & + \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(##Lset(i), u, y, x) \ is_wfrec(##Lset(i), iterates_MH(##Lset(i), is_formula_functor(##Lset(i)), 0), memsn, u, y))]" @@ -268,16 +268,16 @@ lemma formula_replacement2_Reflects: "REFLECTS - [\x. \u[L]. u \ B & u \ nat & + [\x. \u[L]. u \ B \ u \ nat \ is_iterates(L, is_formula_functor(L), 0, u, x), - \i x. \u \ Lset(i). u \ B & u \ nat & + \i x. \u \ Lset(i). u \ B \ u \ nat \ is_iterates(##Lset(i), is_formula_functor(##Lset(i)), 0, u, x)]" by (intro FOL_reflections is_iterates_reflection formula_functor_reflection) lemma formula_replacement2: "strong_replacement(L, - \n y. n\nat & is_iterates(L, is_formula_functor(L), 0, n, y))" + \n y. n\nat \ is_iterates(L, is_formula_functor(L), 0, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{B,0,nat}" in gen_separation_multi [OF formula_replacement2_Reflects], @@ -293,7 +293,7 @@ subsubsection\The Formula \<^term>\is_nth\, Internalized\ (* "is_nth(M,n,l,Z) \ - \X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,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) \ @@ -333,9 +333,9 @@ need to expand iterates_replacement and wfrec_replacement*) lemma nth_replacement_Reflects: "REFLECTS - [\x. \u[L]. u \ B & (\y[L]. pair(L,u,y,x) & + [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)), - \i x. \u \ Lset(i). u \ B & (\y \ Lset(i). pair(##Lset(i), u, y, x) & + \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(##Lset(i), u, y, x) \ is_wfrec(##Lset(i), iterates_MH(##Lset(i), is_tl(##Lset(i)), z), memsn, u, y))]" @@ -380,9 +380,9 @@ lemma eclose_replacement1_Reflects: "REFLECTS - [\x. \u[L]. u \ B & (\y[L]. pair(L,u,y,x) & + [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)), - \i x. \u \ Lset(i). u \ B & (\y \ Lset(i). pair(##Lset(i), u, y, x) & + \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(##Lset(i), u, y, x) \ is_wfrec(##Lset(i), iterates_MH(##Lset(i), big_union(##Lset(i)), A), memsn, u, y))]" @@ -403,15 +403,15 @@ lemma eclose_replacement2_Reflects: "REFLECTS - [\x. \u[L]. u \ B & u \ nat & + [\x. \u[L]. u \ B \ u \ nat \ is_iterates(L, big_union(L), A, u, x), - \i x. \u \ Lset(i). u \ B & u \ nat & + \i x. \u \ Lset(i). u \ B \ u \ nat \ is_iterates(##Lset(i), big_union(##Lset(i)), A, u, x)]" by (intro FOL_reflections function_reflections is_iterates_reflection) lemma eclose_replacement2: "L(A) \ strong_replacement(L, - \n y. n\nat & is_iterates(L, big_union(L), A, n, y))" + \n y. n\nat \ is_iterates(L, big_union(L), A, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{A,B,nat}" in gen_separation_multi [OF eclose_replacement2_Reflects], diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Reflection.thy Tue Sep 27 17:03:23 2022 +0100 @@ -35,16 +35,16 @@ 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) & + 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 "ClEx(P,a) \ Limit(a) \ normalize(FF(P),a) = a" begin @@ -54,7 +54,7 @@ text\Awkward: we need a version of \ClEx_def\ as an equality at the level of classes, which do not really exist\ lemma ClEx_eq: - "ClEx(P) \ \a. Limit(a) & normalize(FF(P),a) = a" + "ClEx(P) \ \a. Limit(a) \ normalize(FF(P),a) = a" by (simp add: ClEx_def [symmetric]) @@ -70,26 +70,26 @@ 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))" + \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), + \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem Or_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ - \ Reflects(\a. Cl(a) & C'(a), \x. P(x) | P'(x), - \a x. Q(a,x) | Q'(a,x))" + \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), + \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem Imp_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ - \ Reflects(\a. Cl(a) & C'(a), + \ Reflects(\a. Cl(a) \ C'(a), \x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x))" by (simp add: Reflects_def Closed_Unbounded_Int, blast) theorem Iff_reflection [intro]: "\Reflects(Cl,P,Q); Reflects(C',P',Q')\ - \ Reflects(\a. Cl(a) & C'(a), + \ 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) @@ -97,33 +97,32 @@ subsection\Reflection for Existential Quantifiers\ 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 + "\y\Mset(a); Ord(a); M(z); P()\ \ \z\Mset(F0(P,y)). P()" + unfolding F0_def M_def + apply clarify + apply (rule LeastI2) + apply (blast intro: Mset_mono [THEN subsetD]) + apply (blast intro: lt_Ord2, blast) + done lemma Ord_F0 [intro,simp]: "Ord(F0(P,y))" -by (simp add: F0_def) + by (simp add: F0_def) lemma Ord_FF [intro,simp]: "Ord(FF(P,y))" -by (simp add: FF_def) + by (simp add: FF_def) lemma cont_Ord_FF: "cont_Ord(FF(P))" -apply (insert Mset_cont) -apply (simp add: cont_Ord_def FF_def, blast) -done + using Mset_cont by (simp add: cont_Ord_def FF_def, blast) text\Recall that \<^term>\F0\ depends upon \<^term>\y\Mset(a)\, while \<^term>\FF\ depends only upon \<^term>\a\.\ lemma FF_works: - "\M(z); y\Mset(a); P(); 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 + "\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 FFN_works: "\M(z); y\Mset(a); P(); Ord(a)\ @@ -156,7 +155,7 @@ lemma ClEx_upward: "\z\Mset(a); y\Mset(a); Q(a,); Cl(a); ClEx(P,a)\ - \ \z. M(z) & P()" + \ \z. M(z) \ P()" apply (simp add: ClEx_def M_def) apply (blast dest: Cl_reflects intro: Limit_is_Ord Pair_in_Mset) @@ -165,7 +164,7 @@ text\Class \ClEx\ indeed consists of reflecting ordinals...\ lemma ZF_ClEx_iff: "\y\Mset(a); Cl(a); ClEx(P,a)\ - \ (\z. M(z) & P()) \ (\z\Mset(a). Q(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\ @@ -187,7 +186,7 @@ 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,))" + \ (\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, @@ -215,8 +214,8 @@ lemma Ex_reflection_0: "Reflects(Cl,P0,Q0) - \ Reflects(\a. Cl(a) & ClEx(P0,a), - \x. \z. M(z) & P0(), + \ Reflects(\a. Cl(a) \ ClEx(P0,a), + \x. \z. M(z) \ P0(), \a x. \z\Mset(a). Q0(a,))" apply (simp add: Reflects_def) apply (intro conjI Closed_Unbounded_Int) @@ -227,7 +226,7 @@ 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) @@ -237,15 +236,15 @@ theorem Ex_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) - \ Reflects(\a. Cl(a) & ClEx(\x. P(fst(x),snd(x)), a), - \x. \z. M(z) & P(x,z), + \ Reflects(\a. Cl(a) \ ClEx(\x. P(fst(x),snd(x)), a), + \x. \z. M(z) \ P(x,z), \a x. \z\Mset(a). Q(a,x,z))" by (rule Ex_reflection_0 [of _ " \x. P(fst(x),snd(x))" "\a x. Q(a,fst(x),snd(x))", simplified]) theorem All_reflection [intro]: "Reflects(Cl, \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))) - \ Reflects(\a. Cl(a) & ClEx(\x. \P(fst(x),snd(x)), a), + \ 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))" @@ -255,14 +254,14 @@ 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 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) @@ -278,7 +277,7 @@ proof state.\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & x \ y, + \x. \y. M(y) \ x \ y, \a x. \y\Mset(a). x \ y)" by fast @@ -286,8 +285,8 @@ in the class of reflecting ordinals. The \<^term>\Ord(a)\ is redundant, though harmless.\ lemma - "Reflects(\a. Ord(a) & ClEx(\x. fst(x) \ snd(x), a), - \x. \y. M(y) & x \ y, + "Reflects(\a. Ord(a) \ ClEx(\x. fst(x) \ snd(x), a), + \x. \y. M(y) \ x \ y, \a x. \y\Mset(a). x \ y)" by fast @@ -295,31 +294,31 @@ text\Example 2\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & (\z. M(z) \ z \ x \ z \ y), + \x. \y. M(y) \ (\z. M(z) \ z \ x \ z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x \ z \ y)" by fast text\Example 2'. We give the reflecting class explicitly.\ lemma "Reflects - (\a. (Ord(a) & - ClEx(\x. \ (snd(x) \ fst(fst(x)) \ snd(x) \ snd(fst(x))), a)) & + (\a. (Ord(a) \ + ClEx(\x. \ (snd(x) \ fst(fst(x)) \ snd(x) \ snd(fst(x))), a)) \ ClEx(\x. \z. M(z) \ z \ fst(x) \ z \ snd(x), a), - \x. \y. M(y) & (\z. M(z) \ z \ x \ z \ y), + \x. \y. M(y) \ (\z. M(z) \ z \ x \ z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x \ z \ y)" by fast text\Example 2''. We expand the subset relation.\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & (\z. M(z) \ (\w. M(w) \ w\z \ w\x) \ z\y), + \x. \y. M(y) \ (\z. M(z) \ (\w. M(w) \ w\z \ w\x) \ z\y), \a x. \y\Mset(a). \z\Mset(a). (\w\Mset(a). w\z \ w\x) \ z\y)" by fast text\Example 2'''. Single-step version, to reveal the reflecting class.\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & (\z. M(z) \ z \ x \ z \ y), + \x. \y. M(y) \ (\z. M(z) \ z \ x \ z \ y), \a x. \y\Mset(a). \z\Mset(a). z \ x \ z \ y)" apply (rule Ex_reflection) txt\ @@ -339,21 +338,21 @@ if \<^term>\P\ is quantifier-free, since it is not being relativized.\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & (\z. M(z) \ z \ y \ z \ x & P(z)), - \a x. \y\Mset(a). \z\Mset(a). z \ y \ z \ x & P(z))" + \x. \y. M(y) \ (\z. M(z) \ z \ y \ z \ x \ P(z)), + \a x. \y\Mset(a). \z\Mset(a). z \ y \ z \ x \ P(z))" by fast text\Example 3'\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & y = Collect(x,P), + \x. \y. M(y) \ y = Collect(x,P), \a x. \y\Mset(a). y = Collect(x,P))" by fast text\Example 3''\ schematic_goal "Reflects(?Cl, - \x. \y. M(y) & y = Replace(x,P), + \x. \y. M(y) \ y = Replace(x,P), \a x. \y\Mset(a). y = Replace(x,P))" by fast @@ -361,7 +360,7 @@ to be relativized.\ schematic_goal "Reflects(?Cl, - \A. 0\A \ (\f. M(f) & f \ (\X \ A. X)), + \A. 0\A \ (\f. M(f) \ f \ (\X \ A. X)), \a A. 0\A \ (\f\Mset(a). f \ (\X \ A. X)))" by fast diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Relative.thy Tue Sep 27 17:03:23 2022 +0100 @@ -20,21 +20,21 @@ 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) & - (\y[M]. upair(M,a,b,y) & upair(M,x,y,z))" + "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 @@ -42,15 +42,15 @@ 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 @@ -58,84 +58,84 @@ 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) \ - (A=0 \ z=0) & + (A=0 \ z=0) \ (A\0 \ (\x[M]. x \ z \ (\y[M]. y\A \ x \ y)))" definition cartprod :: "[i=>o,i,i,i] => o" where "cartprod(M,A,B,z) \ - \u[M]. u \ z \ (\x[M]. x\A & (\y[M]. y\B & pair(M,x,y,u)))" + \u[M]. u \ z \ (\x[M]. x\A \ (\y[M]. y\B \ pair(M,x,y,u)))" definition is_sum :: "[i=>o,i,i,i] => o" where "is_sum(M,A,B,Z) \ \A0[M]. \n1[M]. \s1[M]. \B1[M]. - number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) & - cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" + 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) \ \x[M]. x \ z \ - (\w[M]. w\r & (\u[M]. \v[M]. pair(M,u,v,w) & pair(M,v,u,x)))" + (\w[M]. w\r \ (\u[M]. \v[M]. pair(M,u,v,w) \ pair(M,v,u,x)))" definition pre_image :: "[i=>o,i,i,i] => o" where "pre_image(M,r,A,z) \ - \x[M]. x \ z \ (\w[M]. w\r & (\y[M]. y\A & pair(M,x,y,w)))" + \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. y\A \ pair(M,x,y,w)))" definition is_domain :: "[i=>o,i,i] => o" where "is_domain(M,r,z) \ - \x[M]. x \ z \ (\w[M]. w\r & (\y[M]. pair(M,x,y,w)))" + \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. pair(M,x,y,w)))" definition image :: "[i=>o,i,i,i] => o" where "image(M,r,A,z) \ - \y[M]. y \ z \ (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w)))" + \y[M]. y \ z \ (\w[M]. w\r \ (\x[M]. x\A \ pair(M,x,y,w)))" definition is_range :: "[i=>o,i,i] => o" where \ \the cleaner - \<^term>\\r'[M]. is_converse(M,r,r') & is_domain(M,r',z)\ + \<^term>\\r'[M]. is_converse(M,r,r') \ is_domain(M,r',z)\ unfortunately needs an instance of separation in order to prove \<^term>\M(converse(r))\.\ "is_range(M,r,z) \ - \y[M]. y \ z \ (\w[M]. w\r & (\x[M]. pair(M,x,y,w)))" + \y[M]. y \ z \ (\w[M]. w\r \ (\x[M]. pair(M,x,y,w)))" definition is_field :: "[i=>o,i,i] => o" where "is_field(M,r,z) \ - \dr[M]. \rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) & + \dr[M]. \rr[M]. is_domain(M,r,dr) \ is_range(M,r,rr) \ union(M,dr,rr,z)" definition @@ -153,12 +153,12 @@ fun_apply :: "[i=>o,i,i,i] => o" where "fun_apply(M,f,x,y) \ (\xs[M]. \fxs[M]. - upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" + upair(M,x,x,xs) \ image(M,f,xs,fxs) \ big_union(M,fxs,y))" definition typed_function :: "[i=>o,i,i,i] => o" where "typed_function(M,A,B,r) \ - is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) & + 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 @@ -171,30 +171,30 @@ "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)" + pair(M,x,z,p) \ pair(M,x,y,xy) \ pair(M,y,z,yz) \ + xy \ s \ yz \ r)" definition injection :: "[i=>o,i,i,i] => o" where "injection(M,A,B,f) \ - typed_function(M,A,B,f) & + typed_function(M,A,B,f) \ (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x',y,p') \ p\f \ p'\f \ x=x')" definition surjection :: "[i=>o,i,i,i] => o" where "surjection(M,A,B,f) \ - typed_function(M,A,B,f) & - (\y[M]. y\B \ (\x[M]. x\A & fun_apply(M,f,x,y)))" + 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) \ - \x[M]. x \ z \ (x \ r & (\u[M]. u\A & (\v[M]. pair(M,u,v,x))))" + \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 @@ -203,43 +203,43 @@ 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) & - (\x[M]. x\a \ (\y[M]. y\a & successor(M,x,y)))" + 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)" + 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) & + 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) \ - (empty(M,k) \ z=a) & - (\m[M]. successor(M,m,k) \ is_b(m,z)) & - (is_quasinat(M,k) | empty(M,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 @@ -301,7 +301,7 @@ to \M\. We do not have separation as a scheme; every instance that we need must be assumed (and later proved) separately.\ "separation(M,P) \ - \z[M]. \y[M]. \x[M]. x \ y \ x \ z & P(x)" + \z[M]. \y[M]. \x[M]. x \ y \ x \ z \ P(x)" definition upair_ax :: "(i=>o) => o" where @@ -318,24 +318,24 @@ definition univalent :: "[i=>o, i, [i,i]=>o] => o" where "univalent(M,A,P) \ - \x[M]. x\A \ (\y[M]. \z[M]. P(x,y) & P(x,z) \ y=z)" + \x[M]. x\A \ (\y[M]. \z[M]. P(x,y) \ P(x,z) \ y=z)" definition replacement :: "[i=>o, [i,i]=>o] => o" where "replacement(M,P) \ \A[M]. univalent(M,A,P) \ - (\Y[M]. \b[M]. (\x[M]. x\A & P(x,b)) \ b \ Y)" + (\Y[M]. \b[M]. (\x[M]. x\A \ P(x,b)) \ b \ Y)" definition strong_replacement :: "[i=>o, [i,i]=>o] => o" where "strong_replacement(M,P) \ \A[M]. univalent(M,A,P) \ - (\Y[M]. \b[M]. b \ Y \ (\x[M]. x\A & P(x,b)))" + (\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))" + \x[M]. (\y[M]. y\x) \ (\y[M]. y\x \ \(\z[M]. z\x \ z \ y))" subsection\A trivial consistency proof for $V_\omega$\ @@ -354,7 +354,7 @@ 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\\ @@ -373,7 +373,7 @@ 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\ @@ -487,16 +487,16 @@ lemma replacementD: "\replacement(M,P); M(A); univalent(M,A,P)\ - \ \Y[M]. (\b[M]. ((\x[M]. x\A & P(x,b)) \ b \ Y))" + \ \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))))" + \ \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) @@ -505,7 +505,7 @@ definition order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where "order_isomorphism(M,A,r,B,s,f) \ - bijection(M,A,B,f) & + bijection(M,A,B,f) \ (\x[M]. x\A \ (\y[M]. y\A \ (\p[M]. \fx[M]. \fy[M]. \q[M]. pair(M,x,y,p) \ fun_apply(M,f,x,fx) \ fun_apply(M,f,y,fy) \ @@ -514,12 +514,12 @@ definition pred_set :: "[i=>o,i,i,i,i] => o" where "pred_set(M,A,x,r,B) \ - \y[M]. y \ B \ (\p[M]. p\r & y \ A & pair(M,y,x,p))" + \y[M]. y \ B \ (\p[M]. p\r \ y \ A \ pair(M,y,x,p))" definition membership :: "[i=>o,i,i] => o" where \ \membership relation\ "membership(M,A,r) \ - \p[M]. p \ r \ (\x[M]. x\A & (\y[M]. y\A & x\y & pair(M,x,y,p)))" + \p[M]. p \ r \ (\x[M]. x\A \ (\y[M]. y\A \ x\y \ pair(M,x,y,p)))" subsection\Introducing a Transitive Class Model\ @@ -566,12 +566,12 @@ 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))) \ - (\x\A. P(x)) & (\x. P(x) \ M(x) \ x\A)" + (\x\A. P(x)) \ (\x. P(x) \ M(x) \ x\A)" by (blast intro: transM) text\Simplifies proofs of equalities when there's an iff-equality @@ -604,15 +604,15 @@ done lemma (in M_trivial) upair_in_MI [intro!]: - " M(a) & M(b) \ M({a,b})" + " M(a) \ M(b) \ M({a,b})" by (insert upair_ax, simp add: upair_ax_def) lemma (in M_trans) upair_in_MD [dest!]: - "M({a,b}) \ M(a) & M(b)" + "M({a,b}) \ M(a) \ M(b)" by (blast intro: transM) lemma (in M_trivial) upair_in_M_iff [simp]: - "M({a,b}) \ M(a) & M(b)" + "M({a,b}) \ M(a) \ M(b)" by blast lemma (in M_trivial) singleton_in_MI [intro!]: @@ -634,19 +634,19 @@ done lemma (in M_trans) pair_in_MD [dest!]: - "M() \ M(a) & M(b)" + "M() \ M(a) \ M(b)" by (simp add: Pair_def, blast intro: transM) lemma (in M_trivial) pair_in_MI [intro!]: - "M(a) & M(b) \ M()" + "M(a) \ M(b) \ M()" by (simp add: Pair_def) lemma (in M_trivial) pair_in_M_iff [simp]: - "M() \ M(a) & M(b)" + "M() \ M(a) \ M(b)" 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]: @@ -746,7 +746,7 @@ 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))" + \ u \ Replace(A,P) \ (\x. x\A \ P(x,u))" unfolding Replace_iff univalent_def by (blast dest: transM) @@ -788,13 +788,13 @@ apply (simp add: RepFun_def) done -lemma Replace_conj_eq: "{y . x \ A, x\A & y=f(x)} = {y . x\A, y=f(x)}" +lemma Replace_conj_eq: "{y . x \ A, x\A \ y=f(x)} = {y . x\A, y=f(x)}" by simp text\Better than \RepFun_closed\ when having the formula \<^term>\x\A\ makes relativization easier.\ lemma (in M_trans) RepFun_closed2: - "\strong_replacement(M, \x y. x\A & y = f(x)); M(A); \x\A. M(f(x))\ + "\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) @@ -807,7 +807,7 @@ is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where "is_lambda(M, A, is_b, z) \ \p[M]. p \ z \ - (\u[M]. \v[M]. u\A & pair(M,u,v,p) & is_b(u,v))" + (\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))\ @@ -816,7 +816,7 @@ 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)\); + "\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) @@ -951,7 +951,7 @@ 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 @@ -1006,7 +1006,7 @@ primrec "natnumber_aux(M,0) = (\x\nat. if empty(M,x) then 1 else 0)" "natnumber_aux(M,succ(n)) = - (\x\nat. if (\y[M]. natnumber_aux(M,n)`y=1 & successor(M,y,x)) + (\x\nat. if (\y[M]. natnumber_aux(M,n)`y=1 \ successor(M,y,x)) then 1 else 0)" definition @@ -1026,36 +1026,36 @@ "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)))" + \ 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)))" + \ separation(M, \y. \p[M]. p\r \ (\x[M]. x\A \ pair(M,x,y,p)))" and converse_separation: "M(r) \ separation(M, - \z. \p[M]. p\r & (\x[M]. \y[M]. pair(M,x,y,p) & pair(M,y,x,z)))" + \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]. - pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) & - xy\s & yz\r)" + 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)" + "separation(M, \z. \x[M]. \y[M]. pair(M,x,y,z) \ x \ y)" and funspace_succ_replacement: "M(n) \ strong_replacement(M, \p z. \f[M]. \b[M]. \nb[M]. \cnbf[M]. - pair(M,f,b,p) & pair(M,n,b,nb) & is_cons(M,nb,f,cnbf) & + pair(M,f,b,p) \ pair(M,n,b,nb) \ is_cons(M,nb,f,cnbf) \ upair(M,cnbf,cnbf,z))" and is_recfun_separation: \ \for well-founded recursion: used to prove \is_recfun_equal\\ "\M(r); M(f); M(g); M(a); M(b)\ \ separation(M, \x. \xa[M]. \xb[M]. - pair(M,x,a,xa) & xa \ r & pair(M,x,b,xb) & xb \ r & - (\fx[M]. \gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) & + pair(M,x,a,xa) \ xa \ r \ pair(M,x,b,xb) \ xb \ r \ + (\fx[M]. \gx[M]. fun_apply(M,f,x,fx) \ fun_apply(M,g,x,gx) \ fx \ gx))" and power_ax: "power_ax(M)" @@ -1075,7 +1075,7 @@ lemma (in M_basic) cartprod_iff: "\M(A); M(B); M(C)\ \ cartprod(M,A,B,C) \ - (\p1[M]. \p2[M]. powerset(M,A \ B,p1) & powerset(M,p1,p2) & + (\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) defer 1 @@ -1138,7 +1138,7 @@ "M(r) \ converse(r) = {z \ \(\(r)) * \(\(r)). - \p\r. \x[M]. \y[M]. p = \x,y\ & z = \y,x\}" + \p\r. \x[M]. \y[M]. p = \x,y\ \ z = \y,x\}" apply (rule equalityI) prefer 2 apply (blast dest: transM, clarify, simp) apply (simp add: Pair_def) @@ -1266,7 +1266,7 @@ "\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}" + \x[M]. \y[M]. \z[M]. xz = \x,z\ \ \x,y\ \ s \ \y,z\ \ r}" apply (simp add: comp_def) apply (rule equalityI) apply clarify @@ -1342,13 +1342,13 @@ 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]) (*???equalities*) lemma Collect_Un_Collect_eq: - "Collect(A,P) \ Collect(A,Q) = Collect(A, %x. P(x) | Q(x))" + "Collect(A,P) \ Collect(A,Q) = Collect(A, %x. P(x) \ Q(x))" by blast lemma Diff_Collect_eq: @@ -1362,7 +1362,7 @@ 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]) @@ -1405,7 +1405,7 @@ lemma (in M_basic) succ_fun_eq2: "\M(B); M(n->B)\ \ succ(n) -> B = - \{z. p \ (n->B)*B, \f[M]. \b[M]. p = & z = {cons(, f)}}" + \{z. p \ (n->B)*B, \f[M]. \b[M]. p = \ z = {cons(, f)}}" apply (simp add: succ_fun_eq) apply (blast dest: transM) done @@ -1430,22 +1430,22 @@ 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)" @@ -1487,12 +1487,12 @@ 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)" @@ -1501,7 +1501,7 @@ 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)" +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]: @@ -1511,11 +1511,11 @@ 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 @@ -1527,9 +1527,9 @@ is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where \ \Returns 0 for non-lists\ "is_list_case(M, a, is_b, xs, z) \ - (is_Nil(M,xs) \ z=a) & - (\x[M]. \l[M]. is_Cons(M,x,l,xs) \ is_b(x,l,z)) & - (is_quasilist(M,xs) | empty(M,z))" + (is_Nil(M,xs) \ z=a) \ + (\x[M]. \l[M]. is_Cons(M,x,l,xs) \ is_b(x,l,z)) \ + (is_quasilist(M,xs) \ empty(M,z))" definition hd' :: "i => i" where @@ -1546,17 +1546,17 @@ \ \\<^term>\hd([]) = 0\ no constraints if not a list. Avoiding implication prevents the simplifier's looping.\ "is_hd(M,xs,H) \ - (is_Nil(M,xs) \ empty(M,H)) & - (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | H=x) & - (is_quasilist(M,xs) | empty(M,H))" + (is_Nil(M,xs) \ empty(M,H)) \ + (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) \ H=x) \ + (is_quasilist(M,xs) \ empty(M,H))" definition is_tl :: "[i=>o,i,i] => o" where \ \\<^term>\tl([]) = []\; see comments about \<^term>\is_hd\\ "is_tl(M,xs,T) \ - (is_Nil(M,xs) \ T=xs) & - (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | T=l) & - (is_quasilist(M,xs) | empty(M,T))" + (is_Nil(M,xs) \ T=xs) \ + (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) \ T=l) \ + (is_quasilist(M,xs) \ empty(M,T))" subsubsection\\<^term>\quasilist\: For Case-Splitting with \<^term>\list_case'\\ diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Satisfies_absolute.thy --- a/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 17:03:23 2022 +0100 @@ -14,8 +14,8 @@ (* "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" *) + 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) \ @@ -60,10 +60,10 @@ (* is_formula_case :: "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" "is_formula_case(M, is_a, is_b, is_c, is_d, v, z) \ - (\x[M]. \y[M]. x\nat \ y\nat \ is_Member(M,x,y,v) \ is_a(x,y,z)) & - (\x[M]. \y[M]. x\nat \ y\nat \ is_Equal(M,x,y,v) \ is_b(x,y,z)) & + (\x[M]. \y[M]. x\nat \ y\nat \ is_Member(M,x,y,v) \ is_a(x,y,z)) \ + (\x[M]. \y[M]. x\nat \ y\nat \ is_Equal(M,x,y,v) \ is_b(x,y,z)) \ (\x[M]. \y[M]. x\formula \ y\formula \ - is_Nand(M,x,y,v) \ is_c(x,y,z)) & + is_Nand(M,x,y,v) \ is_c(x,y,z)) \ (\x[M]. x\formula \ is_Forall(M,x,v) \ is_d(x,z))" *) definition @@ -179,8 +179,8 @@ \ \Merely a useful abbreviation for the sequel.\ "is_depth_apply(M,h,p,z) \ \dp[M]. \sdp[M]. \hsdp[M]. - finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) & - fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)" + 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)\ @@ -206,7 +206,7 @@ is_lambda(M, lA, \env z. is_bool_of_o(M, \nx[M]. \ny[M]. - is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \ ny, z), + is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ nx \ ny, z), zz)" definition @@ -222,7 +222,7 @@ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. is_bool_of_o(M, - \nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z), + \nx[M]. is_nth(M,x,env,nx) \ is_nth(M,y,env,nx), z), zz)" definition @@ -234,9 +234,9 @@ "satisfies_is_c(M,A,h) \ \p q zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \hp[M]. \hq[M]. - (\rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & - (\rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & - (\pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)), + (\rp[M]. is_depth_apply(M,h,p,rp) \ fun_apply(M,rp,env,hp)) \ + (\rq[M]. is_depth_apply(M,h,q,rq) \ fun_apply(M,rq,env,hq)) \ + (\pq[M]. is_and(M,hp,hq,pq) \ is_not(M,pq,z)), zz)" definition @@ -249,7 +249,7 @@ "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) & + \env z. \rp[M]. is_depth_apply(M,h,p,rp) \ is_bool_of_o(M, \x[M]. \xenv[M]. \hp[M]. x\A \ is_Cons(M,x,env,xenv) \ @@ -294,36 +294,36 @@ "\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) & + env \ list(A) \ is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ + is_bool_of_o(M, nx \ ny, bo) \ pair(M, env, bo, z))" and Equal_replacement: "\M(A); x \ nat; y \ nat\ \ strong_replacement (M, \env z. \bo[M]. \nx[M]. \ny[M]. - env \ list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & - is_bool_of_o(M, nx = ny, bo) & + env \ list(A) \ is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ + is_bool_of_o(M, nx = ny, bo) \ pair(M, env, bo, z))" and Nand_replacement: "\M(A); M(rp); M(rq)\ \ strong_replacement (M, \env z. \rpe[M]. \rqe[M]. \andpq[M]. \notpq[M]. - fun_apply(M,rp,env,rpe) & fun_apply(M,rq,env,rqe) & - is_and(M,rpe,rqe,andpq) & is_not(M,andpq,notpq) & - env \ list(A) & pair(M, env, notpq, z))" + fun_apply(M,rp,env,rpe) \ fun_apply(M,rq,env,rqe) \ + is_and(M,rpe,rqe,andpq) \ is_not(M,andpq,notpq) \ + env \ list(A) \ pair(M, env, notpq, z))" and Forall_replacement: "\M(A); M(rp)\ \ strong_replacement (M, \env z. \bo[M]. - env \ list(A) & + env \ list(A) \ is_bool_of_o (M, \a[M]. \co[M]. \rpco[M]. a\A \ is_Cons(M,a,env,co) \ fun_apply(M,rp,co,rpco) \ number1(M, rpco), - bo) & + bo) \ pair(M,env,bo,z))" and formula_rec_replacement: @@ -334,39 +334,39 @@ \ \For the \\-abstraction\ in the \<^term>\transrec\ body\ "\M(g); M(A)\ \ strong_replacement (M, - \x y. mem_formula(M,x) & + \x y. mem_formula(M,x) \ (\c[M]. is_formula_case(M, satisfies_is_a(M,A), satisfies_is_b(M,A), satisfies_is_c(M,A,g), - satisfies_is_d(M,A,g), x, c) & + satisfies_is_d(M,A,g), x, c) \ pair(M, x, c, y)))" lemma (in M_satisfies) Member_replacement': "\M(A); x \ nat; y \ nat\ \ strong_replacement - (M, \env z. env \ list(A) & + (M, \env z. env \ list(A) \ z = \env, bool_of_o(nth(x, env) \ nth(y, env))\)" by (insert Member_replacement, simp) lemma (in M_satisfies) Equal_replacement': "\M(A); x \ nat; y \ nat\ \ strong_replacement - (M, \env z. env \ list(A) & + (M, \env z. env \ list(A) \ z = \env, bool_of_o(nth(x, env) = nth(y, env))\)" by (insert Equal_replacement, simp) lemma (in M_satisfies) Nand_replacement': "\M(A); M(rp); M(rq)\ \ strong_replacement - (M, \env z. env \ list(A) & z = \env, not(rp`env and rq`env)\)" + (M, \env z. env \ list(A) \ z = \env, not(rp`env and rq`env)\)" by (insert Nand_replacement, simp) lemma (in M_satisfies) Forall_replacement': "\M(A); M(rp)\ \ strong_replacement (M, \env z. - env \ list(A) & + env \ list(A) \ z = \env, bool_of_o (\a\A. rp ` Cons(a,env) = 1)\)" by (insert Forall_replacement, simp) @@ -447,7 +447,7 @@ lemma (in M_satisfies) fr_lam_replace: "\M(g); M(A)\ \ - strong_replacement (M, \x y. x \ formula & + strong_replacement (M, \x y. x \ formula \ y = \x, formula_rec_case(satisfies_a(A), satisfies_b(A), @@ -512,8 +512,8 @@ (* 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) *) + 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) \ @@ -555,7 +555,7 @@ is_lambda(M, lA, \env z. is_bool_of_o(M, \nx[M]. \ny[M]. - is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \ ny, z), + is_nth(M,x,env,nx) \ is_nth(M,y,env,ny) \ nx \ ny, z), zz) *) definition @@ -607,7 +607,7 @@ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. is_bool_of_o(M, - \nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z), + \nx[M]. is_nth(M,x,env,nx) \ is_nth(M,y,env,nx), z), zz) *) definition @@ -655,9 +655,9 @@ (* satisfies_is_c(M,A,h) \ \p q zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \hp[M]. \hq[M]. - (\rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & - (\rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & - (\pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)), + (\rp[M]. is_depth_apply(M,h,p,rp) \ fun_apply(M,rp,env,hp)) \ + (\rq[M]. is_depth_apply(M,h,q,rq) \ fun_apply(M,rq,env,hq)) \ + (\pq[M]. is_and(M,hp,hq,pq) \ is_not(M,pq,z)), zz) *) definition @@ -706,7 +706,7 @@ (* 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) & + \env z. \rp[M]. is_depth_apply(M,h,p,rp) \ is_bool_of_o(M, \x[M]. \xenv[M]. \hp[M]. x\A \ is_Cons(M,x,env,xenv) \ @@ -836,8 +836,8 @@ "\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) & + env \ list(A) \ is_nth(L,x,env,nx) \ is_nth(L,y,env,ny) \ + is_bool_of_o(L, nx \ ny, bo) \ pair(L, env, bo, z))" apply (rule strong_replacementI) apply (rule_tac u="{list(A),B,x,y}" @@ -866,8 +866,8 @@ "\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) & + env \ list(A) \ is_nth(L,x,env,nx) \ is_nth(L,y,env,ny) \ + is_bool_of_o(L, nx = ny, bo) \ pair(L, env, bo, z))" apply (rule strong_replacementI) apply (rule_tac u="{list(A),B,x,y}" @@ -898,9 +898,9 @@ "\L(A); L(rp); L(rq)\ \ strong_replacement (L, \env z. \rpe[L]. \rqe[L]. \andpq[L]. \notpq[L]. - fun_apply(L,rp,env,rpe) & fun_apply(L,rq,env,rqe) & - is_and(L,rpe,rqe,andpq) & is_not(L,andpq,notpq) & - env \ list(A) & pair(L, env, notpq, z))" + fun_apply(L,rp,env,rpe) \ fun_apply(L,rq,env,rqe) \ + is_and(L,rpe,rqe,andpq) \ is_not(L,andpq,notpq) \ + env \ list(A) \ pair(L, env, notpq, z))" apply (rule strong_replacementI) apply (rule_tac u="{list(A),B,rp,rq}" in gen_separation_multi [OF Nand_Reflects], @@ -933,12 +933,12 @@ "\L(A); L(rp)\ \ strong_replacement (L, \env z. \bo[L]. - env \ list(A) & + env \ list(A) \ is_bool_of_o (L, \a[L]. \co[L]. \rpco[L]. a\A \ is_Cons(L,a,env,co) \ fun_apply(L,rp,co,rpco) \ number1(L, rpco), - bo) & + bo) \ pair(L,env,bo,z))" apply (rule strong_replacementI) apply (rule_tac u="{A,list(A),B,rp}" @@ -974,20 +974,20 @@ subsubsection\The Lambda Replacement Case\ lemma formula_rec_lambda_replacement_Reflects: - "REFLECTS [\x. \u[L]. u \ B & - mem_formula(L,u) & + "REFLECTS [\x. \u[L]. u \ B \ + mem_formula(L,u) \ (\c[L]. is_formula_case (L, satisfies_is_a(L,A), satisfies_is_b(L,A), satisfies_is_c(L,A,g), satisfies_is_d(L,A,g), - u, c) & + u, c) \ pair(L,u,c,x)), - \i x. \u \ Lset(i). u \ B & mem_formula(##Lset(i),u) & + \i x. \u \ Lset(i). u \ B \ mem_formula(##Lset(i),u) \ (\c \ Lset(i). is_formula_case (##Lset(i), satisfies_is_a(##Lset(i),A), satisfies_is_b(##Lset(i),A), satisfies_is_c(##Lset(i),A,g), satisfies_is_d(##Lset(i),A,g), - u, c) & + u, c) \ pair(##Lset(i),u,c,x))]" by (intro FOL_reflections function_reflections mem_formula_reflection is_formula_case_reflection satisfies_is_a_reflection @@ -998,11 +998,11 @@ \ \For the \<^term>\transrec\\ "\L(g); L(A)\ \ strong_replacement (L, - \x y. mem_formula(L,x) & + \x y. mem_formula(L,x) \ (\c[L]. is_formula_case(L, satisfies_is_a(L,A), satisfies_is_b(L,A), satisfies_is_c(L,A,g), - satisfies_is_d(L,A,g), x, c) & + satisfies_is_d(L,A,g), x, c) \ pair(L, x, c, y)))" apply (rule strong_replacementI) apply (rule_tac u="{B,A,g}" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Separation.thy --- a/src/ZF/Constructible/Separation.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Separation.thy Tue Sep 27 17:03:23 2022 +0100 @@ -18,12 +18,12 @@ 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) \ 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))" + \ {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) @@ -43,7 +43,7 @@ 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))}") +apply (subgoal_tac "{x \ z. P(x)} = {x \ Lset(j). x \ z \ (Q(x))}") prefer 2 apply (blast dest: mem_Lset_imp_subset_Lset) apply (simp add: Lset_succ Collect_conj_in_DPow_Lset) @@ -118,14 +118,14 @@ subsection\Separation for Cartesian Product\ lemma cartprod_Reflects: - "REFLECTS[\z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)), - \i z. \x\Lset(i). x\A & (\y\Lset(i). y\B & + "REFLECTS[\z. \x[L]. x\A \ (\y[L]. y\B \ pair(L,x,y,z)), + \i z. \x\Lset(i). x\A \ (\y\Lset(i). y\B \ pair(##Lset(i),x,y,z))]" 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)))" + \ 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)+ @@ -134,13 +134,13 @@ subsection\Separation for Image\ lemma image_Reflects: - "REFLECTS[\y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)), - \i y. \p\Lset(i). p\r & (\x\Lset(i). x\A & pair(##Lset(i),x,y,p))]" + "REFLECTS[\y. \p[L]. p\r \ (\x[L]. x\A \ pair(L,x,y,p)), + \i y. \p\Lset(i). p\r \ (\x\Lset(i). x\A \ pair(##Lset(i),x,y,p))]" 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)))" + \ 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)+ @@ -150,14 +150,14 @@ subsection\Separation for Converse\ lemma converse_Reflects: - "REFLECTS[\z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)), - \i z. \p\Lset(i). p\r & (\x\Lset(i). \y\Lset(i). - pair(##Lset(i),x,y,p) & pair(##Lset(i),y,x,z))]" + "REFLECTS[\z. \p[L]. p\r \ (\x[L]. \y[L]. pair(L,x,y,p) \ pair(L,y,x,z)), + \i z. \p\Lset(i). p\r \ (\x\Lset(i). \y\Lset(i). + pair(##Lset(i),x,y,p) \ pair(##Lset(i),y,x,z))]" by (intro FOL_reflections function_reflections) lemma converse_separation: "L(r) \ separation(L, - \z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)))" + \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) apply (rule sep_rules | simp)+ @@ -167,12 +167,12 @@ subsection\Separation for Restriction\ lemma restrict_Reflects: - "REFLECTS[\z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)), - \i z. \x\Lset(i). x\A & (\y\Lset(i). pair(##Lset(i),x,y,z))]" + "REFLECTS[\z. \x[L]. x\A \ (\y[L]. pair(L,x,y,z)), + \i z. \x\Lset(i). x\A \ (\y\Lset(i). pair(##Lset(i),x,y,z))]" 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)+ @@ -183,18 +183,18 @@ lemma comp_Reflects: "REFLECTS[\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, + pair(L,x,z,xz) \ pair(L,x,y,xy) \ pair(L,y,z,yz) \ + xy\s \ yz\r, \i xz. \x\Lset(i). \y\Lset(i). \z\Lset(i). \xy\Lset(i). \yz\Lset(i). - pair(##Lset(i),x,z,xz) & pair(##Lset(i),x,y,xy) & - pair(##Lset(i),y,z,yz) & xy\s & yz\r]" + pair(##Lset(i),x,z,xz) \ pair(##Lset(i),x,y,xy) \ + pair(##Lset(i),y,z,yz) \ xy\s \ yz\r]" 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]. - pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & - xy\s & yz\r)" + 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) txt\Subgoals after applying general ``separation'' rule: @{subgoals[display,indent=0,margin=65]}\ @@ -208,12 +208,12 @@ subsection\Separation for Predecessors in an Order\ lemma pred_Reflects: - "REFLECTS[\y. \p[L]. p\r & pair(L,y,x,p), - \i y. \p \ Lset(i). p\r & pair(##Lset(i),y,x,p)]" + "REFLECTS[\y. \p[L]. p\r \ pair(L,y,x,p), + \i y. \p \ Lset(i). p\r \ pair(##Lset(i),y,x,p)]" 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)+ @@ -223,12 +223,12 @@ subsection\Separation for the Membership Relation\ lemma Memrel_Reflects: - "REFLECTS[\z. \x[L]. \y[L]. pair(L,x,y,z) & x \ y, - \i z. \x \ Lset(i). \y \ Lset(i). pair(##Lset(i),x,y,z) & x \ y]" + "REFLECTS[\z. \x[L]. \y[L]. pair(L,x,y,z) \ x \ y, + \i z. \x \ Lset(i). \y \ Lset(i). pair(##Lset(i),x,y,z) \ x \ y]" by (intro FOL_reflections function_reflections) lemma Memrel_separation: - "separation(L, \z. \x[L]. \y[L]. pair(L,x,y,z) & x \ y)" + "separation(L, \z. \x[L]. \y[L]. pair(L,x,y,z) \ x \ y)" apply (rule gen_separation [OF Memrel_Reflects nonempty]) apply (rule_tac env="[]" in DPow_LsetI) apply (rule sep_rules | simp)+ @@ -238,19 +238,19 @@ subsection\Replacement for FunSpace\ lemma funspace_succ_Reflects: - "REFLECTS[\z. \p[L]. p\A & (\f[L]. \b[L]. \nb[L]. \cnbf[L]. - pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & + "REFLECTS[\z. \p[L]. p\A \ (\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)), - \i z. \p \ Lset(i). p\A & (\f \ Lset(i). \b \ Lset(i). + \i z. \p \ Lset(i). p\A \ (\f \ Lset(i). \b \ Lset(i). \nb \ Lset(i). \cnbf \ Lset(i). - pair(##Lset(i),f,b,p) & pair(##Lset(i),n,b,nb) & - is_cons(##Lset(i),nb,f,cnbf) & upair(##Lset(i),cnbf,cnbf,z))]" + pair(##Lset(i),f,b,p) \ pair(##Lset(i),n,b,nb) \ + is_cons(##Lset(i),nb,f,cnbf) \ upair(##Lset(i),cnbf,cnbf,z))]" by (intro FOL_reflections function_reflections) lemma funspace_succ_replacement: "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) & + pair(L,f,b,p) \ pair(L,n,b,nb) \ is_cons(L,nb,f,cnbf) \ upair(L,cnbf,cnbf,z))" apply (rule strong_replacementI) apply (rule_tac u="{n,B}" in gen_separation_multi [OF funspace_succ_Reflects], @@ -264,13 +264,13 @@ lemma is_recfun_reflects: "REFLECTS[\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) & + 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) \ fx \ gx), \i x. \xa \ Lset(i). \xb \ Lset(i). - pair(##Lset(i),x,a,xa) & xa \ r & pair(##Lset(i),x,b,xb) & xb \ r & - (\fx \ Lset(i). \gx \ Lset(i). fun_apply(##Lset(i),f,x,fx) & - fun_apply(##Lset(i),g,x,gx) & fx \ gx)]" + pair(##Lset(i),x,a,xa) \ xa \ r \ pair(##Lset(i),x,b,xb) \ xb \ r \ + (\fx \ Lset(i). \gx \ Lset(i). fun_apply(##Lset(i),f,x,fx) \ + fun_apply(##Lset(i),g,x,gx) \ fx \ gx)]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma is_recfun_separation: @@ -278,8 +278,8 @@ "\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) & + 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) \ fx \ gx))" apply (rule gen_separation_multi [OF is_recfun_reflects, of "{r,f,g,a,b}"], auto) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/WF_absolute.thy --- a/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 17:03:23 2022 +0100 @@ -12,7 +12,7 @@ rtrancl_alt :: "[i,i]=>i" where "rtrancl_alt(A,r) \ {p \ A*A. \n\nat. \f \ succ(n) -> A. - (\x y. p = & f`0 = x & f`n = y) & + (\x y. p = \ f`0 = x \ f`n = y) \ (\i\n. \ r)}" lemma alt_rtrancl_lemma1 [rule_format]: @@ -63,14 +63,14 @@ \ \The property of belonging to \rtran_closure(r)\\ "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) & - (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) & empty(M,zero) & - fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & + omega(M,nnat) \ n\nnat \ successor(M,n,n') \ + (\f[M]. typed_function(M,n',A,f) \ + (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) \ empty(M,zero) \ + fun_apply(M,f,zero,x) \ fun_apply(M,f,n,y)) \ (\j[M]. j\n \ (\fj[M]. \sj[M]. \fsj[M]. \ffp[M]. - fun_apply(M,f,j,fj) & successor(M,j,sj) & - fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \ r)))" + fun_apply(M,f,j,fj) \ successor(M,j,sj) \ + fun_apply(M,f,sj,fsj) \ pair(M,fj,fsj,ffp) \ ffp \ r)))" definition rtran_closure :: "[i=>o,i,i] => o" where @@ -81,7 +81,7 @@ definition tran_closure :: "[i=>o,i,i] => o" where "tran_closure(M,r,t) \ - \s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" + \s[M]. rtran_closure(M,r,s) \ composition(M,r,s,t)" locale M_trancl = M_basic + assumes rtrancl_separation: @@ -90,15 +90,15 @@ "\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)" + 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) \ - (\n[M]. n\nat & - (\f[M]. f \ succ(n) -> A & - (\x[M]. \y[M]. p = & f`0 = x & f`n = y) & + (\n[M]. n\nat \ + (\f[M]. f \ succ(n) -> A \ + (\x[M]. \y[M]. p = \ f`0 = x \ f`n = y) \ (\i\n. \ r)))" apply (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) done @@ -137,7 +137,7 @@ 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 @@ -145,7 +145,7 @@ 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) +apply (drule_tac x = "{x\A. \w. \w,x\ \ r^+ \ w \ Z}" in spec) apply (blast elim: tranclE) done @@ -155,10 +155,10 @@ apply (simp add: wellfounded_on_def) apply (safe intro!: equalityI) apply (rename_tac Z x) -apply (subgoal_tac "M({x\A. \w[M]. w \ Z & \w,x\ \ r^+})") +apply (subgoal_tac "M({x\A. \w[M]. w \ Z \ \w,x\ \ r^+})") prefer 2 apply (blast intro: wellfounded_trancl_separation') -apply (drule_tac x = "{x\A. \w[M]. w \ Z & \w,x\ \ r^+}" in rspec, safe) +apply (drule_tac x = "{x\A. \w[M]. w \ Z \ \w,x\ \ r^+}" in rspec, safe) apply (blast dest: transM, simp) apply (rename_tac y w) apply (drule_tac x=w in bspec, assumption, clarify) @@ -183,12 +183,12 @@ lemma (in M_trancl) wfrec_relativize: "\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) & + 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 \ - (\f[M]. is_recfun(r^+, a, \x f. H(x, restrict(f, r -`` {x})), f) & + (\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) apply (simp add: wftrec_def wfrec_def, safe) @@ -207,7 +207,7 @@ "\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))" + \ 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 @@ -227,7 +227,7 @@ wfrec_replacement(M,MH,r); relation2(M,MH,H); \x[M]. \g[M]. function(g) \ M(H(x,g))\ \ y = \ - (\f[M]. is_recfun(r,x,H,f) & y = )" + (\f[M]. is_recfun(r,x,H,f) \ y = )" apply safe apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) txt\converse direction\ @@ -255,9 +255,9 @@ text\Eliminates one instance of replacement.\ lemma (in M_trancl) wfrec_replacement_iff: "strong_replacement(M, \x z. - \y[M]. pair(M,x,y,z) & (\g[M]. is_recfun(r,x,H,g) & y = H(x,g))) \ + \y[M]. pair(M,x,y,z) \ (\g[M]. is_recfun(r,x,H,g) \ y = H(x,g))) \ strong_replacement(M, - \x y. \f[M]. is_recfun(r,x,H,f) & y = )" + \x y. \f[M]. is_recfun(r,x,H,f) \ y = )" apply simp apply (rule strong_replacement_cong, blast) done @@ -278,12 +278,12 @@ lemma (in M_trancl) eq_pair_wfrec_iff: "\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) & + 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 = \ - (\f[M]. is_recfun(r^+, x, \x f. H(x, restrict(f, r -`` {x})), f) & + (\f[M]. is_recfun(r^+, x, \x f. H(x, restrict(f, r -`` {x})), f) \ y = )" apply safe apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/WFrec.thy Tue Sep 27 17:03:23 2022 +0100 @@ -21,7 +21,7 @@ text\Expresses \is_recfun\ as a recursion equation\ lemma is_recfun_iff_equation: "is_recfun(r,a,H,f) \ - f \ r -`` {a} \ range(f) & + f \ r -`` {a} \ range(f) \ (\x \ r-``{a}. f`x = H(x, restrict(f, r-``{x})))" apply (rule iffI) apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, @@ -135,7 +135,7 @@ "\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 = ))" + (\x[M]. \ r \ z = ))" apply (simp add: is_recfun_def lam_def) apply (safe intro!: equalityI) apply (drule equalityD1 [THEN subsetD], assumption) @@ -174,8 +174,8 @@ \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[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" apply (subgoal_tac "\y \ r-``{x}. \z. :Y \ :f") @@ -201,7 +201,7 @@ lemma (in M_basic) univalent_is_recfun: "\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)))" + \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) done @@ -213,7 +213,7 @@ "\\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)); + \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)" apply (drule_tac A="r-``{a1}" in strong_replacementD) @@ -246,7 +246,7 @@ "\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)); + \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)" @@ -257,7 +257,7 @@ lemma (in M_basic) wf_exists_is_recfun [rule_format]: "\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)); + \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))" apply (rule wf_induct, assumption+) @@ -275,20 +275,20 @@ "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) & - pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) & - xa \ r & MH(x, f_r_sx, y))" + pair(M,x,y,z) \ pair(M,x,a,xa) \ upair(M,x,x,sx) \ + pre_image(M,r,sx,r_sx) \ restriction(M,f,r_sx,f_r_sx) \ + xa \ r \ MH(x, f_r_sx, y))" definition is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where "is_wfrec(M,MH,r,a,z) \ - \f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,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) \ strong_replacement(M, - \x z. \y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))" + \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); @@ -309,7 +309,7 @@ "\\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))" + (\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\ @@ -318,7 +318,7 @@ \x[M]. \g[M]. function(g) \ M(H(x,g)); 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)))" + 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]: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Constructible/Wellorderings.thy --- a/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 17:03:23 2022 +0100 @@ -34,18 +34,18 @@ 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))" + \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))" + \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) \ - transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)" + transitive_rel(M,A,r) \ linear_rel(M,A,r) \ wellfounded_on(M,A,r)" subsubsection \Trivial absoluteness proofs\ @@ -108,7 +108,7 @@ (*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)\ + \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) @@ -118,7 +118,7 @@ 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)\ + \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) @@ -188,7 +188,7 @@ done lemma (in M_basic) M_Memrel_iff: - "M(A) \ Memrel(A) = {z \ A*A. \x[M]. \y[M]. z = \x,y\ & x \ y}" + "M(A) \ Memrel(A) = {z \ A*A. \x[M]. \y[M]. z = \x,y\ \ x \ y}" unfolding Memrel_def by (blast dest: transM) lemma (in M_basic) Memrel_closed [intro,simp]: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Epsilon.thy Tue Sep 27 17:03:23 2022 +0100 @@ -359,8 +359,8 @@ lemma def_transrec2: "(\x. f(x)\transrec2(x,a,b)) - \ f(0) = a & - f(succ(i)) = b(i, f(i)) & + \ f(0) = a \ + f(succ(i)) = b(i, f(i)) \ (Limit(K) \ f(K) = (\j : 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: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Finite.thy --- a/src/ZF/Finite.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Finite.thy Tue Sep 27 17:03:23 2022 +0100 @@ -183,7 +183,7 @@ lemma FiniteFun_Collect_iff: "f \ FiniteFun(A, {y \ B. P(y)}) - \ f \ FiniteFun(A,B) & (\x\domain(f). P(f`x))" + \ f \ FiniteFun(A,B) \ (\x\domain(f). P(f`x))" apply auto apply (blast intro: FiniteFun_mono [THEN [2] rev_subsetD]) apply (blast dest: Pair_mem_PiD FiniteFun_is_fun) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Fixedpt.thy --- a/src/ZF/Fixedpt.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Fixedpt.thy Tue Sep 27 17:03:23 2022 +0100 @@ -10,7 +10,7 @@ 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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/IMP/Denotation.thy --- a/src/ZF/IMP/Denotation.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/IMP/Denotation.thy Tue Sep 27 17:03:23 2022 +0100 @@ -58,7 +58,7 @@ 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" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/Binary_Trees.thy --- a/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 17:03:23 2022 +0100 @@ -20,7 +20,7 @@ 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'" +lemma Br_iff: "Br(a, l, r) = Br(a', l', r') \ a = a' \ l = l' \ r = r'" \ \Proving a freeness theorem.\ by (fast elim!: bt.free_elims) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/Comb.thy --- a/src/ZF/Induct/Comb.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/Comb.thy Tue Sep 27 17:03:23 2022 +0100 @@ -78,14 +78,14 @@ definition diamond :: "i \ o" where "diamond(r) \ - \x y. \r \ (\y'. \r \ (\z. \r & \ r))" + \x y. \r \ (\y'. \r \ (\z. \r \ \ r))" subsection \Transitive closure preserves the Church-Rosser property\ lemma diamond_strip_lemmaD [rule_format]: "\diamond(r); :r^+\ \ - \y'. :r \ (\z. : r^+ & : r)" + \y'. :r \ (\z. : r^+ \ : r)" apply (unfold diamond_def) apply (erule trancl_induct) apply (blast intro: r_into_trancl) @@ -153,7 +153,7 @@ 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" @@ -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)" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/FoldSet.thy --- a/src/ZF/Induct/FoldSet.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/FoldSet.thy Tue Sep 27 17:03:23 2022 +0100 @@ -38,7 +38,7 @@ 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" + \ B - {y} = C-{x} \ x\C \ y\B" apply (auto elim: equalityE) done @@ -57,7 +57,7 @@ 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 @@ -172,7 +172,7 @@ lemma (in fold_typing) fold_cons_lemma [rule_format]: "\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))" + (\y. \ fold_set(C, B, f, e) \ v = f(c, y))" apply auto prefer 2 apply (blast intro: fold_set_imp_cons) apply (frule_tac Fin.consI [of c, THEN FinD, THEN fold_set_mono, THEN subsetD], assumption+) @@ -296,7 +296,7 @@ prefer 2 apply blast apply (subgoal_tac "C (x) \ (\i\B. C (i)) = 0") prefer 2 apply blast -apply (subgoal_tac "Finite (\i\B. C (i)) & Finite (C (x)) & Finite (B) ") +apply (subgoal_tac "Finite (\i\B. C (i)) \ Finite (C (x)) \ Finite (B) ") apply (simp (no_asm_simp) add: setsum_Un_disjoint) apply (auto intro: Finite_Union Finite_RepFun) done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/ListN.thy --- a/src/ZF/Induct/ListN.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/ListN.thy Tue Sep 27 17:03:23 2022 +0100 @@ -24,7 +24,7 @@ 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" +lemma listn_iff: " \ listn(A) \ l \ list(A) \ length(l)=n" apply (rule iffI) apply (erule listn.induct) apply auto diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/Multiset.thy Tue Sep 27 17:03:23 2022 +0100 @@ -26,7 +26,7 @@ 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 @@ -42,7 +42,7 @@ (*convert a function to a multiset by eliminating 0*) normalize :: "i => i" where "normalize(f) \ - if (\A. f \ A -> nat & Finite(A)) then + if (\A. f \ A -> nat \ Finite(A)) then funrestrict(f, {x \ mset_of(f). 0 < f`x}) else 0" @@ -87,7 +87,7 @@ "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)}" + N=M0 +# {#a#} \ M=M0 +# K \ (\b \ mset_of(K). \ r)}" definition multirel :: "[i, i] => i" where @@ -97,15 +97,15 @@ 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\ @@ -142,7 +142,7 @@ text\A useful simplification rule\ lemma multiset_fun_iff: - "(f \ A -> nat-{0}) \ f \ A->nat&(\a \ A. f`a \ nat & 0 < f`a)" + "(f \ A -> nat-{0}) \ f \ A->nat\(\a \ A. f`a \ nat \ 0 < f`a)" apply safe apply (rule_tac [4] B1 = "range (f) " in Pi_mono [THEN subsetD]) apply (auto intro!: Ord_0_lt @@ -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) @@ -168,7 +168,7 @@ apply (blast intro: Fin_into_Finite) done -lemma Mult_iff_multiset: "M \ Mult(A) \ multiset(M) & mset_of(M)\A" +lemma Mult_iff_multiset: "M \ Mult(A) \ multiset(M) \ mset_of(M)\A" by (blast dest: Mult_into_multiset intro: multiset_into_Mult) lemma multiset_iff_Mult_mset_of: "multiset(M) \ M \ Mult(mset_of(M))" @@ -205,7 +205,7 @@ (* msingle *) -lemma msingle_not_0 [iff]: "{#a#} \ 0 & 0 \ {#a#}" +lemma msingle_not_0 [iff]: "{#a#} \ 0 \ 0 \ {#a#}" by (simp add: msingle_def) lemma msingle_eq_iff [iff]: "({#a#} = {#b#}) \ (a = b)" @@ -223,7 +223,7 @@ lemma normalize_idem [simp]: "normalize(normalize(f)) = normalize(f)" apply (simp add: normalize_def funrestrict_def mset_of_def) -apply (case_tac "\A. f \ A -> nat & Finite (A) ") +apply (case_tac "\A. f \ A -> nat \ Finite (A) ") apply clarify apply (drule_tac x = "{x \ domain (f) . 0 < f ` x}" in spec) apply auto @@ -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 @@ -424,10 +424,10 @@ (** 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 @@ -440,12 +440,12 @@ "\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#})" + \ (M +# N = {#a#}) \ (M={#a#} \ N=0) | (M = 0 \ N = {#a#})" apply (simp (no_asm_simp) add: multiset_equality) apply safe apply simp_all @@ -465,7 +465,7 @@ done lemma msingle_is_union: "\multiset(M); multiset(N)\ - \ ({#a#} = M +# N) \ ({#a#} = M & N=0 | M = 0 & {#a#} = 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 @@ -663,7 +663,7 @@ by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def) lemma MCollect_mem_iff [iff]: - "x \ mset_of({#x \ M. P(x)#}) \ x \ mset_of(M) & P(x)" + "x \ mset_of({#x \ M. P(x)#}) \ x \ mset_of(M) \ P(x)" by (simp add: MCollect_def mset_of_def) lemma mcount_MCollect [simp]: @@ -680,8 +680,8 @@ (* and more algebraic laws on multisets *) lemma munion_eq_conv_diff: "\multiset(M); multiset(N)\ - \ (M +# {#a#} = N +# {#b#}) \ (M = N & a = b | - M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})" + \ (M +# {#a#} = N +# {#b#}) \ (M = N \ a = b | + M = N -# {#a#} +# {#b#} \ N = M -# {#b#} +# {#a#})" apply (simp del: mcount_single add: multiset_equality) apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE) apply (case_tac "a=b", auto) @@ -695,7 +695,7 @@ lemma melem_diff_single: "multiset(M) \ - k \ mset_of(M -# {#a#}) \ (k=a & 1 < mcount(M,a)) | (k\ a & k \ mset_of(M))" + k \ mset_of(M -# {#a#}) \ (k=a \ 1 < mcount(M,a)) | (k\ a \ k \ mset_of(M))" apply (simp add: multiset_def) apply (simp add: normalize_def mset_of_def msingle_def mdiff_def mcount_def) apply (auto dest: domain_type intro: zero_less_diff [THEN iffD1] @@ -708,7 +708,7 @@ lemma munion_eq_conv_exist: "\M \ Mult(A); N \ Mult(A)\ \ (M +# {#a#} = N +# {#b#}) \ - (M=N & a=b | (\K \ Mult(A). M= K +# {#b#} & N=K +# {#a#}))" + (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) @@ -727,9 +727,9 @@ lemma multirel1_iff: " \ multirel1(A, r) \ - (\a. a \ A & - (\M0. M0 \ Mult(A) & (\K. K \ Mult(A) & - M=M0 +# {#a#} & N=M0 +# K & (\b \ mset_of(K). \ r))))" + (\a. a \ A \ + (\M0. M0 \ Mult(A) \ (\K. K \ Mult(A) \ + M=M0 +# {#a#} \ N=M0 +# K \ (\b \ mset_of(K). \ r))))" by (auto simp add: multirel1_def Mult_iff_multiset Bex_def) @@ -766,8 +766,8 @@ by (auto simp add: multirel1_def Mult_iff_multiset) 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)" + (\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]) apply (simp add: multirel1_iff) apply (auto simp add: munion_eq_conv_exist) @@ -917,8 +917,8 @@ " \ multirel(A, r) \ trans[A](r) \ (\I J K. - I \ Mult(A) & J \ Mult(A) & K \ Mult(A) & - N = I +# J & M = I +# K & J \ 0 & + I \ Mult(A) \ J \ Mult(A) \ K \ Mult(A) \ + N = I +# J \ M = I +# K \ J \ 0 \ (\k \ mset_of(K). \j \ mset_of(J). \ r))" apply (simp add: multirel_def Ball_def Bex_def) apply (erule converse_trancl_induct) @@ -969,7 +969,7 @@ 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" + \ \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) @@ -983,8 +983,8 @@ lemma one_step_implies_multirel_lemma [rule_format (no_asm)]: "n \ nat \ (\I J K. - I \ Mult(A) & J \ Mult(A) & K \ Mult(A) & - (msize(J) = $# n & J \0 & (\k \ mset_of(K). \j \ mset_of(J). \ r)) + I \ Mult(A) \ J \ Mult(A) \ K \ Mult(A) \ + (msize(J) = $# n \ J \0 \ (\k \ mset_of(K). \j \ mset_of(J). \ r)) \ \ multirel(A, r))" apply (simp add: Mult_iff_multiset) apply (erule nat_induct, clarify) @@ -1126,7 +1126,7 @@ lemma munion_multirel_mono: "\ \ multirel(A, r); \ multirel(A, r)\ \ \ multirel(A, r)" -apply (subgoal_tac "M \ Mult(A) & N \ Mult(A) & K \ Mult(A) & L \ Mult(A) ") +apply (subgoal_tac "M \ Mult(A) \ N \ Mult(A) \ K \ Mult(A) \ L \ Mult(A) ") prefer 2 apply (blast dest: multirel_type [THEN subsetD]) apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2) done @@ -1235,7 +1235,7 @@ apply (blast intro: mless_trans) done -lemma mless_le_iff: "M <# N \ (M <#= N & M \ N)" +lemma mless_le_iff: "M <# N \ (M <#= N \ M \ N)" by (simp add: mle_def, auto) (** Monotonicity of mless **) @@ -1253,7 +1253,7 @@ 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" @@ -1264,7 +1264,7 @@ (* <#= *) -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" @@ -1278,7 +1278,7 @@ 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))) ") +apply (subgoal_tac "\i. Ord (i) \ M \ Mult(field(Memrel(i))) ") prefer 2 apply (simp add: omultiset_def) apply (case_tac "M=0", simp_all, clarify) apply (subgoal_tac "<0 +# 0, 0 +# M> \ multirel(field (Memrel(i)), Memrel(i))") diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/Mutil.thy --- a/src/ZF/Induct/Mutil.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/Mutil.thy Tue Sep 27 17:03:23 2022 +0100 @@ -38,7 +38,7 @@ subsection \Basic properties of evnodd\ -lemma evnodd_iff: ": evnodd(A,b) \ : A & (i#+j) mod 2 = b" +lemma evnodd_iff: ": evnodd(A,b) \ : A \ (i#+j) mod 2 = b" by (unfold evnodd_def) blast lemma evnodd_subset: "evnodd(A, b) \ A" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/PropLog.thy --- a/src/ZF/Induct/PropLog.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/PropLog.thy Tue Sep 27 17:03:23 2022 +0100 @@ -1,5 +1,5 @@ (* Title: ZF/Induct/PropLog.thy - Author: Tobias Nipkow & Lawrence C Paulson + Author: Tobias Nipkow \ Lawrence C Paulson Copyright 1993 University of Cambridge *) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Induct/Term.thy --- a/src/ZF/Induct/Term.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Induct/Term.thy Tue Sep 27 17:03:23 2022 +0100 @@ -110,7 +110,7 @@ \ \\<^term>\map\ works correctly on the underlying list of terms.\ apply (induct set: list) apply simp - apply (subgoal_tac "rank (a) rank (l) < i") apply (simp add: rank_of_Ord) apply (simp add: list.con_defs) apply (blast dest: rank_rls [THEN lt_trans]) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/InfDatatype.thy --- a/src/ZF/InfDatatype.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/InfDatatype.thy Tue Sep 27 17:03:23 2022 +0100 @@ -12,7 +12,7 @@ lemma fun_Vcsucc_lemma: assumes f: "f \ D -> Vfrom(A,csucc(K))" and DK: "|D| \ K" and ICK: "InfCard(K)" - shows "\j. f \ D -> Vfrom(A,j) & j < csucc(K)" + shows "\j. f \ D -> Vfrom(A,j) \ j < csucc(K)" proof (rule exI, rule conjI) show "f \ D \ Vfrom(A, \z\D. \ i. f`z \ Vfrom (A,i))" proof (rule Pi_type [OF f]) @@ -46,7 +46,7 @@ lemma subset_Vcsucc: "\D \ Vfrom(A,csucc(K)); |D| \ K; InfCard(K)\ - \ \j. D \ Vfrom(A,j) & j < csucc(K)" + \ \j. D \ Vfrom(A,j) \ j < csucc(K)" by (simp add: subset_iff_id fun_Vcsucc_lemma) (*Version for arbitrary index sets*) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Int.thy --- a/src/ZF/Int.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Int.thy Tue Sep 27 17:03:23 2022 +0100 @@ -10,7 +10,7 @@ definition intrel :: i where "intrel \ {p \ (nat*nat)*(nat*nat). - \x1 y1 x2 y2. p=<,> & x1#+y2 = x2#+y1}" + \x1 y1 x2 y2. p=<,> \ x1#+y2 = x2#+y1}" definition int :: i where @@ -34,7 +34,7 @@ definition znegative :: "i=>o" where - "znegative(z) \ \x y. xnat & \z" + "znegative(z) \ \x y. x y\nat \ \z" definition iszero :: "i=>o" where @@ -52,8 +52,8 @@ 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) | - (znegative(z) & $- z = $# m))" + THE m. m\nat \ ((\ znegative(z) \ z = $# m) | + (znegative(z) \ $- z = $# m))" definition raw_zmult :: "[i,i]=>i" where @@ -99,7 +99,7 @@ lemma intrel_iff [simp]: "<,>: intrel \ - x1\nat & y1\nat & x2\nat & y2\nat & x1#+y2 = x2#+y1" + x1\nat \ y1\nat \ x2\nat \ y2\nat \ x1#+y2 = x2#+y1" by (simp add: intrel_def) lemma intrelI [intro!]: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/IntDiv.thy --- a/src/ZF/IntDiv.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/IntDiv.thy Tue Sep 27 17:03:23 2022 +0100 @@ -36,8 +36,8 @@ definition quorem :: "[i,i] => o" where "quorem \ % . - a = b$*q $+ r & - (#0$r & r$(#0$ #0)" + a = b$*q $+ r \ + (#0$ #0$\r \ r$(#0$ b$ r $\ #0)" definition adjust :: "[i,i] => i" where @@ -298,7 +298,7 @@ lemma zmult_zless_lemma: "\k \ int; m \ int; n \ int\ - \ (m$*k $< n$*k) \ ((#0 $< k & m$ (m$*k $< n$*k) \ ((#0 $< k \ m$ n$ ((#0 $< k & m$ ((#0 $< k \ m$ n$ ((#0 $< k & m$ ((#0 $< k \ m$ n$ n$*k) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" + "(m$*k $\ n$*k) \ ((#0 $< k \ m$\n) \ (k $< #0 \ n$\m))" by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2) lemma zmult_zle_cancel1: - "(k$*m $\ k$*n) \ ((#0 $< k \ m$\n) & (k $< #0 \ n$\m))" + "(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 @@ -462,7 +462,7 @@ (*FIXME: use intify in integ_of so that we always have @{term"integ_of w \ int"}. then this rewrite can work for all constants\*) -lemma intify_eq_0_iff_zle: "intify(m) = #0 \ (m $\ #0 & #0 $\ m)" +lemma intify_eq_0_iff_zle: "intify(m) = #0 \ (m $\ #0 \ #0 $\ m)" by (simp add: int_eq_iff_zle) @@ -482,7 +482,7 @@ lemma int_0_less_lemma: "\x \ int; y \ int\ - \ (#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" + \ (#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) @@ -496,30 +496,30 @@ done lemma int_0_less_mult_iff: - "(#0 $< x $* y) \ (#0 $< x & #0 $< y | x $< #0 & y $< #0)" + "(#0 $< x $* y) \ (#0 $< x \ #0 $< y | x $< #0 \ y $< #0)" apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma) apply auto done lemma int_0_le_lemma: "\x \ int; y \ int\ - \ (#0 $\ x $* y) \ (#0 $\ x & #0 $\ y | x $\ #0 & y $\ #0)" + \ (#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: - "(#0 $\ x $* y) \ ((#0 $\ x & #0 $\ y) | (x $\ #0 & y $\ #0))" + "(#0 $\ x $* y) \ ((#0 $\ x \ #0 $\ y) | (x $\ #0 \ y $\ #0))" apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma) apply auto done lemma zmult_less_0_iff: - "(x $* y $< #0) \ (#0 $< x & y $< #0 | x $< #0 & #0 $< y)" + "(x $* y $< #0) \ (#0 $< x \ y $< #0 | x $< #0 \ #0 $< y)" apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym]) apply (auto dest: zless_not_sym simp add: not_zle_iff_zless) done lemma zmult_le_0_iff: - "(x $* y $\ #0) \ (#0 $\ x & y $\ #0 | x $\ #0 & #0 $\ y)" + "(x $* y $\ #0) \ (#0 $\ x \ y $\ #0 | x $\ #0 \ #0 $\ y)" by (auto dest: zless_not_sym simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym]) @@ -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) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/List.thy --- a/src/ZF/List.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/List.thy Tue Sep 27 17:03:23 2022 +0100 @@ -136,11 +136,11 @@ (*An elimination rule, for type-checking*) inductive_cases ConsE: "Cons(a,l) \ list(A)" -lemma Cons_type_iff [simp]: "Cons(a,l) \ list(A) \ a \ A & l \ list(A)" +lemma Cons_type_iff [simp]: "Cons(a,l) \ list(A) \ a \ A \ l \ list(A)" by (blast elim: ConsE) (*Proving freeness results*) -lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' & l=l'" +lemma Cons_iff: "Cons(a,l)=Cons(a',l') \ a=a' \ l=l'" by auto lemma Nil_Cons_iff: "\ Nil=Cons(a,l)" @@ -480,7 +480,7 @@ 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 i 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]: @@ -554,7 +554,7 @@ (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff [rule_format]: "\xs:list(A); ys:list(A); zs:list(A)\ \ - length(ys)=length(zs) \ (xs@ys=zs \ (xs=Nil & ys=zs))" + length(ys)=length(zs) \ (xs@ys=zs \ (xs=Nil \ ys=zs))" apply (erule list.induct) apply (auto simp add: length_app) done @@ -562,14 +562,14 @@ (*TOO SLOW as a default simprule!*) lemma append_left_is_Nil_iff2 [rule_format]: "\xs:list(A); ys:list(A); zs:list(A)\ \ - length(ys)=length(zs) \ (zs=ys@xs \ (xs=Nil & ys=zs))" + length(ys)=length(zs) \ (zs=ys@xs \ (xs=Nil \ ys=zs))" apply (erule list.induct) apply (auto simp add: length_app) done lemma append_eq_append_iff [rule_format]: "xs:list(A) \ \ys \ list(A). - length(xs)=length(ys) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" + length(xs)=length(ys) \ (xs@us = ys@vs) \ (xs=ys \ us=vs)" apply (erule list.induct) apply (simp (no_asm_simp)) apply clarify @@ -580,7 +580,7 @@ lemma append_eq_append [rule_format]: "xs:list(A) \ \ys \ list(A). \us \ list(A). \vs \ list(A). - length(us) = length(vs) \ (xs@us = ys@vs) \ (xs=ys & us=vs)" + length(us) = length(vs) \ (xs@us = ys@vs) \ (xs=ys \ us=vs)" apply (induct_tac "xs") apply (force simp add: length_app, clarify) apply (erule_tac a = ys in list.cases, simp) @@ -590,7 +590,7 @@ 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@us = ys@vs \ (xs=ys \ us=vs)" apply (rule iffI) apply (rule append_eq_append, auto) done @@ -606,7 +606,7 @@ (* 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) @@ -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 x = nth(i,xs)}" apply (induct_tac "xs", simp_all) apply (rule equalityI, auto) apply (rule_tac x = 0 in bexI, auto) @@ -968,7 +968,7 @@ "\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)}" + \ 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 **) @@ -1192,7 +1192,7 @@ lemma sublist_shift_lemma: "\xs:list(B); i \ nat\ \ map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) = - map(fst, filter(%p. snd(p):nat & snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" + map(fst, filter(%p. snd(p):nat \ snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" apply (erule list_append_induct) apply (simp (no_asm_simp)) apply (auto simp add: add_commute length_app filter_append map_app_distrib) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Nat.thy --- a/src/ZF/Nat.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Nat.thy Tue Sep 27 17:03:23 2022 +0100 @@ -18,7 +18,7 @@ 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 @@ -244,7 +244,7 @@ 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 @@ -291,7 +291,7 @@ apply (blast intro: lt_trans) done -lemma Le_iff [iff]: " \ Le \ x \ y & x \ nat & y \ nat" +lemma Le_iff [iff]: " \ Le \ x \ y \ x \ nat \ y \ nat" by (force simp add: Le_def) end diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/OrdQuant.thy Tue Sep 27 17:03:23 2022 +0100 @@ -15,7 +15,7 @@ definition oex :: "[i, i => o] => o" where - "oex(A, P) \ \x. x \x. x P(x)" definition (* Ordinal Union *) @@ -44,12 +44,12 @@ lemma [simp]: "\(\x<0. P(x))" by (simp add: oex_def) -lemma [simp]: "(\x (Ord(i) \ P(i) & (\xx (Ord(i) \ P(i) \ (\xx (Ord(i) & (P(i) | (\xx (Ord(i) \ (P(i) | (\xo, 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) @@ -261,14 +261,14 @@ declare atomize_rall [symmetric, rulify] lemma rall_simps1: - "(\x[M]. P(x) & Q) <-> (\x[M]. P(x)) & ((\x[M]. False) | Q)" + "(\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))" by blast+ lemma rall_simps2: - "(\x[M]. P & Q(x)) <-> ((\x[M]. False) | P) & (\x[M]. Q(x))" + "(\x[M]. P \ Q(x)) <-> ((\x[M]. False) | P) \ (\x[M]. Q(x))" "(\x[M]. P | Q(x)) <-> (P | (\x[M]. Q(x)))" "(\x[M]. P \ Q(x)) <-> (P \ (\x[M]. Q(x)))" by blast+ @@ -276,19 +276,19 @@ lemmas rall_simps [simp] = rall_simps1 rall_simps2 lemma rall_conj_distrib: - "(\x[M]. P(x) & Q(x)) <-> ((\x[M]. P(x)) & (\x[M]. Q(x)))" + "(\x[M]. P(x) \ Q(x)) <-> ((\x[M]. P(x)) \ (\x[M]. Q(x)))" by blast lemma rex_simps1: - "(\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) \ 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))" by blast+ lemma rex_simps2: - "(\x[M]. P & Q(x)) <-> (P & (\x[M]. Q(x)))" - "(\x[M]. P | Q(x)) <-> ((\x[M]. True) & P) | (\x[M]. Q(x))" + "(\x[M]. P \ Q(x)) <-> (P \ (\x[M]. Q(x)))" + "(\x[M]. P | Q(x)) <-> ((\x[M]. True) \ P) | (\x[M]. Q(x))" "(\x[M]. P \ Q(x)) <-> (((\x[M]. False) | P) \ (\x[M]. Q(x)))" by blast+ @@ -307,10 +307,10 @@ lemma rex_triv_one_point2 [simp]: "(\x[M]. a=x) <-> ( M(a))" by blast -lemma rex_one_point1 [simp]: "(\x[M]. x=a & P(x)) <-> ( M(a) & P(a))" +lemma rex_one_point1 [simp]: "(\x[M]. x=a \ P(x)) <-> ( M(a) \ P(a))" by blast -lemma rex_one_point2 [simp]: "(\x[M]. a=x & P(x)) <-> ( M(a) & P(a))" +lemma rex_one_point2 [simp]: "(\x[M]. a=x \ P(x)) <-> ( M(a) \ P(a))" by blast lemma rall_one_point1 [simp]: "(\x[M]. x=a \ P(x)) <-> ( M(a) \ P(a))" @@ -349,7 +349,7 @@ text \Setting up the one-point-rule simproc\ -simproc_setup defined_rex ("\x[M]. P(x) & Q(x)") = \ +simproc_setup defined_rex ("\x[M]. P(x) \ Q(x)") = \ fn _ => Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms rex_def}) \ diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Order.thy --- a/src/ZF/Order.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Order.thy Tue Sep 27 17:03:23 2022 +0100 @@ -17,7 +17,7 @@ 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 @@ -25,7 +25,7 @@ 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,7 +41,7 @@ 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 @@ -64,7 +64,7 @@ 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\ @@ -102,7 +102,7 @@ (** Derived rules for pred(A,x,r) **) -lemma pred_iff: "y \ pred(A,x,r) \ :r & y \ A" +lemma pred_iff: "y \ pred(A,x,r) \ :r \ y \ A" by (unfold pred_def, blast) lemmas predI = conjI [THEN pred_iff [THEN iffD2]] @@ -474,7 +474,7 @@ f \ ord_iso(A,r, B,s); g \ ord_iso(A,r, B,s)\ \ f = g" apply (rule fun_extension) apply (erule ord_iso_is_bij [THEN bij_is_fun])+ -apply (subgoal_tac "f`x \ B & g`x \ B & linear(B,s)") +apply (subgoal_tac "f`x \ B \ g`x \ B \ linear(B,s)") apply (simp add: linear_def) apply (blast dest: well_ord_iso_unique_lemma) apply (blast intro: ord_iso_is_bij bij_is_fun apply_funtype @@ -517,7 +517,7 @@ apply (unfold mono_map_def) apply (simp (no_asm_simp) add: ord_iso_map_fun) apply safe -apply (subgoal_tac "x \ A & ya:A & y \ B & yb:B") +apply (subgoal_tac "x \ A \ ya:A \ y \ B \ yb:B") apply (simp add: apply_equality [OF _ ord_iso_map_fun]) apply (unfold ord_iso_map_def) apply (blast intro: well_ord_iso_preserving, blast) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/OrderArith.thy --- a/src/ZF/OrderArith.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/OrderArith.thy Tue Sep 27 17:03:23 2022 +0100 @@ -13,21 +13,21 @@ "radd(A,r,B,s) \ {z: (A+B) * (A+B). (\x y. z = ) | - (\x' x. z = & :r) | - (\y' y. z = & :s)}" + (\x' x. z = \ :r) | + (\y' y. z = \ :s)}" definition (*lexicographic product of two relations; underlies ordinal multiplication*) rmult :: "[i,i,i,i]=>i" where "rmult(A,r,B,s) \ {z: (A*B) * (A*B). - \x' y' x y. z = <, > & - (: r | (x'=x & : s))}" + \x' y' x y. z = <, > \ + (: r | (x'=x \ : s))}" 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 @@ -39,15 +39,15 @@ subsubsection\Rewrite rules. Can be used to obtain introduction rules\ lemma radd_Inl_Inr_iff [iff]: - " \ radd(A,r,B,s) \ a \ A & b \ B" + " \ radd(A,r,B,s) \ a \ A \ b \ B" by (unfold radd_def, blast) lemma radd_Inl_iff [iff]: - " \ radd(A,r,B,s) \ a':A & a \ A & :r" + " \ radd(A,r,B,s) \ a':A \ a \ A \ :r" by (unfold radd_def, blast) lemma radd_Inr_iff [iff]: - " \ radd(A,r,B,s) \ b':B & b \ B & :s" + " \ radd(A,r,B,s) \ b':B \ b \ B \ :s" by (unfold radd_def, blast) lemma radd_Inr_Inl_iff [simp]: @@ -164,8 +164,8 @@ lemma rmult_iff [iff]: "<, > \ rmult(A,r,B,s) \ - (: r & a':A & a \ A & b': B & b \ B) | - (: s & a'=a & a \ A & b': B & b \ B)" + (: r \ a':A \ a \ A \ b': B \ b \ B) | + (: s \ a'=a \ a \ A \ b': B \ b \ B)" by (unfold rmult_def, blast) @@ -307,7 +307,7 @@ subsubsection\Rewrite rule\ -lemma rvimage_iff: " \ rvimage(A,f,r) \ : r & a \ A & b \ A" +lemma rvimage_iff: " \ rvimage(A,f,r) \ : r \ a \ A \ b \ A" by (unfold rvimage_def, blast) subsubsection\Type checking\ @@ -361,7 +361,7 @@ lemma wf_rvimage [intro!]: "wf(r) \ wf(rvimage(A,f,r))" apply (simp (no_asm_use) add: rvimage_def wf_eq_minimal) apply clarify -apply (subgoal_tac "\w. w \ {w: {f`x. x \ Q}. \x. x \ Q & (f`x = w) }") +apply (subgoal_tac "\w. w \ {w: {f`x. x \ Q}. \x. x \ Q \ (f`x = w) }") apply (erule allE) apply (erule impE) apply assumption @@ -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,7 +450,7 @@ 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]) @@ -496,7 +496,7 @@ lemma wf_measure [iff]: "wf(measure(A,f))" by (simp (no_asm) add: measure_eq_rvimage_Memrel wf_Memrel wf_rvimage) -lemma measure_iff [iff]: " \ measure(A,f) \ x \ A & y \ A & f(x) \ measure(A,f) \ x \ A \ y \ A \ f(x)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)" + \ (\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 e44d86131648 -r 0c18df79b1c8 src/ZF/OrderType.thy --- a/src/ZF/OrderType.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/OrderType.thy Tue Sep 27 17:03:23 2022 +0100 @@ -22,7 +22,7 @@ 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*) @@ -600,7 +600,7 @@ 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) @@ -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 Limit(k)" apply (rule_tac x="i ++ nat" in exI) apply (blast intro: oadd_LimitI oadd_lt_self Limit_nat [THEN Limit_has_0]) done -lemma Ord2_imp_greater_Limit: "\Ord(i); Ord(j)\ \ \k. iOrd(i); Ord(j)\ \ \k. i j Limit(k)" apply (insert Ord_Un [of i j, THEN Ord_imp_greater_Limit]) apply (simp add: Un_least_lt_iff) done @@ -796,7 +796,7 @@ lemma lt_omult: "\Ord(i); Ord(j); k - \ \j' i'. k = j**i' ++ j' & j' \j' i'. k = j**i' ++ j' \ j' 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 @@ -17,15 +17,15 @@ 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 - "i i\j & Ord(j)" + "i i\j \ Ord(j)" definition Limit :: "i=>o" where - "Limit(i) \ Ord(i) & 0y. y succ(y) Ord(i) \ 0 (\y. y succ(y)\\ 50) where @@ -50,11 +50,11 @@ 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 @@ -232,7 +232,7 @@ text\Recall that \<^term>\i \ j\ abbreviates \<^term>\i \\ -lemma le_iff: "i \ j <-> i j <-> i Ord(j))" by (unfold lt_def, blast) (*Equivalently, i i < succ(j)*) @@ -247,7 +247,7 @@ 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: @@ -267,7 +267,7 @@ subsection\Natural Deduction Rules for Memrel\ (*The lemmas MemrelI/E give better speed than [iff] here*) -lemma Memrel_iff [simp]: " \ Memrel(A) <-> a\b & a\A & b\A" +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)" @@ -313,7 +313,7 @@ (*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) @@ -428,10 +428,10 @@ 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)" +lemma le_subset_iff: "j \ i <-> j<=i \ Ord(i) \ Ord(j)" by (blast dest: subset_imp_le le_imp_subset elim: ltE) -lemma le_succ_iff: "i \ succ(j) <-> i \ j | i=succ(j) & Ord(i)" +lemma le_succ_iff: "i \ succ(j) <-> i \ j | i=succ(j) \ Ord(i)" apply (simp (no_asm) add: le_iff) apply blast done @@ -476,7 +476,7 @@ lemma lt_imp_0_lt: "j 0 i j" +lemma succ_lt_iff: "succ(i) < j <-> i succ(i) \ j" apply auto apply (blast intro: lt_trans le_refl dest: lt_Ord) apply (frule lt_Ord) @@ -506,14 +506,14 @@ 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 jOrd(i); Ord(j); Ord(k)\ \ 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 @@ -697,7 +697,7 @@ 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: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Perm.thy --- a/src/ZF/Perm.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Perm.thy Tue Sep 27 17:03:23 2022 +0100 @@ -16,7 +16,7 @@ (*composition of relations and functions; NOT Suppes's relative product*) comp :: "[i,i]=>i" (infixr \O\ 60) where "r O s \ {xz \ domain(s)*range(r) . - \x y z. xz= & :s & :r}" + \x y z. xz= \ :s \ :r}" definition (*the identity function for A*) @@ -192,7 +192,7 @@ done text\\<^term>\id\ as the identity relation\ -lemma id_iff [simp]: " \ id(A) \ x=y & y \ A" +lemma id_iff [simp]: " \ id(A) \ x=y \ y \ A" by auto diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/QPair.thy --- a/src/ZF/QPair.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/QPair.thy Tue Sep 27 17:03:23 2022 +0100 @@ -38,7 +38,7 @@ 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 @@ -77,7 +77,7 @@ lemma QPair_empty [simp]: "<0;0> = 0" by (simp add: QPair_def) -lemma QPair_iff [simp]: " = \ a=c & b=d" +lemma QPair_iff [simp]: " = \ a=c \ b=d" apply (simp add: QPair_def) apply (rule sum_equal_iff) done @@ -263,13 +263,13 @@ (** <+> is itself injective... who cares?? **) lemma qsum_iff: - "u \ A <+> B \ (\x. x \ A & u=QInl(x)) | (\y. y \ B & u=QInr(y))" + "u \ A <+> B \ (\x. x \ A \ u=QInl(x)) | (\y. y \ B \ u=QInr(y))" by blast -lemma qsum_subset_iff: "A <+> B \ C <+> D \ A<=C & B<=D" +lemma qsum_subset_iff: "A <+> B \ C <+> D \ A<=C \ B<=D" by blast -lemma qsum_equal_iff: "A <+> B = C <+> D \ A=C & B=D" +lemma qsum_equal_iff: "A <+> B = C <+> D \ A=C \ B=D" apply (simp (no_asm) add: extension qsum_subset_iff) apply blast done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/QUniv.thy --- a/src/ZF/QUniv.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/QUniv.thy Tue Sep 27 17:03:23 2022 +0100 @@ -27,7 +27,7 @@ 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 @@ -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]) @@ -130,7 +130,7 @@ lemmas quniv_QPair_E = quniv_QPair_D [THEN conjE] -lemma quniv_QPair_iff: " \ quniv(A) \ a: quniv(A) & b: quniv(A)" +lemma quniv_QPair_iff: " \ quniv(A) \ a: quniv(A) \ b: quniv(A)" by (blast intro: QPair_in_quniv dest: quniv_QPair_D) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Resid/Confluence.thy --- a/src/ZF/Resid/Confluence.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Resid/Confluence.thy Tue Sep 27 17:03:23 2022 +0100 @@ -8,12 +8,12 @@ definition confluence :: "i=>o" where "confluence(R) \ - \x y. \ R \ (\z. \ R \ (\u. \ R & \ 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)))" + (\z.(x =1=> z) \ (\u.(y =1=> u) \ (z=\u)))" (* ------------------------------------------------------------------------- *) @@ -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 e44d86131648 -r 0c18df79b1c8 src/ZF/Resid/Reduction.thy --- a/src/ZF/Resid/Reduction.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Resid/Reduction.thy Tue Sep 27 17:03:23 2022 +0100 @@ -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+) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Resid/Residuals.thy --- a/src/ZF/Resid/Residuals.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Resid/Residuals.thy Tue Sep 27 17:03:23 2022 +0100 @@ -199,8 +199,8 @@ subsection\paving theorem\ 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)" + \uv vu. (w|>u) |> vu = (w|>v) |> uv \ (w|>u) \ vu \ + regular(vu) \ (w|>v) \ uv \ regular(uv)" apply (subgoal_tac "u \ v") apply (safe intro!: exI) apply (rule cube) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Sum.thy --- a/src/ZF/Sum.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Sum.thy Tue Sep 27 17:03:23 2022 +0100 @@ -28,7 +28,7 @@ subsection\Rules for the \<^term>\Part\ Primitive\ lemma Part_iff: - "a \ Part(A,h) \ a \ A & (\y. a=h(y))" + "a \ Part(A,h) \ a \ A \ (\y. a=h(y))" apply (unfold Part_def) apply (rule separation) done @@ -106,7 +106,7 @@ 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))" +lemma sum_iff: "u \ A+B \ (\x. x \ A \ u=Inl(x)) | (\y. y \ B \ u=Inr(y))" by blast lemma Inl_in_sum_iff [simp]: "(Inl(x) \ A+B) \ (x \ A)" @@ -115,10 +115,10 @@ lemma Inr_in_sum_iff [simp]: "(Inr(y) \ A+B) \ (y \ B)" by auto -lemma sum_subset_iff: "A+B \ C+D \ A<=C & B<=D" +lemma sum_subset_iff: "A+B \ C+D \ A<=C \ B<=D" by blast -lemma sum_equal_iff: "A+B = C+D \ A=C & B=D" +lemma sum_equal_iff: "A+B = C+D \ A=C \ B=D" by (simp add: extension sum_subset_iff, blast) lemma sum_eq_2_times: "A+A = 2*A" @@ -142,7 +142,7 @@ lemma expand_case: "u \ A+B \ R(case(c,d,u)) \ - ((\x\A. u = Inl(x) \ R(c(x))) & + ((\x\A. u = Inl(x) \ R(c(x))) \ (\y\B. u = Inr(y) \ R(d(y))))" by auto diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Trancl.thy --- a/src/ZF/Trancl.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Trancl.thy Tue Sep 27 17:03:23 2022 +0100 @@ -46,7 +46,7 @@ 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\ @@ -185,7 +185,7 @@ "\ \ r^*; (a=b) \ P; \y.\ \ r^*; \ r\ \ P\ \ P" -apply (subgoal_tac "a = b | (\y. \ r^* & \ r) ") +apply (subgoal_tac "a = b | (\y. \ r^* \ \ r) ") (*see HOL/trancl*) apply blast apply (erule rtrancl_induct, blast+) @@ -257,7 +257,7 @@ \ r \ P; \y.\ \ r^+; \ r\ \ P \ \ P" -apply (subgoal_tac " \ r | (\y. \ r^+ & \ r) ") +apply (subgoal_tac " \ r | (\y. \ r^+ \ \ r) ") apply blast apply (rule compEpair) apply (unfold trancl_def, assumption) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/AllocBase.thy --- a/src/ZF/UNITY/AllocBase.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/AllocBase.thy Tue Sep 27 17:03:23 2022 +0100 @@ -71,12 +71,12 @@ subsection\Various simple lemmas\ -lemma Nclients_NbT_gt_0 [simp]: "0 < Nclients & 0 < NbT" +lemma Nclients_NbT_gt_0 [simp]: "0 < Nclients \ 0 < NbT" apply (cut_tac Nclients_pos NbT_pos) apply (auto intro: Ord_0_lt) done -lemma Nclients_NbT_not_0 [simp]: "Nclients \ 0 & NbT \ 0" +lemma Nclients_NbT_not_0 [simp]: "Nclients \ 0 \ NbT \ 0" by (cut_tac Nclients_pos NbT_pos, auto) lemma Nclients_type [simp,TC]: "Nclients\nat" @@ -94,7 +94,7 @@ (\i\nat. i f(i) $\ g(i)) \ setsum(f, n) $\ setsum(g,n)" apply (induct_tac "n", simp_all) -apply (subgoal_tac "Finite(x) & x\x") +apply (subgoal_tac "Finite(x) \ x\x") prefer 2 apply (simp add: nat_into_Finite mem_not_refl, clarify) apply (simp (no_asm_simp) add: succ_def) apply (subgoal_tac "\i\nat. i f(i) $\ g(i) ") @@ -135,7 +135,7 @@ 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 @@ -279,7 +279,7 @@ lemma all_distinct_Cons [simp]: "all_distinct(Cons(a,l)) \ - (a\set_of_list(l) \ False) & (a \ set_of_list(l) \ all_distinct(l))" + (a\set_of_list(l) \ False) \ (a \ set_of_list(l) \ all_distinct(l))" apply (unfold all_distinct_def) apply (auto elim: list.cases) done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/AllocImpl.thy --- a/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 17:03:23 2022 +0100 @@ -18,31 +18,31 @@ axiomatization where alloc_type_assumes [simp]: - "type_of(NbR) = nat & type_of(available_tok)=nat" and + "type_of(NbR) = nat \ type_of(available_tok)=nat" and alloc_default_val_assumes [simp]: - "default_val(NbR) = 0 & default_val(available_tok)=0" + "default_val(NbR) = 0 \ default_val(available_tok)=0" definition "alloc_giv_act \ { \ state*state. - \k. k = length(s`giv) & + \k. k = length(s`giv) \ t = s(giv := s`giv @ [nth(k, s`ask)], - available_tok := s`available_tok #- nth(k, s`ask)) & - k < length(s`ask) & nth(k, s`ask) \ s`available_tok}" + available_tok := s`available_tok #- nth(k, s`ask)) \ + k < length(s`ask) \ nth(k, s`ask) \ s`available_tok}" definition "alloc_rel_act \ { \ state*state. t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel), - NbR := succ(s`NbR)) & + NbR := succ(s`NbR)) \ s`NbR < length(s`rel)}" definition (*The initial condition s`giv=[] is missing from the original definition: S. O. Ehmety *) "alloc_prog \ - mk_program({s:state. s`available_tok=NbT & s`NbR=0 & s`giv=Nil}, + 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)) \ preserves(lift(NbR)) \ @@ -74,8 +74,8 @@ lemma alloc_prog_ok_iff: "\G \ program. (alloc_prog ok G) \ - (G \ preserves(lift(giv)) & G \ preserves(lift(available_tok)) & - G \ preserves(lift(NbR)) & alloc_prog \ Allowed(G))" + (G \ preserves(lift(giv)) \ G \ preserves(lift(available_tok)) \ + G \ preserves(lift(NbR)) \ alloc_prog \ Allowed(G))" by (auto simp add: ok_iff_Allowed alloc_prog_def [THEN def_prg_Allowed]) @@ -151,7 +151,7 @@ apply (subgoal_tac "G \ preserves (fun_pair (lift (available_tok), fun_pair (lift (NbR), lift (giv))))") apply (rotate_tac -1) apply (cut_tac A = "nat * nat * list(nat)" - and P = "% y. n \ length(y) & + and P = "% y. n \ length(y) \ m #+ tokens(l) = NbT #+ tokens(take(n,y))" and g = "lift(rel)" and F = alloc_prog in stable_Join_Stable) @@ -440,7 +440,7 @@ by auto (*needed?*) -lemma single_state_Diff_eq [simp]: "{s}-{x \ state. P(x)} = (if s\state & P(s) then 0 else {s})" +lemma single_state_Diff_eq [simp]: "{s}-{x \ state. P(x)} = (if s\state \ P(s) then 0 else {s})" by auto @@ -494,9 +494,9 @@ lemma (in alloc_progress) length_giv_disj: "\k \ tokbag; n \ nat\ \ alloc_prog \ G \ - {s\state. length(s`giv) = n & tokens(s`giv) = k} + {s\state. length(s`giv) = n \ tokens(s`giv) = k} \w - {s\state. (length(s`giv) = n & tokens(s`giv) = k & + {s\state. (length(s`giv) = n \ tokens(s`giv) = k \ k \ tokens(take(s`NbR, s`rel))) | n < length(s`giv)}" apply (rule single_LeadsTo_I, safe) apply (rule_tac a1 = "s`giv" in Increasing_imp_Stable [THEN PSP_StableI]) @@ -514,9 +514,9 @@ lemma (in alloc_progress) length_giv_disj2: "\k \ tokbag; n \ nat\ \ alloc_prog \ G \ - {s\state. length(s`giv) = n & tokens(s`giv) = k} + {s\state. length(s`giv) = n \ tokens(s`giv) = k} \w - {s\state. (length(s`giv) = n & NbT \ s`available_tok) | + {s\state. (length(s`giv) = n \ NbT \ s`available_tok) | n < length(s`giv)}" apply (rule LeadsTo_weaken_R) apply (rule Always_LeadsToD [OF alloc_prog_Always_lemma length_giv_disj], auto) @@ -529,7 +529,7 @@ \ alloc_prog \ G \ {s\state. length(s`giv) = n} \w - {s\state. (length(s`giv) = n & NbT \ s`available_tok) | + {s\state. (length(s`giv) = n \ NbT \ s`available_tok) | n < length(s`giv)}" apply (rule LeadsTo_weaken_L) apply (rule_tac I = nat in LeadsTo_UN) @@ -542,9 +542,9 @@ lemma (in alloc_progress) length_ask_giv: "\k \ nat; n < k\ \ alloc_prog \ G \ - {s\state. length(s`ask) = k & length(s`giv) = n} + {s\state. length(s`ask) = k \ length(s`giv) = n} \w - {s\state. (NbT \ s`available_tok & length(s`giv) < length(s`ask) & + {s\state. (NbT \ s`available_tok \ length(s`giv) < length(s`ask) \ length(s`giv) = n) | n < length(s`giv)}" apply (rule single_LeadsTo_I, safe) @@ -565,10 +565,10 @@ lemma (in alloc_progress) length_ask_giv2: "\k \ nat; n < k\ \ alloc_prog \ G \ - {s\state. length(s`ask) = k & length(s`giv) = n} + {s\state. length(s`ask) = k \ length(s`giv) = n} \w - {s\state. (nth(length(s`giv), s`ask) \ s`available_tok & - length(s`giv) < length(s`ask) & length(s`giv) = n) | + {s\state. (nth(length(s`giv), s`ask) \ s`available_tok \ + length(s`giv) < length(s`ask) \ length(s`giv) = n) | n < length(s`giv)}" apply (rule LeadsTo_weaken_R) apply (rule Always_LeadsToD [OF safety length_ask_giv], assumption+, clarify) @@ -582,7 +582,7 @@ lemma (in alloc_progress) extend_giv: "\k \ nat; n < k\ \ alloc_prog \ G \ - {s\state. length(s`ask) = k & length(s`giv) = n} + {s\state. length(s`ask) = k \ length(s`giv) = n} \w {s\state. n < length(s`giv)}" apply (rule LeadsTo_Un_duplicate) apply (rule LeadsTo_cancel1) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/ClientImpl.thy --- a/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 17:03:23 2022 +0100 @@ -14,11 +14,11 @@ axiomatization where type_assumes: - "type_of(ask) = list(tokbag) & type_of(giv) = list(tokbag) & - type_of(rel) = list(tokbag) & type_of(tok) = nat" and + "type_of(ask) = list(tokbag) \ type_of(giv) = list(tokbag) \ + type_of(rel) = list(tokbag) \ type_of(tok) = nat" and default_val_assumes: - "default_val(ask) = Nil & default_val(giv) = Nil & - default_val(rel) = Nil & default_val(tok) = 0" + "default_val(ask) = Nil \ default_val(giv) = Nil \ + default_val(rel) = Nil \ default_val(tok) = 0" (*Array indexing is translated to list indexing as A[n] \ nth(n-1,A). *) @@ -27,9 +27,9 @@ (** Release some client_tokens **) "client_rel_act \ { \ state*state. - \nrel \ nat. nrel = length(s`rel) & - t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) & - nrel < length(s`giv) & + \nrel \ nat. nrel = length(s`rel) \ + t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) \ + nrel < length(s`giv) \ nth(nrel, s`ask) \ nth(nrel, s`giv)}" (** Choose a new token requirement **) @@ -44,8 +44,8 @@ definition "client_prog \ - mk_program({s \ state. s`tok \ NbT & s`giv = Nil & - s`ask = Nil & s`rel = Nil}, + mk_program({s \ state. s`tok \ NbT \ s`giv = Nil \ + s`ask = Nil \ s`rel = Nil}, {client_rel_act, client_tok_act, client_ask_act}, \G \ preserves(lift(rel)) Int preserves(lift(ask)) Int @@ -92,8 +92,8 @@ lemma client_prog_ok_iff: "\G \ program. (client_prog ok G) \ - (G \ preserves(lift(rel)) & G \ preserves(lift(ask)) & - G \ preserves(lift(tok)) & client_prog \ Allowed(G))" + (G \ preserves(lift(rel)) \ G \ preserves(lift(ask)) \ + G \ preserves(lift(tok)) \ client_prog \ Allowed(G))" by (auto simp add: ok_iff_Allowed client_prog_def [THEN def_prg_Allowed]) lemma client_prog_preserves: @@ -186,8 +186,8 @@ lemma transient_lemma: "client_prog \ - transient({s \ state. s`rel = k & \ strict_prefix(nat) - & \ prefix(nat) & h pfixGe s`ask})" + transient({s \ state. s`rel = k \ \ strict_prefix(nat) + \ \ prefix(nat) \ h pfixGe s`ask})" apply (rule_tac act = client_rel_act in transientI) apply (simp (no_asm) add: client_prog_def [THEN def_prg_Acts]) apply (simp (no_asm) add: client_rel_act_def [THEN def_act_eq, THEN act_subset]) @@ -209,7 +209,7 @@ done lemma strict_prefix_is_prefix: - " \ strict_prefix(A) \ \ prefix(A) & xs\ys" + " \ strict_prefix(A) \ \ prefix(A) \ xs\ys" apply (unfold strict_prefix_def id_def lam_def) apply (auto dest: prefix_type [THEN subsetD]) done @@ -217,11 +217,11 @@ lemma induct_lemma: "\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} + {s \ state. s`rel = k \ \ strict_prefix(nat) + \ \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ strict_prefix(nat) - & \ prefix(nat) & - \ prefix(nat) & + \ \ prefix(nat) \ + \ prefix(nat) \ h pfixGe s`ask}" apply (rule single_LeadsTo_I) prefer 2 apply simp @@ -247,7 +247,7 @@ "\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} + \ \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ prefix(nat)}" apply (rule_tac f = "\x \ state. length(h) #- length(x`rel)" in LessThan_induct) @@ -275,7 +275,7 @@ lemma progress_lemma: "\client_prog \ G \ Incr(lift(giv)); client_prog ok G; G \ program\ \ client_prog \ G - \ {s \ state. \ prefix(nat) & h pfixGe s`ask} + \ {s \ state. \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ prefix(nat)}" apply (rule client_prog_Join_Always_rel_le_giv [THEN Always_LeadsToI], assumption) @@ -290,7 +290,7 @@ (*Progress property: all tokens that are given will be released*) lemma client_prog_progress: "client_prog \ Incr(lift(giv)) guarantees - (\h \ list(nat). {s \ state. \ prefix(nat) & + (\h \ list(nat). {s \ state. \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ prefix(nat)})" apply (rule guaranteesI) apply (blast intro: progress_lemma, auto) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Comp.thy --- a/src/ZF/UNITY/Comp.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Comp.thy Tue Sep 27 17:03:23 2022 +0100 @@ -23,16 +23,16 @@ 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 *) @@ -61,7 +61,7 @@ lemma component_eq_subset: "G \ program \ (F component G) \ - (Init(G) \ Init(F) & Acts(F) \ Acts(G) & AllowedActs(G) \ AllowedActs(F))" + (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) @@ -121,7 +121,7 @@ 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) @@ -129,7 +129,7 @@ lemma Join_component_iff: "H \ program \ - ((F \ G) component H) \ (F component H & G component H)" + ((F \ G) component H) \ (F component H \ G component H)" apply (simp (no_asm_simp) add: component_eq_subset) apply blast done @@ -158,13 +158,13 @@ lemma preserves_imp_eq: "\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") +apply (subgoal_tac "s \ state \ t \ state") prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) done lemma Join_preserves [iff]: "(F \ G \ preserves(v)) \ - (programify(F) \ preserves(v) & programify(G) \ preserves(v))" + (programify(F) \ preserves(v) \ programify(G) \ preserves(v))" by (auto simp add: preserves_def INT_iff) lemma JOIN_preserves [iff]: @@ -297,7 +297,7 @@ lemma Increasing_preserves_Stable: "\F \ stable({s \ state. :r}); G \ preserves(f); F \ G \ Increasing(A, r, g); - \x \ state. f(x):A & g(x):A\ + \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]) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Constrains.thy Tue Sep 27 17:03:23 2022 +0100 @@ -344,7 +344,7 @@ 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] diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Distributor.thy --- a/src/ZF/UNITY/Distributor.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Distributor.thy Tue Sep 27 17:03:23 2022 +0100 @@ -22,7 +22,7 @@ (\n \ Nclients. lift(Out(n)) Fols - (%s. sublist(s`In, {k \ nat. k nat. k nth(k, s`iIn) = n})) Wrt prefix(A)/list(A))" definition @@ -43,13 +43,13 @@ and A \ \the type of items being distributed\ and D assumes - var_assumes [simp]: "In \ var & iIn \ var & (\n. Out(n):var)" + var_assumes [simp]: "In \ var \ iIn \ var \ (\n. Out(n):var)" and all_distinct_vars: "\n. all_distinct([In, iIn, Out(n)])" - and type_assumes [simp]: "type_of(In)=list(A) & type_of(iIn)=list(nat) & + and type_assumes [simp]: "type_of(In)=list(A) \ type_of(iIn)=list(nat) \ (\n. type_of(Out(n))=list(A))" and default_val_assumes [simp]: - "default_val(In)=Nil & - default_val(iIn)=Nil & + "default_val(In)=Nil \ + default_val(iIn)=Nil \ (\n. default_val(Out(n))=Nil)" and distr_spec: "D \ distr_spec(A, In, iIn, Out)" @@ -79,7 +79,7 @@ lemma (in distr) D_ok_iff: "G \ program \ - D ok G \ ((\n \ nat. G \ preserves(lift(Out(n)))) & D \ Allowed(G))" + 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 Allowed_def ok_iff_Allowed) @@ -106,7 +106,7 @@ \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) & + {k \ nat. k < length(s`iIn) \ nth(k, s`iIn)= n})), Nclients, A) = bag_of(sublist(s`In, length(s`iIn)))})" @@ -122,7 +122,7 @@ apply (simp (no_asm_simp)) apply (simp add: set_of_list_conv_nth [of _ nat]) apply (subgoal_tac - "(\i \ Nclients. {k\nat. k < length(s`iIn) & nth(k, s`iIn) = i}) = + "(\i \ Nclients. {k\nat. k < length(s`iIn) \ nth(k, s`iIn) = i}) = length(s`iIn) ") apply (simp (no_asm_simp)) apply (rule equalityI) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Follows.thy --- a/src/ZF/UNITY/Follows.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Follows.thy Tue Sep 27 17:03:23 2022 +0100 @@ -51,7 +51,7 @@ 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) @@ -59,7 +59,7 @@ lemma imp_Always_comp: "\F \ Always({x \ state. \ r}); - 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\ \ F \ Always({x \ state. <(h comp f)(x), (h comp g)(x)> \ s})" by (blast intro: subset_Always_comp [THEN subsetD]) @@ -68,7 +68,7 @@ "\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\ + \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) @@ -78,7 +78,7 @@ lemma subset_LeadsTo_comp: "\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\ \ (\j \ A. {s \ state. \ r} \w {s \ state. \ r}) <= (\k \ B. {x \ state. \ s} \w {x \ state. \ s})" @@ -99,7 +99,7 @@ lemma imp_LeadsTo_comp: "\F:(\j \ A. {s \ state. \ r} \w {s \ state. \ r}); mono1(A, r, B, s, h); refl(A,r); trans[B](s); - \x \ state. f(x):A & g(x):A\ \ + \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 @@ -108,7 +108,7 @@ "\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) @@ -131,7 +131,7 @@ "\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) @@ -156,7 +156,7 @@ \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\ + \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) @@ -174,7 +174,7 @@ lemma FollowsD: "F \ Follows(A, r, f, g)\ - F \ program & (\a. a \ A) & (\x \ state. f(x):A & g(x):A)" + F \ program \ (\a. a \ A) \ (\x \ state. f(x):A \ g(x):A)" apply (unfold Follows_def) apply (blast dest: IncreasingD) done @@ -440,7 +440,7 @@ lemma Always_munion: "\F \ Always({s \ state. \ MultLe(A,r)}); F \ Always({s \ state. \ MultLe(A,r)}); - \x \ state. f1(x):Mult(A)&f(x):Mult(A) & g1(x):Mult(A) & g(x):Mult(A)\ + \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) @@ -456,8 +456,8 @@ lemma Follows_msetsum_UN: "\f. \\i \ I. F \ Follows(Mult(A), MultLe(A, r), f'(i), f(i)); - \s. \i \ I. multiset(f'(i, s)) & mset_of(f'(i, s))<=A & - multiset(f(i, s)) & mset_of(f(i, s))<=A ; + \s. \i \ I. multiset(f'(i, s)) \ mset_of(f'(i, s))<=A \ + multiset(f(i, s)) \ mset_of(f(i, s))<=A ; Finite(I); F \ program\ \ F \ Follows(Mult(A), MultLe(A, r), %x. msetsum(%i. f'(i, x), I, A), diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/GenPrefix.thy --- a/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 17:03:23 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" @@ -80,8 +80,8 @@ lemma Cons_gen_prefix_aux: "\ \ gen_prefix(A, r)\ \ (\x xs. x \ A \ xs'= Cons(x,xs) \ - (\y ys. y \ A & ys' = Cons(y,ys) & - :r & \ gen_prefix(A, r)))" + (\y ys. y \ A \ ys' = Cons(y,ys) \ + :r \ \ gen_prefix(A, r)))" apply (erule gen_prefix.induct) prefer 3 apply (force intro: gen_prefix.append, auto) done @@ -97,7 +97,7 @@ lemma Cons_gen_prefix_Cons: "( \ gen_prefix(A, r)) - \ (x \ A & y \ A & :r & \ gen_prefix(A, r))" + \ (x \ A \ y \ A \ :r \ \ gen_prefix(A, r))" apply (auto intro: gen_prefix.prepend) done declare Cons_gen_prefix_Cons [iff] @@ -213,7 +213,7 @@ prefer 2 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) apply (drule_tac psi = " \ gen_prefix (A,r) " in asm_rl) apply simp -apply (subgoal_tac "length (ys @ zs) = length (ys) #+ length (zs) &ys \ list (A) &xs \ list (A) ") +apply (subgoal_tac "length (ys @ zs) = length (ys) #+ length (zs) \ys \ list (A) \xs \ list (A) ") prefer 2 apply (blast intro: length_app dest: gen_prefix.dom_subset [THEN subsetD]) apply (drule gen_prefix_length_le)+ apply clarify @@ -240,7 +240,7 @@ 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)))" + (xs=[] | (\z zs. xs=Cons(z,zs) \ z \ A \ :r \ \ gen_prefix(A,r)))" apply (induct_tac "xs", auto) done @@ -276,7 +276,7 @@ (* Append case is hardest *) apply (frule gen_prefix_length_le [THEN le_iff [THEN iffD1]]) apply (frule gen_prefix.dom_subset [THEN subsetD], clarify) -apply (subgoal_tac "length (xs) :nat&length (ys) :nat &length (zs) :nat") +apply (subgoal_tac "length (xs) :nat\length (ys) :nat \length (zs) :nat") prefer 2 apply (blast intro: length_type, clarify) apply (simp_all add: nth_append length_type length_app) apply (rule conjI) @@ -325,7 +325,7 @@ done lemma gen_prefix_iff_nth: "( \ gen_prefix(A,r)) \ - (xs \ list(A) & ys \ list(A) & length(xs) \ length(ys) & + (xs \ list(A) \ ys \ list(A) \ length(xs) \ length(ys) \ (\i. i < length(xs) \ : r))" apply (rule iffI) apply (frule gen_prefix.dom_subset [THEN subsetD]) @@ -367,7 +367,7 @@ apply (unfold prefix_def) apply (erule gen_prefix.induct) -apply (subgoal_tac [3] "xs \ list (A) &ys \ list (A) ") +apply (subgoal_tac [3] "xs \ list (A) \ys \ list (A) ") prefer 4 apply (blast dest: gen_prefix.dom_subset [THEN subsetD]) apply (auto simp add: set_of_list_append) done @@ -392,7 +392,7 @@ declare prefix_Nil [iff] lemma Cons_prefix_Cons: -" \ prefix(A) \ (x=y & \ prefix(A) & y \ A)" +" \ prefix(A) \ (x=y \ \ prefix(A) \ y \ A)" apply (unfold prefix_def, auto) done declare Cons_prefix_Cons [iff] @@ -422,7 +422,7 @@ lemma prefix_Cons: "\xs \ list(A); ys \ list(A); y \ A\ \ \ prefix(A) \ - (xs=[] | (\zs. xs=Cons(y,zs) & \ prefix(A)))" + (xs=[] | (\zs. xs=Cons(y,zs) \ \ prefix(A)))" apply (unfold prefix_def) apply (auto simp add: gen_prefix_Cons) done @@ -457,7 +457,7 @@ " \ 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)") +apply (subgoal_tac [!] "ys \ list(A) \ xs \ list(A)") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: length_type) apply (subgoal_tac "length (zs) =0") @@ -475,7 +475,7 @@ (*Equivalence to the definition used in Lex/Prefix.thy*) lemma prefix_iff: - " \ prefix(A) \ (\ys \ list(A). zs = xs@ys) & xs \ list(A)" + " \ prefix(A) \ (\ys \ list(A). zs = xs@ys) \ xs \ list(A)" apply (unfold prefix_def) apply (auto simp add: gen_prefix_iff_nth lt_nat_in_nat nth_append nth_type app_type length_app) apply (subgoal_tac "drop (length (xs), zs) \ list (A) ") @@ -505,7 +505,7 @@ 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)))" + ( \ prefix(A) | (\us. xs = ys@us \ \ prefix(A)))" apply (erule list_append_induct, force, clarify) apply (rule iffI) apply (simp add: add: app_assoc [symmetric]) @@ -646,7 +646,7 @@ apply (erule natE, auto) done -lemma prefix_take_iff: " \ prefix(A) \ (xs=take(length(xs), ys) & xs \ list(A) & ys \ list(A))" +lemma prefix_take_iff: " \ prefix(A) \ (xs=take(length(xs), ys) \ xs \ list(A) \ ys \ list(A))" apply (rule iffI) apply (frule prefix_type [THEN subsetD]) apply (blast intro: prefix_imp_take, clarify) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Guar.thy --- a/src/ZF/UNITY/Guar.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Guar.thy Tue Sep 27 17:03:23 2022 +0100 @@ -38,23 +38,23 @@ 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 & - (SKIP \ X & (\F \ program. \G \ program. F ok G \ F \ X & G \ X \ (F \ G) \ X))" + "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 & - (SKIP \ X & (\F \ program. \G \ program. F ok G \(F \ X & G \ X) \ (F \ G \ X)))" + "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 @@ -69,7 +69,7 @@ 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 "\"*) @@ -79,7 +79,7 @@ definition refines :: "[i, i, i] => o" (\(3_ refines _ wrt _)\ [10,10,10] 10) where "G refines F wrt X \ - \H \ program. (F ok H & G ok H & F \ H \ welldef \ X) + \H \ program. (F ok H \ G ok H \ F \ H \ welldef \ X) \ (G \ H \ welldef \ X)" definition @@ -111,11 +111,11 @@ apply (auto intro: ok_sym) done -(*Chandy & Sanders take this as a definition*) +(*Chandy \ Sanders take this as a definition*) lemma ex_prop_finite: - "ex_prop(X) \ (X\program & - (\GG \ Fin(program). GG \ X \ 0 & OK(GG,(%G. G))\(\G \ GG. G) \ X))" + "ex_prop(X) \ (X\program \ + (\GG \ Fin(program). GG \ X \ 0 \ OK(GG,(%G. G))\(\G \ GG. G) \ X))" apply auto apply (blast intro: ex1 ex2 dest: ex_imp_subset_program)+ done @@ -123,7 +123,7 @@ (* Equivalent definition of ex_prop given at the end of section 3*) lemma ex_prop_equiv: "ex_prop(X) \ - X\program & (\G \ program. (G \ X \ (\H \ program. (G component_of H) \ H \ X)))" + X\program \ (\G \ program. (G \ X \ (\H \ program. (G component_of H) \ H \ X)))" apply (unfold ex_prop_def component_of_def, safe, force, force, blast) apply (subst Join_commute) apply (blast intro: ok_sym) @@ -138,7 +138,7 @@ lemma uv1 [rule_format]: "GG \ Fin(program) \ - (uv_prop(X)\ GG \ X & OK(GG, (%G. G)) \ (\G \ GG. G) \ X)" + (uv_prop(X)\ GG \ X \ OK(GG, (%G. G)) \ (\G \ GG. G) \ X)" apply (unfold uv_prop_def) apply (erule Fin_induct) apply (auto simp add: OK_cons_iff) @@ -146,7 +146,7 @@ lemma uv2 [rule_format]: "X\program \ - (\GG \ Fin(program). GG \ X & OK(GG,(%G. G)) \ (\G \ GG. G) \ X) + (\GG \ Fin(program). GG \ X \ OK(GG,(%G. G)) \ (\G \ GG. G) \ X) \ uv_prop(X)" apply (unfold uv_prop_def, auto) apply (drule_tac x = 0 in bspec, simp+) @@ -154,10 +154,10 @@ apply (force dest: ok_sym simp add: OK_iff_ok) done -(*Chandy & Sanders take this as a definition*) +(*Chandy \ Sanders take this as a definition*) lemma uv_prop_finite: -"uv_prop(X) \ X\program & - (\GG \ Fin(program). GG \ X & OK(GG, %G. G) \ (\G \ GG. G) \ X)" +"uv_prop(X) \ X\program \ + (\GG \ Fin(program). GG \ X \ OK(GG, %G. G) \ (\G \ GG. G) \ X)" apply auto apply (blast dest: uv_imp_subset_program) apply (blast intro: uv1) @@ -388,7 +388,7 @@ lemma refines_refl: "F refines F wrt X" by (unfold refines_def, blast) -(* More results on guarantees, added by Sidi Ehmety from Chandy & Sander, section 6 *) +(* More results on guarantees, added by Sidi Ehmety from Chandy \ Sander, section 6 *) lemma wg_type: "wg(F, X) \ program" by (unfold wg_def, auto) @@ -396,14 +396,14 @@ 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 lemma guarantees_equiv: "(F \ X guarantees Y) \ - F \ program & (\H \ program. H \ X \ (F component_of H \ H \ Y))" + F \ program \ (\H \ program. H \ X \ (F component_of H \ H \ Y))" by (unfold guar_def component_of_def, force) lemma wg_weakest: @@ -415,7 +415,7 @@ lemma wg_equiv: "H \ wg(F,X) \ - ((F component_of H \ H \ X) & F \ program & H \ program)" + ((F component_of H \ H \ X) \ F \ program \ H \ program)" apply (simp add: wg_def guarantees_equiv) apply (rule iffI, safe) apply (rule_tac [4] x = "{H}" in bexI) @@ -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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Increasing.thy --- a/src/ZF/UNITY/Increasing.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Increasing.thy Tue Sep 27 17:03:23 2022 +0100 @@ -13,13 +13,13 @@ definition increasing :: "[i, i, i=>i] => i" (\increasing[_]'(_, _')\) where "increasing[A](r, f) \ - {F \ program. (\k \ A. F \ stable({s \ state. \ r})) & + {F \ program. (\k \ A. F \ stable({s \ state. \ r})) \ (\x \ state. f(x):A)}" definition Increasing :: "[i, i, i=>i] => i" (\Increasing[_]'(_, _')\) where "Increasing[A](r, f) \ - {F \ program. (\k \ A. F \ Stable({s \ state. \ r})) & + {F \ program. (\k \ A. F \ Stable({s \ state. \ r})) \ (\x \ state. f(x):A)}" abbreviation (input) @@ -40,14 +40,14 @@ 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) done lemma increasing_constant [simp]: - "F \ increasing[A](r, %s. c) \ F \ program & c \ A" + "F \ increasing[A](r, %s. c) \ F \ program \ c \ A" apply (unfold increasing_def stable_def) apply (subgoal_tac "\x. x \ state") apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state) @@ -109,14 +109,14 @@ 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) done lemma Increasing_constant [simp]: - "F \ Increasing[A](r, %s. c) \ F \ program & (c \ A)" + "F \ Increasing[A](r, %s. c) \ F \ program \ (c \ A)" apply (subgoal_tac "\x. x \ state") apply (auto dest!: IncreasingD intro: st0_in_state increasing_imp_Increasing) done @@ -127,7 +127,7 @@ apply (unfold Increasing_def Stable_def Constrains_def part_order_def constrains_def mono1_def metacomp_def, safe) apply (simp_all add: ActsD) -apply (subgoal_tac "xb \ state & xa \ state") +apply (subgoal_tac "xb \ state \ xa \ state") prefer 2 apply (simp add: ActsD) apply (subgoal_tac ":r") prefer 2 apply (force simp add: refl_def) @@ -168,9 +168,9 @@ part_order_def constrains_def mono2_def, clarify, simp) apply clarify apply (rename_tac xa xb) -apply (subgoal_tac "xb \ state & xa \ state") +apply (subgoal_tac "xb \ state \ xa \ state") prefer 2 apply (blast dest!: ActsD) -apply (subgoal_tac ":r & :s") +apply (subgoal_tac ":r \ :s") prefer 2 apply (force simp add: refl_def) apply (rotate_tac 6) apply (drule_tac x = "f (xb) " in bspec) @@ -199,9 +199,9 @@ apply (unfold Increasing_def stable_def part_order_def constrains_def mono2_def Stable_def Constrains_def, safe) apply (simp_all add: ActsD) -apply (subgoal_tac "xa \ state & x \ state") +apply (subgoal_tac "xa \ state \ x \ state") prefer 2 apply (blast dest!: ActsD) -apply (subgoal_tac ":r & :s") +apply (subgoal_tac ":r \ :s") prefer 2 apply (force simp add: refl_def) apply (rotate_tac 6) apply (drule_tac x = "f (xa) " in bspec) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Merge.thy --- a/src/ZF/UNITY/Merge.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Merge.thy Tue Sep 27 17:03:23 2022 +0100 @@ -38,7 +38,7 @@ (\n \ Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A)) guarantees (\n \ Nclients. - (%s. sublist(s`Out, {k \ nat. k < length(s`iOut) & + (%s. sublist(s`Out, {k \ nat. k < length(s`iOut) \ nth(k, s`iOut) = n})) Fols lift(In(n)) Wrt prefix(A)/list(A))" @@ -70,16 +70,16 @@ and A \ \the type of items being merged\ and M assumes var_assumes [simp]: - "(\n. In(n):var) & Out \ var & iOut \ var" + "(\n. In(n):var) \ Out \ var \ iOut \ var" and all_distinct_vars: "\n. all_distinct([In(n), Out, iOut])" and type_assumes [simp]: - "(\n. type_of(In(n))=list(A)) & - type_of(Out)=list(A) & + "(\n. type_of(In(n))=list(A)) \ + type_of(Out)=list(A) \ type_of(iOut)=list(nat)" and default_val_assumes [simp]: - "(\n. default_val(In(n))=Nil) & - default_val(Out)=Nil & + "(\n. default_val(In(n))=Nil) \ + default_val(Out)=Nil \ default_val(iOut)=Nil" and merge_spec: "M \ merge_spec(A, In, Out, iOut)" @@ -114,8 +114,8 @@ lemma (in merge) M_ok_iff: "G \ program \ - M ok G \ (G \ preserves(lift(Out)) & - G \ preserves(lift(iOut)) & M \ Allowed(G))" + M ok G \ (G \ preserves(lift(Out)) \ + G \ preserves(lift(iOut)) \ M \ Allowed(G))" apply (cut_tac merge_spec) apply (auto simp add: merge_Allowed ok_iff_Allowed) done @@ -147,14 +147,14 @@ 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})), + {k \ nat. k < length(s`iOut) \ nth(k, s`iOut)=i})), Nclients, A) = bag_of(s`Out)})" apply (rule Always_Diff_Un_eq [THEN iffD1]) apply (rule_tac [2] state_AlwaysI [THEN Always_weaken]) apply (rule Always_Int_I [OF merge_Always_Out_eq_iOut merge_Bounded], auto) apply (subst bag_of_sublist_UN_disjoint [symmetric]) apply (auto simp add: nat_into_Finite set_of_list_conv_nth [OF iOut_value_type]) -apply (subgoal_tac " (\i \ Nclients. {k \ nat. k < length (x`iOut) & nth (k, x`iOut) = i}) = length (x`iOut) ") +apply (subgoal_tac " (\i \ Nclients. {k \ nat. k < length (x`iOut) \ nth (k, x`iOut) = i}) = length (x`iOut) ") apply (auto simp add: sublist_upt_eq_take [OF Out_value_type] length_type [OF iOut_value_type] take_all [OF _ Out_value_type] diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Monotonicity.thy --- a/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 17:03:23 2022 +0100 @@ -14,7 +14,7 @@ definition mono1 :: "[i, i, i, i, i=>i] => o" where "mono1(A, r, B, s, f) \ - (\x \ A. \y \ A. \ r \ \ s) & (\x \ A. f(x) \ B)" + (\x \ A. \y \ A. \ r \ \ s) \ (\x \ A. f(x) \ B)" (* monotonicity of a 2-place meta-function f *) @@ -22,7 +22,7 @@ mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o" where "mono2(A, r, B, s, C, t, f) \ (\x \ A. \y \ A. \u \ B. \v \ B. - \ r & \ s \ \ t) & + \ r \ \ s \ \ t) \ (\x \ A. \y \ B. f(x,y) \ C)" (* Internalized relations on sets and multisets *) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/MultisetSum.thy --- a/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 17:03:23 2022 +0100 @@ -11,7 +11,7 @@ definition lcomm :: "[i, i, [i,i]=>i]=>o" where "lcomm(A, B, f) \ - (\x \ A. \y \ A. \z \ B. f(x, f(y, z))= f(y, f(x, z))) & + (\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 @@ -57,7 +57,7 @@ (** msetsum **) lemma multiset_general_setsum: - "\C \ Fin(A); \x \ A. multiset(g(x))& mset_of(g(x))\B\ + "\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) @@ -68,7 +68,7 @@ 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\ + "\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) @@ -81,20 +81,20 @@ by (simp add: msetsum_def) lemma mset_of_msetsum: - "\C \ Fin(A); \x \ A. multiset(g(x)) & mset_of(g(x))\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\ + "\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) ") apply (auto simp add: msetsum_multiset) -apply (subgoal_tac "y \ D \ Fin (A) & y \ D \ Fin (A) ") +apply (subgoal_tac "y \ D \ Fin (A) \ y \ D \ Fin (A) ") apply clarify apply (case_tac "x \ D") apply (subgoal_tac [2] "cons (x, y) \ D = y \ D") @@ -107,7 +107,7 @@ lemma msetsum_Un_disjoint: "\C \ Fin(A); D \ Fin(A); C \ D = 0; - \x \ A. multiset(g(x)) & mset_of(g(x))\B\ + \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) @@ -119,7 +119,7 @@ lemma msetsum_UN_disjoint [rule_format (no_asm)]: "\I \ Fin(K); \i \ K. C(i) \ Fin(A)\ \ - (\x \ A. multiset(f(x)) & mset_of(f(x))\B) \ + (\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)" apply (erule Fin_induct, auto) @@ -127,10 +127,10 @@ prefer 2 apply blast apply (subgoal_tac "C(x) \ (\i \ y. C (i)) = 0") prefer 2 apply blast -apply (subgoal_tac " (\i \ y. C (i)):Fin (A) & C(x) :Fin (A) ") +apply (subgoal_tac " (\i \ y. C (i)):Fin (A) \ C(x) :Fin (A) ") prefer 2 apply (blast intro: UN_Fin_lemma dest: FinD, clarify) apply (simp (no_asm_simp) add: msetsum_Un_disjoint) -apply (subgoal_tac "\x \ K. multiset (msetsum (f, C(x), B)) & mset_of (msetsum (f, C(x), B)) \ B") +apply (subgoal_tac "\x \ K. multiset (msetsum (f, C(x), B)) \ mset_of (msetsum (f, C(x), B)) \ B") apply (simp (no_asm_simp)) apply clarify apply (drule_tac x = xa in bspec) @@ -139,10 +139,10 @@ lemma msetsum_addf: "\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(f(x)) \ mset_of(f(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 (subgoal_tac "\x \ A. multiset (f(x) +# g(x)) \ mset_of (f(x) +# g(x)) \ B") apply (erule Fin_induct) apply (auto simp add: munion_ac) done @@ -156,10 +156,10 @@ by (simp add: multiset_equality) lemma msetsum_Un: "\C \ Fin(A); D \ Fin(A); - \x \ A. multiset(f(x)) & mset_of(f(x)) \ B\ + \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 (subgoal_tac "C \ D \ Fin (A) \ C \ D \ Fin (A) ") apply clarify apply (subst msetsum_Un_Int [symmetric]) apply (simp_all (no_asm_simp) add: msetsum_multiset multiset_union_diff) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Mutex.thy --- a/src/ZF/UNITY/Mutex.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Mutex.thy Tue Sep 27 17:03:23 2022 +0100 @@ -28,65 +28,65 @@ abbreviation "v \ Var([1,0])" axiomatization where \ \Type declarations\ - p_type: "type_of(p)=bool & default_val(p)=0" and - m_type: "type_of(m)=int & default_val(m)=#0" and - n_type: "type_of(n)=int & default_val(n)=#0" and - u_type: "type_of(u)=bool & default_val(u)=0" and - v_type: "type_of(v)=bool & default_val(v)=0" + p_type: "type_of(p)=bool \ default_val(p)=0" and + m_type: "type_of(m)=int \ default_val(m)=#0" and + n_type: "type_of(n)=int \ default_val(n)=#0" and + u_type: "type_of(u)=bool \ default_val(u)=0" and + v_type: "type_of(v)=bool \ default_val(v)=0" 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)) - & (s`m = #3 \ s`p=0)}" + "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)) - & (s`n = #3 \ s`p=1)}" + "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))& - (#3 $\ s`m & s`m $\ #4 \ s`p=0)}" + "bad_IU \ {s:state. (s`u = 1 \ (#1 $\ s`m \ s`m $\ #3))\ + (#3 $\ s`m \ s`m $\ #4 \ s`p=0)}" (** Variables' types **) @@ -174,7 +174,7 @@ 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 @@ -206,10 +206,10 @@ by (unfold op_Unless_def Mutex_def, safety) lemma U_F1: - "Mutex \ {s \ state. s`m=#1} \w {s \ state. s`p = s`v & s`m = #2}" + "Mutex \ {s \ state. s`m=#1} \w {s \ state. s`p = s`v \ s`m = #2}" by (unfold Mutex_def, ensures U1) -lemma U_F2: "Mutex \ {s \ state. s`p =0 & s`m = #2} \w {s \ state. s`m = #3}" +lemma U_F2: "Mutex \ {s \ state. s`p =0 \ s`m = #2} \w {s \ state. s`m = #3}" apply (cut_tac IU) apply (unfold Mutex_def, ensures U2) 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) @@ -243,7 +243,7 @@ done -lemma U_lemma123: "Mutex \ {s \ state. #1 $\ s`m & s`m $\ #3} \w {s \ state. s`p=1}" +lemma U_lemma123: "Mutex \ {s \ state. #1 $\ s`m \ s`m $\ #3} \w {s \ state. s`p=1}" by (simp add: eq_123 Collect_disj_eq LeadsTo_Un_distrib U_lemma1 U_lemma2 U_F3) @@ -257,10 +257,10 @@ lemma V_F0: "Mutex \ {s \ state. s`n=#2} Unless {s \ state. s`n=#3}" by (unfold op_Unless_def Mutex_def, safety) -lemma V_F1: "Mutex \ {s \ state. s`n=#1} \w {s \ state. s`p = not(s`u) & s`n = #2}" +lemma V_F1: "Mutex \ {s \ state. s`n=#1} \w {s \ state. s`p = not(s`u) \ s`n = #2}" by (unfold Mutex_def, ensures "V1") -lemma V_F2: "Mutex \ {s \ state. s`p=1 & s`n = #2} \w {s \ state. s`n = #3}" +lemma V_F2: "Mutex \ {s \ state. s`p=1 \ s`n = #2} \w {s \ state. s`n = #3}" apply (cut_tac IV) apply (unfold Mutex_def, ensures "V2") done @@ -282,7 +282,7 @@ lemma V_lemma1: "Mutex \ {s \ state. s`n = #1} \w {s \ state. s`p = 0}" by (rule LeadsTo_Trans [OF V_F1 [THEN LeadsTo_weaken_R] V_lemma2], blast) -lemma V_lemma123: "Mutex \ {s \ state. #1 $\ s`n & s`n $\ #3} \w {s \ state. s`p = 0}" +lemma V_lemma123: "Mutex \ {s \ state. #1 $\ s`n \ s`n $\ #3} \w {s \ state. s`p = 0}" by (simp add: eq_123 Collect_disj_eq LeadsTo_Un_distrib V_lemma1 V_lemma2 V_F3) (*Misra's F4*) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/State.thy --- a/src/ZF/UNITY/State.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/State.thy Tue Sep 27 17:03:23 2022 +0100 @@ -51,7 +51,7 @@ (* Union *) -lemma st_set_Un_iff [iff]: "st_set(A \ B) \ st_set(A) & st_set(B)" +lemma st_set_Un_iff [iff]: "st_set(A \ B) \ st_set(A) \ st_set(B)" by (simp add: st_set_def, auto) lemma st_set_Union_iff [iff]: "st_set(\(S)) \ (\A \ S. st_set(A))" @@ -99,7 +99,7 @@ lemma st_set_compl [simp]: "st_set(st_compl(A))" by (simp add: st_compl_def, auto) -lemma st_compl_iff [simp]: "x \ st_compl(A) \ x \ state & x \ A" +lemma st_compl_iff [simp]: "x \ st_compl(A) \ x \ state \ x \ A" by (simp add: st_compl_def) lemma st_compl_Collect [simp]: @@ -108,7 +108,7 @@ (*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 e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/SubstAx.thy Tue Sep 27 17:03:23 2022 +0100 @@ -153,14 +153,14 @@ done (** Distributive laws **) -lemma LeadsTo_Un_distrib: "(F \ (A \ B) \w C) \ (F \ A \w C & F \ B \w C)" +lemma LeadsTo_Un_distrib: "(F \ (A \ B) \w C) \ (F \ A \w C \ F \ B \w C)" by (blast intro: LeadsTo_Un LeadsTo_weaken_L) -lemma LeadsTo_UN_distrib: "(F \ (\i \ I. A(i)) \w B) \ (\i \ I. F \ A(i) \w B) & F \ program" +lemma LeadsTo_UN_distrib: "(F \ (\i \ I. A(i)) \w B) \ (\i \ I. F \ A(i) \w B) \ F \ program" by (blast dest: LeadsTo_type [THEN subsetD] intro: LeadsTo_UN LeadsTo_weaken_L) -lemma LeadsTo_Union_distrib: "(F \ \(S) \w B) \ (\A \ S. F \ A \w B) & F \ program" +lemma LeadsTo_Union_distrib: "(F \ \(S) \w B) \ (\A \ S. F \ A \w B) \ F \ program" by (blast dest: LeadsTo_type [THEN subsetD] intro: LeadsTo_Union LeadsTo_weaken_L) @@ -353,7 +353,7 @@ ORELSE (*subgoal may involve LeadsTo, leadsTo or ensures*) REPEAT (ares_tac ctxt [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis}, @{thm EnsuresI}, @{thm ensuresI}] 1), - (*now there are two subgoals: co & transient*) + (*now there are two subgoals: co \ transient*) simp_tac (ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\program\)) 2, Rule_Insts.res_inst_tac ctxt [((("act", 0), Position.none), sact)] [] @{thm transientI} 2, diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/UNITY.thy --- a/src/ZF/UNITY/UNITY.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/UNITY.thy Tue Sep 27 17:03:23 2022 +0100 @@ -16,7 +16,7 @@ program :: i where "program \ {: Pow(state) * Pow(Pow(state*state)) * Pow(Pow(state*state)). - id(state) \ acts & id(state) \ allowed}" + id(state) \ acts \ id(state) \ allowed}" definition mk_program :: "[i,i,i]=>i" where @@ -69,7 +69,7 @@ "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\ @@ -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\, @@ -307,7 +307,7 @@ lemma program_equality_iff: "\F \ program; G \ program\ \(F=G) \ - (Init(F) = Init(G) & Acts(F) = Acts(G) & AllowedActs(F) = AllowedActs(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\ @@ -328,8 +328,8 @@ lemma def_prg_simps: "\F \ mk_program (init,acts,allowed)\ - \ Init(F) = init \ state & - Acts(F) = cons(id(state), acts \ Pow(state*state)) & + \ Init(F) = init \ state \ + Acts(F) = cons(id(state), acts \ Pow(state*state)) \ AllowedActs(F) = cons(id(state), allowed \ Pow(state*state))" by auto @@ -337,7 +337,7 @@ 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'))" by auto text\A set is expanded only if an element is being tested against it\ @@ -359,21 +359,21 @@ "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" by (force simp add: constrains_def st_set_def) -lemma constrains_empty2 [iff]: "(F \ A co 0) \ (A=0 & F \ program)" +lemma constrains_empty2 [iff]: "(F \ A co 0) \ (A=0 \ F \ program)" by (force simp add: constrains_def st_set_def) -lemma constrains_state [iff]: "(F \ state co B) \ (state\B & F \ program)" +lemma constrains_state [iff]: "(F \ state co B) \ (state\B \ F \ program)" apply (cut_tac F = F in Acts_type) apply (force simp add: constrains_def st_set_def) done -lemma constrains_state2 [iff]: "F \ A co state \ (F \ program & st_set(A))" +lemma constrains_state2 [iff]: "F \ A co state \ (F \ program \ st_set(A))" apply (cut_tac F = F in Acts_type) apply (force simp add: constrains_def st_set_def) done @@ -500,7 +500,7 @@ 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" @@ -571,10 +571,10 @@ 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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/Union.thy --- a/src/ZF/UNITY/Union.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/Union.thy Tue Sep 27 17:03:23 2022 +0100 @@ -16,7 +16,7 @@ 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 @@ -37,8 +37,8 @@ definition (*Characterizes safety properties. Used with specifying AllowedActs*) safety_prop :: "i => o" where - "safety_prop(X) \ X\program & - SKIP \ X & (\G \ program. Acts(G) \ (\F \ X. Acts(F)) \ G \ X)" + "safety_prop(X) \ X\program \ + SKIP \ X \ (\G \ program. Acts(G) \ (\F \ X. Acts(F)) \ G \ X)" syntax "_JOIN1" :: "[pttrns, i] => i" (\(3\_./ _)\ 10) @@ -71,7 +71,7 @@ subsection\SKIP and safety properties\ -lemma SKIP_in_constrains_iff [iff]: "(SKIP \ A co B) \ (A\B & st_set(A))" +lemma SKIP_in_constrains_iff [iff]: "(SKIP \ A co B) \ (A\B \ st_set(A))" by (unfold constrains_def st_set_def, auto) lemma SKIP_in_Constrains_iff [iff]: "(SKIP \ A Co B)\ (state \ A\B)" @@ -231,12 +231,12 @@ done lemma Join_constrains [iff]: - "(F \ G \ A co B) \ (programify(F) \ A co B & programify(G) \ A co B)" + "(F \ G \ A co B) \ (programify(F) \ A co B \ programify(G) \ A co B)" by (auto simp add: constrains_def) lemma Join_unless [iff]: "(F \ G \ A unless B) \ - (programify(F) \ A unless B & programify(G) \ A unless B)" + (programify(F) \ A unless B \ programify(G) \ A unless B)" by (simp add: Join_constrains unless_def) @@ -247,7 +247,7 @@ lemma Join_constrains_weaken: "\F \ A co A'; G \ B co B'\ \ F \ G \ (A \ B) co (A' \ B')" -apply (subgoal_tac "st_set (A) & st_set (B) & F \ program & G \ program") +apply (subgoal_tac "st_set (A) \ st_set (B) \ F \ program \ G \ program") prefer 2 apply (blast dest: constrainsD2, simp) apply (blast intro: constrains_weaken) done @@ -267,7 +267,7 @@ done lemma JOIN_stable: - "(\i \ I. F(i)) \ stable(A) \ ((\i \ I. programify(F(i)) \ stable(A)) & st_set(A))" + "(\i \ I. F(i)) \ stable(A) \ ((\i \ I. programify(F(i)) \ stable(A)) \ st_set(A))" apply (auto simp add: stable_def constrains_def JOIN_def) apply (cut_tac F = "F (i) " in Acts_type) apply (drule_tac x = act in bspec, auto) @@ -298,7 +298,7 @@ lemma Join_stable [iff]: " (F \ G \ stable(A)) \ - (programify(F) \ stable(A) & programify(G) \ stable(A))" + (programify(F) \ stable(A) \ programify(G) \ stable(A))" by (simp add: stable_def) lemma initially_JoinI [intro!]: @@ -308,7 +308,7 @@ lemma invariant_JoinI: "\F \ invariant(A); G \ invariant(A)\ \ F \ G \ invariant(A)" -apply (subgoal_tac "F \ program&G \ program") +apply (subgoal_tac "F \ program\G \ program") prefer 2 apply (blast dest: invariantD2) apply (simp add: invariant_def) apply (auto intro: Join_in_program) @@ -347,14 +347,14 @@ lemma JOIN_ensures: "i \ I \ (\i \ I. F(i)) \ A ensures B \ - ((\i \ I. programify(F(i)) \ (A-B) co (A \ B)) & + ((\i \ I. programify(F(i)) \ (A-B) co (A \ B)) \ (\i \ I. programify(F(i)) \ A ensures B))" by (auto simp add: ensures_def JOIN_constrains JOIN_transient) lemma Join_ensures: "F \ G \ A ensures B \ - (programify(F) \ (A-B) co (A \ B) & programify(G) \ (A-B) co (A \ B) & + (programify(F) \ (A-B) co (A \ B) \ programify(G) \ (A-B) co (A \ B) \ (programify(F) \ transient (A-B) | programify(G) \ transient (A-B)))" apply (unfold ensures_def) @@ -373,7 +373,7 @@ weaker than G \ stable A *) lemma stable_Join_Always1: "\F \ stable(A); G \ invariant(A)\ \ F \ G \ Always(A)" -apply (subgoal_tac "F \ program & G \ program & st_set (A) ") +apply (subgoal_tac "F \ program \ G \ program \ st_set (A) ") prefer 2 apply (blast dest: invariantD2 stableD2) apply (simp add: Always_def invariant_def initially_def Stable_eq_stable) apply (force intro: stable_Int) @@ -390,7 +390,7 @@ lemma stable_Join_ensures1: "\F \ stable(A); G \ A ensures B\ \ F \ G \ A ensures B" -apply (subgoal_tac "F \ program & G \ program & st_set (A) ") +apply (subgoal_tac "F \ program \ G \ program \ st_set (A) ") prefer 2 apply (blast dest: stableD2 ensures_type [THEN subsetD]) apply (simp (no_asm_simp) add: Join_ensures) apply (simp add: stable_def ensures_def) @@ -414,7 +414,7 @@ by (auto dest: Acts_type [THEN subsetD] simp add: ok_def) lemma ok_Join_commute: - "(F ok G & (F \ G) ok H) \ (G ok H & F ok (G \ H))" + "(F ok G \ (F \ G) ok H) \ (G ok H \ F ok (G \ H))" by (auto simp add: ok_def) lemma ok_commute: "(F ok G) \(G ok F)" @@ -422,14 +422,14 @@ lemmas ok_sym = ok_commute [THEN iffD1] -lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, snd) \ (F ok G & (F \ G) ok H)" +lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, snd) \ (F ok G \ (F \ G) ok H)" by (simp add: ok_def Join_def OK_def Int_assoc cons_absorb Int_Un_distrib2 Ball_def, safe, force+) -lemma ok_Join_iff1 [iff]: "F ok (G \ H) \ (F ok G & F ok H)" +lemma ok_Join_iff1 [iff]: "F ok (G \ H) \ (F ok G \ F ok H)" by (auto simp add: ok_def) -lemma ok_Join_iff2 [iff]: "(G \ H) ok F \ (G ok F & H ok F)" +lemma ok_Join_iff2 [iff]: "(G \ H) ok F \ (G ok F \ H ok F)" by (auto simp add: ok_def) (*useful? Not with the previous two around*) @@ -456,7 +456,7 @@ lemma OK_cons_iff: "OK(cons(i, I), F) \ - (i \ I & OK(I, F)) | (i\I & OK(I, F) & F(i) ok JOIN(I,F))" + (i \ I \ OK(I, F)) | (i\I \ OK(I, F) \ F(i) ok JOIN(I,F))" apply (simp add: OK_iff_ok) apply (blast intro: ok_sym) done @@ -480,7 +480,7 @@ done lemma ok_iff_Allowed: - "F ok G \ (programify(F) \ Allowed(programify(G)) & + "F ok G \ (programify(F) \ Allowed(programify(G)) \ programify(G) \ Allowed(programify(F)))" by (simp add: ok_def Allowed_def) @@ -525,7 +525,7 @@ (*For safety_prop to hold, the property must be satisfiable!*) lemma safety_prop_constrains [iff]: - "safety_prop(A co B) \ (A \ B & st_set(A))" + "safety_prop(A co B) \ (A \ B \ st_set(A))" by (simp add: safety_prop_def constrains_def st_set_def, blast) (* To be used with resolution *) @@ -566,7 +566,7 @@ 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 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 e44d86131648 -r 0c18df79b1c8 src/ZF/UNITY/WFair.thy --- a/src/ZF/UNITY/WFair.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/UNITY/WFair.thy Tue Sep 27 17:03:23 2022 +0100 @@ -17,8 +17,8 @@ (* This definition specifies weak fairness. The rest of the theory is generic to all forms of fairness.*) transient :: "i=>i" where - "transient(A) \{F \ program. (\act\Acts(F). A<=domain(act) & - act``A \ state-A) & st_set(A)}" + "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 @@ -65,7 +65,7 @@ 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 @@ -127,7 +127,7 @@ 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'" @@ -164,7 +164,7 @@ 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 @@ -240,14 +240,14 @@ 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)" +lemma leadsTo_refl_iff: "F \ A \ A \ F \ program \ st_set(A)" by (auto intro: leadsTo_refl dest: leadsToD2) -lemma empty_leadsTo: "F \ 0 \ B \ (F \ program & st_set(B))" +lemma empty_leadsTo: "F \ 0 \ B \ (F \ program \ st_set(B))" by (auto intro: subset_imp_leadsTo dest: leadsToD2) declare empty_leadsTo [iff] -lemma leadsTo_state: "F \ A \ state \ (F \ program & st_set(A))" +lemma leadsTo_state: "F \ A \ state \ (F \ program \ st_set(A))" by (auto intro: subset_imp_leadsTo dest: leadsToD2 st_setD) declare leadsTo_state [iff] @@ -272,15 +272,15 @@ done (*Distributes over binary unions*) -lemma leadsTo_Un_distrib: "F:(A \ B) \ C \ (F \ A \ C & F \ B \ C)" +lemma leadsTo_Un_distrib: "F:(A \ B) \ C \ (F \ A \ C \ F \ B \ C)" by (blast intro: leadsTo_Un leadsTo_weaken_L) lemma leadsTo_UN_distrib: -"(F:(\i \ I. A(i)) \ B)\ ((\i \ I. F \ A(i) \ B) & F \ program & st_set(B))" +"(F:(\i \ I. A(i)) \ B)\ ((\i \ I. F \ A(i) \ B) \ F \ program \ st_set(B))" apply (blast dest: leadsToD2 intro: leadsTo_UN leadsTo_weaken_L) done -lemma leadsTo_Union_distrib: "(F \ \(S) \ B) \ (\A \ S. F \ A \ B) & F \ program & st_set(B)" +lemma leadsTo_Union_distrib: "(F \ \(S) \ B) \ (\A \ S. F \ A \ B) \ F \ program \ st_set(B)" by (blast dest: leadsToD2 intro: leadsTo_Union leadsTo_weaken_L) text\Set difference: maybe combine with \leadsTo_weaken_L\??\ @@ -298,14 +298,14 @@ (*Binary union version*) lemma leadsTo_Un_Un: "\F \ A \ A'; F \ B \ B'\ \ F \ (A \ B) \ (A' \ B')" -apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') ") +apply (subgoal_tac "st_set (A) \ st_set (A') \ st_set (B) \ st_set (B') ") prefer 2 apply (blast dest: leadsToD2) apply (blast intro: leadsTo_Un leadsTo_weaken_R) done (** The cancellation law **) lemma leadsTo_cancel2: "\F \ A \ (A' \ B); F \ B \ B'\ \ F \ A \ (A' \ B')" -apply (subgoal_tac "st_set (A) & st_set (A') & st_set (B) & st_set (B') &F \ program") +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 @@ -397,9 +397,9 @@ "\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)) + \S. \A \ S. F \ A \ zb \ P(A) \ st_set(A) \ P(\(S)) \ \ P(za)" -apply (subgoal_tac " (F \ za \ zb) & P (za) ") +apply (subgoal_tac " (F \ za \ zb) \ P (za) ") apply (erule conjunct2) apply (frule leadsToD2) apply (erule leadsTo_induct_pre_aux) @@ -449,7 +449,7 @@ lemma psp: "\F \ A \ A'; F \ B co B'; st_set(B')\\ F:(A \ B') \ ((A' \ B) \ (B' - B))" -apply (subgoal_tac "F \ program & st_set (A) & st_set (A') & st_set (B) ") +apply (subgoal_tac "F \ program \ st_set (A) \ st_set (A') \ st_set (B) ") prefer 2 apply (blast dest!: constrainsD2 leadsToD2) apply (erule leadsTo_induct) prefer 3 apply (blast intro: leadsTo_Union_Int) @@ -471,7 +471,7 @@ "\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') ") +apply (subgoal_tac "st_set (A) \st_set (A') ") prefer 2 apply (blast dest: leadsToD2) apply (drule psp, assumption, blast) apply (blast intro: leadsTo_weaken) @@ -556,7 +556,7 @@ done declare wlt_st_set [iff] -lemma wlt_leadsTo_iff: "F \ wlt(F, B) \ B \ (F \ program & st_set(B))" +lemma wlt_leadsTo_iff: "F \ wlt(F, B) \ B \ (F \ program \ st_set(B))" apply (unfold wlt_def) apply (blast dest: leadsToD2 intro!: leadsTo_Union) done @@ -571,7 +571,7 @@ done (*Misra's property W2*) -lemma leadsTo_eq_subset_wlt: "F \ A \ B \ (A \ wlt(F,B) & F \ program & st_set(B))" +lemma leadsTo_eq_subset_wlt: "F \ A \ B \ (A \ wlt(F,B) \ F \ program \ st_set(B))" apply auto apply (blast dest: leadsToD2 leadsTo_subset intro: leadsTo_weaken_L wlt_leadsTo)+ done @@ -594,7 +594,7 @@ (*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\ @@ -605,7 +605,7 @@ apply (blast intro: leadsTo_123_aux leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate, blast) txt\Union\ apply (clarify dest!: ball_conj_distrib [THEN iffD1]) -apply (subgoal_tac "\y. y \ Pi (S, %A. {Ba \ Pow (state) . A<=Ba & F \ Ba \ B & F \ Ba - B co Ba \ B}) ") +apply (subgoal_tac "\y. y \ Pi (S, %A. {Ba \ Pow (state) . A<=Ba \ F \ Ba \ B \ F \ Ba - B co Ba \ B}) ") defer 1 apply (rule AC_ball_Pi, safe) apply (rotate_tac 1) @@ -634,7 +634,7 @@ F \ A \ (A' \ C); F \ A' co (A' \ C); F \ B \ (B' \ C); F \ B' co (B' \ C)\ \ F \ (A \ B) \ ((A' \ B') \ C)" -apply (subgoal_tac "st_set (C) &st_set (W) &st_set (W-C) &st_set (A') &st_set (A) & st_set (B) & st_set (B') & F \ program") +apply (subgoal_tac "st_set (C) \st_set (W) \st_set (W-C) \st_set (A') \st_set (A) \ st_set (B) \ st_set (B') \ F \ program") prefer 2 apply simp apply (blast dest!: leadsToD2) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Univ.thy --- a/src/ZF/Univ.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Univ.thy Tue Sep 27 17:03:23 2022 +0100 @@ -283,7 +283,7 @@ 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: @@ -317,7 +317,7 @@ 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 + \ \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) @@ -677,7 +677,7 @@ 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) \ jThis 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 \ Vfrom(X,i) \ b \ Vfrom(X,i)" by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD], assumption, fast) @@ -776,7 +776,7 @@ lemma Pair_in_Vfrom_D: "\ \ Vfrom(X,succ(i)); Transset(X)\ - \ a \ Vfrom(X,i) & b \ Vfrom(X,i)" + \ a \ Vfrom(X,i) \ b \ Vfrom(X,i)" apply (unfold Pair_def) apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D) done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ZF_Base.thy --- a/src/ZF/ZF_Base.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ZF_Base.thy Tue Sep 27 17:03:23 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{_ ./ _ \ _, _})\) @@ -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{_ \ _ ./ _})\) @@ -98,7 +98,7 @@ 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" @@ -157,7 +157,7 @@ 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 '(_,_,_')\) @@ -240,7 +240,7 @@ 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)" @@ -347,7 +347,7 @@ by (simp add: Bex_def, blast) (*We do not even have @{term"(\x\A. True) <-> True"} unless @{term"A" is nonempty\*) -lemma bex_triv [simp]: "(\x\A. P) <-> ((\x. x\A) & P)" +lemma bex_triv [simp]: "(\x\A. P) <-> ((\x. x\A) \ P)" by (simp add: Bex_def) lemma bex_cong [cong]: @@ -429,7 +429,7 @@ subsection\Rules for Replace -- the derived form of replacement\ lemma Replace_iff: - "b \ {y. x\A, P(x,y)} <-> (\x\A. P(x,b) & (\y. P(x,y) \ y=b))" + "b \ {y. x\A, P(x,y)} <-> (\x\A. P(x,b) \ (\y. P(x,y) \ y=b))" apply (unfold Replace_def) apply (rule replacement [THEN iff_trans], blast+) done @@ -493,7 +493,7 @@ subsection\Rules for Collect -- forming a subset by separation\ (*Separation is derivable from Replacement*) -lemma separation [simp]: "a \ {x\A. P(x)} <-> a\A & P(a)" +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)}" @@ -586,7 +586,7 @@ subsection\Rules for Inter\ (*Not obviously useful for proving InterI, InterD, InterE*) -lemma Inter_iff: "A \ \(C) <-> (\x\C. A: x) & C\0" +lemma Inter_iff: "A \ \(C) <-> (\x\C. A: x) \ C\0" by (simp add: Inter_def Ball_def, blast) (* Intersection is well-behaved only if the family is non-empty! *) @@ -609,7 +609,7 @@ (* @{term"\x\A. B(x)"} abbreviates @{term"\({B(x). x\A})"} *) -lemma INT_iff: "b \ (\x\A. B(x)) <-> (\x\A. b \ B(x)) & A\0" +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))" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/Zorn.thy --- a/src/ZF/Zorn.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/Zorn.thy Tue Sep 27 17:03:23 2022 +0100 @@ -12,7 +12,7 @@ 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 @@ -20,7 +20,7 @@ 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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/equalities.thy --- a/src/ZF/equalities.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/equalities.thy Tue Sep 27 17:03:23 2022 +0100 @@ -22,7 +22,7 @@ The following are not added to the default simpset because (a) they duplicate the body and (b) there are no similar rules for \Int\.\ -lemma ball_Un: "(\x \ A\B. P(x)) \ (\x \ A. P(x)) & (\x \ B. P(x))" +lemma ball_Un: "(\x \ A\B. P(x)) \ (\x \ A. P(x)) \ (\x \ B. P(x))" by blast lemma bex_Un: "(\x \ A\B. P(x)) \ (\x \ A. P(x)) | (\x \ B. P(x))" @@ -76,7 +76,7 @@ lemma subset_consI: "B \ cons(a,B)" by blast -lemma cons_subset_iff [iff]: "cons(a,B)\C \ a\C & B\C" +lemma cons_subset_iff [iff]: "cons(a,B)\C \ a\C \ B\C" by blast (*A safe special case of subset elimination, adding no new variables @@ -86,7 +86,7 @@ lemma subset_empty_iff: "A\0 \ A=0" by blast -lemma subset_cons_iff: "C\cons(a,B) \ C\B | (a\C & C-{a} \ B)" +lemma subset_cons_iff: "C\cons(a,B) \ C\B | (a\C \ C-{a} \ B)" by blast (* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*) @@ -134,7 +134,7 @@ "\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)" +lemma succ_subset_iff: "succ(a) \ B \ (a \ B \ a \ B)" by (unfold succ_def, blast) @@ -142,7 +142,7 @@ (** Intersection is the greatest lower bound of two sets **) -lemma Int_subset_iff: "C \ A \ B \ C \ A & C \ B" +lemma Int_subset_iff: "C \ A \ B \ C \ A \ C \ B" by blast lemma Int_lower1: "A \ B \ A" @@ -211,7 +211,7 @@ (** Union is the least upper bound of two sets *) -lemma Un_subset_iff: "A \ B \ C \ A \ C & B \ C" +lemma Un_subset_iff: "A \ B \ C \ A \ C \ B \ C" by blast lemma Un_upper1: "A \ A \ B" @@ -259,7 +259,7 @@ lemma subset_Un_iff2: "A\B \ B \ A = B" by (blast elim!: equalityE) -lemma Un_empty [iff]: "(A \ B = 0) \ (A = 0 & B = 0)" +lemma Un_empty [iff]: "(A \ B = 0) \ (A = 0 \ B = 0)" by blast lemma Un_eq_Union: "A \ B = \({A, B})" @@ -273,7 +273,7 @@ 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" +lemma subset_Diff_cons_iff: "B \ A - cons(c,C) \ B\A-C \ c \ B" by blast lemma Diff_cancel: "A - A = 0" @@ -571,7 +571,7 @@ by blast lemma times_subset_iff: - "(A'*B' \ A*B) \ (A' = 0 | B' = 0 | (A'\A) & (B'\B))" + "(A'*B' \ A*B) \ (A' = 0 | B' = 0 | (A'\A) \ (B'\B))" by blast lemma Int_Sigma_eq: @@ -765,7 +765,7 @@ by blast lemma Collect_image_eq: - "{z \ Sigma(A,B). P(z)} `` C = (\x \ A. {y \ B(x). x \ C & P()})" + "{z \ Sigma(A,B). P(z)} `` C = (\x \ A. {y \ B(x). x \ C \ P()})" by blast lemma image_Int_subset: "r``(A \ B) \ (r``A) \ (r``B)" @@ -941,11 +941,11 @@ by blast lemma Collect_Collect_eq [simp]: - "Collect(Collect(A,P), Q) = Collect(A, %x. P(x) & Q(x))" + "Collect(Collect(A,P), Q) = Collect(A, %x. P(x) \ Q(x))" by blast lemma Collect_Int_Collect_eq: - "Collect(A,P) \ Collect(A,Q) = Collect(A, %x. P(x) & Q(x))" + "Collect(A,P) \ Collect(A,Q) = Collect(A, %x. P(x) \ Q(x))" by blast lemma Collect_Union_eq [simp]: @@ -961,7 +961,7 @@ lemma Collect_disj_eq: "{x\A. P(x) | Q(x)} = Collect(A, P) \ Collect(A, Q)" by blast -lemma Collect_conj_eq: "{x\A. P(x) & Q(x)} = Collect(A, P) \ Collect(A, Q)" +lemma Collect_conj_eq: "{x\A. P(x) \ Q(x)} = Collect(A, P) \ Collect(A, Q)" by blast lemmas subset_SIs = subset_refl cons_subsetI subset_consI diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/CoUnit.thy --- a/src/ZF/ex/CoUnit.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/CoUnit.thy Tue Sep 27 17:03:23 2022 +0100 @@ -55,7 +55,7 @@ inductive_cases Con2E: "Con2(x, y) \ counit2" -lemma Con2_iff: "Con2(x, y) = Con2(x', y') \ x = x' & y = y'" +lemma Con2_iff: "Con2(x, y) = Con2(x', y') \ x = x' \ y = y'" \ \Proving freeness results.\ by (fast elim!: counit2.free_elims) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/Commutation.thy --- a/src/ZF/ex/Commutation.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/Commutation.thy Tue Sep 27 17:03:23 2022 +0100 @@ -1,5 +1,5 @@ (* Title: ZF/ex/Commutation.thy - Author: Tobias Nipkow & Sidi Ould Ehmety + Author: Tobias Nipkow \ Sidi Ould Ehmety Copyright 1995 TU Muenchen Commutation theory for proving the Church Rosser theorem. @@ -10,7 +10,7 @@ definition square :: "[i, i, i, i] => o" where "square(r,s,t,u) \ - (\a b. \ r \ (\c. \ s \ (\x. \ t & \ u)))" + (\a b. \ r \ (\c. \ s \ (\x. \ t \ \ u)))" definition commute :: "[i, i] => o" where @@ -27,7 +27,7 @@ definition Church_Rosser :: "i => o" where "Church_Rosser(r) \ (\x y. \ (r \ converse(r))^* \ - (\z. \ r^* & \ r^*))" + (\z. \ r^* \ \ r^*))" definition confluent :: "i=>o" where diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/Group.thy --- a/src/ZF/ex/Group.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/Group.thy Tue Sep 27 17:03:23 2022 +0100 @@ -35,7 +35,7 @@ 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]: @@ -87,7 +87,7 @@ locale group = monoid + assumes inv_ex: - "\x. x \ carrier(G) \ \y \ carrier(G). y \ x = \ & x \ y = \" + "\x. x \ carrier(G) \ \y \ carrier(G). y \ x = \ \ x \ y = \" lemma (in group) is_group [simp]: "group(G)" by (rule group_axioms) @@ -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)" @@ -143,7 +143,7 @@ by (simp add: m_assoc [symmetric] l_inv r_one) with x y have r_inv: "x \ y = \" by simp - from x y show "\y \ carrier(G). y \ x = \ & x \ y = \" + from x y show "\y \ carrier(G). y \ x = \ \ x \ y = \" by (fast intro: l_inv r_inv) qed show ?thesis @@ -152,7 +152,7 @@ qed lemma (in group) inv [simp]: - "x \ carrier(G) \ inv x \ carrier(G) & inv x \ x = \ & x \ inv x = \" + "x \ carrier(G) \ inv x \ carrier(G) \ inv x \ x = \ \ x \ inv x = \" apply (frule inv_ex) apply (unfold Bex_def m_inv_def) apply (erule exE) @@ -681,7 +681,7 @@ text\Alternative characterization of normal subgroups\ lemma (in group) normal_inv_iff: "(N \ G) \ - (subgroup(N,G) & (\x \ carrier(G). \h \ N. x \ h \ (inv x) \ N))" + (subgroup(N,G) \ (\x \ carrier(G). \h \ N. x \ h \ (inv x) \ N))" (is "_ \ ?rhs") proof assume N: "N \ G" @@ -739,7 +739,7 @@ assume "\h\H. y = x \ h" and x: "x \ carrier(G)" and sb: "subgroup(H,G)" - then obtain h' where h': "h' \ H & x \ h' = y" by blast + then obtain h' where h': "h' \ H \ x \ h' = y" by blast show "\h\H. x = y \ h" proof show "x = y \ inv h'" using h' x sb diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/LList.thy --- a/src/ZF/ex/LList.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/LList.thy Tue Sep 27 17:03:23 2022 +0100 @@ -74,7 +74,7 @@ inductive_cases LConsE: "LCons(a,l) \ llist(A)" (*Proving freeness results*) -lemma LCons_iff: "LCons(a,l)=LCons(a',l') \ a=a' & l=l'" +lemma LCons_iff: "LCons(a,l)=LCons(a',l') \ a=a' \ l=l'" by auto lemma LNil_LCons_iff: "\ LNil=LCons(a,l)" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/Limit.thy --- a/src/ZF/ex/Limit.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/Limit.thy Tue Sep 27 17:03:23 2022 +0100 @@ -30,23 +30,23 @@ definition po :: "i=>o" where "po(D) \ - (\x \ set(D). rel(D,x,x)) & + (\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)) & + rel(D,x,y) \ rel(D,y,z) \ rel(D,x,z)) \ (\x \ set(D). \y \ set(D). rel(D,x,y) \ rel(D,y,x) \ x = y)" 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 @@ -54,15 +54,15 @@ 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 @@ -97,16 +97,16 @@ definition matrix :: "[i,i]=>o" where "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))) & + 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))) \ (\n \ nat. \m \ nat. rel(D,M`n`m,M`succ(n)`succ(m)))" definition projpair :: "[i,i,i,i]=>o" where "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)))" + 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 @@ -133,18 +133,18 @@ definition subcpo :: "[i,i]=>o" where "subcpo(D,E) \ - set(D) \ set(E) & - (\x \ set(D). \y \ set(D). rel(D,x,y) \ rel(E,x,y)) & + 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) \ - (\n \ nat. cpo(DD`n)) & (\n \ nat. emb(DD`n,DD`succ(n),ee`n))" + (\n \ nat. cpo(DD`n)) \ (\n \ nat. emb(DD`n,DD`succ(n),ee`n))" definition Dinf :: "[i,i]=>i" where @@ -181,12 +181,12 @@ definition commute :: "[i,i,i,i=>i]=>o" where "commute(DD,ee,E,r) \ - (\n \ nat. emb(DD`n,E,r(n))) & + (\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] @@ -342,7 +342,7 @@ 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 @@ -901,7 +901,7 @@ 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))" @@ -2293,7 +2293,7 @@ mediating (Dinf(DD,ee),G,rho_emb(DD,ee),f, lub(cf(Dinf(DD,ee),G), - \n \ nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))) & + \n \ nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))) \ (\t. mediating(Dinf(DD,ee),G,rho_emb(DD,ee),f,t) \ t = lub(cf(Dinf(DD,ee),G), \n \ nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))))" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/NatSum.thy --- a/src/ZF/ex/NatSum.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/NatSum.thy Tue Sep 27 17:03:23 2022 +0100 @@ -1,5 +1,5 @@ (* Title: ZF/ex/NatSum.thy - Author: Tobias Nipkow & Lawrence C Paulson + Author: Tobias Nipkow \ Lawrence C Paulson A summation operator. sum(f,n+1) is the sum of all f(i), i=0...n. diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/Primes.thy --- a/src/ZF/ex/Primes.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/Primes.thy Tue Sep 27 17:03:23 2022 +0100 @@ -9,12 +9,12 @@ 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)) & - (\d\nat. (d dvd m) & (d dvd n) \ d dvd p)" + "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 @@ -28,12 +28,12 @@ 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: @@ -176,7 +176,7 @@ 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]) @@ -217,7 +217,7 @@ done lemma gcd_greatest_iff [simp]: "\k \ nat; m \ nat; n \ nat\ - \ (k dvd gcd (m, n)) \ (k dvd m & k dvd n)" + \ (k dvd gcd (m, n)) \ (k dvd m \ k dvd n)" by (blast intro!: gcd_greatest gcd_dvd1 gcd_dvd2 intro: dvd_trans) @@ -387,7 +387,7 @@ lemma reduction: "\k#*k = p#*(j#*j); p \ prime; 0 < k; j \ nat; k \ nat\ - \ k < p#*j & 0 < j" + \ k < p#*j \ 0 < j" apply (rule ccontr) apply (simp add: not_lt_iff_le prime_into_nat) apply (erule disjE) diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/Ramsey.thy --- a/src/ZF/ex/Ramsey.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/Ramsey.thy Tue Sep 27 17:03:23 2022 +0100 @@ -36,17 +36,17 @@ 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) \ - (\C. Clique(C,V,E) & Atleast(i,C)) | - (\I. Indept(I,V,E) & Atleast(j,I))" + "Ramsey(n,i,j) \ \V E. Symmetric(E) \ Atleast(n,V) \ + (\C. Clique(C,V,E) \ Atleast(i,C)) | + (\I. Indept(I,V,E) \ Atleast(j,I))" (*** Cliques and Independent sets ***) @@ -147,7 +147,7 @@ lemma Indept_succ: "\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))" + Indept(cons(a,I), V, E) \ Atleast(succ(j), cons(a,I))" apply (unfold Symmetric_def Indept_def) apply (blast intro!: Atleast_succI) done @@ -156,7 +156,7 @@ lemma Clique_succ: "\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))" + Clique(cons(a,C), V, E) \ Atleast(succ(j), cons(a,C))" apply (unfold Symmetric_def Clique_def) apply (blast intro!: Atleast_succI) done diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/Ring.thy --- a/src/ZF/ex/Ring.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/Ring.thy Tue Sep 27 17:03:23 2022 +0100 @@ -234,9 +234,9 @@ ring_hom :: "[i,i] => i" where "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) & - h ` (x \\<^bsub>R\<^esub> y) = (h ` x) \\<^bsub>S\<^esub> (h ` y)) & + (\x y. x \ carrier(R) \ y \ carrier(R) \ + h ` (x \\<^bsub>R\<^esub> y) = (h ` x) \\<^bsub>S\<^esub> (h ` y) \ + h ` (x \\<^bsub>R\<^esub> y) = (h ` x) \\<^bsub>S\<^esub> (h ` y)) \ h ` \\<^bsub>R\<^esub> = \\<^bsub>S\<^esub>}" lemma ring_hom_memI: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/ex/misc.thy --- a/src/ZF/ex/misc.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/ex/misc.thy Tue Sep 27 17:03:23 2022 +0100 @@ -36,15 +36,15 @@ text\These two are cited in Benzmueller and Kohlhase's system description of LEO, CADE-15, 1998 (page 139-143) as theorems LEO could not prove.\ -lemma "(X = Y \ Z) \ (Y \ X & Z \ X & (\V. Y \ V & Z \ V \ X \ V))" +lemma "(X = Y \ Z) \ (Y \ X \ Z \ X \ (\V. Y \ V \ Z \ V \ X \ V))" by (blast intro!: equalityI) text\the dual of the previous one\ -lemma "(X = Y \ Z) \ (X \ Y & X \ Z & (\V. V \ Y & V \ Z \ V \ X))" +lemma "(X = Y \ Z) \ (X \ Y \ X \ Z \ (\V. V \ Y \ V \ Z \ V \ X))" 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!\ @@ -63,7 +63,7 @@ by best text\A characterization of functions suggested by Tobias Nipkow\ -lemma "r \ domain(r)->B \ r \ domain(r)*B & (\X. r `` (r -`` X) \ X)" +lemma "r \ domain(r)->B \ r \ domain(r)*B \ (\X. r `` (r -`` X) \ X)" by (unfold Pi_def function_def, best) @@ -80,17 +80,17 @@ (*Force helps prove conditions of rewrites such as comp_fun_apply, since rewriting does not instantiate Vars.*) lemma "(\A f B g. hom(A,f,B,g) = - {H \ A->B. f \ A*A->A & g \ B*B->B & + {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) \ + J \ hom(A,f,B,g) \ K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" by force text\Another version, with meta-level rewriting\ lemma "(\A f B g. hom(A,f,B,g) \ - {H \ A->B. f \ A*A->A & g \ B*B->B & + {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 diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/func.thy --- a/src/ZF/func.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/func.thy Tue Sep 27 17:03:23 2022 +0100 @@ -20,12 +20,12 @@ by (simp add: restrict_def relation_def, blast) lemma Pi_iff: - "f \ Pi(A,B) \ function(f) & f<=Sigma(A,B) & A<=domain(f)" + "f \ Pi(A,B) \ function(f) \ f<=Sigma(A,B) \ A<=domain(f)" by (unfold Pi_def, blast) (*For upward compatibility with the former definition*) lemma Pi_iff_old: - "f \ Pi(A,B) \ f<=Sigma(A,B) & (\x\A. \!y. : f)" + "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)" @@ -96,7 +96,7 @@ 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 @@ -110,7 +110,7 @@ (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*) lemma Pi_Collect_iff: "(f \ Pi(A, %x. {y \ B(x). P(x,y)})) - \ f \ Pi(A,B) & (\x\A. P(x, f`x))" + \ f \ Pi(A,B) \ (\x\A. P(x, f`x))" by (blast intro: Pi_type dest: apply_type) lemma Pi_weaken_type: @@ -126,7 +126,7 @@ 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\ @@ -196,7 +196,7 @@ lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)" by (unfold Pi_def function_def, force) -lemma fun_space_empty_iff [iff]: "(A->X)=0 \ X=0 & (A \ 0)" +lemma fun_space_empty_iff [iff]: "(A->X)=0 \ X=0 \ (A \ 0)" apply auto apply (fast intro!: equals0I intro: lam_type) done @@ -296,7 +296,7 @@ lemma restrict_empty [simp]: "restrict(f,0) = 0" by (unfold restrict_def, simp) -lemma restrict_iff: "z \ restrict(r,A) \ z \ r & (\x\A. \y. z = \x, y\)" +lemma restrict_iff: "z \ restrict(r,A) \ z \ r \ (\x\A. \y. z = \x, y\)" by (simp add: restrict_def) lemma restrict_restrict [simp]: diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/pair.thy --- a/src/ZF/pair.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/pair.thy Tue Sep 27 17:03:23 2022 +0100 @@ -18,7 +18,7 @@ ML \val ZF_ss = simpset_of \<^context>\ -simproc_setup defined_Bex ("\x\A. P(x) & Q(x)") = \ +simproc_setup defined_Bex ("\x\A. P(x) \ Q(x)") = \ fn _ => Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}) \ @@ -34,10 +34,10 @@ lemma singleton_eq_iff [iff]: "{a} = {b} \ a=b" by (rule extension [THEN iff_trans], blast) -lemma doubleton_eq_iff: "{a,b} = {c,d} \ (a=c & b=d) | (a=d & b=c)" +lemma doubleton_eq_iff: "{a,b} = {c,d} \ (a=c \ b=d) | (a=d \ b=c)" by (rule extension [THEN iff_trans], blast) -lemma Pair_iff [simp]: " = \ a=c & b=d" +lemma Pair_iff [simp]: " = \ a=c \ b=d" by (simp add: Pair_def doubleton_eq_iff, blast) lemmas Pair_inject = Pair_iff [THEN iffD1, THEN conjE, elim!] @@ -77,7 +77,7 @@ text\Generalizes Cartesian product\ -lemma Sigma_iff [simp]: ": Sigma(A,B) \ a \ A & b \ B(a)" +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)" diff -r e44d86131648 -r 0c18df79b1c8 src/ZF/upair.thy --- a/src/ZF/upair.thy Tue Sep 27 16:51:35 2022 +0100 +++ b/src/ZF/upair.thy Tue Sep 27 17:03:23 2022 +0100 @@ -65,7 +65,7 @@ subsection\Rules for Binary Intersection, Defined via \<^term>\Upair\\ -lemma Int_iff [simp]: "c \ A \ B \ (c \ A & c \ B)" +lemma Int_iff [simp]: "c \ A \ B \ (c \ A \ c \ B)" apply (unfold Int_def) apply (blast intro: UpairI1 UpairI2 elim: UpairE) done @@ -85,7 +85,7 @@ subsection\Rules for Set Difference, Defined via \<^term>\Upair\\ -lemma Diff_iff [simp]: "c \ A-B \ (c \ A & c\B)" +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" @@ -223,7 +223,7 @@ 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,7 +238,7 @@ 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]: @@ -247,7 +247,7 @@ (** 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 @@ -324,19 +324,19 @@ subsection\Miniscoping of the Bounded Universal Quantifier\ lemma ball_simps1: - "(\x\A. P(x) & Q) \ (\x\A. P(x)) & (A=0 | Q)" + "(\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\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))" + "(\x\succ(i).P(x)) \ P(i) \ (\x\i. P(x))" + "(\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))" by blast+ lemma ball_simps2: - "(\x\A. P & Q(x)) \ (A=0 | P) & (\x\A. Q(x))" + "(\x\A. P \ Q(x)) \ (A=0 | P) \ (\x\A. Q(x))" "(\x\A. P | Q(x)) \ (P | (\x\A. Q(x)))" "(\x\A. P \ Q(x)) \ (P \ (\x\A. Q(x)))" by blast+ @@ -348,16 +348,16 @@ lemmas ball_simps [simp] = ball_simps1 ball_simps2 ball_simps3 lemma ball_conj_distrib: - "(\x\A. P(x) & Q(x)) \ ((\x\A. P(x)) & (\x\A. Q(x)))" + "(\x\A. P(x) \ Q(x)) \ ((\x\A. P(x)) \ (\x\A. Q(x)))" by blast subsection\Miniscoping of the Bounded Existential Quantifier\ lemma bex_simps1: - "(\x\A. P(x) & Q) \ ((\x\A. P(x)) & Q)" - "(\x\A. P(x) | Q) \ (\x\A. P(x)) | (A\0 & Q)" - "(\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)) | (A\0 \ Q)" + "(\x\A. P(x) \ Q) \ ((\x\A. P(x)) \ (A\0 \ Q))" "(\x\0.P(x)) \ False" "(\x\succ(i).P(x)) \ P(i) | (\x\i. P(x))" "(\x\cons(a,B).P(x)) \ P(a) | (\x\B. P(x))" @@ -367,13 +367,13 @@ by blast+ lemma bex_simps2: - "(\x\A. P & Q(x)) \ (P & (\x\A. Q(x)))" - "(\x\A. P | Q(x)) \ (A\0 & P) | (\x\A. Q(x))" + "(\x\A. P \ Q(x)) \ (P \ (\x\A. Q(x)))" + "(\x\A. P | Q(x)) \ (A\0 \ P) | (\x\A. Q(x))" "(\x\A. P \ Q(x)) \ ((A=0 | P) \ (\x\A. Q(x)))" by blast+ lemma bex_simps3: - "(\x\Collect(A,Q).P(x)) \ (\x\A. Q(x) & P(x))" + "(\x\Collect(A,Q).P(x)) \ (\x\A. Q(x) \ P(x))" by blast lemmas bex_simps [simp] = bex_simps1 bex_simps2 bex_simps3 @@ -391,10 +391,10 @@ lemma bex_triv_one_point2 [simp]: "(\x\A. a=x) \ (a \ A)" by blast -lemma bex_one_point1 [simp]: "(\x\A. x=a & P(x)) \ (a \ A & P(a))" +lemma bex_one_point1 [simp]: "(\x\A. x=a \ P(x)) \ (a \ A \ P(a))" by blast -lemma bex_one_point2 [simp]: "(\x\A. a=x & P(x)) \ (a \ A & P(a))" +lemma bex_one_point2 [simp]: "(\x\A. a=x \ P(x)) \ (a \ A \ P(a))" by blast lemma ball_one_point1 [simp]: "(\x\A. x=a \ P(x)) \ (a \ A \ P(a))"