# HG changeset patch # User paulson # Date 1664297212 -3600 # Node ID a642599ffdea7b0df43f7b210c08e6f62b58378e # Parent 0c18df79b1c86dc0b85150ef68421b0b4b6f76c0 More syntactic cleanup. LaTeX markup working diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC.thy --- a/src/ZF/AC.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC.thy Tue Sep 27 17:46:52 2022 +0100 @@ -26,13 +26,13 @@ done lemma AC_Pi_Pow: "\f. f \ (\X \ Pow(C)-{0}. X)" -apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE]) +apply (rule_tac B1 = "\x. x" in AC_Pi [THEN exE]) apply (erule_tac [2] exI, blast) done lemma AC_func: "\\x. x \ A \ (\y. y \ x)\ \ \f \ A->\(A). \x \ A. f`x \ x" -apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE]) +apply (rule_tac B1 = "\x. x" in AC_Pi [THEN exE]) prefer 2 apply (blast dest: apply_type intro: Pi_type, blast) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/AC15_WO6.thy --- a/src/ZF/AC/AC15_WO6.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/AC15_WO6.thy Tue Sep 27 17:46:52 2022 +0100 @@ -34,7 +34,7 @@ lemma lepoll_Sigma: "A\0 \ B \ A*B" apply (unfold lepoll_def) apply (erule not_emptyE) -apply (rule_tac x = "\z \ B. " in exI) +apply (rule_tac x = "\z \ B. \x,z\" in exI) apply (fast intro!: snd_conv lam_injective) done @@ -69,7 +69,7 @@ apply (erule exE) apply (rule_tac x = "\x \ RepFun(A,P). \ j. \a\A. x=P(a) \ f`a=j" in exI) -apply (rule_tac d = "%y. P (converse (f) `y) " in lam_injective) +apply (rule_tac d = "\y. P (converse (f) `y) " in lam_injective) apply (erule RepFunE) apply (frule inj_is_fun [THEN apply_type], assumption) apply (fast intro: LeastI2 elim!: Ord_in_Ord inj_is_fun [THEN apply_type]) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/AC16_WO4.thy --- a/src/ZF/AC/AC16_WO4.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/AC16_WO4.thy Tue Sep 27 17:46:52 2022 +0100 @@ -79,7 +79,7 @@ lemma succ_not_lepoll_lemma: "\\(\x \ A. f`x=y); f \ inj(A, B); y \ B\ \ (\a \ succ(A). if(a=A, y, f`a)) \ inj(succ(A), B)" -apply (rule_tac d = "%z. if (z=y, A, converse (f) `z) " in lam_injective) +apply (rule_tac d = "\z. if (z=y, A, converse (f) `z) " in lam_injective) apply (force simp add: inj_is_fun [THEN apply_type]) (*this preliminary simplification prevents looping somehow*) apply (simp (no_asm_simp)) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/AC16_lemmas.thy --- a/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/AC16_lemmas.thy Tue Sep 27 17:46:52 2022 +0100 @@ -74,7 +74,7 @@ apply (unfold lepoll_def) apply (rule_tac x = "\y \ {y \ Pow(x) . y\succ (succ (n))}. <\ i. i \ y, y-{\ i. i \ y}>" in exI) -apply (rule_tac d = "%z. cons (fst(z), snd(z))" in lam_injective) +apply (rule_tac d = "\z. cons (fst(z), snd(z))" in lam_injective) apply (blast intro!: Diff_sing_eqpoll intro: InfCard_Least_in) apply (simp, blast intro: InfCard_Least_in) done @@ -147,7 +147,7 @@ apply (unfold lepoll_def) apply (rule_tac x = "\z \ {y\Pow(x). y\succ(n)}. cons(succ(\(z)), z)" in exI) -apply (rule_tac d = "%z. z-{\(z) }" in lam_injective) +apply (rule_tac d = "\z. z-{\(z) }" in lam_injective) apply (blast intro!: succ_Union_in_x succ_Union_not_mem intro: cons_eqpoll_succ Ord_in_Ord dest!: InfCard_is_Card [THEN Card_is_Ord]) @@ -185,7 +185,7 @@ apply (unfold eqpoll_def) apply (erule exE) apply (rule_tac x = "\X \ {Y \ Pow (A) . \f. f \ bij (Y, n) }. f``X" in exI) -apply (rule_tac d = "%Z. converse (f) ``Z" in lam_bijective) +apply (rule_tac d = "\Z. converse (f) ``Z" in lam_bijective) apply (fast intro!: bij_is_inj [THEN restrict_bij, THEN bij_converse_bij, THEN comp_bij] elim!: bij_is_fun [THEN fun_is_rel, THEN image_subset]) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/AC18_AC19.thy --- a/src/ZF/AC/AC18_AC19.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/AC18_AC19.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,7 +9,7 @@ begin definition - uu :: "i => i" where + uu :: "i \ i" where "uu(a) \ {c \ {0}. c \ a}" @@ -48,7 +48,7 @@ theorem (in AC18) AC19 apply (unfold AC19_def) apply (intro allI impI) -apply (rule AC18 [of _ "%x. x", THEN mp], blast) +apply (rule AC18 [of _ "\x. x", THEN mp], blast) done (* ********************************************************************** *) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/AC7_AC9.thy --- a/src/ZF/AC/AC7_AC9.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/AC7_AC9.thy Tue Sep 27 17:46:52 2022 +0100 @@ -86,7 +86,7 @@ (* ********************************************************************** *) lemma AC1_AC8_lemma1: - "\B \ A. \B1 B2. B= \ B1 \ B2 + "\B \ A. \B1 B2. B=\B1,B2\ \ B1 \ B2 \ 0 \ { bij(fst(B),snd(B)). B \ A }" apply (unfold eqpoll_def, auto) done @@ -110,7 +110,7 @@ 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\ \ B1 \ B2" by fast lemma AC8_AC9: "AC8 \ AC9" @@ -151,7 +151,7 @@ lemma AC9_AC1_lemma2: "\B1 \ {(F*B)*N. B \ A} \ {cons(0,(F*B)*N). B \ A}. \B2 \ {(F*B)*N. B \ A} \ {cons(0,(F*B)*N). B \ A}. - f` \ bij(B1, B2) + f`\B1,B2\ \ bij(B1, B2) \ (\B \ A. snd(fst((f`)`0))) \ (\X \ A. X)" apply (intro lam_type snd_type fst_type) apply (rule apply_type [OF _ consI1]) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/AC_Equiv.thy --- a/src/ZF/AC/AC_Equiv.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/AC_Equiv.thy Tue Sep 27 17:46:52 2022 +0100 @@ -46,11 +46,11 @@ definition (* Auxiliary concepts needed below *) - pairwise_disjoint :: "i => o" where + pairwise_disjoint :: "i \ o" where "pairwise_disjoint(A) \ \A1 \ A. \A2 \ A. A1 \ A2 \ 0 \ A1=A2" definition - sets_of_size_between :: "[i, i, i] => o" where + sets_of_size_between :: "[i, i, i] \ o" where "sets_of_size_between(A,m,n) \ \B \ A. m \ B \ B \ n" @@ -80,12 +80,12 @@ "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\ \ B1\B2) \ (\f. \B \ A. f`B \ bij(fst(B),snd(B)))" definition "AC9 \ \A. (\B1 \ A. \B2 \ A. B1\B2) \ - (\f. \B1 \ A. \B2 \ A. f` \ bij(B1,B2))" + (\f. \B1 \ A. \B2 \ A. f`\B1,B2\ \ bij(B1,B2))" definition "AC10(n) \ \A. (\B \ A. \Finite(B)) \ @@ -140,21 +140,21 @@ lemma rvimage_id: "rvimage(A,id(A),r) = r \ A*A" apply (unfold rvimage_def) apply (rule equalityI, safe) -apply (drule_tac P = "%a. :r" in id_conv [THEN subst], +apply (drule_tac P = "\a. :r" in id_conv [THEN subst], assumption) -apply (drule_tac P = "%a. :r" in id_conv [THEN subst], (assumption+)) +apply (drule_tac P = "\a. \a,ya\:r" in id_conv [THEN subst], (assumption+)) apply (fast intro: id_conv [THEN ssubst]) done (* used only in Hartog.ML *) lemma ordertype_Int: "well_ord(A,r) \ ordertype(A, r \ A*A) = ordertype(A,r)" -apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst]) +apply (rule_tac P = "\a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst]) apply (erule id_bij [THEN bij_ordertype_vimage]) done lemma lam_sing_bij: "(\x \ A. {x}) \ bij(A, {{x}. x \ A})" -apply (rule_tac d = "%z. THE x. z={x}" in lam_bijective) +apply (rule_tac d = "\z. THE x. z={x}" in lam_bijective) apply (auto simp add: the_equality) done @@ -196,9 +196,9 @@ "\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. \x,y\ \ u \ f`\x,y\ = i" in exI) -apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective) +apply (rule_tac d = "\y. fst (converse(f) ` y) " in lam_injective) apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] inj_is_fun [THEN apply_type]) apply (erule domainE) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/Cardinal_aux.thy --- a/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/Cardinal_aux.thy Tue Sep 27 17:46:52 2022 +0100 @@ -101,7 +101,7 @@ lemma UN_sing_lepoll: "Ord(a) \ (\x \ a. {P(x)}) \ a" apply (unfold lepoll_def) apply (rule_tac x = "\z \ (\x \ a. {P (x) }) . (\ i. P (i) =z) " in exI) -apply (rule_tac d = "%z. P (z) " in lam_injective) +apply (rule_tac d = "\z. P (z) " in lam_injective) apply (fast intro!: Least_in_Ord) apply (fast intro: LeastI elim!: Ord_in_Ord) done @@ -146,9 +146,9 @@ apply (rule subsetI) apply (erule UN_E) apply (rule UN_I) - apply (rule_tac P = "%z. x \ F (z) " in Least_in_Ord, (assumption+)) + apply (rule_tac P = "\z. x \ F (z) " in Least_in_Ord, (assumption+)) apply (rule DiffI, best intro: Ord_in_Ord LeastI, clarify) -apply (erule_tac P = "%z. x \ F (z) " and i = c in less_LeastE) +apply (erule_tac P = "\z. x \ F (z) " and i = c in less_LeastE) apply (blast intro: Ord_Least ltI) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/DC.thy --- a/src/ZF/AC/DC.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/DC.thy Tue Sep 27 17:46:52 2022 +0100 @@ -11,7 +11,7 @@ lemma RepFun_lepoll: "Ord(a) \ {P(b). b \ a} \ a" apply (unfold lepoll_def) apply (rule_tac x = "\z \ RepFun (a,P) . \ i. z=P (i) " in exI) -apply (rule_tac d="%z. P (z)" in lam_injective) +apply (rule_tac d="\z. P (z)" in lam_injective) apply (fast intro!: Least_in_Ord) apply (rule sym) apply (fast intro: LeastI Ord_in_Ord) @@ -21,7 +21,7 @@ lemma image_Ord_lepoll: "\f \ X->Y; Ord(X)\ \ f``X \ X" apply (unfold lepoll_def) apply (rule_tac x = "\x \ f``X. \ y. f`y = x" in exI) -apply (rule_tac d = "%z. f`z" in lam_injective) +apply (rule_tac d = "\z. f`z" in lam_injective) apply (fast intro!: Least_in_Ord apply_equality, clarify) apply (rule LeastI) apply (erule apply_equality, assumption+) @@ -29,35 +29,35 @@ done lemma range_subset_domain: - "\R \ X*X; \g. g \ X \ \u. \ R\ + "\R \ X*X; \g. g \ X \ \u. \g,u\ \ R\ \ range(R) \ domain(R)" by blast -lemma cons_fun_type: "g \ n->X \ cons(, g) \ succ(n) -> cons(x, X)" +lemma cons_fun_type: "g \ n->X \ cons(\n,x\, g) \ succ(n) -> cons(x, X)" apply (unfold succ_def) apply (fast intro!: fun_extend elim!: mem_irrefl) done lemma cons_fun_type2: - "\g \ n->X; x \ X\ \ cons(, g) \ succ(n) -> X" + "\g \ n->X; x \ X\ \ cons(\n,x\, g) \ succ(n) -> X" by (erule cons_absorb [THEN subst], erule cons_fun_type) -lemma cons_image_n: "n \ nat \ cons(, g)``n = g``n" +lemma cons_image_n: "n \ nat \ cons(\n,x\, g)``n = g``n" by (fast elim!: mem_irrefl) -lemma cons_val_n: "g \ n->X \ cons(, g)`n = x" +lemma cons_val_n: "g \ n->X \ cons(\n,x\, g)`n = x" by (fast intro!: apply_equality elim!: cons_fun_type) -lemma cons_image_k: "k \ n \ cons(, g)``k = g``k" +lemma cons_image_k: "k \ n \ cons(\n,x\, g)``k = g``k" by (fast elim: mem_asym) -lemma cons_val_k: "\k \ n; g \ n->X\ \ cons(, g)`k = g`k" +lemma cons_val_k: "\k \ n; g \ n->X\ \ cons(\n,x\, g)`k = g`k" by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair) -lemma domain_cons_eq_succ: "domain(f)=x \ domain(cons(, f)) = succ(x)" +lemma domain_cons_eq_succ: "domain(f)=x \ domain(cons(\x,y\, f)) = succ(x)" by (simp add: domain_cons succ_def) -lemma restrict_cons_eq: "g \ n->X \ restrict(cons(, g), n) = g" +lemma restrict_cons_eq: "g \ n->X \ restrict(cons(\n,x\, g), n) = g" apply (simp add: restrict_def Pi_iff) apply (blast intro: elim: mem_irrefl) done @@ -80,9 +80,9 @@ definition - DC :: "i => o" where + DC :: "i \ o" where "DC(a) \ \X R. R \ Pow(X)*X \ - (\Y \ Pow(X). Y \ a \ (\x \ X. \ R)) + (\Y \ Pow(X). Y \ a \ (\x \ X. \Y,x\ \ R)) \ (\f \ a->X. \b \ R)" definition @@ -91,18 +91,18 @@ \ (\f \ nat->domain(R). \n \ nat. :R)" definition - ff :: "[i, i, i, i] => i" where + ff :: "[i, i, i, i] \ i" where "ff(b, X, Q, R) \ - transrec(b, %c r. THE x. first(x, {x \ X. \ R}, Q))" + transrec(b, \c r. THE x. first(x, {x \ X. \ R}, Q))" locale DC0_imp = fixes XX and RR and X and R - assumes all_ex: "\Y \ Pow(X). Y \ nat \ (\x \ X. \ R)" + assumes all_ex: "\Y \ Pow(X). Y \ nat \ (\x \ X. \Y, x\ \ R)" defines XX_def: "XX \ (\n \ nat. {f \ n->X. \k \ n. \ R})" - and RR_def: "RR \ {:XX*XX. domain(z2)=succ(domain(z1)) + and RR_def: "RR \ {\z1,z2\:XX*XX. domain(z2)=succ(domain(z1)) \ restrict(z2, domain(z1)) = z1}" begin @@ -141,7 +141,7 @@ apply (erule_tac [2] notE [OF _ empty_subsetI [THEN PowI]]) apply (erule_tac impE [OF _ nat_0I [THEN n_lesspoll_nat]]) apply (erule bexE) -apply (rule_tac a = "<0, {<0, x>}>" in not_emptyI) +apply (rule_tac a = "<0, {\0, x\}>" in not_emptyI) apply (rule CollectI) apply (rule SigmaI) apply (rule nat_0I [THEN UN_I]) @@ -161,7 +161,7 @@ [OF _ nat_into_Ord] n_lesspoll_nat]], assumption+) apply (erule bexE) -apply (rule_tac x = "cons (, g) " in exI) +apply (rule_tac x = "cons (\n,x\, g) " in exI) apply (rule CollectI) apply (force elim!: cons_fun_type2 simp add: cons_image_n cons_val_n cons_image_k cons_val_k) @@ -266,12 +266,12 @@ XX = (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R}) - RR = {:Fin(XX)*XX. + RR = {\z1,z2\:Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(z2, domain(f)) = f)) | (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(g, domain(f)) = f)) \ - z2={<0,x>})} + z2={\0,x\})} Then XX and RR satisfy the hypotheses of DC(omega). So applying DC: @@ -287,7 +287,7 @@ ************************************************************************* *) lemma singleton_in_funs: - "x \ X \ {<0,x>} \ + "x \ X \ {\0,x\} \ (\n \ nat. {f \ succ(n)->X. \k \ n. \ R})" apply (rule nat_0I [THEN UN_I]) apply (force simp add: singleton_0 [symmetric] @@ -300,15 +300,15 @@ defines XX_def: "XX \ (\n \ nat. {f \ succ(n)->domain(R). \k \ n. \ R})" and RR_def: - "RR \ {:Fin(XX)*XX. + "RR \ {\z1,z2\:Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. restrict(z2, domain(f)) = f)) | (\ (\g \ XX. domain(g)=succ(\f \ z1. domain(f)) - \ (\f \ z1. restrict(g, domain(f)) = f)) \ z2={<0,x>})}" + \ (\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)) + {\z1,z2\\Fin(XX)*XX. (domain(z2)=succ(\f \ z1. domain(f)) \ (\f \ z1. domain(f)) = b \ (\f \ z1. restrict(z2,domain(f)) = f))}" begin @@ -316,7 +316,7 @@ lemma lemma4: "\range(R) \ domain(R); x \ domain(R)\ \ RR \ Pow(XX)*XX \ - (\Y \ Pow(XX). Y \ nat \ (\x \ XX. :RR))" + (\Y \ Pow(XX). Y \ nat \ (\x \ XX. \Y,x\:RR))" apply (rule conjI) apply (force dest!: FinD [THEN PowI] simp add: RR_def) apply (rule impI [THEN ballI]) @@ -368,7 +368,7 @@ lemma restrict_cons_eq_restrict: "\restrict(h, domain(u))=u; h \ n->X; domain(u) \ n\ - \ restrict(cons(, h), domain(u)) = u" + \ restrict(cons(\n, y\, h), domain(u)) = u" apply (unfold restrict_def) apply (simp add: restrict_def Pi_iff) apply (erule sym [THEN trans, symmetric]) @@ -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 = "{\z1,z2\ \ Pow (A) *A . z1 \ Hartog (A) \ z2 \ z1}" in allE) apply simp apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE]) @@ -549,15 +549,15 @@ by (fast intro!: lam_type RepFunI) lemma lemmaX: - "\\Y \ Pow(X). Y \ K \ (\x \ X. \ R); + "\\Y \ Pow(X). Y \ K \ (\x \ X. \Y, x\ \ R); b \ K; Z \ Pow(X); Z \ K\ - \ {x \ X. \ R} \ 0" + \ {x \ X. \Z,x\ \ R} \ 0" by blast lemma WO1_DC_lemma: "\Card(K); well_ord(X,Q); - \Y \ Pow(X). Y \ K \ (\x \ X. \ R); b \ K\ + \Y \ Pow(X). Y \ K \ (\x \ X. \Y, x\ \ R); b \ K\ \ ff(b, X, Q, R) \ {x \ X. <(\c \ b. ff(c, X, Q, R))``b, x> \ R}" apply (rule_tac P = "b \ K" in impE, (erule_tac [2] asm_rl)+) apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct], diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/HH.thy --- a/src/ZF/AC/HH.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/HH.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,8 +12,8 @@ begin definition - HH :: "[i, i, i] => i" where - "HH(f,x,a) \ transrec(a, %b r. let z = x - (\c \ b. r`c) + HH :: "[i, i, i] \ i" where + "HH(f,x,a) \ transrec(a, \b r. let z = x - (\c \ b. r`c) in if f`z \ Pow(z)-{0} then f`z else {x})" subsection\Lemmas useful in each of the three proofs\ @@ -57,7 +57,7 @@ prefer 2 apply assumption+ apply (rule leI [THEN le_imp_subset, THEN subset_imp_Diff_eq, THEN ssubst], assumption) -apply (rule_tac t = "%z. z-X" for X in subst_context) +apply (rule_tac t = "\z. z-X" for X in subst_context) apply (rule Diff_UN_eq_self) apply (drule Ord_DiffE, assumption) apply (fast elim: ltE, auto) @@ -82,7 +82,7 @@ lemma HH_eq_arg_lt: "\HH(f,x,v)=HH(f,x,w); HH(f,x,v) \ Pow(x)-{0}; v \ w\ \ P" -apply (frule_tac P = "%y. y \ Pow (x) -{0}" in subst, assumption) +apply (frule_tac P = "\y. y \ Pow (x) -{0}" in subst, assumption) apply (drule_tac a = w in HH_subset_x_imp_subset_Diff_UN) apply (drule subst_elem, assumption) apply (fast intro!: singleton_iff [THEN iffD2] equals0I) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/Hartog.thy --- a/src/ZF/AC/Hartog.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/Hartog.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,7 +9,7 @@ begin definition - Hartog :: "i => i" where + Hartog :: "i \ i" where "Hartog(X) \ \ i. \ i \ X" lemma Ords_in_set: "\a. Ord(a) \ a \ X \ P" @@ -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=\Y,R\ \ ordertype(Y,R)=b}" apply (intro allI impI) apply (elim allE impE, assumption) apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/AC/WO2_AC16.thy --- a/src/ZF/AC/WO2_AC16.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/AC/WO2_AC16.thy Tue Sep 27 17:46:52 2022 +0100 @@ -18,10 +18,10 @@ (**** A recursive definition used in the proof of WO2 \ AC16 ****) definition - recfunAC16 :: "[i,i,i,i] => i" where + recfunAC16 :: "[i,i,i,i] \ i" where "recfunAC16(f,h,i,a) \ transrec2(i, 0, - %g r. if (\y \ r. h`g \ y) then r + \g r. if (\y \ r. h`g \ y) then r else r \ {f`(\ i. h`g \ f`i \ (\b f`i \ (\t \ r. \ h`b \ t))))})" @@ -224,7 +224,7 @@ empty_subsetI [THEN subset_imp_lepoll]) (*succ case*) apply (simp add: recfunAC16_succ - Diff_UN_succ_empty [of _ "%j. recfunAC16(f,g,j,a)"] + Diff_UN_succ_empty [of _ "\j. recfunAC16(f,g,j,a)"] empty_subsetI [THEN subset_imp_lepoll]) apply (best intro: Diff_UN_succ_subset [THEN subset_imp_lepoll] singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] lepoll_trans) @@ -563,7 +563,7 @@ apply (elim exE) apply (rename_tac h) apply (rule_tac x = "\j i" where + NN :: "i \ i" where "NN(y) \ {m \ nat. \a. \f. Ord(a) \ domain(f)=a \ (\b (\b m)}" definition - uu :: "[i, i, i, i] => i" where + uu :: "[i, i, i, i] \ i" where "uu(f, beta, gamma, delta) \ (f`beta * f`gamma) \ f`delta" (** Definitions for case 1 **) definition - vv1 :: "[i, i, i] => i" where + vv1 :: "[i, i, i] \ i" where "vv1(f,m,b) \ let g = \ g. (\d. Ord(d) \ (domain(uu(f,b,g,d)) \ 0 \ domain(uu(f,b,g,d)) \ m)); @@ -36,26 +36,26 @@ in if f`b \ 0 then domain(uu(f,b,g,d)) else 0" definition - ww1 :: "[i, i, i] => i" where + ww1 :: "[i, i, i] \ i" where "ww1(f,m,b) \ f`b - vv1(f,m,b)" definition - gg1 :: "[i, i, i] => i" where + gg1 :: "[i, i, i] \ i" where "gg1(f,a,m) \ \b \ a++a. if b i" where + vv2 :: "[i, i, i, i] \ i" where "vv2(f,b,g,s) \ if f`g \ 0 then {uu(f, b, g, \ d. uu(f,b,g,d) \ 0)`s} else 0" definition - ww2 :: "[i, i, i, i] => i" where + ww2 :: "[i, i, i, i] \ i" where "ww2(f,b,g,s) \ f`g - vv2(f,b,g,s)" definition - gg2 :: "[i, i, i, i] => i" where + gg2 :: "[i, i, i, i] \ i" where "gg2(f,a,b,s) \ \g \ a++a. if g a. \x. Ord(x) \ P(a, x))\ \ P(Least_a, \ b. P(Least_a, b))" apply (erule ssubst) -apply (rule_tac Q = "%z. P (z, \ b. P (z, b))" in LeastI2) +apply (rule_tac Q = "\z. P (z, \ b. P (z, b))" in LeastI2) apply (fast elim!: LeastI)+ done lemmas nested_Least_instance = - nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \ 0 \ + 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: @@ -409,26 +409,26 @@ (used only in the following two lemmas) *) lemma z_n_subset_z_succ_n: - "\n \ nat. rec(n, x, %k r. r \ r*r) \ rec(succ(n), x, %k r. r \ r*r)" + "\n \ nat. rec(n, x, \k r. r \ r*r) \ rec(succ(n), x, \k r. r \ r*r)" by (fast intro: rec_succ [THEN ssubst]) lemma le_subsets: "\\n \ nat. f(n)<=f(succ(n)); n\m; n \ nat; m \ nat\ \ f(n)<=f(m)" apply (erule_tac P = "n\m" in rev_mp) -apply (rule_tac P = "%z. n\z \ f (n) \ f (z) " in nat_induct) +apply (rule_tac P = "\z. n\z \ f (n) \ f (z) " in nat_induct) apply (auto simp add: le_iff) done lemma le_imp_rec_subset: "\n\m; m \ nat\ - \ rec(n, x, %k r. r \ r*r) \ rec(m, x, %k r. r \ r*r)" + \ rec(n, x, \k r. r \ r*r) \ rec(m, x, \k r. r \ r*r)" apply (rule z_n_subset_z_succ_n [THEN le_subsets]) apply (blast intro: lt_nat_in_nat)+ done lemma lemma_iv: "\y. x \ y*y \ y" -apply (rule_tac x = "\n \ nat. rec (n, x, %k r. r \ r*r) " in exI) +apply (rule_tac x = "\n \ nat. rec (n, x, \k r. r \ r*r) " in exI) apply safe apply (rule nat_0I [THEN UN_I], simp) apply (rule_tac a = "succ (n \ na) " in UN_I) @@ -477,7 +477,7 @@ apply (rule_tac x = a in exI) apply (rule_tac x = "\x \ y. \ i. f`i = {x}" in exI) apply (rule conjI, assumption) -apply (rule_tac d = "%i. THE x. x \ f`i" in lam_injective) +apply (rule_tac d = "\i. THE x. x \ f`i" in lam_injective) apply (drule lemma1, assumption+) apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord) apply (rule lemma2 [THEN ssubst], assumption+, blast) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Arith.thy --- a/src/ZF/Arith.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Arith.thy Tue Sep 27 17:46:52 2022 +0100 @@ -6,7 +6,7 @@ (*"Difference" is subtraction of natural numbers. There are no negative numbers; we have m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. - Also, rec(m, 0, %z w.z) is pred(m). + Also, rec(m, 0, \z w.z) is pred(m). *) section\Arithmetic Operators and Their Definitions\ @@ -16,18 +16,18 @@ text\Proofs about elementary arithmetic: addition, multiplication, etc.\ definition - pred :: "i=>i" (*inverse of succ*) where - "pred(y) \ nat_case(0, %x. x, y)" + pred :: "i\i" (*inverse of succ*) where + "pred(y) \ nat_case(0, \x. x, y)" definition - natify :: "i=>i" (*coerces non-nats to nats*) where - "natify \ Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a)) + natify :: "i\i" (*coerces non-nats to nats*) where + "natify \ Vrecursor(\f a. if a = succ(pred(a)) then succ(f`pred(a)) else 0)" consts - raw_add :: "[i,i]=>i" - raw_diff :: "[i,i]=>i" - raw_mult :: "[i,i]=>i" + raw_add :: "[i,i]\i" + raw_diff :: "[i,i]\i" + raw_mult :: "[i,i]\i" primrec "raw_add (0, n) = n" @@ -36,40 +36,40 @@ primrec raw_diff_0: "raw_diff(m, 0) = m" raw_diff_succ: "raw_diff(m, succ(n)) = - nat_case(0, %x. x, raw_diff(m, n))" + nat_case(0, \x. x, raw_diff(m, n))" primrec "raw_mult(0, n) = 0" "raw_mult(succ(m), n) = raw_add (n, raw_mult(m, n))" definition - add :: "[i,i]=>i" (infixl \#+\ 65) where + add :: "[i,i]\i" (infixl \#+\ 65) where "m #+ n \ raw_add (natify(m), natify(n))" definition - diff :: "[i,i]=>i" (infixl \#-\ 65) where + diff :: "[i,i]\i" (infixl \#-\ 65) where "m #- n \ raw_diff (natify(m), natify(n))" definition - mult :: "[i,i]=>i" (infixl \#*\ 70) where + mult :: "[i,i]\i" (infixl \#*\ 70) where "m #* n \ raw_mult (natify(m), natify(n))" definition - raw_div :: "[i,i]=>i" where + raw_div :: "[i,i]\i" where "raw_div (m, n) \ - transrec(m, %j f. if jj f. if ji" where + raw_mod :: "[i,i]\i" where "raw_mod (m, n) \ - transrec(m, %j f. if jj f. if ji" (infixl \div\ 70) where + div :: "[i,i]\i" (infixl \div\ 70) where "m div n \ raw_div (natify(m), natify(n))" definition - mod :: "[i,i]=>i" (infixl \mod\ 70) where + mod :: "[i,i]\i" (infixl \mod\ 70) where "m mod n \ raw_mod (natify(m), natify(n))" declare rec_type [simp] @@ -337,7 +337,7 @@ text\\\\ monotonicity, 1st argument\ lemma add_le_mono1: "\i \ j; j\nat\ \ i#+k \ j#+k" -apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck) +apply (rule_tac f = "\j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck) apply (blast intro: add_lt_mono1 add_type [THEN nat_into_Ord])+ done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Bin.thy --- a/src/ZF/Bin.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Bin.thy Tue Sep 27 17:46:52 2022 +0100 @@ -26,13 +26,13 @@ | Bit ("w \ bin", "b \ bool") (infixl \BIT\ 90) consts - integ_of :: "i=>i" - NCons :: "[i,i]=>i" - bin_succ :: "i=>i" - bin_pred :: "i=>i" - bin_minus :: "i=>i" - bin_adder :: "i=>i" - bin_mult :: "[i,i]=>i" + integ_of :: "i\i" + NCons :: "[i,i]\i" + bin_succ :: "i\i" + bin_pred :: "i\i" + bin_minus :: "i\i" + bin_adder :: "i\i" + bin_mult :: "[i,i]\i" primrec integ_of_Pls: "integ_of (Pls) = $# 0" @@ -74,7 +74,7 @@ "bin_adder (v BIT x) = (\w\bin. bin_case (v BIT x, bin_pred(v BIT x), - %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w), + \w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w), x xor y), w))" @@ -87,7 +87,7 @@ *) definition - bin_add :: "[i,i]=>i" where + bin_add :: "[i,i]\i" where "bin_add(v,w) \ bin_adder(v)`w" @@ -114,8 +114,8 @@ "#-2" \ "CONST integ_of(CONST Min BIT 0)" syntax - "_Int" :: "num_token => i" (\#_\ 1000) - "_Neg_Int" :: "num_token => i" (\#-_\ 1000) + "_Int" :: "num_token \ i" (\#_\ 1000) + "_Neg_Int" :: "num_token \ i" (\#-_\ 1000) ML_file \Tools/numeral_syntax.ML\ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Bool.thy --- a/src/ZF/Bool.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Bool.thy Tue Sep 27 17:46:52 2022 +0100 @@ -24,15 +24,15 @@ definition "not(b) \ cond(b,0,1)" definition - "and" :: "[i,i]=>i" (infixl \and\ 70) where + "and" :: "[i,i]\i" (infixl \and\ 70) where "a and b \ cond(a,b,0)" definition - or :: "[i,i]=>i" (infixl \or\ 65) where + or :: "[i,i]\i" (infixl \or\ 65) where "a or b \ cond(a,1,b)" definition - xor :: "[i,i]=>i" (infixl \xor\ 65) where + xor :: "[i,i]\i" (infixl \xor\ 65) where "a xor b \ cond(a,not(b),b)" @@ -152,7 +152,7 @@ definition - bool_of_o :: "o=>i" where + bool_of_o :: "o\i" where "bool_of_o(P) \ (if P then 1 else 0)" lemma [simp]: "bool_of_o(True) = 1" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Cardinal.thy --- a/src/ZF/Cardinal.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Cardinal.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,31 +9,31 @@ definition (*least ordinal operator*) - Least :: "(i=>o) => i" (binder \\ \ 10) where + Least :: "(i\o) \ i" (binder \\ \ 10) where "Least(P) \ THE i. Ord(i) \ P(i) \ (\j. j \P(j))" definition - eqpoll :: "[i,i] => o" (infixl \\\ 50) where + eqpoll :: "[i,i] \ o" (infixl \\\ 50) where "A \ B \ \f. f \ bij(A,B)" definition - lepoll :: "[i,i] => o" (infixl \\\ 50) where + lepoll :: "[i,i] \ o" (infixl \\\ 50) where "A \ B \ \f. f \ inj(A,B)" definition - lesspoll :: "[i,i] => o" (infixl \\\ 50) where + lesspoll :: "[i,i] \ o" (infixl \\\ 50) where "A \ B \ A \ B \ \(A \ B)" definition - cardinal :: "i=>i" (\|_|\) where + cardinal :: "i\i" (\|_|\) where "|A| \ (\ i. i \ A)" definition - Finite :: "i=>o" where + Finite :: "i\o" where "Finite(A) \ \n\nat. A \ n" definition - Card :: "i=>o" where + Card :: "i\o" where "Card(i) \ (i = |i|)" @@ -42,14 +42,14 @@ (** Lemma: Banach's Decomposition Theorem **) -lemma decomp_bnd_mono: "bnd_mono(X, %W. X - g``(Y - f``W))" +lemma decomp_bnd_mono: "bnd_mono(X, \W. X - g``(Y - f``W))" by (rule bnd_monoI, blast+) lemma Banach_last_equation: "g \ Y->X - \ g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) = - X - lfp(X, %W. X - g``(Y - f``W))" -apply (rule_tac P = "%u. v = X-u" for v + \ g``(Y - f`` lfp(X, \W. X - g``(Y - f``W))) = + X - lfp(X, \W. X - g``(Y - f``W))" +apply (rule_tac P = "\u. v = X-u" for v in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst]) apply (simp add: double_complement fun_is_rel [THEN image_subset]) done @@ -738,7 +738,7 @@ "\A \ B; b \ B\ \ cons(a,A) \ cons(b,B)" apply (unfold lepoll_def, safe) apply (rule_tac x = "\y\cons (a,A) . if y=a then b else f`y" in exI) -apply (rule_tac d = "%z. if z \ B then converse (f) `z else a" in lam_injective) +apply (rule_tac d = "\z. if z \ B then converse (f) `z else a" in lam_injective) apply (safe elim!: consE') apply simp_all apply (blast intro: inj_is_fun [THEN apply_type])+ @@ -769,7 +769,7 @@ lemma not_0_is_lepoll_1: "A \ 0 \ 1 \ A" apply (erule not_emptyE) apply (rule_tac a = "cons (x, A-{x}) " in subst) -apply (rule_tac [2] a = "cons(0,0)" and P= "%y. y \ cons (x, A-{x})" in subst) +apply (rule_tac [2] a = "cons(0,0)" and P= "\y. y \ cons (x, A-{x})" in subst) prefer 3 apply (blast intro: cons_lepoll_cong subset_imp_lepoll, auto) done @@ -796,8 +796,8 @@ "\f \ inj(A,B); A \ B = 0\ \ A \ (B - range(f)) \ B" apply (unfold eqpoll_def) apply (rule exI) -apply (rule_tac c = "%x. if x \ A then f`x else x" - and d = "%y. if y \ range (f) then converse (f) `y else y" +apply (rule_tac c = "\x. if x \ A then f`x else x" + and d = "\y. if y \ range (f) then converse (f) `y else y" in lam_bijective) apply (blast intro!: if_type inj_is_fun [THEN apply_type]) apply (simp (no_asm_simp) add: inj_converse_fun [THEN apply_funtype]) @@ -848,7 +848,7 @@ lemma Un_lepoll_sum: "A \ B \ A+B" apply (unfold lepoll_def) apply (rule_tac x = "\x\A \ B. if x\A then Inl (x) else Inr (x)" in exI) -apply (rule_tac d = "%z. snd (z)" in lam_injective) +apply (rule_tac d = "\z. snd (z)" in lam_injective) apply force apply (simp add: Inl_def Inr_def) done @@ -862,7 +862,7 @@ lemma disj_Un_eqpoll_sum: "A \ B = 0 \ A \ B \ A + B" apply (unfold eqpoll_def) apply (rule_tac x = "\a\A \ B. if a \ A then Inl (a) else Inr (a)" in exI) -apply (rule_tac d = "%z. case (%x. x, %x. x, z)" in lam_bijective) +apply (rule_tac d = "\z. case (\x. x, \x. x, z)" in lam_bijective) apply auto done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/CardinalArith.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,33 +8,33 @@ theory CardinalArith imports Cardinal OrderArith ArithSimp Finite begin definition - InfCard :: "i=>o" where + InfCard :: "i\o" where "InfCard(i) \ Card(i) \ nat \ i" definition - cmult :: "[i,i]=>i" (infixl \\\ 70) where + cmult :: "[i,i]\i" (infixl \\\ 70) where "i \ j \ |i*j|" definition - cadd :: "[i,i]=>i" (infixl \\\ 65) where + cadd :: "[i,i]\i" (infixl \\\ 65) where "i \ j \ |i+j|" definition - csquare_rel :: "i=>i" where + csquare_rel :: "i\i" where "csquare_rel(K) \ rvimage(K*K, - lam :K*K. y, x, y>, + lam \x,y\:K*K. y, x, y>, rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))" definition - jump_cardinal :: "i=>i" where + jump_cardinal :: "i\i" where \ \This definition is more complex than Kunen's but it more easily proved to be a cardinal\ "jump_cardinal(K) \ \X\Pow(K). {z. r \ Pow(K*K), well_ord(X,r) \ z = ordertype(X,r)}" definition - csucc :: "i=>i" where + csucc :: "i\i" where \ \needed because \<^term>\jump_cardinal(K)\ might not be the successor of \<^term>\K\\ "csucc(K) \ \ L. Card(L) \ KA \ C; B \ D\ \ A + B \ C + D" apply (unfold lepoll_def) apply (elim exE) -apply (rule_tac x = "\z\A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI) -apply (rule_tac d = "case (%w. Inl(converse(f) `w), %y. Inr(converse(fa) ` y))" +apply (rule_tac x = "\z\A+B. case (\w. Inl(f`w), \y. Inr(fa`y), z)" in exI) +apply (rule_tac d = "case (\w. Inl(converse(f) `w), \y. Inr(converse(fa) ` y))" in lam_injective) apply (typecheck add: inj_is_fun, auto) done @@ -183,8 +183,8 @@ lemma sum_succ_eqpoll: "succ(A)+B \ succ(A+B)" apply (unfold eqpoll_def) apply (rule exI) -apply (rule_tac c = "%z. if z=Inl (A) then A+B else z" - and d = "%z. if z=A+B then Inl (A) else z" in lam_bijective) +apply (rule_tac c = "\z. if z=Inl (A) then A+B else z" + and d = "\z. if z=A+B then Inl (A) else z" in lam_bijective) apply simp_all apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+ done @@ -221,7 +221,7 @@ lemma prod_commute_eqpoll: "A*B \ B*A" apply (unfold eqpoll_def) apply (rule exI) -apply (rule_tac c = "%." and d = "%." in lam_bijective, +apply (rule_tac c = "\\x,y\.\y,x\" and d = "\\x,y\.\y,x\" in lam_bijective, auto) done @@ -301,7 +301,7 @@ lemma prod_square_lepoll: "A \ A*A" apply (unfold lepoll_def inj_def) -apply (rule_tac x = "\x\A. " in exI, simp) +apply (rule_tac x = "\x\A. \x,x\" in exI, simp) done (*Could probably weaken the premise to well_ord(K,r), or remove using AC*) @@ -318,7 +318,7 @@ lemma prod_lepoll_self: "b \ B \ A \ A*B" apply (unfold lepoll_def inj_def) -apply (rule_tac x = "\x\A. " in exI, simp) +apply (rule_tac x = "\x\A. \x,b\" in exI, simp) done (*Could probably weaken the premises to well_ord(K,r), or removing using AC*) @@ -337,8 +337,8 @@ "\A \ C; B \ D\ \ A * B \ C * D" apply (unfold lepoll_def) apply (elim exE) -apply (rule_tac x = "lam :A*B. " in exI) -apply (rule_tac d = "%. " +apply (rule_tac x = "lam \w,y\:A*B. " in exI) +apply (rule_tac d = "\\w,y\. " in lam_injective) apply (typecheck add: inj_is_fun, auto) done @@ -357,8 +357,8 @@ lemma prod_succ_eqpoll: "succ(A)*B \ B + A*B" apply (unfold eqpoll_def) apply (rule exI) -apply (rule_tac c = "%. if x=A then Inl (y) else Inr ()" - and d = "case (%y. , %z. z)" in lam_bijective) +apply (rule_tac c = "\\x,y\. if x=A then Inl (y) else Inr (\x,y\)" + and d = "case (\y. \A,y\, \z. z)" in lam_bijective) apply safe apply (simp_all add: succI2 if_type mem_imp_not_eq) done @@ -399,7 +399,7 @@ (*This proof is modelled upon one assuming nat<=A, with injection \z\cons(u,A). if z=u then 0 else if z \ nat then succ(z) else z - and inverse %y. if y \ nat then nat_case(u, %z. z, y) else y. \ + and inverse \y. if y \ nat then nat_case(u, \z. z, y) else y. \ If f \ inj(nat,A) then range(f) behaves like the natural numbers.*) lemma nat_cons_lepoll: "nat \ A \ cons(u,A) \ A" apply (unfold lepoll_def) @@ -410,7 +410,7 @@ else if z \ range (f) then f`succ (converse (f) `z) else z" in exI) apply (rule_tac d = - "%y. if y \ range(f) then nat_case (u, %z. f`z, converse(f) `y) + "\y. if y \ range(f) then nat_case (u, \z. f`z, converse(f) `y) else y" in lam_injective) apply (fast intro!: if_type apply_type intro: inj_is_fun inj_converse_fun) @@ -491,7 +491,7 @@ subsubsection\Characterising initial segments of the well-ordering\ lemma csquareD: - "\<, > \ csquare_rel(K); x \ x \ z \ y \ z" + "\<\x,y\, \z,z\> \ csquare_rel(K); x \ x \ z \ y \ z" apply (unfold csquare_rel_def) apply (erule rev_mp) apply (elim ltE) @@ -501,14 +501,14 @@ done lemma pred_csquare_subset: - "z Order.pred(K*K, , csquare_rel(K)) \ succ(z)*succ(z)" + "z Order.pred(K*K, \z,z\, csquare_rel(K)) \ succ(z)*succ(z)" apply (unfold Order.pred_def) apply (safe del: SigmaI dest!: csquareD) apply (unfold lt_def, auto) done lemma csquare_ltI: - "\x \ <, > \ csquare_rel(K)" + "\x \ <\x,y\, \z,z\> \ 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 \ <\x,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)) ` " + ordermap(K*K, csquare_rel(K)) ` \x,y\ < + ordermap(K*K, csquare_rel(K)) ` \z,z\" apply (subgoal_tac "z well_ord (K*K, csquare_rel (K))") prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ Limit_is_Ord [THEN well_ord_csquare], clarify) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Cardinal_AC.thy --- a/src/ZF/Cardinal_AC.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Cardinal_AC.thy Tue Sep 27 17:46:52 2022 +0100 @@ -118,7 +118,7 @@ lemma surj_implies_inj: assumes f: "f \ surj(X,Y)" shows "\g. g \ inj(Y,X)" proof - - from f AC_Pi [of Y "%y. f-``{y}"] + from f AC_Pi [of Y "\y. f-``{y}"] obtain z where z: "z \ (\y\Y. f -`` {y})" by (auto simp add: surj_def) (fast dest: apply_Pair) show ?thesis @@ -161,8 +161,8 @@ proof (unfold lepoll_def) show "\f. f \ inj(\RepFun(K, X), K \ K)" apply (rule exI) - apply (rule_tac c = "%z. \\ i. z \ X(i), f ` (\ i. z \ X(i)) ` z\" - and d = "%\i,j\. converse (f`i) ` j" in lam_injective) + apply (rule_tac c = "\z. \\ i. z \ X(i), f ` (\ i. z \ X(i)) ` z\" + and d = "\\i,j\. converse (f`i) ` j" in lam_injective) apply (force intro: f inj_is_fun mems apply_type Perm.left_inverse)+ done qed diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Coind/ECR.thy --- a/src/ZF/Coind/ECR.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Coind/ECR.thy Tue Sep 27 17:46:52 2022 +0100 @@ -26,7 +26,7 @@ (* Pointwise extension to environments *) definition - hastyenv :: "[i,i] => o" where + hastyenv :: "[i,i] \ o" where "hastyenv(ve,te) \ ve_dom(ve) = te_dom(te) \ (\x \ ve_dom(ve). \ HasTyRel)" @@ -55,7 +55,7 @@ HasTyRel.dom_subset [THEN subsetD, THEN SigmaD1, THEN ValNEE] lemma hastyenv_owr: - "\ve \ ValEnv; te \ TyEnv; hastyenv(ve,te); \ HasTyRel\ + "\ve \ ValEnv; te \ TyEnv; hastyenv(ve,te); \v,t\ \ HasTyRel\ \ hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))" by (auto simp add: hastyenv_def ve_app_owr HasTyRel_non_zero) @@ -97,7 +97,7 @@ "\ve \ ValEnv; x \ ExVar; e \ Exp; f \ ExVar; cl \ Val; v_clos(x,e,ve_owr(ve,f,cl)) = cl; hastyenv(ve,te); \ ElabRel\ \ - \ HasTyRel" + \cl,t\ \ HasTyRel" apply (unfold hastyenv_def) apply (erule elab_fixE, safe) apply hypsubst_thin @@ -130,12 +130,12 @@ \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; \ EvalRel; - \t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel; + \t te. hastyenv(ve,te) \ \ ElabRel \ \v2,t\ \ HasTyRel; \ EvalRel; \t te. hastyenv(ve_owr(vem,xm,v2),te) \ - \ ElabRel \ \ HasTyRel; + \ ElabRel \ \v,t\ \ HasTyRel; hastyenv(ve,te); \ ElabRel\ - \ \ HasTyRel" + \ \v,t\ \ HasTyRel" apply (erule elab_appE) apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+) apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+) @@ -151,7 +151,7 @@ lemma consistency [rule_format]: " \ EvalRel - \ (\t te. hastyenv(ve,te) \ \ ElabRel \ \ HasTyRel)" + \ (\t te. hastyenv(ve,te) \ \ ElabRel \ \v,t\ \ HasTyRel)" apply (erule EvalRel.induct) apply (simp_all add: consistency_const consistency_var consistency_fn consistency_fix consistency_app1) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Coind/Language.thy --- a/src/ZF/Coind/Language.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Coind/Language.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,7 +10,7 @@ axiomatization Const :: i and (* Abstract type of constants *) - c_app :: "[i,i] => i" (* Abstract constructor for fun application*) + c_app :: "[i,i] \ i" (* Abstract constructor for fun application*) where constNEE: "c \ Const \ c \ 0" and c_appI: "\c1 \ Const; c2 \ Const\ \ c_app(c1,c2) \ Const" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Coind/Map.thy --- a/src/ZF/Coind/Map.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Coind/Map.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,11 +8,11 @@ theory Map imports ZF begin definition - TMap :: "[i,i] => i" where + TMap :: "[i,i] \ i" where "TMap(A,B) \ {m \ Pow(A*\(B)).\a \ A. m``{a} \ B}" definition - PMap :: "[i,i] => i" where + PMap :: "[i,i] \ i" where "PMap(A,B) \ TMap(A,cons(0,B))" (* Note: 0 \ B \ TMap(A,B) = PMap(A,B) *) @@ -22,11 +22,11 @@ "map_emp \ 0" definition - map_owr :: "[i,i,i]=>i" where + map_owr :: "[i,i,i]\i" where "map_owr(m,a,b) \ \x \ {a} \ domain(m). if x=a then b else m``{x}" definition - map_app :: "[i,i]=>i" where + map_app :: "[i,i]\i" where "map_app(m,a) \ m``{a}" lemma "{0,1} \ {1} \ TMap(I, {0,1})" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Coind/Static.thy --- a/src/ZF/Coind/Static.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Coind/Static.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,13 +9,13 @@ parameter of the proof. A concrete version could be defined inductively. ***) -axiomatization isof :: "[i,i] => o" +axiomatization isof :: "[i,i] \ o" where isof_app: "\isof(c1,t_fun(t1,t2)); isof(c2,t1)\ \ isof(c_app(c1,c2),t2)" (*Its extension to environments*) definition - isofenv :: "[i,i] => o" where + isofenv :: "[i,i] \ o" where "isofenv(ve,te) \ ve_dom(ve) = te_dom(te) \ (\x \ ve_dom(ve). diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Coind/Types.thy --- a/src/ZF/Coind/Types.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Coind/Types.thy Tue Sep 27 17:46:52 2022 +0100 @@ -24,8 +24,8 @@ | te_owr ("te \ TyEnv","x \ ExVar","t \ Ty") consts - te_dom :: "i => i" - te_app :: "[i,i] => i" + te_dom :: "i \ i" + te_app :: "[i,i] \ i" primrec (*domain of the type environment*) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Coind/Values.thy --- a/src/ZF/Coind/Values.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Coind/Values.thy Tue Sep 27 17:46:52 2022 +0100 @@ -21,9 +21,9 @@ type_intros A_into_univ mapQU consts - ve_owr :: "[i,i,i] => i" - ve_dom :: "i=>i" - ve_app :: "[i,i] => i" + ve_owr :: "[i,i,i] \ i" + ve_dom :: "i\i" + ve_app :: "[i,i] \ i" primrec "ve_owr(ve_mk(m), x, v) = ve_mk(map_owr(m,x,v))" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/AC_in_L.thy --- a/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/AC_in_L.thy Tue Sep 27 17:46:52 2022 +0100 @@ -11,7 +11,7 @@ text\This could be moved into a library.\ consts - rlist :: "[i,i]=>i" + rlist :: "[i,i]\i" inductive domains "rlist(A,r)" \ "list(A) * list(A)" @@ -159,20 +159,20 @@ assumes fn_inj: "fn \ inj(nat*nat, nat)" -consts enum :: "[i,i]=>i" +consts enum :: "[i,i]\i" primrec - "enum(f, Member(x,y)) = f ` <0, f ` >" - "enum(f, Equal(x,y)) = f ` <1, f ` >" + "enum(f, Member(x,y)) = f ` <0, f ` \x,y\>" + "enum(f, Equal(x,y)) = f ` <1, f ` \x,y\>" "enum(f, Nand(p,q)) = f ` <2, f ` >" "enum(f, Forall(p)) = f ` " lemma (in Nat_Times_Nat) fn_type [TC,simp]: - "\x \ nat; y \ nat\ \ fn` \ nat" + "\x \ nat; y \ nat\ \ fn`\x,y\ \ nat" by (blast intro: inj_is_fun [OF fn_inj] apply_funtype) lemma (in Nat_Times_Nat) fn_iff: "\x \ nat; y \ nat; u \ nat; v \ nat\ - \ (fn` = fn`) \ (x=u \ y=v)" + \ (fn`\x,y\ = fn`\u,v\) \ (x=u \ y=v)" by (blast dest: inj_apply_equality [OF fn_inj]) lemma (in Nat_Times_Nat) enum_type [TC,simp]: @@ -230,34 +230,34 @@ (environment, formula) pairs into the ordinals. For each member of \<^term>\DPow(A)\, we take the minimum such ordinal.\ definition - env_form_r :: "[i,i,i]=>i" where + env_form_r :: "[i,i,i]\i" where \ \wellordering on (environment, formula) pairs\ "env_form_r(f,r,A) \ rmult(list(A), rlist(A, r), formula, measure(formula, enum(f)))" definition - env_form_map :: "[i,i,i,i]=>i" where + env_form_map :: "[i,i,i,i]\i" where \ \map from (environment, formula) pairs to ordinals\ "env_form_map(f,r,A,z) \ ordermap(list(A) * formula, env_form_r(f,r,A)) ` z" definition - DPow_ord :: "[i,i,i,i,i]=>o" where + DPow_ord :: "[i,i,i,i,i]\o" where \ \predicate that holds if \<^term>\k\ is a valid index for \<^term>\X\\ "DPow_ord(f,r,A,X,k) \ \env \ list(A). \p \ formula. arity(p) \ succ(length(env)) \ X = {x\A. sats(A, p, Cons(x,env))} \ - env_form_map(f,r,A,) = k" + env_form_map(f,r,A,\env,p\) = k" definition - DPow_least :: "[i,i,i,i]=>i" where + DPow_least :: "[i,i,i,i]\i" where \ \function yielding the smallest index for \<^term>\X\\ "DPow_least(f,r,A,X) \ \ k. DPow_ord(f,r,A,X,k)" definition - DPow_r :: "[i,i,i]=>i" where + DPow_r :: "[i,i,i]\i" where \ \a wellordering on \<^term>\DPow(A)\\ "DPow_r(f,r,A) \ measure(DPow(A), DPow_least(f,r,A))" @@ -329,7 +329,7 @@ of wellorderings for smaller ordinals.\ definition - rlimit :: "[i,i=>i]=>i" where + rlimit :: "[i,i\i]\i" where \ \Expresses the wellordering at limit ordinals. The conditional lets us remove the premise \<^term>\Limit(i)\ from some theorems.\ "rlimit(i,r) \ @@ -341,7 +341,7 @@ else 0" definition - Lset_new :: "i=>i" where + Lset_new :: "i\i" where \ \This constant denotes the set of elements introduced at level \<^term>\succ(i)\\ "Lset_new(i) \ {x \ Lset(succ(i)). lrank(x) = i}" @@ -413,8 +413,8 @@ subsection\Transfinite Definition of the Wellordering on \<^term>\L\\ definition - L_r :: "[i, i] => i" where - "L_r(f) \ %i. + L_r :: "[i, i] \ i" where + "L_r(f) \ \i. transrec3(i, 0, \x r. DPow_r(f, r, Lset(x)), \x r. rlimit(x, \y. r`y))" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/DPow_absolute.thy --- a/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/DPow_absolute.thy Tue Sep 27 17:46:52 2022 +0100 @@ -15,7 +15,7 @@ text\The three arguments of \<^term>\p\ are always 2, 1, 0. It is buried within 11 quantifiers\\ -(* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" +(* 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) \ 2 1 0 @@ -23,7 +23,7 @@ *) definition - formula_rec_fm :: "[i, i, i]=>i" where + formula_rec_fm :: "[i, i, i]\i" where "formula_rec_fm(mh,p,z) \ Exists(Exists(Exists( And(finite_ordinal_fm(2), @@ -81,7 +81,7 @@ (* is_satisfies(M,A,p,z) \ is_formula_rec (M, satisfies_MH(M,A), p, z) *) definition - satisfies_fm :: "[i,i,i]=>i" where + satisfies_fm :: "[i,i,i]\i" where "satisfies_fm(x) \ formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))" lemma is_satisfies_type [TC]: @@ -113,14 +113,14 @@ 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 = \env,p\ \ z = {x\A. sats(A, p, Cons(x,env))}}" by (simp add: DPow'_def, blast) text\Relativize the use of \<^term>\sats\ within \<^term>\DPow'\ (the comprehension).\ definition - is_DPow_sats :: "[i=>o,i,i,i,i] => o" where + is_DPow_sats :: "[i\o,i,i,i,i] \ o" where "is_DPow_sats(M,A,env,p,x) \ \n1[M]. \e[M]. \sp[M]. is_satisfies(M,A,p,sp) \ is_Cons(M,x,env,e) \ @@ -149,7 +149,7 @@ fun_apply(M, sp, e, n1) \ number1(M, n1) *) definition - DPow_sats_fm :: "[i,i,i,i]=>i" where + DPow_sats_fm :: "[i,i,i,i]\i" where "DPow_sats_fm(A,env,p,x) \ Forall(Forall(Forall( Implies(satisfies_fm(A#+3,p#+3,0), @@ -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 = \env,p\ \ z = {x \ A . sats(A, p, Cons(x, env))})" by (insert rep [of A], simp add: Collect_DPow_sats_abs) @@ -220,7 +220,7 @@ text\Relativization of the Operator \<^term>\DPow'\\ definition - is_DPow' :: "[i=>o,i,i] => o" where + is_DPow' :: "[i\o,i,i] \ o" where "is_DPow'(M,A,Z) \ \X[M]. X \ Z \ subset(M,X,A) \ @@ -308,11 +308,11 @@ text\The formula \<^term>\is_P\ has one free variable, 0, and it is enclosed within a single quantifier.\ -(* is_Collect :: "[i=>o,i,i=>o,i] => o" +(* is_Collect :: "[i\o,i,i\o,i] \ o" "is_Collect(M,A,P,z) \ \x[M]. x \ z \ x \ A \ P(x)" *) definition - Collect_fm :: "[i, i, i]=>i" where + Collect_fm :: "[i, i, i]\i" where "Collect_fm(A,is_P,z) \ Forall(Iff(Member(0,succ(z)), And(Member(0,succ(A)), is_P)))" @@ -359,11 +359,11 @@ text\BEWARE! The formula \<^term>\is_P\ has free variables 0, 1 and not the usual 1, 0! It is enclosed within two quantifiers.\ -(* is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" +(* 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))" *) definition - Replace_fm :: "[i, i, i]=>i" where + Replace_fm :: "[i, i, i]\i" where "Replace_fm(A,is_P,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,A#+2), is_P))))" @@ -417,7 +417,7 @@ is_Collect(M, A, is_DPow_sats(M,A,env,p), X))" *) definition - DPow'_fm :: "[i,i]=>i" where + DPow'_fm :: "[i,i]\i" where "DPow'_fm(A,Z) \ Forall( Iff(Member(0,succ(Z)), @@ -456,7 +456,7 @@ subsection\A Locale for Relativizing the Operator \<^term>\Lset\\ definition - transrec_body :: "[i=>o,i,i,i,i] => o" where + 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)" @@ -508,11 +508,11 @@ text\Relativization of the Operator \<^term>\Lset\\ definition - is_Lset :: "[i=>o, i, i] => o" where + is_Lset :: "[i\o, i, i] \ o" where \ \We can use the term language below because \<^term>\is_Lset\ will not have to be internalized: it isn't used in any instance of separation.\ - "is_Lset(M,a,z) \ is_transrec(M, %x f u. u = (\y\x. DPow'(f`y)), a, z)" + "is_Lset(M,a,z) \ is_transrec(M, \x f u. u = (\y\x. DPow'(f`y)), a, z)" lemma (in M_Lset) Lset_abs: "\Ord(i); M(i); M(z)\ @@ -613,7 +613,7 @@ subsection\The Notion of Constructible Set\ definition - constructible :: "[i=>o,i] => o" where + constructible :: "[i\o,i] \ o" where "constructible(M,x) \ \i[M]. \Li[M]. ordinal(M,i) \ is_Lset(M,i,Li) \ x \ Li" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Datatype_absolute.thy --- a/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Datatype_absolute.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,11 +10,11 @@ subsection\The lfp of a continuous function can be expressed as a union\ definition - directed :: "i=>o" where + directed :: "i\o" where "directed(A) \ A\0 \ (\x\A. \y\A. x \ y \ A)" definition - contin :: "(i=>i) => o" where + contin :: "(i\i) \ o" where "contin(h) \ (\A. directed(A) \ h(\A) = (\X\A. h(X)))" lemma bnd_mono_iterates_subset: "\bnd_mono(D, h); n \ nat\ \ h^n (0) \ D" @@ -114,19 +114,19 @@ subsection \Absoluteness for "Iterates"\ definition - iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where + 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), n, z)" definition - is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o" where + 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) \ is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)" definition - iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where + iterates_replacement :: "[i\o, [i,i]\o, i] \ o" where "iterates_replacement(M,isF,v) \ \n[M]. n\nat \ wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))" @@ -209,7 +209,7 @@ definition - is_list_functor :: "[i=>o,i,i,i] => o" where + 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)" @@ -262,7 +262,7 @@ definition - is_formula_functor :: "[i=>o,i,i] => o" where + 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) \ @@ -280,7 +280,7 @@ subsection\\<^term>\M\ Contains the List and Formula Datatypes\ definition - list_N :: "[i,i] => i" where + list_N :: "[i,i] \ i" where "list_N(A,n) \ (\X. {0} + A * X)^n (0)" lemma Nil_in_list_N [simp]: "[] \ list_N(A,succ(n))" @@ -341,24 +341,24 @@ done definition - is_list_N :: "[i=>o,i,i,i] => o" where + is_list_N :: "[i\o,i,i,i] \ o" where "is_list_N(M,A,n,Z) \ \zero[M]. empty(M,zero) \ is_iterates(M, is_list_functor(M,A), zero, n, Z)" definition - mem_list :: "[i=>o,i,i] => o" where + mem_list :: "[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" definition - is_list :: "[i=>o,i,i] => o" where + is_list :: "[i\o,i,i] \ o" where "is_list(M,A,Z) \ \l[M]. l \ Z \ mem_list(M,A,l)" subsubsection\Towards Absoluteness of \<^term>\formula_rec\\ -consts depth :: "i=>i" +consts depth :: "i\i" primrec "depth(Member(x,y)) = 0" "depth(Equal(x,y)) = 0" @@ -370,7 +370,7 @@ definition - formula_N :: "i => i" where + formula_N :: "i \ i" where "formula_N(n) \ (\X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)" lemma Member_in_formula_N [simp]: @@ -445,20 +445,20 @@ done definition - is_formula_N :: "[i=>o,i,i] => o" where + is_formula_N :: "[i\o,i,i] \ o" where "is_formula_N(M,n,Z) \ \zero[M]. empty(M,zero) \ is_iterates(M, is_formula_functor(M), zero, n, Z)" definition - mem_formula :: "[i=>o,i] => o" where + mem_formula :: "[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" definition - is_formula :: "[i=>o,i] => o" where + is_formula :: "[i\o,i] \ o" where "is_formula(M,Z) \ \p[M]. p \ Z \ mem_formula(M,p)" locale M_datatypes = M_trancl + @@ -473,7 +473,7 @@ "strong_replacement(M, \n y. n\nat \ is_iterates(M, is_formula_functor(M), 0, n, y))" and nth_replacement: - "M(l) \ iterates_replacement(M, %l t. is_tl(M,l,t), l)" + "M(l) \ iterates_replacement(M, \l t. is_tl(M,l,t), l)" subsubsection\Absoluteness of the List Construction\ @@ -588,17 +588,17 @@ done definition - is_eclose_n :: "[i=>o,i,i,i] => o" where + is_eclose_n :: "[i\o,i,i,i] \ o" where "is_eclose_n(M,A,n,Z) \ is_iterates(M, big_union(M), A, n, Z)" definition - mem_eclose :: "[i=>o,i,i] => o" where + 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" definition - is_eclose :: "[i=>o,i,i] => o" where + is_eclose :: "[i\o,i,i] \ o" where "is_eclose(M,A,Z) \ \u[M]. u \ Z \ mem_eclose(M,A,u)" @@ -650,14 +650,14 @@ text\\<^prop>\transrec(a,H) \ wfrec(Memrel(eclose({a})), a, H)\\ definition - is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where + 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) \ is_wfrec(M,MH,mesa,a,z)" definition - transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where + 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) \ @@ -698,7 +698,7 @@ text\But it is never used.\ definition - is_length :: "[i=>o,i,i,i] => o" where + 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 \ @@ -746,7 +746,7 @@ done definition - is_nth :: "[i=>o,i,i,i] => o" where + 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)" @@ -764,7 +764,7 @@ subsection\Relativization and Absoluteness for the \<^term>\formula\ Constructors\ definition - is_Member :: "[i=>o,i,i,i] => o" where + 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)" @@ -778,7 +778,7 @@ by (simp add: Member_def) definition - is_Equal :: "[i=>o,i,i,i] => o" where + 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)" @@ -791,7 +791,7 @@ by (simp add: Equal_def) definition - is_Nand :: "[i=>o,i,i,i] => o" where + 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)" @@ -804,7 +804,7 @@ by (simp add: Nand_def) definition - is_Forall :: "[i=>o,i,i] => o" where + 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)" @@ -820,7 +820,7 @@ subsection \Absoluteness for \<^term>\formula_rec\\ definition - formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where + formula_rec_case :: "[[i,i]\i, [i,i]\i, [i,i,i,i]\i, [i,i]\i, i, i] \ i" where \ \the instance of \<^term>\formula_case\ in \<^term>\formula_rec\\ "formula_rec_case(a,b,c,d,h) \ formula_case (a, b, @@ -854,7 +854,7 @@ subsubsection\Absoluteness for the Formula Operator \<^term>\depth\\ definition - is_depth :: "[i=>o,i,i] => o" where + 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 \ @@ -880,7 +880,7 @@ definition is_formula_case :: - "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where + "[i\o, [i,i,i]\o, [i,i,i]\o, [i,i,i]\o, [i,i]\o, i, i] \ o" where \ \no constraint on non-formulas\ "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) \ (\x[M]. \y[M]. finite_ordinal(M,x) \ finite_ordinal(M,y) \ @@ -914,7 +914,7 @@ subsubsection \Absoluteness for \<^term>\formula_rec\: Final Results\ definition - is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where + 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) \ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Formula.thy --- a/src/ZF/Constructible/Formula.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Formula.thy Tue Sep 27 17:46:52 2022 +0100 @@ -21,27 +21,27 @@ declare formula.intros [TC] definition - Neg :: "i=>i" where + Neg :: "i\i" where "Neg(p) \ Nand(p,p)" definition - And :: "[i,i]=>i" where + And :: "[i,i]\i" where "And(p,q) \ Neg(Nand(p,q))" definition - Or :: "[i,i]=>i" where + Or :: "[i,i]\i" where "Or(p,q) \ Nand(Neg(p),Neg(q))" definition - Implies :: "[i,i]=>i" where + Implies :: "[i,i]\i" where "Implies(p,q) \ Nand(p,Neg(q))" definition - Iff :: "[i,i]=>i" where + Iff :: "[i,i]\i" where "Iff(p,q) \ And(Implies(p,q), Implies(q,p))" definition - Exists :: "i=>i" where + Exists :: "i\i" where "Exists(p) \ Neg(Forall(Neg(p)))" lemma Neg_type [TC]: "p \ formula \ Neg(p) \ formula" @@ -65,7 +65,7 @@ by (simp add: Exists_def) -consts satisfies :: "[i,i]=>i" +consts satisfies :: "[i,i]\i" primrec (*explicit lambda is required because the environment varies*) "satisfies(A,Member(x,y)) = (\env \ list(A). bool_of_o (nth(x,env) \ nth(y,env)))" @@ -84,7 +84,7 @@ by (induct set: formula) simp_all abbreviation - sats :: "[i,i,i] => o" where + sats :: "[i,i,i] \ o" where "sats(A,p,env) \ satisfies(A,p)`env = 1" lemma sats_Member_iff [simp]: @@ -195,7 +195,7 @@ subsection\Arity of a Formula: Maximum Free de Bruijn Index\ -consts arity :: "i=>i" +consts arity :: "i\i" primrec "arity(Member(x,y)) = succ(x) \ succ(y)" @@ -250,7 +250,7 @@ subsection\Renaming Some de Bruijn Variables\ definition - incr_var :: "[i,i]=>i" where + incr_var :: "[i,i]\i" where "incr_var(x,nq) \ if x incr_var(x,nq) = x" @@ -261,7 +261,7 @@ apply (blast dest: lt_trans1) done -consts incr_bv :: "i=>i" +consts incr_bv :: "i\i" primrec "incr_bv(Member(x,y)) = (\nq \ nat. Member (incr_var(x,nq), incr_var(y,nq)))" @@ -339,7 +339,7 @@ subsection\Renaming all but the First de Bruijn Variable\ definition - incr_bv1 :: "i => i" where + incr_bv1 :: "i \ i" where "incr_bv1(p) \ incr_bv(p)`1" @@ -391,7 +391,7 @@ text\The definable powerset operation: Kunen's definition VI 1.1, page 165.\ definition - DPow :: "i => i" where + DPow :: "i \ i" where "DPow(A) \ {X \ Pow(A). \env \ list(A). \p \ formula. arity(p) \ succ(length(env)) \ @@ -514,7 +514,7 @@ subsubsection\The subset relation\ definition - subset_fm :: "[i,i]=>i" where + subset_fm :: "[i,i]\i" where "subset_fm(x,y) \ Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))" lemma subset_type [TC]: "\x \ nat; y \ nat\ \ subset_fm(x,y) \ formula" @@ -535,7 +535,7 @@ subsubsection\Transitive sets\ definition - transset_fm :: "i=>i" where + transset_fm :: "i\i" where "transset_fm(x) \ Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))" lemma transset_type [TC]: "x \ nat \ transset_fm(x) \ formula" @@ -556,7 +556,7 @@ subsubsection\Ordinals\ definition - ordinal_fm :: "i=>i" where + ordinal_fm :: "i\i" where "ordinal_fm(x) \ And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))" @@ -588,11 +588,11 @@ subsection\Constant Lset: Levels of the Constructible Universe\ definition - Lset :: "i=>i" where - "Lset(i) \ transrec(i, %x f. \y\x. DPow(f`y))" + Lset :: "i\i" where + "Lset(i) \ transrec(i, \x f. \y\x. DPow(f`y))" definition - L :: "i=>o" where \ \Kunen's definition VI 1.5, page 167\ + L :: "i\o" where \ \Kunen's definition VI 1.5, page 167\ "L(x) \ \i. Ord(i) \ x \ Lset(i)" text\NOT SUITABLE FOR REWRITING -- RECURSIVE!\ @@ -798,7 +798,7 @@ by (simp add: Lset_succ empty_in_DPow cons_in_DPow) lemma Pair_in_Lset: - "\a \ Lset(i); b \ Lset(i); Ord(i)\ \ \ Lset(succ(succ(i)))" + "\a \ Lset(i); b \ Lset(i); Ord(i)\ \ \a,b\ \ Lset(succ(succ(i)))" apply (unfold Pair_def) apply (blast intro: doubleton_in_Lset) done @@ -822,7 +822,7 @@ done lemma Pair_in_LLimit: - "\a \ Lset(i); b \ Lset(i); Limit(i)\ \ \ Lset(i)" + "\a \ Lset(i); b \ Lset(i); Limit(i)\ \ \a,b\ \ Lset(i)" txt\Infer that a, b occur at ordinals x,xa < i.\ apply (erule Limit_LsetE, assumption) apply (erule Limit_LsetE, assumption) @@ -835,7 +835,7 @@ text\The rank function for the constructible universe\ definition - lrank :: "i=>i" where \ \Kunen's definition VI 1.7\ + lrank :: "i\i" where \ \Kunen's definition VI 1.7\ "lrank(x) \ \ i. x \ Lset(succ(i))" lemma L_I: "\x \ Lset(i); Ord(i)\ \ L(x)" @@ -995,7 +995,7 @@ text\A simpler version of \<^term>\DPow\: no arity check!\ definition - DPow' :: "i => i" where + DPow' :: "i \ i" where "DPow'(A) \ {X \ Pow(A). \env \ list(A). \p \ formula. X = {x\A. sats(A, p, Cons(x,env))}}" @@ -1022,7 +1022,7 @@ text\And thus we can relativize \<^term>\Lset\ without bothering with \<^term>\arity\ and \<^term>\length\\ -lemma Lset_eq_transrec_DPow': "Lset(i) = transrec(i, %x f. \y\x. DPow'(f`y))" +lemma Lset_eq_transrec_DPow': "Lset(i) = transrec(i, \x f. \y\x. DPow'(f`y))" apply (rule_tac a=i in eps_induct) apply (subst Lset) apply (subst transrec) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Internalize.thy --- a/src/ZF/Constructible/Internalize.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Internalize.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,7 +10,7 @@ (* is_Inl(M,a,z) \ \zero[M]. empty(M,zero) \ pair(M,zero,a,z) *) definition - Inl_fm :: "[i,i]=>i" where + Inl_fm :: "[i,i]\i" where "Inl_fm(a,z) \ Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))" lemma Inl_type [TC]: @@ -40,7 +40,7 @@ (* is_Inr(M,a,z) \ \n1[M]. number1(M,n1) \ pair(M,n1,a,z) *) definition - Inr_fm :: "[i,i]=>i" where + Inr_fm :: "[i,i]\i" where "Inr_fm(a,z) \ Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))" lemma Inr_type [TC]: @@ -71,7 +71,7 @@ (* is_Nil(M,xs) \ \zero[M]. empty(M,zero) \ is_Inl(M,zero,xs) *) definition - Nil_fm :: "i=>i" where + Nil_fm :: "i\i" where "Nil_fm(x) \ Exists(And(empty_fm(0), Inl_fm(0,succ(x))))" lemma Nil_type [TC]: "x \ nat \ Nil_fm(x) \ formula" @@ -100,7 +100,7 @@ (* "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 :: "[i,i,i]\i" where "Cons_fm(a,l,Z) \ Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))" @@ -132,7 +132,7 @@ (* is_quasilist(M,xs) \ is_Nil(M,z) | (\x[M]. \l[M]. is_Cons(M,x,l,z))" *) definition - quasilist_fm :: "i=>i" where + quasilist_fm :: "i\i" where "quasilist_fm(x) \ Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))" @@ -167,7 +167,7 @@ (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | H=x) \ (is_quasilist(M,xs) | empty(M,H))" *) definition - hd_fm :: "[i,i]=>i" where + hd_fm :: "[i,i]\i" where "hd_fm(xs,H) \ And(Implies(Nil_fm(xs), empty_fm(H)), And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(H#+2,1)))), @@ -204,7 +204,7 @@ (\x[M]. \l[M]. \ is_Cons(M,x,l,xs) | T=l) \ (is_quasilist(M,xs) | empty(M,T))" *) definition - tl_fm :: "[i,i]=>i" where + tl_fm :: "[i,i]\i" where "tl_fm(xs,T) \ And(Implies(Nil_fm(xs), Equal(T,xs)), And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))), @@ -236,12 +236,12 @@ subsubsection\The Operator \<^term>\is_bool_of_o\\ -(* is_bool_of_o :: "[i=>o, o, i] => 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))" *) text\The formula \<^term>\p\ has no free variables.\ definition - bool_of_o_fm :: "[i, i]=>i" where + bool_of_o_fm :: "[i, i]\i" where "bool_of_o_fm(p,z) \ Or(And(p,number1_fm(z)), And(Neg(p),empty_fm(z)))" @@ -279,12 +279,12 @@ text\The two arguments of \<^term>\p\ are always 1, 0. Remember that \<^term>\p\ will be enclosed by three quantifiers.\ -(* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" +(* 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))" *) definition - lambda_fm :: "[i, i, i]=>i" where + lambda_fm :: "[i, i, i]\i" where "lambda_fm(p,A,z) \ Forall(Iff(Member(0,succ(z)), Exists(Exists(And(Member(1,A#+3), @@ -324,7 +324,7 @@ (* "is_Member(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) \ is_Inl(M,p,u) \ is_Inl(M,u,Z)" *) definition - Member_fm :: "[i,i,i]=>i" where + Member_fm :: "[i,i,i]\i" where "Member_fm(x,y,Z) \ Exists(Exists(And(pair_fm(x#+2,y#+2,1), And(Inl_fm(1,0), Inl_fm(0,Z#+2)))))" @@ -357,7 +357,7 @@ (* "is_Equal(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) \ is_Inr(M,p,u) \ is_Inl(M,u,Z)" *) definition - Equal_fm :: "[i,i,i]=>i" where + Equal_fm :: "[i,i,i]\i" where "Equal_fm(x,y,Z) \ Exists(Exists(And(pair_fm(x#+2,y#+2,1), And(Inr_fm(1,0), Inl_fm(0,Z#+2)))))" @@ -390,7 +390,7 @@ (* "is_Nand(M,x,y,Z) \ \p[M]. \u[M]. pair(M,x,y,p) \ is_Inl(M,p,u) \ is_Inr(M,u,Z)" *) definition - Nand_fm :: "[i,i,i]=>i" where + Nand_fm :: "[i,i,i]\i" where "Nand_fm(x,y,Z) \ Exists(Exists(And(pair_fm(x#+2,y#+2,1), And(Inl_fm(1,0), Inr_fm(0,Z#+2)))))" @@ -422,7 +422,7 @@ (* "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 :: "[i,i]\i" where "Forall_fm(x,Z) \ Exists(And(Inr_fm(succ(x),0), Inr_fm(0,succ(Z))))" @@ -455,7 +455,7 @@ (* 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 :: "[i,i,i]\i" where "and_fm(a,b,z) \ Or(And(number1_fm(a), Equal(z,b)), And(Neg(number1_fm(a)),empty_fm(z)))" @@ -490,7 +490,7 @@ (\number1(M,a) \ z=b) *) definition - or_fm :: "[i,i,i]=>i" where + or_fm :: "[i,i,i]\i" where "or_fm(a,b,z) \ Or(And(number1_fm(a), number1_fm(z)), And(Neg(number1_fm(a)), Equal(z,b)))" @@ -525,7 +525,7 @@ (* 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 :: "[i,i]\i" where "not_fm(a,z) \ Or(And(number1_fm(a), empty_fm(z)), And(Neg(number1_fm(a)), number1_fm(z)))" @@ -578,7 +578,7 @@ done -(* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" +(* M_is_recfun :: "[i\o, [i,i,i]\o, i, i, i] \ o" "M_is_recfun(M,MH,r,a,f) \ \z[M]. z \ f \ 2 1 0 @@ -592,7 +592,7 @@ text\The three arguments of \<^term>\p\ are always 2, 1, 0 and z\ definition - is_recfun_fm :: "[i, i, i, i]=>i" where + is_recfun_fm :: "[i, i, i, i]\i" where "is_recfun_fm(p,r,a,f) \ Forall(Iff(Member(0,succ(f)), Exists(Exists(Exists( @@ -651,11 +651,11 @@ text\The three arguments of \<^term>\p\ are always 2, 1, 0; \<^term>\p\ is enclosed by 5 quantifiers.\ -(* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o" +(* 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)" *) definition - is_wfrec_fm :: "[i, i, i, i]=>i" where + is_wfrec_fm :: "[i, i, i, i]\i" where "is_wfrec_fm(p,r,a,z) \ Exists(And(is_recfun_fm(p, succ(r), succ(a), 0), Exists(Exists(Exists(Exists( @@ -714,7 +714,7 @@ subsubsection\Binary Products, Internalized\ definition - cartprod_fm :: "[i,i,i]=>i" where + 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)))" *) "cartprod_fm(A,B,z) \ @@ -755,7 +755,7 @@ 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 :: "[i,i,i]\i" where "sum_fm(A,B,Z) \ Exists(Exists(Exists(Exists( And(number1_fm(2), @@ -791,7 +791,7 @@ (* "is_quasinat(M,z) \ empty(M,z) | (\m[M]. successor(M,m,z))" *) definition - quasinat_fm :: "i=>i" where + quasinat_fm :: "i\i" where "quasinat_fm(z) \ Or(empty_fm(z), Exists(succ_fm(0,succ(z))))" lemma quasinat_type [TC]: @@ -822,14 +822,14 @@ \<^term>\is_b\ takes two arguments. Instead it must be a formula where 1 and 0 stand for \<^term>\m\ and \<^term>\b\, respectively.\ -(* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" +(* 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)) \ (is_quasinat(M,k) | empty(M,z))" *) text\The formula \<^term>\is_b\ has free variables 1 and 0.\ definition - is_nat_case_fm :: "[i, i, i, i]=>i" where + is_nat_case_fm :: "[i, i, i, i]\i" where "is_nat_case_fm(a,is_b,k,z) \ And(Implies(empty_fm(k), Equal(z,a)), And(Forall(Implies(succ_fm(0,succ(k)), @@ -880,12 +880,12 @@ subsection\The Operator \<^term>\iterates_MH\, Needed for Iteration\ -(* iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" +(* 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), n, z)" *) definition - iterates_MH_fm :: "[i, i, i, i, i]=>i" where + iterates_MH_fm :: "[i, i, i, i, i]\i" where "iterates_MH_fm(isF,v,n,g,z) \ is_nat_case_fm(v, Exists(And(fun_apply_fm(succ(succ(succ(g))),2,0), @@ -951,7 +951,7 @@ 1 0 is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"*) definition - is_iterates_fm :: "[i, i, i, i]=>i" where + is_iterates_fm :: "[i, i, i, i]\i" where "is_iterates_fm(p,v,n,Z) \ Exists(Exists( And(succ_fm(n#+2,1), @@ -1022,7 +1022,7 @@ (* is_eclose_n(M,A,n,Z) \ is_iterates(M, big_union(M), A, n, Z) *) definition - eclose_n_fm :: "[i,i,i]=>i" where + eclose_n_fm :: "[i,i,i]\i" where "eclose_n_fm(A,n,Z) \ is_iterates_fm(big_union_fm(1,0), A, n, Z)" lemma eclose_n_fm_type [TC]: @@ -1059,7 +1059,7 @@ \n[M]. \eclosen[M]. finite_ordinal(M,n) \ is_eclose_n(M,A,n,eclosen) \ l \ eclosen *) definition - mem_eclose_fm :: "[i,i]=>i" where + mem_eclose_fm :: "[i,i]\i" where "mem_eclose_fm(x,y) \ Exists(Exists( And(finite_ordinal_fm(1), @@ -1092,7 +1092,7 @@ (* is_eclose(M,A,Z) \ \l[M]. l \ Z \ mem_eclose(M,A,l) *) definition - is_eclose_fm :: "[i,i]=>i" where + is_eclose_fm :: "[i,i]\i" where "is_eclose_fm(A,Z) \ Forall(Iff(Member(0,succ(Z)), mem_eclose_fm(succ(A),0)))" @@ -1122,7 +1122,7 @@ subsubsection\The List Functor, Internalized\ definition - list_functor_fm :: "[i,i,i]=>i" where + 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)" *) @@ -1163,7 +1163,7 @@ is_iterates(M, is_list_functor(M,A), zero, n, Z)" *) definition - list_N_fm :: "[i,i,i]=>i" where + list_N_fm :: "[i,i,i]\i" where "list_N_fm(A,n,Z) \ Exists( And(empty_fm(0), @@ -1204,7 +1204,7 @@ \n[M]. \listn[M]. finite_ordinal(M,n) \ is_list_N(M,A,n,listn) \ l \ listn *) definition - mem_list_fm :: "[i,i]=>i" where + mem_list_fm :: "[i,i]\i" where "mem_list_fm(x,y) \ Exists(Exists( And(finite_ordinal_fm(1), @@ -1237,7 +1237,7 @@ (* is_list(M,A,Z) \ \l[M]. l \ Z \ mem_list(M,A,l) *) definition - is_list_fm :: "[i,i]=>i" where + is_list_fm :: "[i,i]\i" where "is_list_fm(A,Z) \ Forall(Iff(Member(0,succ(Z)), mem_list_fm(succ(A),0)))" @@ -1266,7 +1266,7 @@ subsubsection\The Formula Functor, Internalized\ -definition formula_functor_fm :: "[i,i]=>i" where +definition formula_functor_fm :: "[i,i]\i" where (* "is_formula_functor(M,X,Z) \ \nat'[M]. \natnat[M]. \natnatsum[M]. \XX[M]. \X3[M]. 4 3 2 1 0 @@ -1313,7 +1313,7 @@ \zero[M]. empty(M,zero) \ is_iterates(M, is_formula_functor(M), zero, n, Z)" *) definition - formula_N_fm :: "[i,i]=>i" where + formula_N_fm :: "[i,i]\i" where "formula_N_fm(n,Z) \ Exists( And(empty_fm(0), @@ -1354,7 +1354,7 @@ \n[M]. \formn[M]. finite_ordinal(M,n) \ is_formula_N(M,n,formn) \ p \ formn *) definition - mem_formula_fm :: "i=>i" where + mem_formula_fm :: "i\i" where "mem_formula_fm(x) \ Exists(Exists( And(finite_ordinal_fm(1), @@ -1387,7 +1387,7 @@ (* is_formula(M,Z) \ \p[M]. p \ Z \ mem_formula(M,p) *) definition - is_formula_fm :: "i=>i" where + is_formula_fm :: "i\i" where "is_formula_fm(Z) \ Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))" lemma is_formula_type [TC]: @@ -1419,14 +1419,14 @@ We call \<^term>\p\ with arguments a, f, z by equating them with the corresponding quantified variables with de Bruijn indices 2, 1, 0.\ -(* is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" +(* is_transrec :: "[i\o, [i,i,i]\o, i, i] \ o" "is_transrec(M,MH,a,z) \ \sa[M]. \esa[M]. \mesa[M]. 2 1 0 upair(M,a,a,sa) \ is_eclose(M,sa,esa) \ membership(M,esa,mesa) \ is_wfrec(M,MH,mesa,a,z)" *) definition - is_transrec_fm :: "[i, i, i]=>i" where + is_transrec_fm :: "[i, i, i]\i" where "is_transrec_fm(p,a,z) \ Exists(Exists(Exists( And(upair_fm(a#+3,a#+3,2), diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/L_axioms.thy Tue Sep 27 17:46:52 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) @@ -126,22 +126,22 @@ text\instances of locale constants\ definition - L_F0 :: "[i=>o,i] => i" where - "L_F0(P,y) \ \ b. (\z. L(z) \ P()) \ (\z\Lset(b). P())" + L_F0 :: "[i\o,i] \ i" where + "L_F0(P,y) \ \ b. (\z. L(z) \ P(\y,z\)) \ (\z\Lset(b). P(\y,z\))" definition - L_FF :: "[i=>o,i] => i" where + L_FF :: "[i\o,i] \ i" where "L_FF(P) \ \a. \y\Lset(a). L_F0(P,y)" definition - L_ClEx :: "[i=>o,i] => o" where + L_ClEx :: "[i\o,i] \ o" where "L_ClEx(P) \ \a. Limit(a) \ normalize(L_FF(P),a) = a" text\We must use the meta-existential quantifier; otherwise the reflection terms become enormous!\ definition - L_Reflects :: "[i=>o,[i,i]=>o] => prop" (\(3REFLECTS/ [_,/ _])\) where + L_Reflects :: "[i\o,[i,i]\o] \ prop" (\(3REFLECTS/ [_,/ _])\) where "REFLECTS[P,Q] \ (\Cl. Closed_Unbounded(Cl) \ (\a. Cl(a) \ (\x \ Lset(a). P(x) \ Q(a,x))))" @@ -303,7 +303,7 @@ subsubsection\The Empty Set, Internalized\ definition - empty_fm :: "i=>i" where + empty_fm :: "i\i" where "empty_fm(x) \ Forall(Neg(Member(0,succ(x))))" lemma empty_type [TC]: @@ -342,7 +342,7 @@ subsubsection\Unordered Pairs, Internalized\ definition - upair_fm :: "[i,i,i]=>i" where + upair_fm :: "[i,i,i]\i" where "upair_fm(x,y,z) \ And(Member(x,z), And(Member(y,z), @@ -385,7 +385,7 @@ subsubsection\Ordered pairs, Internalized\ definition - pair_fm :: "[i,i,i]=>i" where + pair_fm :: "[i,i,i]\i" where "pair_fm(x,y,z) \ Exists(And(upair_fm(succ(x),succ(x),0), Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0), @@ -418,7 +418,7 @@ subsubsection\Binary Unions, Internalized\ definition - union_fm :: "[i,i,i]=>i" where + union_fm :: "[i,i,i]\i" where "union_fm(x,y,z) \ Forall(Iff(Member(0,succ(z)), Or(Member(0,succ(x)),Member(0,succ(y)))))" @@ -450,7 +450,7 @@ subsubsection\Set ``Cons,'' Internalized\ definition - cons_fm :: "[i,i,i]=>i" where + cons_fm :: "[i,i,i]\i" where "cons_fm(x,y,z) \ Exists(And(upair_fm(succ(x),succ(x),0), union_fm(0,succ(y),succ(z))))" @@ -483,7 +483,7 @@ subsubsection\Successor Function, Internalized\ definition - succ_fm :: "[i,i]=>i" where + succ_fm :: "[i,i]\i" where "succ_fm(x,y) \ cons_fm(x,x,y)" lemma succ_type [TC]: @@ -514,7 +514,7 @@ (* "number1(M,a) \ (\x[M]. empty(M,x) \ successor(M,x,a))" *) definition - number1_fm :: "i=>i" where + number1_fm :: "i\i" where "number1_fm(a) \ Exists(And(empty_fm(0), succ_fm(0,succ(a))))" lemma number1_type [TC]: @@ -544,7 +544,7 @@ (* "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 :: "[i,i]\i" where "big_union_fm(A,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), Member(1,0)))))" @@ -625,7 +625,7 @@ subsubsection\Membership Relation, Internalized\ definition - Memrel_fm :: "[i,i]=>i" where + Memrel_fm :: "[i,i]\i" where "Memrel_fm(A,r) \ Forall(Iff(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), @@ -659,7 +659,7 @@ subsubsection\Predecessor Set, Internalized\ definition - pred_set_fm :: "[i,i,i,i]=>i" where + pred_set_fm :: "[i,i,i,i]\i" where "pred_set_fm(A,x,r,B) \ Forall(Iff(Member(0,succ(B)), Exists(And(Member(0,succ(succ(r))), @@ -698,7 +698,7 @@ (* "is_domain(M,r,z) \ \x[M]. (x \ z \ (\w[M]. w\r \ (\y[M]. pair(M,x,y,w))))" *) definition - domain_fm :: "[i,i]=>i" where + domain_fm :: "[i,i]\i" where "domain_fm(r,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), @@ -733,7 +733,7 @@ (* "is_range(M,r,z) \ \y[M]. (y \ z \ (\w[M]. w\r \ (\x[M]. pair(M,x,y,w))))" *) definition - range_fm :: "[i,i]=>i" where + range_fm :: "[i,i]\i" where "range_fm(r,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), @@ -769,7 +769,7 @@ \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 :: "[i,i]\i" where "field_fm(r,z) \ Exists(And(domain_fm(succ(r),0), Exists(And(range_fm(succ(succ(r)),0), @@ -805,7 +805,7 @@ (* "image(M,r,A,z) \ \y[M]. (y \ z \ (\w[M]. w\r \ (\x[M]. x\A \ pair(M,x,y,w))))" *) definition - image_fm :: "[i,i,i]=>i" where + image_fm :: "[i,i,i]\i" where "image_fm(r,A,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), @@ -841,7 +841,7 @@ (* "pre_image(M,r,A,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. y\A \ pair(M,x,y,w)))" *) definition - pre_image_fm :: "[i,i,i]=>i" where + pre_image_fm :: "[i,i,i]\i" where "pre_image_fm(r,A,z) \ Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), @@ -878,7 +878,7 @@ (\xs[M]. \fxs[M]. upair(M,x,x,xs) \ image(M,f,xs,fxs) \ big_union(M,fxs,y))" *) definition - fun_apply_fm :: "[i,i,i]=>i" where + fun_apply_fm :: "[i,i,i]\i" where "fun_apply_fm(f,x,y) \ Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1), And(image_fm(succ(succ(f)), 1, 0), @@ -914,7 +914,7 @@ (* "is_relation(M,r) \ (\z[M]. z\r \ (\x[M]. \y[M]. pair(M,x,y,z)))" *) definition - relation_fm :: "i=>i" where + relation_fm :: "i\i" where "relation_fm(r) \ Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))" @@ -947,7 +947,7 @@ \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x,y',p') \ p\r \ p'\r \ y=y'" *) definition - function_fm :: "i=>i" where + function_fm :: "i\i" where "function_fm(r) \ Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,3,1), @@ -985,7 +985,7 @@ (\u[M]. u\r \ (\x[M]. \y[M]. pair(M,x,y,u) \ y\B))" *) definition - typed_function_fm :: "[i,i,i]=>i" where + typed_function_fm :: "[i,i,i]\i" where "typed_function_fm(A,B,r) \ And(function_fm(r), And(relation_fm(r), @@ -1045,7 +1045,7 @@ 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 :: "[i,i,i]\i" where "composition_fm(r,s,t) \ Forall(Iff(Member(0,succ(t)), Exists(Exists(Exists(Exists(Exists( @@ -1085,7 +1085,7 @@ (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x',y,p') \ p\f \ p'\f \ x=x')" *) definition - injection_fm :: "[i,i,i]=>i" where + injection_fm :: "[i,i,i]\i" where "injection_fm(A,B,f) \ And(typed_function_fm(A,B,f), Forall(Forall(Forall(Forall(Forall( @@ -1121,12 +1121,12 @@ subsubsection\Surjections, Internalized\ -(* surjection :: "[i=>o,i,i,i] => o" +(* surjection :: "[i\o,i,i,i] \ o" "surjection(M,A,B,f) \ typed_function(M,A,B,f) \ (\y[M]. y\B \ (\x[M]. x\A \ fun_apply(M,f,x,y)))" *) definition - surjection_fm :: "[i,i,i]=>i" where + surjection_fm :: "[i,i,i]\i" where "surjection_fm(A,B,f) \ And(typed_function_fm(A,B,f), Forall(Implies(Member(0,succ(B)), @@ -1160,10 +1160,10 @@ subsubsection\Bijections, Internalized\ -(* bijection :: "[i=>o,i,i,i] => o" +(* bijection :: "[i\o,i,i,i] \ o" "bijection(M,A,B,f) \ injection(M,A,B,f) \ surjection(M,A,B,f)" *) definition - bijection_fm :: "[i,i,i]=>i" where + bijection_fm :: "[i,i,i]\i" where "bijection_fm(A,B,f) \ And(injection_fm(A,B,f), surjection_fm(A,B,f))" lemma bijection_type [TC]: @@ -1196,7 +1196,7 @@ (* "restriction(M,r,A,z) \ \x[M]. x \ z \ (x \ r \ (\u[M]. u\A \ (\v[M]. pair(M,u,v,x))))" *) definition - restriction_fm :: "[i,i,i]=>i" where + restriction_fm :: "[i,i,i]\i" where "restriction_fm(r,A,z) \ Forall(Iff(Member(0,succ(z)), And(Member(0,succ(r)), @@ -1228,7 +1228,7 @@ subsubsection\Order-Isomorphisms, Internalized\ -(* order_isomorphism :: "[i=>o,i,i,i,i,i] => o" +(* order_isomorphism :: "[i\o,i,i,i,i,i] \ o" "order_isomorphism(M,A,r,B,s,f) \ bijection(M,A,B,f) \ (\x[M]. x\A \ (\y[M]. y\A \ @@ -1238,7 +1238,7 @@ *) definition - order_isomorphism_fm :: "[i,i,i,i,i]=>i" where + order_isomorphism_fm :: "[i,i,i,i,i]\i" where "order_isomorphism_fm(A,r,B,s,f) \ And(bijection_fm(A,B,f), Forall(Implies(Member(0,succ(A)), @@ -1286,7 +1286,7 @@ (\x[M]. x\a \ (\y[M]. y\a \ successor(M,x,y)))" *) definition - limit_ordinal_fm :: "i=>i" where + limit_ordinal_fm :: "i\i" where "limit_ordinal_fm(x) \ And(ordinal_fm(x), And(Neg(empty_fm(x)), @@ -1323,7 +1323,7 @@ ordinal(M,a) \ \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x))" *) definition - finite_ordinal_fm :: "i=>i" where + finite_ordinal_fm :: "i\i" where "finite_ordinal_fm(x) \ And(ordinal_fm(x), And(Neg(limit_ordinal_fm(x)), @@ -1357,7 +1357,7 @@ (* omega(M,a) \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x)) *) definition - omega_fm :: "i=>i" where + omega_fm :: "i\i" where "omega_fm(x) \ And(limit_ordinal_fm(x), Forall(Implies(Member(0,succ(x)), diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Normal.thy --- a/src/ZF/Constructible/Normal.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Normal.thy Tue Sep 27 17:46:52 2022 +0100 @@ -18,60 +18,60 @@ subsection \Closed and Unbounded (c.u.) Classes of Ordinals\ definition - Closed :: "(i=>o) => o" where + Closed :: "(i\o) \ o" where "Closed(P) \ \I. I \ 0 \ (\i\I. Ord(i) \ P(i)) \ P(\(I))" definition - Unbounded :: "(i=>o) => o" where + Unbounded :: "(i\o) \ o" where "Unbounded(P) \ \i. Ord(i) \ (\j. i P(j))" definition - Closed_Unbounded :: "(i=>o) => o" where + Closed_Unbounded :: "(i\o) \ o" where "Closed_Unbounded(P) \ Closed(P) \ Unbounded(P)" subsubsection\Simple facts about c.u. classes\ lemma ClosedI: - "\\I. \I \ 0; \i\I. Ord(i) \ P(i)\ \ P(\(I))\ + "\\I. \I \ 0; \i\I. Ord(i) \ P(i)\ \ P(\(I))\ \ Closed(P)" -by (simp add: Closed_def) + by (simp add: Closed_def) lemma ClosedD: - "\Closed(P); I \ 0; \i. i\I \ Ord(i); \i. i\I \ P(i)\ + "\Closed(P); I \ 0; \i. i\I \ Ord(i); \i. i\I \ P(i)\ \ P(\(I))" -by (simp add: Closed_def) + by (simp add: Closed_def) lemma UnboundedD: - "\Unbounded(P); Ord(i)\ \ \j. i P(j)" -by (simp add: Unbounded_def) + "\Unbounded(P); Ord(i)\ \ \j. i P(j)" + by (simp add: Unbounded_def) lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) \ Unbounded(C)" -by (simp add: Closed_Unbounded_def) + by (simp add: Closed_Unbounded_def) text\The universal class, V, is closed and unbounded. A bit odd, since C. U. concerns only ordinals, but it's used below!\ theorem Closed_Unbounded_V [simp]: "Closed_Unbounded(\x. True)" -by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast) + by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast) text\The class of ordinals, \<^term>\Ord\, is closed and unbounded.\ theorem Closed_Unbounded_Ord [simp]: "Closed_Unbounded(Ord)" -by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast) + by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast) text\The class of limit ordinals, \<^term>\Limit\, is closed and unbounded.\ theorem Closed_Unbounded_Limit [simp]: "Closed_Unbounded(Limit)" -apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def Limit_Union, - clarify) -apply (rule_tac x="i++nat" in exI) -apply (blast intro: oadd_lt_self oadd_LimitI Limit_has_0) -done + apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def Limit_Union, + clarify) + apply (rule_tac x="i++nat" in exI) + apply (blast intro: oadd_lt_self oadd_LimitI Limit_has_0) + done text\The class of cardinals, \<^term>\Card\, is closed and unbounded.\ theorem Closed_Unbounded_Card [simp]: "Closed_Unbounded(Card)" -apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def) -apply (blast intro: lt_csucc Card_csucc) -done + apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def) + apply (blast intro: lt_csucc Card_csucc) + done subsubsection\The intersection of any set-indexed family of c.u. classes is @@ -87,160 +87,163 @@ and A_non0: "A\0" defines "next_greater(a,x) \ \ y. x P(a,y)" and "sup_greater(x) \ \a\A. next_greater(a,x)" - + +begin text\Trivial that the intersection is closed.\ -lemma (in cub_family) Closed_INT: "Closed(\x. \i\A. P(i,x))" -by (blast intro: ClosedI ClosedD [OF closed]) +lemma Closed_INT: "Closed(\x. \i\A. P(i,x))" + by (blast intro: ClosedI ClosedD [OF closed]) text\All remaining effort goes to show that the intersection is unbounded.\ -lemma (in cub_family) Ord_sup_greater: - "Ord(sup_greater(x))" -by (simp add: sup_greater_def next_greater_def) +lemma Ord_sup_greater: + "Ord(sup_greater(x))" + by (simp add: sup_greater_def next_greater_def) -lemma (in cub_family) Ord_next_greater: - "Ord(next_greater(a,x))" -by (simp add: next_greater_def) +lemma Ord_next_greater: + "Ord(next_greater(a,x))" + by (simp add: next_greater_def) text\\<^term>\next_greater\ works as expected: it returns a larger value and one that belongs to class \<^term>\P(a)\.\ -lemma (in cub_family) next_greater_lemma: - "\Ord(x); a\A\ \ P(a, next_greater(a,x)) \ x < next_greater(a,x)" -apply (simp add: next_greater_def) -apply (rule exE [OF UnboundedD [OF unbounded]]) - apply assumption+ -apply (blast intro: LeastI2 lt_Ord2) -done +lemma next_greater_lemma: + "\Ord(x); a\A\ \ P(a, next_greater(a,x)) \ x < next_greater(a,x)" + apply (simp add: next_greater_def) + apply (rule exE [OF UnboundedD [OF unbounded]]) + apply assumption+ + apply (blast intro: LeastI2 lt_Ord2) + done -lemma (in cub_family) next_greater_in_P: - "\Ord(x); a\A\ \ P(a, next_greater(a,x))" -by (blast dest: next_greater_lemma) +lemma next_greater_in_P: + "\Ord(x); a\A\ \ P(a, next_greater(a,x))" + by (blast dest: next_greater_lemma) -lemma (in cub_family) next_greater_gt: - "\Ord(x); a\A\ \ x < next_greater(a,x)" -by (blast dest: next_greater_lemma) +lemma next_greater_gt: + "\Ord(x); a\A\ \ x < next_greater(a,x)" + by (blast dest: next_greater_lemma) -lemma (in cub_family) sup_greater_gt: - "Ord(x) \ x < sup_greater(x)" -apply (simp add: sup_greater_def) -apply (insert A_non0) -apply (blast intro: UN_upper_lt next_greater_gt Ord_next_greater) -done +lemma sup_greater_gt: + "Ord(x) \ x < sup_greater(x)" + apply (simp add: sup_greater_def) + apply (insert A_non0) + apply (blast intro: UN_upper_lt next_greater_gt Ord_next_greater) + done -lemma (in cub_family) next_greater_le_sup_greater: - "a\A \ next_greater(a,x) \ sup_greater(x)" -apply (simp add: sup_greater_def) -apply (blast intro: UN_upper_le Ord_next_greater) -done +lemma next_greater_le_sup_greater: + "a\A \ next_greater(a,x) \ sup_greater(x)" + apply (simp add: sup_greater_def) + apply (blast intro: UN_upper_le Ord_next_greater) + done -lemma (in cub_family) omega_sup_greater_eq_UN: - "\Ord(x); a\A\ +lemma omega_sup_greater_eq_UN: + "\Ord(x); a\A\ \ sup_greater^\ (x) = (\n\nat. next_greater(a, sup_greater^n (x)))" -apply (simp add: iterates_omega_def) -apply (rule le_anti_sym) -apply (rule le_implies_UN_le_UN) -apply (blast intro: leI next_greater_gt Ord_iterates Ord_sup_greater) -txt\Opposite bound: + apply (simp add: iterates_omega_def) + apply (rule le_anti_sym) + apply (rule le_implies_UN_le_UN) + apply (blast intro: leI next_greater_gt Ord_iterates Ord_sup_greater) + txt\Opposite bound: @{subgoals[display,indent=0,margin=65]} \ -apply (rule UN_least_le) -apply (blast intro: Ord_iterates Ord_sup_greater) -apply (rule_tac a="succ(n)" in UN_upper_le) -apply (simp_all add: next_greater_le_sup_greater) -apply (blast intro: Ord_iterates Ord_sup_greater) -done + apply (rule UN_least_le) + apply (blast intro: Ord_iterates Ord_sup_greater) + apply (rule_tac a="succ(n)" in UN_upper_le) + apply (simp_all add: next_greater_le_sup_greater) + apply (blast intro: Ord_iterates Ord_sup_greater) + done -lemma (in cub_family) P_omega_sup_greater: - "\Ord(x); a\A\ \ P(a, sup_greater^\ (x))" -apply (simp add: omega_sup_greater_eq_UN) -apply (rule ClosedD [OF closed]) -apply (blast intro: ltD, auto) -apply (blast intro: Ord_iterates Ord_next_greater Ord_sup_greater) -apply (blast intro: next_greater_in_P Ord_iterates Ord_sup_greater) -done +lemma P_omega_sup_greater: + "\Ord(x); a\A\ \ P(a, sup_greater^\ (x))" + apply (simp add: omega_sup_greater_eq_UN) + apply (rule ClosedD [OF closed]) + apply (blast intro: ltD, auto) + apply (blast intro: Ord_iterates Ord_next_greater Ord_sup_greater) + apply (blast intro: next_greater_in_P Ord_iterates Ord_sup_greater) + done -lemma (in cub_family) omega_sup_greater_gt: - "Ord(x) \ x < sup_greater^\ (x)" -apply (simp add: iterates_omega_def) -apply (rule UN_upper_lt [of 1], simp_all) - apply (blast intro: sup_greater_gt) -apply (blast intro: Ord_iterates Ord_sup_greater) -done +lemma omega_sup_greater_gt: + "Ord(x) \ x < sup_greater^\ (x)" + apply (simp add: iterates_omega_def) + apply (rule UN_upper_lt [of 1], simp_all) + apply (blast intro: sup_greater_gt) + apply (blast intro: Ord_iterates Ord_sup_greater) + done -lemma (in cub_family) Unbounded_INT: "Unbounded(\x. \a\A. P(a,x))" -apply (unfold Unbounded_def) -apply (blast intro!: omega_sup_greater_gt P_omega_sup_greater) -done +lemma Unbounded_INT: "Unbounded(\x. \a\A. P(a,x))" + apply (unfold Unbounded_def) + apply (blast intro!: omega_sup_greater_gt P_omega_sup_greater) + done -lemma (in cub_family) Closed_Unbounded_INT: - "Closed_Unbounded(\x. \a\A. P(a,x))" -by (simp add: Closed_Unbounded_def Closed_INT Unbounded_INT) +lemma Closed_Unbounded_INT: + "Closed_Unbounded(\x. \a\A. P(a,x))" + by (simp add: Closed_Unbounded_def Closed_INT Unbounded_INT) + +end theorem Closed_Unbounded_INT: - "(\a. a\A \ Closed_Unbounded(P(a))) + "(\a. a\A \ Closed_Unbounded(P(a))) \ Closed_Unbounded(\x. \a\A. P(a, x))" -apply (case_tac "A=0", simp) -apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro]) -apply (simp_all add: Closed_Unbounded_def) -done + apply (case_tac "A=0", simp) + apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro]) + apply (simp_all add: Closed_Unbounded_def) + done lemma Int_iff_INT2: - "P(x) \ Q(x) \ (\i\2. (i=0 \ P(x)) \ (i=1 \ Q(x)))" -by auto + "P(x) \ Q(x) \ (\i\2. (i=0 \ P(x)) \ (i=1 \ Q(x)))" + by auto theorem Closed_Unbounded_Int: - "\Closed_Unbounded(P); Closed_Unbounded(Q)\ + "\Closed_Unbounded(P); Closed_Unbounded(Q)\ \ Closed_Unbounded(\x. P(x) \ Q(x))" -apply (simp only: Int_iff_INT2) -apply (rule Closed_Unbounded_INT, auto) -done + apply (simp only: Int_iff_INT2) + apply (rule Closed_Unbounded_INT, auto) + done subsection \Normal Functions\ definition - mono_le_subset :: "(i=>i) => o" where + mono_le_subset :: "(i\i) \ o" where "mono_le_subset(M) \ \i j. i\j \ M(i) \ M(j)" definition - mono_Ord :: "(i=>i) => o" where + mono_Ord :: "(i\i) \ o" where "mono_Ord(F) \ \i j. i F(i) < F(j)" definition - cont_Ord :: "(i=>i) => o" where + cont_Ord :: "(i\i) \ o" where "cont_Ord(F) \ \l. Limit(l) \ F(l) = (\ii) => o" where + Normal :: "(i\i) \ o" where "Normal(F) \ mono_Ord(F) \ cont_Ord(F)" subsubsection\Immediate properties of the definitions\ lemma NormalI: - "\\i j. i F(i) < F(j); \l. Limit(l) \ F(l) = (\i + "\\i j. i F(i) < F(j); \l. Limit(l) \ F(l) = (\i \ Normal(F)" -by (simp add: Normal_def mono_Ord_def cont_Ord_def) + by (simp add: Normal_def mono_Ord_def cont_Ord_def) lemma mono_Ord_imp_Ord: "\Ord(i); mono_Ord(F)\ \ Ord(F(i))" -apply (auto simp add: mono_Ord_def) -apply (blast intro: lt_Ord) -done + apply (auto simp add: mono_Ord_def) + apply (blast intro: lt_Ord) + done lemma mono_Ord_imp_mono: "\i \ F(i) < F(j)" -by (simp add: mono_Ord_def) + by (simp add: mono_Ord_def) lemma Normal_imp_Ord [simp]: "\Normal(F); Ord(i)\ \ Ord(F(i))" -by (simp add: Normal_def mono_Ord_imp_Ord) + by (simp add: Normal_def mono_Ord_imp_Ord) lemma Normal_imp_cont: "\Normal(F); Limit(l)\ \ F(l) = (\ii \ F(i) < F(j)" -by (simp add: Normal_def mono_Ord_def) + by (simp add: Normal_def mono_Ord_def) lemma Normal_increasing: assumes i: "Ord(i)" and F: "Normal(F)" shows"i \ F(i)" @@ -272,58 +275,58 @@ text\The proof is from Drake, pages 113--114.\ lemma mono_Ord_imp_le_subset: "mono_Ord(F) \ mono_le_subset(F)" -apply (simp add: mono_le_subset_def, clarify) -apply (subgoal_tac "F(i)\F(j)", blast dest: le_imp_subset) -apply (simp add: le_iff) -apply (blast intro: lt_Ord2 mono_Ord_imp_Ord mono_Ord_imp_mono) -done + apply (simp add: mono_le_subset_def, clarify) + apply (subgoal_tac "F(i)\F(j)", blast dest: le_imp_subset) + apply (simp add: le_iff) + apply (blast intro: lt_Ord2 mono_Ord_imp_Ord mono_Ord_imp_mono) + done text\The following equation is taken for granted in any set theory text.\ lemma cont_Ord_Union: - "\cont_Ord(F); mono_le_subset(F); X=0 \ F(0)=0; \x\X. Ord(x)\ + "\cont_Ord(F); mono_le_subset(F); X=0 \ F(0)=0; \x\X. Ord(x)\ \ F(\(X)) = (\y\X. F(y))" -apply (frule Ord_set_cases) -apply (erule disjE, force) -apply (thin_tac "X=0 \ Q" for Q, auto) - txt\The trival case of \<^term>\\X \ X\\ - apply (rule equalityI, blast intro: Ord_Union_eq_succD) - apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff) - apply (blast elim: equalityE) -txt\The limit case, \<^term>\Limit(\X)\: + apply (frule Ord_set_cases) + apply (erule disjE, force) + apply (thin_tac "X=0 \ Q" for Q, auto) + txt\The trival case of \<^term>\\X \ X\\ + apply (rule equalityI, blast intro: Ord_Union_eq_succD) + apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff) + apply (blast elim: equalityE) + txt\The limit case, \<^term>\Limit(\X)\: @{subgoals[display,indent=0,margin=65]} \ -apply (simp add: OUN_Union_eq cont_Ord_def) -apply (rule equalityI) -txt\First inclusion:\ - apply (rule UN_least [OF OUN_least]) - apply (simp add: mono_le_subset_def, blast intro: leI) -txt\Second inclusion:\ -apply (rule UN_least) -apply (frule Union_upper_le, blast, blast) -apply (erule leE, drule ltD, elim UnionE) - apply (simp add: OUnion_def) - apply blast+ -done + apply (simp add: OUN_Union_eq cont_Ord_def) + apply (rule equalityI) + txt\First inclusion:\ + apply (rule UN_least [OF OUN_least]) + apply (simp add: mono_le_subset_def, blast intro: leI) + txt\Second inclusion:\ + apply (rule UN_least) + apply (frule Union_upper_le, blast, blast) + apply (erule leE, drule ltD, elim UnionE) + apply (simp add: OUnion_def) + apply blast+ + done lemma Normal_Union: - "\X\0; \x\X. Ord(x); Normal(F)\ \ F(\(X)) = (\y\X. F(y))" -apply (simp add: Normal_def) -apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) -done + "\X\0; \x\X. Ord(x); Normal(F)\ \ F(\(X)) = (\y\X. F(y))" + apply (simp add: Normal_def) + apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) + done lemma Normal_imp_fp_Closed: "Normal(F) \ Closed(\i. F(i) = i)" -apply (simp add: Closed_def ball_conj_distrib, clarify) -apply (frule Ord_set_cases) -apply (auto simp add: Normal_Union) -done + apply (simp add: Closed_def ball_conj_distrib, clarify) + apply (frule Ord_set_cases) + apply (auto simp add: Normal_Union) + done lemma iterates_Normal_increasing: - "\n\nat; x < F(x); Normal(F)\ + "\n\nat; x < F(x); Normal(F)\ \ F^n (x) < F^(succ(n)) (x)" -apply (induct n rule: nat_induct) -apply (simp_all add: Normal_imp_mono) -done + apply (induct n rule: nat_induct) + apply (simp_all add: Normal_imp_mono) + done lemma Ord_iterates_Normal: "\n\nat; Normal(F); Ord(x)\ \ Ord(F^n (x))" @@ -331,20 +334,20 @@ text\THIS RESULT IS UNUSED\ lemma iterates_omega_Limit: - "\Normal(F); x < F(x)\ \ Limit(F^\ (x))" -apply (frule lt_Ord) -apply (simp add: iterates_omega_def) -apply (rule increasing_LimitI) - \ \this lemma is @{thm increasing_LimitI [no_vars]}\ - apply (blast intro: UN_upper_lt [of "1"] Normal_imp_Ord - Ord_iterates lt_imp_0_lt - iterates_Normal_increasing, clarify) -apply (rule bexI) - apply (blast intro: Ord_in_Ord [OF Ord_iterates_Normal]) -apply (rule UN_I, erule nat_succI) -apply (blast intro: iterates_Normal_increasing Ord_iterates_Normal - ltD [OF lt_trans1, OF succ_leI, OF ltI]) -done + "\Normal(F); x < F(x)\ \ Limit(F^\ (x))" + apply (frule lt_Ord) + apply (simp add: iterates_omega_def) + apply (rule increasing_LimitI) + \ \this lemma is @{thm increasing_LimitI [no_vars]}\ + apply (blast intro: UN_upper_lt [of "1"] Normal_imp_Ord + Ord_iterates lt_imp_0_lt + iterates_Normal_increasing, clarify) + apply (rule bexI) + apply (blast intro: Ord_in_Ord [OF Ord_iterates_Normal]) + apply (rule UN_I, erule nat_succI) + apply (blast intro: iterates_Normal_increasing Ord_iterates_Normal + ltD [OF lt_trans1, OF succ_leI, OF ltI]) + done lemma iterates_omega_fixedpoint: "\Normal(F); Ord(a)\ \ F(F^\ (a)) = F^\ (a)" @@ -390,7 +393,7 @@ normal function, by @{thm [source] Normal_imp_fp_Unbounded}. \ definition - normalize :: "[i=>i, i] => i" where + normalize :: "[i\i, i] \ i" where "normalize(F,a) \ transrec2(a, F(0), \x r. F(succ(x)) \ succ(r))" @@ -428,11 +431,11 @@ qed theorem Normal_normalize: - "(\x. Ord(x) \ Ord(F(x))) \ Normal(normalize(F))" -apply (rule NormalI) -apply (blast intro!: normalize_increasing) -apply (simp add: def_transrec2 [OF normalize_def]) -done + "(\x. Ord(x) \ Ord(F(x))) \ Normal(normalize(F))" + apply (rule NormalI) + apply (blast intro!: normalize_increasing) + apply (simp add: def_transrec2 [OF normalize_def]) + done theorem le_normalize: assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "\x. Ord(x) \ Ord(F(x))" @@ -456,15 +459,15 @@ numbers.\ definition - Aleph :: "i => i" (\\_\ [90] 90) where + Aleph :: "i \ i" (\\_\ [90] 90) where "Aleph(a) \ transrec2(a, nat, \x r. csucc(r))" lemma Card_Aleph [simp, intro]: - "Ord(a) \ Card(Aleph(a))" -apply (erule trans_induct3) -apply (simp_all add: Card_csucc Card_nat Card_is_Ord - def_transrec2 [OF Aleph_def]) -done + "Ord(a) \ Card(Aleph(a))" + apply (erule trans_induct3) + apply (simp_all add: Card_csucc Card_nat Card_is_Ord + def_transrec2 [OF Aleph_def]) + done lemma Aleph_increasing: assumes ab: "a < b" shows "Aleph(a) < Aleph(b)" @@ -492,9 +495,9 @@ qed theorem Normal_Aleph: "Normal(Aleph)" -apply (rule NormalI) -apply (blast intro!: Aleph_increasing) -apply (simp add: def_transrec2 [OF Aleph_def]) -done + apply (rule NormalI) + apply (blast intro!: Aleph_increasing) + apply (simp add: def_transrec2 [OF Aleph_def]) + done end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Rank.thy --- a/src/ZF/Constructible/Rank.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Rank.thy Tue Sep 27 17:46:52 2022 +0100 @@ -65,7 +65,7 @@ lemma (in M_ordertype) wellordered_iso_pred_eq_lemma: "\f \ \Order.pred(A,y,r), r\ \ \Order.pred(A,x,r), r\; - wellordered(M,A,r); x\A; y\A; M(A); M(f); M(r)\ \ \ r" + wellordered(M,A,r); x\A; y\A; M(A); M(f); M(r)\ \ \x,y\ \ r" apply (frule wellordered_is_trans_on, assumption) apply (rule notI) apply (drule_tac x2=y and x=x and r2=r in @@ -136,21 +136,21 @@ definition - obase :: "[i=>o,i,i] => i" where + obase :: "[i\o,i,i] \ i" where \ \the domain of \om\, eventually shown to equal \A\\ "obase(M,A,r) \ {a\A. \x[M]. \g[M]. Ord(x) \ g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x))}" definition - omap :: "[i=>o,i,i,i] => o" where + omap :: "[i\o,i,i,i] \ o" where \ \the function that maps wosets to order types\ "omap(M,A,r,f) \ \z[M]. - z \ f \ (\a\A. \x[M]. \g[M]. z = \ Ord(x) \ + z \ f \ (\a\A. \x[M]. \g[M]. z = \a,x\ \ Ord(x) \ g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))" definition - otype :: "[i=>o,i,i,i] => o" where \ \the order types themselves\ + otype :: "[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)" @@ -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 = \a,x\ \ Ord(x) \ g \ ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))" apply (simp add: omap_def) apply (rule iffI) @@ -416,16 +416,16 @@ replacement below. Here j is used to define the relation, namely Memrel(succ(j)), while x determines the domain of f.*) definition - is_oadd_fun :: "[i=>o,i,i,i,i] => o" where + is_oadd_fun :: "[i\o,i,i,i,i] \ o" where "is_oadd_fun(M,i,j,x,f) \ (\sj msj. M(sj) \ M(msj) \ successor(M,j,sj) \ membership(M,sj,msj) \ M_is_recfun(M, - %x g y. \gx[M]. image(M,g,x,gx) \ union(M,i,gx,y), + \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 :: "[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) | @@ -437,7 +437,7 @@ definition (*NEEDS RELATIVIZATION*) - omult_eqns :: "[i,i,i,i] => o" where + omult_eqns :: "[i,i,i,i] \ o" where "omult_eqns(i,x,g,z) \ Ord(x) \ (x=0 \ z=0) \ @@ -445,14 +445,14 @@ (Limit(x) \ z = \(g``x))" definition - is_omult_fun :: "[i=>o,i,i,i] => o" where + is_omult_fun :: "[i\o,i,i,i] \ o" where "is_omult_fun(M,i,j,f) \ (\df. M(df) \ is_function(M,f) \ is_domain(M,f,df) \ subset(M, j, df)) \ (\x\j. omult_eqns(i,x,f,f`x))" definition - is_omult :: "[i=>o,i,i,i] => o" where + is_omult :: "[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) \ @@ -470,8 +470,8 @@ 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 = \x,y\ \ + (\g[M]. is_recfun(Memrel(succ(j)),x,\x g. THE z. omult_eqns(i,x,g,z),g) \ y = (THE z. omult_eqns(i, x, g, z))))" @@ -483,7 +483,7 @@ 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"] + relation2_def is_recfun_abs [of "\x g. i \ g``x"] is_recfun_iff_equation Ball_def lt_trans [OF ltI, of _ a] lt_Memrel) apply (simp add: lt_def) @@ -494,18 +494,18 @@ 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 = \x,y\ \ + (\g[M]. is_recfun(Memrel(succ(j)),x,\x g. i \ g``x,g) \ y = i \ g``x))" apply (insert oadd_strong_replacement [of i j]) apply (simp add: is_oadd_fun_def relation2_def - is_recfun_abs [of "%x g. i \ g``x"]) + is_recfun_abs [of "\x g. i \ g``x"]) done lemma (in M_ord_arith) exists_oadd: "\Ord(j); M(i); M(j)\ - \ \f[M]. is_recfun(Memrel(succ(j)), j, %x g. i \ g``x, f)" + \ \f[M]. is_recfun(Memrel(succ(j)), j, \x g. i \ g``x, f)" apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel]) apply (simp_all add: Memrel_type oadd_strong_replacement') done @@ -609,7 +609,7 @@ lemma (in M_ord_arith) exists_omult: "\Ord(j); M(i); M(j)\ - \ \f[M]. is_recfun(Memrel(succ(j)), j, %x g. THE z. omult_eqns(i,x,g,z), f)" + \ \f[M]. is_recfun(Memrel(succ(j)), j, \x g. THE z. omult_eqns(i,x,g,z), f)" apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel]) apply (simp_all add: Memrel_type omult_strong_replacement') apply (blast intro: the_omult_eqns_closed) @@ -678,13 +678,13 @@ "M(r) \ separation (M, \x. \rplus[M]. tran_closure(M,r,rplus) \ - \ (\f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))" + \ (\f[M]. M_is_recfun(M, \x f y. is_range(M,f,y), rplus, x, f)))" and wfrank_strong_replacement: "M(r) \ 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) \ + 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) \ @@ -702,18 +702,18 @@ lemma (in M_wfrank) wfrank_separation': "M(r) \ separation - (M, \x. \ (\f[M]. is_recfun(r^+, x, %x f. range(f), f)))" + (M, \x. \ (\f[M]. is_recfun(r^+, x, \x f. range(f), f)))" apply (insert wfrank_separation [of r]) -apply (simp add: relation2_def is_recfun_abs [of "%x. range"]) +apply (simp add: relation2_def is_recfun_abs [of "\x. range"]) done lemma (in M_wfrank) wfrank_strong_replacement': "M(r) \ strong_replacement(M, \x z. \y[M]. \f[M]. - pair(M,x,y,z) \ is_recfun(r^+, x, %x f. range(f), f) \ + 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"]) +apply (simp add: relation2_def is_recfun_abs [of "\x. range"]) done lemma (in M_wfrank) Ord_wfrank_separation': @@ -721,21 +721,21 @@ separation (M, \x. \ (\f[M]. is_recfun(r^+, x, \x. range, f) \ Ord(range(f))))" apply (insert Ord_wfrank_separation [of r]) -apply (simp add: relation2_def is_recfun_abs [of "%x. range"]) +apply (simp add: relation2_def is_recfun_abs [of "\x. range"]) done text\This function, defined using replacement, is a rank function for well-founded relations within the class M.\ definition - wellfoundedrank :: "[i=>o,i,i] => i" where + wellfoundedrank :: "[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 = \x,y\ \ is_recfun(r^+, x, \x f. range(f), f) \ y = range(f)}" lemma (in M_wfrank) exists_wfrank: "\wellfounded(M,r); M(a); M(r)\ - \ \f[M]. is_recfun(r^+, a, %x f. range(f), f)" + \ \f[M]. is_recfun(r^+, a, \x f. range(f), f)" apply (rule wellfounded_exists_is_recfun) apply (blast intro: wellfounded_trancl) apply (rule trans_trancl) @@ -759,7 +759,7 @@ lemma (in M_wfrank) Ord_wfrank_range [rule_format]: "\wellfounded(M,r); a\A; M(r); M(A)\ - \ \f[M]. is_recfun(r^+, a, %x f. range(f), f) \ Ord(range(f))" + \ \f[M]. is_recfun(r^+, a, \x f. range(f), f) \ Ord(range(f))" apply (drule wellfounded_trancl, assumption) apply (rule wellfounded_induct, assumption, erule (1) transM) apply simp @@ -863,7 +863,7 @@ Ord_in_Ord [OF Ord_range_wellfoundedrank]) lemma (in M_wfrank) wellfoundedrank_eq: - "\is_recfun(r^+, a, %x. range, f); + "\is_recfun(r^+, a, \x. range, f); wellfounded(M,r); a \ A; M(f); M(r); M(A)\ \ wellfoundedrank(M,r,A) ` a = range(f)" apply (rule apply_equality) @@ -882,7 +882,7 @@ lemma (in M_wfrank) wellfoundedrank_lt: - "\ \ r; + "\\a,b\ \ r; wellfounded(M,r); r \ A*A; M(r); M(A)\ \ wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b" apply (frule wellfounded_trancl, assumption) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Rank_Separation.thy --- a/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Rank_Separation.thy Tue Sep 27 17:46:52 2022 +0100 @@ -133,17 +133,17 @@ lemma wfrank_Reflects: "REFLECTS[\x. \rplus[L]. tran_closure(L,r,rplus) \ - \ (\f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)), + \ (\f[L]. M_is_recfun(L, \x f y. is_range(L,f,y), rplus, x, f)), \i x. \rplus \ Lset(i). tran_closure(##Lset(i),r,rplus) \ \ (\f \ Lset(i). - M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), + M_is_recfun(##Lset(i), \x f y. is_range(##Lset(i),f,y), rplus, x, f))]" by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection) lemma wfrank_separation: "L(r) \ separation (L, \x. \rplus[L]. tran_closure(L,r,rplus) \ - \ (\f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))" + \ (\f[L]. M_is_recfun(L, \x f y. is_range(L,f,y), rplus, x, f)))" apply (rule gen_separation [OF wfrank_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ @@ -156,12 +156,12 @@ "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) \ + 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 \ (\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) \ + 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) @@ -171,7 +171,7 @@ 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) \ + 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 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Rec_Separation.thy Tue Sep 27 17:46:52 2022 +0100 @@ -29,7 +29,7 @@ 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 :: "[i,i,i]\i" where "rtran_closure_mem_fm(A,r,p) \ Exists(Exists(Exists( And(omega_fm(2), @@ -87,7 +87,7 @@ \A[M]. is_field(M,r,A) \ (\p[M]. p \ s \ rtran_closure_mem(M,A,r,p))" *) definition - rtran_closure_fm :: "[i,i]=>i" where + rtran_closure_fm :: "[i,i]\i" where "rtran_closure_fm(r,s) \ Forall(Implies(field_fm(succ(r),0), Forall(Iff(Member(0,succ(succ(s))), @@ -122,7 +122,7 @@ (* "tran_closure(M,r,t) \ \s[M]. rtran_closure(M,r,s) \ composition(M,r,s,t)" *) definition - tran_closure_fm :: "[i,i]=>i" where + tran_closure_fm :: "[i,i]\i" where "tran_closure_fm(r,s) \ Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))" @@ -295,7 +295,7 @@ (* "is_nth(M,n,l,Z) \ \X[M]. is_iterates(M, is_tl(M), l, n, X) \ is_hd(M,X,Z)" *) definition - nth_fm :: "[i,i,i]=>i" where + nth_fm :: "[i,i,i]\i" where "nth_fm(n,l,Z) \ Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0), hd_fm(0,succ(Z))))" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Reflection.thy --- a/src/ZF/Constructible/Reflection.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Reflection.thy Tue Sep 27 17:46:52 2022 +0100 @@ -24,7 +24,7 @@ closed under ordered pairing provided \l\ is limit. Possibly this could be avoided: the induction hypothesis \<^term>\Cl_reflects\ (in locale \ex_reflection\) could be weakened to -\<^term>\\y\Mset(a). \z\Mset(a). P() \ Q(a,)\, removing most +\<^term>\\y\Mset(a). \z\Mset(a). P(\y,z\) \ Q(a,\y,z\)\, removing most uses of \<^term>\Pair_in_Mset\. But there isn't much point in doing so, since ultimately the \ex_reflection\ proof is packaged up using the predicate \Reflects\. @@ -34,15 +34,15 @@ assumes Mset_mono_le : "mono_le_subset(Mset)" and Mset_cont : "cont_Ord(Mset)" and Pair_in_Mset : "\x \ Mset(a); y \ Mset(a); Limit(a)\ - \ \ Mset(a)" + \ \x,y\ \ Mset(a)" defines "M(x) \ \a. Ord(a) \ x \ Mset(a)" and "Reflects(Cl,P,Q) \ Closed_Unbounded(Cl) \ (\a. Cl(a) \ (\x\Mset(a). P(x) \ Q(a,x)))" fixes F0 \ \ordinal for a specific value \<^term>\y\\ fixes FF \ \sup over the whole level, \<^term>\y\Mset(a)\\ fixes ClEx \ \Reflecting ordinals for the formula \<^term>\\z. P\\ - defines "F0(P,y) \ \ b. (\z. M(z) \ P()) \ - (\z\Mset(b). P())" + defines "F0(P,y) \ \ b. (\z. M(z) \ P(\y,z\)) \ + (\z\Mset(b). P(\y,z\))" and "FF(P) \ \a. \y\Mset(a). F0(P,y)" and "ClEx(P,a) \ Limit(a) \ normalize(FF(P),a) = a" @@ -97,7 +97,7 @@ subsection\Reflection for Existential Quantifiers\ lemma F0_works: - "\y\Mset(a); Ord(a); M(z); P()\ \ \z\Mset(F0(P,y)). P()" + "\y\Mset(a); Ord(a); M(z); P(\y,z\)\ \ \z\Mset(F0(P,y)). P(\y,z\)" unfolding F0_def M_def apply clarify apply (rule LeastI2) @@ -117,7 +117,7 @@ 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()" + "\M(z); y\Mset(a); P(\y,z\); Ord(a)\ \ \z\Mset(FF(P,a)). P(\y,z\)" apply (simp add: FF_def) apply (simp_all add: cont_Ord_Union [of concl: Mset] Mset_cont Mset_mono_le not_emptyI) @@ -125,8 +125,8 @@ done lemma FFN_works: - "\M(z); y\Mset(a); P(); Ord(a)\ - \ \z\Mset(normalize(FF(P),a)). P()" + "\M(z); y\Mset(a); P(\y,z\); Ord(a)\ + \ \z\Mset(normalize(FF(P),a)). P(\y,z\)" apply (drule FF_works [of concl: P], assumption+) apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD]) done @@ -144,8 +144,8 @@ begin lemma ClEx_downward: - "\M(z); y\Mset(a); P(); Cl(a); ClEx(P,a)\ - \ \z\Mset(a). Q(a,)" + "\M(z); y\Mset(a); P(\y,z\); Cl(a); ClEx(P,a)\ + \ \z\Mset(a). Q(a,\y,z\)" apply (simp add: ClEx_def, clarify) apply (frule Limit_is_Ord) apply (frule FFN_works [of concl: P], assumption+) @@ -154,8 +154,8 @@ done lemma ClEx_upward: - "\z\Mset(a); y\Mset(a); Q(a,); Cl(a); ClEx(P,a)\ - \ \z. M(z) \ P()" + "\z\Mset(a); y\Mset(a); Q(a,\y,z\); Cl(a); ClEx(P,a)\ + \ \z. M(z) \ P(\y,z\)" apply (simp add: ClEx_def M_def) apply (blast dest: Cl_reflects intro: Limit_is_Ord Pair_in_Mset) @@ -164,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(\y,z\)) \ (\z\Mset(a). Q(a,\y,z\))" by (blast intro: dest: ClEx_downward ClEx_upward) text\...and it is closed and unbounded\ @@ -186,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(\y,z\)) \ (\z\Mset(a). Q(a,\y,z\))" 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 +215,8 @@ lemma Ex_reflection_0: "Reflects(Cl,P0,Q0) \ Reflects(\a. Cl(a) \ ClEx(P0,a), - \x. \z. M(z) \ P0(), - \a x. \z\Mset(a). Q0(a,))" + \x. \z. M(z) \ P0(\x,z\), + \a x. \z\Mset(a). Q0(a,\x,z\))" apply (simp add: Reflects_def) apply (intro conjI Closed_Unbounded_Int) apply blast @@ -227,8 +227,8 @@ lemma All_reflection_0: "Reflects(Cl,P0,Q0) \ Reflects(\a. Cl(a) \ ClEx(\x.\P0(x), a), - \x. \z. M(z) \ P0(), - \a x. \z\Mset(a). Q0(a,))" + \x. \z. M(z) \ P0(\x,z\), + \a x. \z\Mset(a). Q0(a,\x,z\))" apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not) apply (rule Not_reflection, drule Not_reflection, simp) apply (erule Ex_reflection_0) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Relative.thy --- a/src/ZF/Constructible/Relative.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Relative.thy Tue Sep 27 17:46:52 2022 +0100 @@ -11,120 +11,120 @@ subsection\Relativized versions of standard set-theoretic concepts\ definition - empty :: "[i=>o,i] => o" where + empty :: "[i\o,i] \ o" where "empty(M,z) \ \x[M]. x \ z" definition - subset :: "[i=>o,i,i] => o" where + subset :: "[i\o,i,i] \ o" where "subset(M,A,B) \ \x[M]. x\A \ x \ B" definition - upair :: "[i=>o,i,i,i] => o" where + upair :: "[i\o,i,i,i] \ o" where "upair(M,a,b,z) \ a \ z \ b \ z \ (\x[M]. x\z \ x = a \ x = b)" definition - pair :: "[i=>o,i,i,i] => o" where + pair :: "[i\o,i,i,i] \ o" where "pair(M,a,b,z) \ \x[M]. upair(M,a,a,x) \ (\y[M]. upair(M,a,b,y) \ upair(M,x,y,z))" definition - union :: "[i=>o,i,i,i] => o" where + union :: "[i\o,i,i,i] \ o" where "union(M,a,b,z) \ \x[M]. x \ z \ x \ a \ x \ b" definition - is_cons :: "[i=>o,i,i,i] => o" where + is_cons :: "[i\o,i,i,i] \ o" where "is_cons(M,a,b,z) \ \x[M]. upair(M,a,a,x) \ union(M,x,b,z)" definition - successor :: "[i=>o,i,i] => o" where + successor :: "[i\o,i,i] \ o" where "successor(M,a,z) \ is_cons(M,a,a,z)" definition - number1 :: "[i=>o,i] => o" where + number1 :: "[i\o,i] \ o" where "number1(M,a) \ \x[M]. empty(M,x) \ successor(M,x,a)" definition - number2 :: "[i=>o,i] => o" where + number2 :: "[i\o,i] \ o" where "number2(M,a) \ \x[M]. number1(M,x) \ successor(M,x,a)" definition - number3 :: "[i=>o,i] => o" where + number3 :: "[i\o,i] \ o" where "number3(M,a) \ \x[M]. number2(M,x) \ successor(M,x,a)" definition - powerset :: "[i=>o,i,i] => o" where + powerset :: "[i\o,i,i] \ o" where "powerset(M,A,z) \ \x[M]. x \ z \ subset(M,x,A)" definition - is_Collect :: "[i=>o,i,i=>o,i] => o" where + is_Collect :: "[i\o,i,i\o,i] \ o" where "is_Collect(M,A,P,z) \ \x[M]. x \ z \ x \ A \ P(x)" definition - is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" where + is_Replace :: "[i\o,i,[i,i]\o,i] \ o" where "is_Replace(M,A,P,z) \ \u[M]. u \ z \ (\x[M]. x\A \ P(x,u))" definition - inter :: "[i=>o,i,i,i] => o" where + inter :: "[i\o,i,i,i] \ o" where "inter(M,a,b,z) \ \x[M]. x \ z \ x \ a \ x \ b" definition - setdiff :: "[i=>o,i,i,i] => o" where + setdiff :: "[i\o,i,i,i] \ o" where "setdiff(M,a,b,z) \ \x[M]. x \ z \ x \ a \ x \ b" definition - big_union :: "[i=>o,i,i] => o" where + big_union :: "[i\o,i,i] \ o" where "big_union(M,A,z) \ \x[M]. x \ z \ (\y[M]. y\A \ x \ y)" definition - big_inter :: "[i=>o,i,i] => o" where + big_inter :: "[i\o,i,i] \ o" where "big_inter(M,A,z) \ (A=0 \ z=0) \ (A\0 \ (\x[M]. x \ z \ (\y[M]. y\A \ x \ y)))" definition - cartprod :: "[i=>o,i,i,i] => o" where + cartprod :: "[i\o,i,i,i] \ o" where "cartprod(M,A,B,z) \ \u[M]. u \ z \ (\x[M]. x\A \ (\y[M]. y\B \ pair(M,x,y,u)))" definition - is_sum :: "[i=>o,i,i,i] => o" where + is_sum :: "[i\o,i,i,i] \ o" where "is_sum(M,A,B,Z) \ \A0[M]. \n1[M]. \s1[M]. \B1[M]. number1(M,n1) \ cartprod(M,n1,A,A0) \ upair(M,n1,n1,s1) \ cartprod(M,s1,B,B1) \ union(M,A0,B1,Z)" definition - is_Inl :: "[i=>o,i,i] => o" where + is_Inl :: "[i\o,i,i] \ o" where "is_Inl(M,a,z) \ \zero[M]. empty(M,zero) \ pair(M,zero,a,z)" definition - is_Inr :: "[i=>o,i,i] => o" where + is_Inr :: "[i\o,i,i] \ o" where "is_Inr(M,a,z) \ \n1[M]. number1(M,n1) \ pair(M,n1,a,z)" definition - is_converse :: "[i=>o,i,i] => o" where + is_converse :: "[i\o,i,i] \ o" where "is_converse(M,r,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\u[M]. \v[M]. pair(M,u,v,w) \ pair(M,v,u,x)))" definition - pre_image :: "[i=>o,i,i,i] => o" where + pre_image :: "[i\o,i,i,i] \ o" where "pre_image(M,r,A,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. y\A \ pair(M,x,y,w)))" definition - is_domain :: "[i=>o,i,i] => o" where + is_domain :: "[i\o,i,i] \ o" where "is_domain(M,r,z) \ \x[M]. x \ z \ (\w[M]. w\r \ (\y[M]. pair(M,x,y,w)))" definition - image :: "[i=>o,i,i,i] => o" where + image :: "[i\o,i,i,i] \ o" where "image(M,r,A,z) \ \y[M]. y \ z \ (\w[M]. w\r \ (\x[M]. x\A \ pair(M,x,y,w)))" definition - is_range :: "[i=>o,i,i] => o" where + is_range :: "[i\o,i,i] \ o" where \ \the cleaner \<^term>\\r'[M]. is_converse(M,r,r') \ is_domain(M,r',z)\ unfortunately needs an instance of separation in order to prove @@ -133,41 +133,41 @@ \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 :: "[i\o,i,i] \ o" where "is_field(M,r,z) \ \dr[M]. \rr[M]. is_domain(M,r,dr) \ is_range(M,r,rr) \ union(M,dr,rr,z)" definition - is_relation :: "[i=>o,i] => o" where + is_relation :: "[i\o,i] \ o" where "is_relation(M,r) \ (\z[M]. z\r \ (\x[M]. \y[M]. pair(M,x,y,z)))" definition - is_function :: "[i=>o,i] => o" where + is_function :: "[i\o,i] \ o" where "is_function(M,r) \ \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x,y',p') \ p\r \ p'\r \ y=y'" definition - fun_apply :: "[i=>o,i,i,i] => o" where + fun_apply :: "[i\o,i,i,i] \ o" where "fun_apply(M,f,x,y) \ (\xs[M]. \fxs[M]. upair(M,x,x,xs) \ image(M,f,xs,fxs) \ big_union(M,fxs,y))" definition - typed_function :: "[i=>o,i,i,i] => o" where + typed_function :: "[i\o,i,i,i] \ o" where "typed_function(M,A,B,r) \ is_function(M,r) \ is_relation(M,r) \ is_domain(M,r,A) \ (\u[M]. u\r \ (\x[M]. \y[M]. pair(M,x,y,u) \ y\B))" definition - is_funspace :: "[i=>o,i,i,i] => o" where + is_funspace :: "[i\o,i,i,i] \ o" where "is_funspace(M,A,B,F) \ \f[M]. f \ F \ typed_function(M,A,B,f)" definition - composition :: "[i=>o,i,i,i] => o" where + composition :: "[i\o,i,i,i] \ o" where "composition(M,r,s,t) \ \p[M]. p \ t \ (\x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. @@ -175,104 +175,104 @@ xy \ s \ yz \ r)" definition - injection :: "[i=>o,i,i,i] => o" where + injection :: "[i\o,i,i,i] \ o" where "injection(M,A,B,f) \ typed_function(M,A,B,f) \ (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) \ pair(M,x',y,p') \ p\f \ p'\f \ x=x')" definition - surjection :: "[i=>o,i,i,i] => o" where + surjection :: "[i\o,i,i,i] \ o" where "surjection(M,A,B,f) \ typed_function(M,A,B,f) \ (\y[M]. y\B \ (\x[M]. x\A \ fun_apply(M,f,x,y)))" definition - bijection :: "[i=>o,i,i,i] => o" where + bijection :: "[i\o,i,i,i] \ o" where "bijection(M,A,B,f) \ injection(M,A,B,f) \ surjection(M,A,B,f)" definition - restriction :: "[i=>o,i,i,i] => o" where + restriction :: "[i\o,i,i,i] \ o" where "restriction(M,r,A,z) \ \x[M]. x \ z \ (x \ r \ (\u[M]. u\A \ (\v[M]. pair(M,u,v,x))))" definition - transitive_set :: "[i=>o,i] => o" where + transitive_set :: "[i\o,i] \ o" where "transitive_set(M,a) \ \x[M]. x\a \ subset(M,x,a)" definition - ordinal :: "[i=>o,i] => o" where + ordinal :: "[i\o,i] \ o" where \ \an ordinal is a transitive set of transitive sets\ "ordinal(M,a) \ transitive_set(M,a) \ (\x[M]. x\a \ transitive_set(M,x))" definition - limit_ordinal :: "[i=>o,i] => o" where + limit_ordinal :: "[i\o,i] \ o" where \ \a limit ordinal is a non-empty, successor-closed ordinal\ "limit_ordinal(M,a) \ ordinal(M,a) \ \ empty(M,a) \ (\x[M]. x\a \ (\y[M]. y\a \ successor(M,x,y)))" definition - successor_ordinal :: "[i=>o,i] => o" where + successor_ordinal :: "[i\o,i] \ o" where \ \a successor ordinal is any ordinal that is neither empty nor limit\ "successor_ordinal(M,a) \ ordinal(M,a) \ \ empty(M,a) \ \ limit_ordinal(M,a)" definition - finite_ordinal :: "[i=>o,i] => o" where + finite_ordinal :: "[i\o,i] \ o" where \ \an ordinal is finite if neither it nor any of its elements are limit\ "finite_ordinal(M,a) \ ordinal(M,a) \ \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x))" definition - omega :: "[i=>o,i] => o" where + omega :: "[i\o,i] \ o" where \ \omega is a limit ordinal none of whose elements are limit\ "omega(M,a) \ limit_ordinal(M,a) \ (\x[M]. x\a \ \ limit_ordinal(M,x))" definition - is_quasinat :: "[i=>o,i] => o" where + is_quasinat :: "[i\o,i] \ o" where "is_quasinat(M,z) \ empty(M,z) \ (\m[M]. successor(M,m,z))" definition - is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" where + is_nat_case :: "[i\o, i, [i,i]\o, i, i] \ o" where "is_nat_case(M, a, is_b, k, z) \ (empty(M,k) \ z=a) \ (\m[M]. successor(M,m,k) \ is_b(m,z)) \ (is_quasinat(M,k) \ empty(M,z))" definition - relation1 :: "[i=>o, [i,i]=>o, i=>i] => o" where + relation1 :: "[i\o, [i,i]\o, i\i] \ o" where "relation1(M,is_f,f) \ \x[M]. \y[M]. is_f(x,y) \ y = f(x)" definition - Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o" where + Relation1 :: "[i\o, i, [i,i]\o, i\i] \ o" where \ \as above, but typed\ "Relation1(M,A,is_f,f) \ \x[M]. \y[M]. x\A \ is_f(x,y) \ y = f(x)" definition - relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o" where + relation2 :: "[i\o, [i,i,i]\o, [i,i]\i] \ o" where "relation2(M,is_f,f) \ \x[M]. \y[M]. \z[M]. is_f(x,y,z) \ z = f(x,y)" definition - Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o" where + Relation2 :: "[i\o, i, i, [i,i,i]\o, [i,i]\i] \ o" where "Relation2(M,A,B,is_f,f) \ \x[M]. \y[M]. \z[M]. x\A \ y\B \ is_f(x,y,z) \ z = f(x,y)" definition - relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o" where + relation3 :: "[i\o, [i,i,i,i]\o, [i,i,i]\i] \ o" where "relation3(M,is_f,f) \ \x[M]. \y[M]. \z[M]. \u[M]. is_f(x,y,z,u) \ u = f(x,y,z)" definition - Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o" where + Relation3 :: "[i\o, i, i, i, [i,i,i,i]\o, [i,i,i]\i] \ o" where "Relation3(M,A,B,C,is_f,f) \ \x[M]. \y[M]. \z[M]. \u[M]. x\A \ y\B \ z\C \ is_f(x,y,z,u) \ u = f(x,y,z)" definition - relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o" where + relation4 :: "[i\o, [i,i,i,i,i]\o, [i,i,i,i]\i] \ o" where "relation4(M,is_f,f) \ \u[M]. \x[M]. \y[M]. \z[M]. \a[M]. is_f(u,x,y,z,a) \ a = f(u,x,y,z)" @@ -290,12 +290,12 @@ subsection \The relativized ZF axioms\ definition - extensionality :: "(i=>o) => o" where + extensionality :: "(i\o) \ o" where "extensionality(M) \ \x[M]. \y[M]. (\z[M]. z \ x \ z \ y) \ x=y" definition - separation :: "[i=>o, i=>o] => o" where + separation :: "[i\o, i\o] \ o" where \ \The formula \P\ should only involve parameters belonging to \M\ and all its quantifiers must be relativized to \M\. We do not have separation as a scheme; every instance @@ -304,36 +304,36 @@ \z[M]. \y[M]. \x[M]. x \ y \ x \ z \ P(x)" definition - upair_ax :: "(i=>o) => o" where + upair_ax :: "(i\o) \ o" where "upair_ax(M) \ \x[M]. \y[M]. \z[M]. upair(M,x,y,z)" definition - Union_ax :: "(i=>o) => o" where + Union_ax :: "(i\o) \ o" where "Union_ax(M) \ \x[M]. \z[M]. big_union(M,x,z)" definition - power_ax :: "(i=>o) => o" where + power_ax :: "(i\o) \ o" where "power_ax(M) \ \x[M]. \z[M]. powerset(M,x,z)" definition - univalent :: "[i=>o, i, [i,i]=>o] => o" where + univalent :: "[i\o, i, [i,i]\o] \ o" where "univalent(M,A,P) \ \x[M]. x\A \ (\y[M]. \z[M]. P(x,y) \ P(x,z) \ y=z)" definition - replacement :: "[i=>o, [i,i]=>o] => o" where + replacement :: "[i\o, [i,i]\o] \ o" where "replacement(M,P) \ \A[M]. univalent(M,A,P) \ (\Y[M]. \b[M]. (\x[M]. x\A \ P(x,b)) \ b \ Y)" definition - strong_replacement :: "[i=>o, [i,i]=>o] => o" where + strong_replacement :: "[i\o, [i,i]\o] \ o" where "strong_replacement(M,P) \ \A[M]. univalent(M,A,P) \ (\Y[M]. \b[M]. b \ Y \ (\x[M]. x\A \ P(x,b)))" definition - foundation_ax :: "(i=>o) => o" where + foundation_ax :: "(i\o) \ o" where "foundation_ax(M) \ \x[M]. (\y[M]. y\x) \ (\y[M]. y\x \ \(\z[M]. z\x \ z \ y))" @@ -360,12 +360,12 @@ text\Congruence rule for separation: can assume the variable is in \M\\ lemma separation_cong [cong]: "(\x. M(x) \ P(x) \ P'(x)) - \ separation(M, %x. P(x)) \ separation(M, %x. P'(x))" + \ separation(M, \x. P(x)) \ separation(M, \x. P'(x))" by (simp add: separation_def) lemma univalent_cong [cong]: "\A=A'; \x y. \x\A; M(x); M(y)\ \ P(x,y) \ P'(x,y)\ - \ univalent(M, A, %x y. P(x,y)) \ univalent(M, A', %x y. P'(x,y))" + \ univalent(M, A, \x y. P(x,y)) \ univalent(M, A', \x y. P'(x,y))" by (simp add: univalent_def) lemma univalent_triv [intro,simp]: @@ -379,8 +379,8 @@ text\Congruence rule for replacement\ lemma strong_replacement_cong [cong]: "\\x y. \M(x); M(y)\ \ P(x,y) \ P'(x,y)\ - \ strong_replacement(M, %x y. P(x,y)) \ - strong_replacement(M, %x y. P'(x,y))" + \ strong_replacement(M, \x y. P(x,y)) \ + strong_replacement(M, \x y. P'(x,y))" by (simp add: strong_replacement_def) text\The extensionality axiom\ @@ -459,13 +459,13 @@ subsection\Lemmas Needed to Reduce Some Set Constructions to Instances of Separation\ -lemma image_iff_Collect: "r `` A = {y \ \(\(r)). \p\r. \x\A. p=}" +lemma image_iff_Collect: "r `` A = {y \ \(\(r)). \p\r. \x\A. p=\x,y\}" apply (rule equalityI, auto) apply (simp add: Pair_def, blast) done lemma vimage_iff_Collect: - "r -`` A = {x \ \(\(r)). \p\r. \y\A. p=}" + "r -`` A = {x \ \(\(r)). \p\r. \y\A. p=\x,y\}" apply (rule equalityI, auto) apply (simp add: Pair_def, blast) done @@ -503,7 +503,7 @@ text\More constants, for order types\ definition - order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where + order_isomorphism :: "[i\o,i,i,i,i,i] \ o" where "order_isomorphism(M,A,r,B,s,f) \ bijection(M,A,B,f) \ (\x[M]. x\A \ (\y[M]. y\A \ @@ -512,12 +512,12 @@ pair(M,fx,fy,q) \ (p\r \ q\s))))" definition - pred_set :: "[i=>o,i,i,i,i] => o" where + pred_set :: "[i\o,i,i,i,i] \ o" where "pred_set(M,A,x,r,B) \ \y[M]. y \ B \ (\p[M]. p\r \ y \ A \ pair(M,y,x,p))" definition - membership :: "[i=>o,i,i] => o" where \ \membership relation\ + membership :: "[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)))" @@ -628,25 +628,25 @@ by blast lemma (in M_trans) pair_abs [simp]: - "M(z) \ pair(M,a,b,z) \ z=" + "M(z) \ pair(M,a,b,z) \ z=\a,b\" apply (simp add: pair_def Pair_def) apply (blast intro: transM) done lemma (in M_trans) pair_in_MD [dest!]: - "M() \ M(a) \ M(b)" + "M(\a,b\) \ M(a) \ M(b)" by (simp add: Pair_def, blast intro: transM) lemma (in M_trivial) pair_in_MI [intro!]: - "M(a) \ M(b) \ M()" + "M(a) \ M(b) \ M(\a,b\)" by (simp add: Pair_def) lemma (in M_trivial) pair_in_M_iff [simp]: - "M() \ M(a) \ M(b)" + "M(\a,b\) \ M(a) \ M(b)" by blast lemma (in M_trans) pair_components_in_M: - "\ \ A; M(A)\ \ M(x) \ M(y)" + "\\x,y\ \ A; M(A)\ \ M(x) \ M(y)" by (blast dest: transM) lemma (in M_trivial) cartprod_abs [simp]: @@ -739,8 +739,8 @@ "\A=A'; \x y. \M(x); M(y)\ \ P(x,y) \ P'(x,y); z=z'\ - \ is_Replace(M, A, %x y. P(x,y), z) \ - is_Replace(M, A', %x y. P'(x,y), z')" + \ is_Replace(M, A, \x y. P(x,y), z) \ + is_Replace(M, A', \x y. P'(x,y), z')" by (simp add: is_Replace_def) lemma (in M_trans) univalent_Replace_iff: @@ -795,7 +795,7 @@ makes relativization easier.\ lemma (in M_trans) RepFun_closed2: "\strong_replacement(M, \x y. x\A \ y = f(x)); M(A); \x\A. M(f(x))\ - \ M(RepFun(A, %x. f(x)))" + \ M(RepFun(A, \x. f(x)))" apply (simp add: RepFun_def) apply (frule strong_replacement_closed, assumption) apply (auto dest: transM simp add: Replace_conj_eq univalent_def) @@ -804,7 +804,7 @@ subsubsection \Absoluteness for \<^term>\Lambda\\ definition - is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where + 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))" @@ -838,8 +838,8 @@ lemma is_lambda_cong [cong]: "\A=A'; z=z'; \x y. \x\A; M(x); M(y)\ \ is_b(x,y) \ is_b'(x,y)\ - \ is_lambda(M, A, %x y. is_b(x,y), z) \ - is_lambda(M, A', %x y. is_b'(x,y), z')" + \ is_lambda(M, A, \x y. is_b(x,y), z) \ + is_lambda(M, A', \x y. is_b'(x,y), z')" by (simp add: is_lambda_def) lemma (in M_trans) image_abs [simp]: @@ -1001,7 +1001,7 @@ equations only hold for x\nat (or in some other set) and not for the whole of the class M. consts - natnumber_aux :: "[i=>o,i] => i" + natnumber_aux :: "[i\o,i] \ i" primrec "natnumber_aux(M,0) = (\x\nat. if empty(M,x) then 1 else 0)" @@ -1010,7 +1010,7 @@ then 1 else 0)" definition - natnumber :: "[i=>o,i,i] => o" + natnumber :: "[i\o,i,i] \ o" "natnumber(M,n,x) \ natnumber_aux(M,n)`x = 1" lemma (in M_trivial) [simp]: @@ -1076,7 +1076,7 @@ "\M(A); M(B); M(C)\ \ cartprod(M,A,B,C) \ (\p1[M]. \p2[M]. powerset(M,A \ B,p1) \ powerset(M,p1,p2) \ - C = {z \ p2. \x\A. \y\B. z = })" + C = {z \ p2. \x\A. \y\B. z = \x,y\})" apply (simp add: Pair_def cartprod_def, safe) defer 1 apply (simp add: powerset_def) @@ -1348,15 +1348,15 @@ (*???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: - "A - Collect(A,P) = Collect(A, %x. \ P(x))" + "A - Collect(A,P) = Collect(A, \x. \ P(x))" by blast lemma (in M_trans) Collect_rall_eq: - "M(Y) \ Collect(A, %x. \y[M]. y\Y \ P(x,y)) = + "M(Y) \ Collect(A, \x. \y[M]. y\Y \ P(x,y)) = (if Y=0 then A else (\y \ Y. {x \ A. P(x,y)}))" by (simp,blast dest: transM) @@ -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 = \f,b\ \ z = {cons(\n,b\, f)}}" apply (simp add: succ_fun_eq) apply (blast dest: transM) done @@ -1429,21 +1429,21 @@ subsection\Relativization and Absoluteness for Boolean Operators\ definition - is_bool_of_o :: "[i=>o, o, i] => o" where + is_bool_of_o :: "[i\o, o, i] \ o" where "is_bool_of_o(M,P,z) \ (P \ number1(M,z)) \ (\P \ empty(M,z))" definition - is_not :: "[i=>o, i, i] => o" where + is_not :: "[i\o, i, i] \ o" where "is_not(M,a,z) \ (number1(M,a) \ empty(M,z)) | (\number1(M,a) \ number1(M,z))" definition - is_and :: "[i=>o, i, i, i] => o" where + is_and :: "[i\o, i, i, i] \ o" where "is_and(M,a,b,z) \ (number1(M,a) \ z=b) | (\number1(M,a) \ empty(M,z))" definition - is_or :: "[i=>o, i, i, i] => o" where + is_or :: "[i\o, i, i, i] \ o" where "is_or(M,a,b,z) \ (number1(M,a) \ number1(M,z)) | (\number1(M,a) \ z=b)" @@ -1485,12 +1485,12 @@ subsection\Relativization and Absoluteness for List Operators\ definition - is_Nil :: "[i=>o, i] => o" where + is_Nil :: "[i\o, i] \ o" where \ \because \<^prop>\[] \ Inl(0)\\ "is_Nil(M,xs) \ \zero[M]. empty(M,zero) \ is_Inl(M,zero,xs)" definition - is_Cons :: "[i=>o,i,i,i] => o" where + 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)" @@ -1510,21 +1510,21 @@ definition - quasilist :: "i => o" where + quasilist :: "i \ o" where "quasilist(xs) \ xs=Nil \ (\x l. xs = Cons(x,l))" definition - is_quasilist :: "[i=>o,i] => o" where + is_quasilist :: "[i\o,i] \ o" where "is_quasilist(M,z) \ is_Nil(M,z) \ (\x[M]. \l[M]. is_Cons(M,x,l,z))" definition - list_case' :: "[i, [i,i]=>i, i] => i" where + list_case' :: "[i, [i,i]\i, i] \ i" where \ \A version of \<^term>\list_case\ that's always defined.\ "list_case'(a,b,xs) \ if quasilist(xs) then list_case(a,b,xs) else 0" definition - is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where + 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) \ @@ -1532,17 +1532,17 @@ (is_quasilist(M,xs) \ empty(M,z))" definition - hd' :: "i => i" where + hd' :: "i \ i" where \ \A version of \<^term>\hd\ that's always defined.\ "hd'(xs) \ if quasilist(xs) then hd(xs) else 0" definition - tl' :: "i => i" where + tl' :: "i \ i" where \ \A version of \<^term>\tl\ that's always defined.\ "tl'(xs) \ if quasilist(xs) then tl(xs) else 0" definition - is_hd :: "[i=>o,i,i] => o" where + is_hd :: "[i\o,i,i] \ o" where \ \\<^term>\hd([]) = 0\ no constraints if not a list. Avoiding implication prevents the simplifier's looping.\ "is_hd(M,xs,H) \ @@ -1551,7 +1551,7 @@ (is_quasilist(M,xs) \ empty(M,H))" definition - is_tl :: "[i=>o,i,i] => o" where + 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) \ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Satisfies_absolute.thy --- a/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Satisfies_absolute.thy Tue Sep 27 17:46:52 2022 +0100 @@ -17,7 +17,7 @@ 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 :: "[i,i]\i" where "depth_fm(p,n) \ Exists(Exists(Exists( And(formula_N_fm(n#+3,1), @@ -58,7 +58,7 @@ will be enclosed by three quantifiers.\ (* is_formula_case :: - "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" + "[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)) \ @@ -67,7 +67,7 @@ (\x[M]. x\formula \ is_Forall(M,x,v) \ is_d(x,z))" *) definition - formula_case_fm :: "[i, i, i, i, i, i]=>i" where + formula_case_fm :: "[i, i, i, i, i, i]\i" where "formula_case_fm(is_a, is_b, is_c, is_d, v, z) \ And(Forall(Forall(Implies(finite_ordinal_fm(1), Implies(finite_ordinal_fm(0), @@ -175,7 +175,7 @@ subsection \Absoluteness for the Function \<^term>\satisfies\\ definition - is_depth_apply :: "[i=>o,i,i,i] => o" where + is_depth_apply :: "[i\o,i,i,i] \ o" where \ \Merely a useful abbreviation for the sequel.\ "is_depth_apply(M,h,p,z) \ \dp[M]. \sdp[M]. \hsdp[M]. @@ -195,12 +195,12 @@ text\These constants let us instantiate the parameters \<^term>\a\, \<^term>\b\, \<^term>\c\, \<^term>\d\, etc., of the locale \Formula_Rec\.\ definition - satisfies_a :: "[i,i,i]=>i" where + satisfies_a :: "[i,i,i]\i" where "satisfies_a(A) \ \x y. \env \ list(A). bool_of_o (nth(x,env) \ nth(y,env))" definition - satisfies_is_a :: "[i=>o,i,i,i,i]=>o" where + satisfies_is_a :: "[i\o,i,i,i,i]\o" where "satisfies_is_a(M,A) \ \x y zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, @@ -210,12 +210,12 @@ zz)" definition - satisfies_b :: "[i,i,i]=>i" where + satisfies_b :: "[i,i,i]\i" where "satisfies_b(A) \ \x y. \env \ list(A). bool_of_o (nth(x,env) = nth(y,env))" definition - satisfies_is_b :: "[i=>o,i,i,i,i]=>o" where + satisfies_is_b :: "[i\o,i,i,i,i]\o" where \ \We simplify the formula to have just \<^term>\nx\ rather than introducing \<^term>\ny\ with \<^term>\nx=ny\\ "satisfies_is_b(M,A) \ @@ -226,11 +226,11 @@ zz)" definition - satisfies_c :: "[i,i,i,i,i]=>i" where + satisfies_c :: "[i,i,i,i,i]\i" where "satisfies_c(A) \ \p q rp rq. \env \ list(A). not(rp ` env and rq ` env)" definition - satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" where + satisfies_is_c :: "[i\o,i,i,i,i,i]\o" where "satisfies_is_c(M,A,h) \ \p q zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, \env z. \hp[M]. \hq[M]. @@ -240,12 +240,12 @@ zz)" definition - satisfies_d :: "[i,i,i]=>i" where + satisfies_d :: "[i,i,i]\i" where "satisfies_d(A) \ \p rp. \env \ list(A). bool_of_o (\x\A. rp ` (Cons(x,env)) = 1)" definition - satisfies_is_d :: "[i=>o,i,i,i,i]=>o" where + satisfies_is_d :: "[i\o,i,i,i,i]\o" where "satisfies_is_d(M,A,h) \ \p zz. \lA[M]. is_list(M,A,lA) \ is_lambda(M, lA, @@ -258,7 +258,7 @@ zz)" definition - satisfies_MH :: "[i=>o,i,i,i,i]=>o" where + satisfies_MH :: "[i\o,i,i,i,i]\o" where \ \The variable \<^term>\u\ is unused, but gives \<^term>\satisfies_MH\ the correct arity.\ "satisfies_MH \ @@ -271,7 +271,7 @@ z)" definition - is_satisfies :: "[i=>o,i,i,i]=>o" where + is_satisfies :: "[i\o,i,i,i]\o" where "is_satisfies(M,A) \ is_formula_rec (M, satisfies_MH(M,A))" @@ -515,7 +515,7 @@ 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 :: "[i,i,i]\i" where "depth_apply_fm(h,p,z) \ Exists(Exists(Exists( And(finite_ordinal_fm(2), @@ -559,7 +559,7 @@ zz) *) definition - satisfies_is_a_fm :: "[i,i,i,i]=>i" where + satisfies_is_a_fm :: "[i,i,i,i]\i" where "satisfies_is_a_fm(A,x,y,z) \ Forall( Implies(is_list_fm(succ(A),0), @@ -611,7 +611,7 @@ zz) *) definition - satisfies_is_b_fm :: "[i,i,i,i]=>i" where + satisfies_is_b_fm :: "[i,i,i,i]\i" where "satisfies_is_b_fm(A,x,y,z) \ Forall( Implies(is_list_fm(succ(A),0), @@ -661,7 +661,7 @@ zz) *) definition - satisfies_is_c_fm :: "[i,i,i,i,i]=>i" where + satisfies_is_c_fm :: "[i,i,i,i,i]\i" where "satisfies_is_c_fm(A,h,p,q,zz) \ Forall( Implies(is_list_fm(succ(A),0), @@ -715,7 +715,7 @@ zz) *) definition - satisfies_is_d_fm :: "[i,i,i,i]=>i" where + satisfies_is_d_fm :: "[i,i,i,i]\i" where "satisfies_is_d_fm(A,h,p,zz) \ Forall( Implies(is_list_fm(succ(A),0), @@ -770,7 +770,7 @@ zz) *) definition - satisfies_MH_fm :: "[i,i,i,i]=>i" where + satisfies_MH_fm :: "[i,i,i,i]\i" where "satisfies_MH_fm(A,u,f,zz) \ Forall( Implies(is_formula_fm(0), diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/WF_absolute.thy --- a/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/WF_absolute.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,10 +9,10 @@ subsection\Transitive closure without fixedpoints\ definition - rtrancl_alt :: "[i,i]=>i" where + 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 = \x,y\ \ f`0 = x \ f`n = y) \ (\i\n. \ r)}" lemma alt_rtrancl_lemma1 [rule_format]: @@ -59,7 +59,7 @@ definition - rtran_closure_mem :: "[i=>o,i,i,i] => o" where + rtran_closure_mem :: "[i\o,i,i,i] \ o" where \ \The property of belonging to \rtran_closure(r)\\ "rtran_closure_mem(M,A,r,p) \ \nnat[M]. \n[M]. \n'[M]. @@ -73,13 +73,13 @@ fun_apply(M,f,sj,fsj) \ pair(M,fj,fsj,ffp) \ ffp \ r)))" definition - rtran_closure :: "[i=>o,i,i] => o" where + rtran_closure :: "[i\o,i,i] \ o" where "rtran_closure(M,r,s) \ \A[M]. is_field(M,r,A) \ (\p[M]. p \ s \ rtran_closure_mem(M,A,r,p))" definition - tran_closure :: "[i=>o,i,i] => o" where + tran_closure :: "[i\o,i,i] \ o" where "tran_closure(M,r,t) \ \s[M]. rtran_closure(M,r,s) \ composition(M,r,s,t)" @@ -98,7 +98,7 @@ \ 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) \ + (\x[M]. \y[M]. p = \x,y\ \ 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 \ \w,x\ \ r^+)" by (insert wellfounded_trancl_separation [of r Z], simp) text\Alternative proof of \wf_on_trancl\; inspiration for the diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/WFrec.thy --- a/src/ZF/Constructible/WFrec.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/WFrec.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,7 +12,7 @@ (*Many of these might be useful in WF.thy*) lemma apply_recfun2: - "\is_recfun(r,a,H,f); :f\ \ i = H(x, restrict(f,r-``{x}))" + "\is_recfun(r,a,H,f); \x,i\:f\ \ i = H(x, restrict(f,r-``{x}))" apply (frule apply_recfun) apply (blast dest: is_recfun_type fun_is_rel) apply (simp add: function_apply_equality [OF _ is_recfun_imp_function]) @@ -36,7 +36,7 @@ by (blast dest: is_recfun_type fun_is_rel) lemma trans_Int_eq: - "\trans(r); \ r\ \ r -`` {x} \ r -`` {y} = r -`` {y}" + "\trans(r); \y,x\ \ r\ \ r -`` {x} \ r -`` {y} = r -`` {y}" by (blast intro: transD) lemma is_recfun_restrict_idem: @@ -91,7 +91,7 @@ "\is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); M(f); M(g); M(r); M(x); M(a); M(b)\ - \ \ r \ \ r \ f`x=g`x" + \ \x,a\ \ r \ \x,b\ \ r \ f`x=g`x" apply (frule_tac f=f in is_recfun_type) apply (frule_tac f=g in is_recfun_type) apply (simp add: is_recfun_def) @@ -103,8 +103,8 @@ apply (erule ssubst)+ apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) apply (rename_tac x1) -apply (rule_tac t="%z. H(x1,z)" in subst_context) -apply (subgoal_tac "\y \ r-``{x1}. \z. \f \ \g") +apply (rule_tac t="\z. H(x1,z)" in subst_context) +apply (subgoal_tac "\y \ r-``{x1}. \z. \y,z\\f \ \y,z\\g") apply (blast intro: transD) apply (simp add: apply_iff) apply (blast intro: transD sym) @@ -113,7 +113,7 @@ lemma (in M_basic) is_recfun_cut: "\is_recfun(r,a,H,f); is_recfun(r,b,H,g); wellfounded(M,r); trans(r); - M(f); M(g); M(r); \ r\ + M(f); M(g); M(r); \b,a\ \ r\ \ restrict(f, r-``{b}) = g" apply (frule_tac f=f in is_recfun_type) apply (rule fun_extension) @@ -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]. \x,a\ \ r \ z = ))" apply (simp add: is_recfun_def lam_def) apply (safe intro!: equalityI) apply (drule equalityD1 [THEN subsetD], assumption) @@ -174,21 +174,21 @@ \x[M]. \g[M]. function(g) \ M(H(x,g)); M(Y); \b[M]. b \ Y \ - (\x[M]. \ r \ + (\x[M]. \x,a1\ \ 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") +apply (subgoal_tac "\y \ r-``{x}. \z. \y,z\:Y \ \y,z\:f") apply (simp (no_asm_simp) add: restrict_def) apply (thin_tac "rall(M,P)" for P)+ \ \essential for efficiency\ apply (frule is_recfun_type [THEN fun_is_rel], blast) apply (frule pair_components_in_M, assumption, clarify) apply (rule iffI) - apply (frule_tac y="" in transM, assumption) + apply (frule_tac y="\y,z\" in transM, assumption) apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff] apply_recfun is_recfun_cut) txt\Opposite inclusion: something in f, show in Y\ -apply (frule_tac y="" in transM, assumption) +apply (frule_tac y="\y,z\" in transM, assumption) apply (simp add: vimage_singleton_iff) apply (rule conjI) apply (blast dest: transD) @@ -271,7 +271,7 @@ subsection\Relativization of the ZF Predicate \<^term>\is_recfun\\ definition - M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where + M_is_recfun :: "[i\o, [i,i,i]\o, i, i, i] \ o" where "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]. @@ -280,12 +280,12 @@ xa \ r \ MH(x, f_r_sx, y))" definition - is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where + 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)" definition - wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where + 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))" @@ -324,8 +324,8 @@ lemma wfrec_replacement_cong [cong]: "\\x y z. \M(x); M(y); M(z)\ \ MH(x,y,z) \ MH'(x,y,z); r=r'\ - \ wfrec_replacement(M, %x y. MH(x,y), r) \ - wfrec_replacement(M, %x y. MH'(x,y), r')" + \ wfrec_replacement(M, \x y. MH(x,y), r) \ + wfrec_replacement(M, \x y. MH'(x,y), r')" by (simp add: is_wfrec_def wfrec_replacement_def) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Constructible/Wellorderings.thy --- a/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Constructible/Wellorderings.thy Tue Sep 27 17:46:52 2022 +0100 @@ -16,33 +16,33 @@ subsection\Wellorderings\ definition - irreflexive :: "[i=>o,i,i]=>o" where - "irreflexive(M,A,r) \ \x[M]. x\A \ \ r" + irreflexive :: "[i\o,i,i]\o" where + "irreflexive(M,A,r) \ \x[M]. x\A \ \x,x\ \ r" definition - transitive_rel :: "[i=>o,i,i]=>o" where + transitive_rel :: "[i\o,i,i]\o" where "transitive_rel(M,A,r) \ \x[M]. x\A \ (\y[M]. y\A \ (\z[M]. z\A \ - \r \ \r \ \r))" + \x,y\\r \ \y,z\\r \ \x,z\\r))" definition - linear_rel :: "[i=>o,i,i]=>o" where + linear_rel :: "[i\o,i,i]\o" where "linear_rel(M,A,r) \ - \x[M]. x\A \ (\y[M]. y\A \ \r | x=y | \r)" + \x[M]. x\A \ (\y[M]. y\A \ \x,y\\r | x=y | \y,x\\r)" definition - wellfounded :: "[i=>o,i]=>o" where + 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 \ \z,y\ \ r))" definition - wellfounded_on :: "[i=>o,i,i]=>o" where + 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 \ \z,y\ \ r))" definition - wellordered :: "[i=>o,i,i]=>o" where + 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)" @@ -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. \y,x\ \ 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. \y,x\ \ 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) @@ -182,7 +182,7 @@ apply (frule transM, assumption) apply blast apply clarify - apply (subgoal_tac "M()", blast) + apply (subgoal_tac "M(\xb,ya\)", blast) apply (blast dest: transM) apply auto done @@ -220,14 +220,14 @@ done lemma (in M_basic) wellfounded_on_asym: - "\wellfounded_on(M,A,r); \r; a\A; x\A; M(A)\ \ \r" + "\wellfounded_on(M,A,r); \a,x\\r; a\A; x\A; M(A)\ \ \x,a\\r" apply (simp add: wellfounded_on_def) apply (drule_tac x="{x,a}" in rspec) apply (blast dest: transM)+ done lemma (in M_basic) wellordered_asym: - "\wellordered(M,A,r); \r; a\A; x\A; M(A)\ \ \r" + "\wellordered(M,A,r); \a,x\\r; a\A; x\A; M(A)\ \ \x,a\\r" by (simp add: wellordered_def, blast dest: wellfounded_on_asym) end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Epsilon.thy --- a/src/ZF/Epsilon.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Epsilon.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,32 +8,32 @@ theory Epsilon imports Nat begin definition - eclose :: "i=>i" where - "eclose(A) \ \n\nat. nat_rec(n, A, %m r. \(r))" + eclose :: "i\i" where + "eclose(A) \ \n\nat. nat_rec(n, A, \m r. \(r))" definition - transrec :: "[i, [i,i]=>i] =>i" where + transrec :: "[i, [i,i]\i] \i" where "transrec(a,H) \ wfrec(Memrel(eclose({a})), a, H)" definition - rank :: "i=>i" where - "rank(a) \ transrec(a, %x f. \y\x. succ(f`y))" + rank :: "i\i" where + "rank(a) \ transrec(a, \x f. \y\x. succ(f`y))" definition - transrec2 :: "[i, i, [i,i]=>i] =>i" where + transrec2 :: "[i, i, [i,i]\i] \i" where "transrec2(k, a, b) \ transrec(k, - %i r. if(i=0, a, + \i r. if(i=0, a, if(\j. i=succ(j), b(THE j. i=succ(j), r`(THE j. i=succ(j))), \ji, i]=>i" where - "recursor(a,b,k) \ transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))" + recursor :: "[i, [i,i]\i, i]\i" where + "recursor(a,b,k) \ transrec(k, \n f. nat_case(a, \m. b(m, f`m), n))" definition - rec :: "[i, i, [i,i]=>i]=>i" where + rec :: "[i, i, [i,i]\i]\i" where "rec(k,a,b) \ recursor(a,b,k)" @@ -85,7 +85,7 @@ (** eclose(A) is the least transitive set including A as a subset. **) lemma eclose_least_lemma: - "\Transset(X); A<=X; n \ nat\ \ nat_rec(n, A, %m r. \(r)) \ X" + "\Transset(X); A<=X; n \ nat\ \ nat_rec(n, A, \m r. \(r)) \ X" apply (unfold Transset_def) apply (erule nat_induct) apply (simp add: nat_rec_0) @@ -290,13 +290,13 @@ apply (erule eclose_rank_lt [THEN succ_leI]) done -lemma rank_pair1: "rank(a) < rank()" +lemma rank_pair1: "rank(a) < rank(\a,b\)" apply (unfold Pair_def) apply (rule consI1 [THEN rank_lt, THEN lt_trans]) apply (rule consI1 [THEN consI2, THEN rank_lt]) done -lemma rank_pair2: "rank(b) < rank()" +lemma rank_pair2: "rank(b) < rank(\a,b\)" apply (unfold Pair_def) apply (rule consI1 [THEN consI2, THEN rank_lt, THEN lt_trans]) apply (rule consI1 [THEN consI2, THEN rank_lt]) @@ -310,7 +310,7 @@ (*The first premise not only fixs i but ensures @{term"f\0"}. The second premise is now essential. Consider otherwise the relation - r = {<0,0>,<0,1>,<0,2>,...}. Then f`0 = \(f``{0}) = \(nat) = nat, + r = {\0,0\,\0,1\,\0,2\,...}. Then f`0 = \(f``{0}) = \(nat) = nat, whose rank equals that of r.*) lemma rank_apply: "\i \ domain(f); function(f)\ \ rank(f`i) < rank(f)" apply clarify diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/EquivClass.thy --- a/src/ZF/EquivClass.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/EquivClass.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,24 +8,24 @@ theory EquivClass imports Trancl Perm begin definition - quotient :: "[i,i]=>i" (infixl \'/'/\ 90) (*set of equiv classes*) where + quotient :: "[i,i]\i" (infixl \'/'/\ 90) (*set of equiv classes*) where "A//r \ {r``{x} . x \ A}" definition - congruent :: "[i,i=>i]=>o" where - "congruent(r,b) \ \y z. :r \ b(y)=b(z)" + congruent :: "[i,i\i]\o" where + "congruent(r,b) \ \y z. \y,z\:r \ b(y)=b(z)" definition - congruent2 :: "[i,i,[i,i]=>i]=>o" where + congruent2 :: "[i,i,[i,i]\i]\o" where "congruent2(r1,r2,b) \ \y1 z1 y2 z2. - :r1 \ :r2 \ b(y1,y2) = b(z1,z2)" + \y1,z1\:r1 \ \y2,z2\:r2 \ b(y1,y2) = b(z1,z2)" abbreviation - RESPECTS ::"[i=>i, i] => o" (infixr \respects\ 80) where + RESPECTS ::"[i\i, i] \ o" (infixr \respects\ 80) where "f respects r \ congruent(r,f)" abbreviation - RESPECTS2 ::"[i=>i=>i, i] => o" (infixr \respects2 \ 80) where + RESPECTS2 ::"[i\i\i, i] \ o" (infixr \respects2 \ 80) where "f respects2 r \ congruent2(r,r,f)" \ \Abbreviation for the common case where the relations are identical\ @@ -54,18 +54,18 @@ "\converse(r) O r = r; domain(r) = A\ \ equiv(A,r)" apply (unfold equiv_def refl_def sym_def trans_def) apply (erule equalityE) -apply (subgoal_tac "\x y. \ r \ \ r", blast+) +apply (subgoal_tac "\x y. \x,y\ \ r \ \y,x\ \ r", blast+) done (** Equivalence classes **) (*Lemma for the next result*) lemma equiv_class_subset: - "\sym(r); trans(r); : r\ \ r``{a} \ r``{b}" + "\sym(r); trans(r); \a,b\: r\ \ r``{a} \ r``{b}" by (unfold trans_def sym_def, blast) lemma equiv_class_eq: - "\equiv(A,r); : r\ \ r``{a} = r``{b}" + "\equiv(A,r); \a,b\: r\ \ r``{a} = r``{b}" apply (unfold equiv_def) apply (safe del: subsetI intro!: equalityI equiv_class_subset) apply (unfold sym_def, blast) @@ -77,26 +77,26 @@ (*Lemma for the next result*) lemma subset_equiv_class: - "\equiv(A,r); r``{b} \ r``{a}; b \ A\ \ : r" + "\equiv(A,r); r``{b} \ r``{a}; b \ A\ \ \a,b\: r" by (unfold equiv_def refl_def, blast) -lemma eq_equiv_class: "\r``{a} = r``{b}; equiv(A,r); b \ A\ \ : r" +lemma eq_equiv_class: "\r``{a} = r``{b}; equiv(A,r); b \ A\ \ \a,b\: r" by (assumption | rule equalityD2 subset_equiv_class)+ (*thus r``{a} = r``{b} as well*) lemma equiv_class_nondisjoint: - "\equiv(A,r); x: (r``{a} \ r``{b})\ \ : r" + "\equiv(A,r); x: (r``{a} \ r``{b})\ \ \a,b\: r" by (unfold equiv_def trans_def sym_def, blast) lemma equiv_type: "equiv(A,r) \ r \ A*A" by (unfold equiv_def, blast) lemma equiv_class_eq_iff: - "equiv(A,r) \ : r \ r``{x} = r``{y} \ x \ A \ y \ A" + "equiv(A,r) \ \x,y\: r \ r``{x} = r``{y} \ x \ A \ y \ A" by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) lemma eq_equiv_class_iff: - "\equiv(A,r); x \ A; y \ A\ \ r``{x} = r``{y} \ : r" + "\equiv(A,r); x \ A; y \ A\ \ r``{x} = r``{y} \ \x,y\: r" by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) (*** Quotients ***) @@ -151,7 +151,7 @@ lemma UN_equiv_class_inject: "\equiv(A,r); b respects r; (\x\X. b(x))=(\y\Y. b(y)); X \ A//r; Y \ A//r; - \x y. \x \ A; y \ A; b(x)=b(y)\ \ :r\ + \x y. \x \ A; y \ A; b(x)=b(y)\ \ \x,y\:r\ \ X=Y" apply (unfold quotient_def, safe) apply (rule equiv_class_eq, assumption) @@ -167,7 +167,7 @@ lemma congruent2_implies_congruent_UN: "\equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a \ A2\ \ - congruent(r1, %x1. \x2 \ r2``{a}. b(x1,x2))" + congruent(r1, \x1. \x2 \ r2``{a}. b(x1,x2))" apply (unfold congruent_def, safe) apply (frule equiv_type [THEN subsetD], assumption) apply clarify @@ -197,8 +197,8 @@ than the direct proof*) lemma congruent2I: "\equiv(A1,r1); equiv(A2,r2); - \y z w. \w \ A2; \ r1\ \ b(y,w) = b(z,w); - \y z w. \w \ A1; \ r2\ \ b(w,y) = b(w,z) + \y z w. \w \ A2; \y,z\ \ r1\ \ b(y,w) = b(z,w); + \y z w. \w \ A1; \y,z\ \ r2\ \ b(w,y) = b(w,z) \ \ congruent2(r1,r2,b)" apply (unfold congruent2_def equiv_def refl_def, safe) apply (blast intro: trans) @@ -207,7 +207,7 @@ lemma congruent2_commuteI: assumes equivA: "equiv(A,r)" and commute: "\y z. \y \ A; z \ A\ \ b(y,z) = b(z,y)" - and congt: "\y z w. \w \ A; : r\ \ b(w,y) = b(w,z)" + and congt: "\y z w. \w \ A; \y,z\: r\ \ b(w,y) = b(w,z)" shows "b respects2 r" apply (insert equivA [THEN equiv_type, THEN subsetD]) apply (rule congruent2I [OF equivA equivA]) @@ -220,9 +220,9 @@ (*Obsolete?*) lemma congruent_commuteI: "\equiv(A,r); Z \ A//r; - \w. \w \ A\ \ congruent(r, %z. b(w,z)); + \w. \w \ A\ \ congruent(r, \z. b(w,z)); \x y. \x \ A; y \ A\ \ b(y,x) = b(x,y) -\ \ congruent(r, %w. \z\Z. b(w,z))" +\ \ congruent(r, \w. \z\Z. b(w,z))" apply (simp (no_asm) add: congruent_def) apply (safe elim!: quotientE) apply (frule equiv_type [THEN subsetD], assumption) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Finite.thy --- a/src/ZF/Finite.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Finite.thy Tue Sep 27 17:46:52 2022 +0100 @@ -18,8 +18,8 @@ consts - Fin :: "i=>i" - FiniteFun :: "[i,i]=>i" (\(_ -||>/ _)\ [61, 60] 60) + Fin :: "i\i" + FiniteFun :: "[i,i]\i" (\(_ -||>/ _)\ [61, 60] 60) inductive domains "Fin(A)" \ "Pow(A)" @@ -34,7 +34,7 @@ intros emptyI: "0 \ A -||> B" consI: "\a \ A; b \ B; h \ A -||> B; a \ domain(h)\ - \ cons(,h) \ A -||> B" + \ cons(\a,b\,h) \ A -||> B" type_intros Fin.intros @@ -200,7 +200,7 @@ subsection\The Contents of a Singleton Set\ definition - contents :: "i=>i" where + contents :: "i\i" where "contents(X) \ THE x. X = {x}" lemma contents_eq [simp]: "contents ({x}) = x" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Fixedpt.thy --- a/src/ZF/Fixedpt.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Fixedpt.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,15 +9,15 @@ definition (*monotone operator from Pow(D) to itself*) - bnd_mono :: "[i,i=>i]=>o" where + bnd_mono :: "[i,i\i]\o" where "bnd_mono(D,h) \ h(D)<=D \ (\W X. W<=X \ X<=D \ h(W) \ h(X))" definition - lfp :: "[i,i=>i]=>i" where + lfp :: "[i,i\i]\i" where "lfp(D,h) \ \({X: Pow(D). h(X) \ X})" definition - gfp :: "[i,i=>i]=>i" where + gfp :: "[i,i\i]\i" where "gfp(D,h) \ \({X: Pow(D). X \ h(X)})" text\The theorem is proved in the lattice of subsets of \<^term>\D\, @@ -285,7 +285,7 @@ (*The version used in the induction/coinduction package*) lemma def_Collect_coinduct: - "\A \ gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w))); + "\A \ gfp(D, \w. Collect(D,P(w))); bnd_mono(D, \w. Collect(D,P(w))); a: X; X \ D; \z. z: X \ P(X \ A, z)\ \ a \ A" apply (rule def_coinduct, assumption+, blast+) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/IMP/Com.thy --- a/src/ZF/IMP/Com.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/IMP/Com.thy Tue Sep 27 17:46:52 2022 +0100 @@ -23,17 +23,17 @@ consts evala :: i abbreviation - evala_syntax :: "[i, i] => o" (infixl \-a->\ 50) - where "p -a-> n \ \ evala" + evala_syntax :: "[i, i] \ o" (infixl \-a->\ 50) + where "p -a-> n \ \p,n\ \ evala" inductive domains "evala" \ "(aexp \ (loc -> nat)) \ nat" intros N: "\n \ nat; sigma \ loc->nat\ \ -a-> n" X: "\x \ loc; sigma \ loc->nat\ \ -a-> sigma`x" - Op1: "\ -a-> n; f \ nat -> nat\ \ -a-> f`n" - Op2: "\ -a-> n0; -a-> n1; f \ (nat\nat) -> nat\ - \ -a-> f`" + Op1: "\\e,sigma\ -a-> n; f \ nat -> nat\ \ -a-> f`n" + Op2: "\\e0,sigma\ -a-> n0; \e1,sigma\ -a-> n1; f \ (nat\nat) -> nat\ + \ -a-> f`\n0,n1\" type_intros aexp.intros apply_funtype @@ -53,20 +53,20 @@ consts evalb :: i abbreviation - evalb_syntax :: "[i,i] => o" (infixl \-b->\ 50) - where "p -b-> b \ \ evalb" + evalb_syntax :: "[i,i] \ o" (infixl \-b->\ 50) + where "p -b-> b \ \p,b\ \ evalb" inductive domains "evalb" \ "(bexp \ (loc -> nat)) \ bool" intros - true: "\sigma \ loc -> nat\ \ -b-> 1" - false: "\sigma \ loc -> nat\ \ -b-> 0" - ROp: "\ -a-> n0; -a-> n1; f \ (nat*nat)->bool\ - \ -b-> f` " - noti: "\ -b-> w\ \ -b-> not(w)" - andi: "\ -b-> w0; -b-> w1\ + true: "\sigma \ loc -> nat\ \ \true,sigma\ -b-> 1" + false: "\sigma \ loc -> nat\ \ \false,sigma\ -b-> 0" + ROp: "\\a0,sigma\ -a-> n0; \a1,sigma\ -a-> n1; f \ (nat*nat)->bool\ + \ -b-> f`\n0,n1\ " + noti: "\\b,sigma\ -b-> w\ \ -b-> not(w)" + andi: "\\b0,sigma\ -b-> w0; \b1,sigma\ -b-> w1\ \ -b-> (w0 and w1)" - ori: "\ -b-> w0; -b-> w1\ + ori: "\\b0,sigma\ -b-> w0; \b1,sigma\ -b-> w1\ \ -b-> (w0 or w1)" type_intros bexp.intros apply_funtype and_type or_type bool_1I bool_0I not_type @@ -87,32 +87,32 @@ consts evalc :: i abbreviation - evalc_syntax :: "[i, i] => o" (infixl \-c->\ 50) - where "p -c-> s \ \ evalc" + evalc_syntax :: "[i, i] \ o" (infixl \-c->\ 50) + where "p -c-> s \ \p,s\ \ evalc" inductive domains "evalc" \ "(com \ (loc -> nat)) \ (loc -> nat)" intros skip: "\sigma \ loc -> nat\ \ <\,sigma> -c-> sigma" - assign: "\m \ nat; x \ loc; -a-> m\ + assign: "\m \ nat; x \ loc; \a,sigma\ -a-> m\ \ a,sigma> -c-> sigma(x:=m)" - semi: "\ -c-> sigma2; -c-> sigma1\ + semi: "\\c0,sigma\ -c-> sigma2; \c1,sigma2\ -c-> sigma1\ \ c1, sigma> -c-> sigma1" if1: "\b \ bexp; c1 \ com; sigma \ loc->nat; - -b-> 1; -c-> sigma1\ + \b,sigma\ -b-> 1; \c0,sigma\ -c-> sigma1\ \ <\ b \ c0 \ c1, sigma> -c-> sigma1" if0: "\b \ bexp; c0 \ com; sigma \ loc->nat; - -b-> 0; -c-> sigma1\ + \b,sigma\ -b-> 0; \c1,sigma\ -c-> sigma1\ \ <\ b \ c0 \ c1, sigma> -c-> sigma1" - while0: "\c \ com; -b-> 0\ + while0: "\c \ com; \b, sigma\ -b-> 0\ \ <\ b \ c,sigma> -c-> sigma" - while1: "\c \ com; -b-> 1; -c-> sigma2; + while1: "\c \ com; \b,sigma\ -b-> 1; \c,sigma\ -c-> sigma2; <\ b \ c, sigma2> -c-> sigma1\ \ <\ b \ c, sigma> -c-> sigma1" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/IMP/Denotation.thy --- a/src/ZF/IMP/Denotation.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/IMP/Denotation.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,12 +9,12 @@ subsection \Definitions\ consts - A :: "i => i => i" - B :: "i => i => i" - C :: "i => i" + A :: "i \ i \ i" + B :: "i \ i \ i" + C :: "i \ i" definition - Gamma :: "[i,i,i] => i" (\\\) where + Gamma :: "[i,i,i] \ i" (\\\) where "\(b,cden) \ (\phi. {io \ (phi O cden). B(b,fst(io))=1} \ {io \ id(loc->nat). B(b,fst(io))=0})" @@ -58,7 +58,7 @@ done lemma C_type_D [dest]: - "\ \ C(c); c \ com\ \ x \ loc->nat \ y \ loc->nat" + "\\x,y\ \ 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 0c18df79b1c8 -r a642599ffdea src/ZF/IMP/Equiv.thy --- a/src/ZF/IMP/Equiv.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/IMP/Equiv.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,7 +8,7 @@ lemma aexp_iff [rule_format]: "\a \ aexp; sigma: loc -> nat\ - \ \n. -a-> n \ A(a,sigma) = n" + \ \n. \a,sigma\ -a-> n \ A(a,sigma) = n" apply (erule aexp.induct) apply (force intro!: evala.intros)+ done @@ -17,8 +17,8 @@ aexp_iff [THEN iffD2, intro!] inductive_cases [elim!]: - " -b-> x" - " -b-> x" + "\true,sigma\ -b-> x" + "\false,sigma\ -b-> x" " -b-> x" " -b-> x" " -b-> x" @@ -27,7 +27,7 @@ lemma bexp_iff [rule_format]: "\b \ bexp; sigma: loc -> nat\ - \ \w. -b-> w \ B(b,sigma) = w" + \ \w. \b,sigma\ -b-> w \ B(b,sigma) = w" apply (erule bexp.induct) apply (auto intro!: evalb.intros) done @@ -35,7 +35,7 @@ declare bexp_iff [THEN iffD1, simp] bexp_iff [THEN iffD2, intro!] -lemma com1: " -c-> sigma' \ \ C(c)" +lemma com1: "\c,sigma\ -c-> sigma' \ \ C(c)" apply (erule evalc.induct) apply (simp_all (no_asm_simp)) txt \\assign\\ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Acc.thy --- a/src/ZF/Induct/Acc.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Acc.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,7 +12,7 @@ \ consts - acc :: "i => i" + acc :: "i \ i" inductive domains "acc(r)" \ "field(r)" @@ -28,15 +28,15 @@ The intended introduction rule: \ -lemma accI: "\\b. :r \ b \ acc(r); a \ field(r)\ \ a \ acc(r)" +lemma accI: "\\b. \b,a\:r \ b \ acc(r); a \ field(r)\ \ a \ acc(r)" by (blast intro: acc.intros) -lemma acc_downward: "\b \ acc(r); : r\ \ a \ acc(r)" +lemma acc_downward: "\b \ acc(r); \a,b\: r\ \ a \ acc(r)" by (erule acc.cases) blast lemma acc_induct [consumes 1, case_names vimage, induct set: acc]: "\a \ acc(r); - \x. \x \ acc(r); \y. :r \ P(y)\ \ P(x) + \x. \x \ acc(r); \y. \y,x\:r \ P(y)\ \ P(x) \ \ P(a)" by (erule acc.induct) (blast intro: acc.intros) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Binary_Trees.thy --- a/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Binary_Trees.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,7 +10,7 @@ subsection \Datatype definition\ consts - bt :: "i => i" + bt :: "i \ i" datatype "bt(A)" = Lf | Br ("a \ A", "t1 \ bt(A)", "t2 \ bt(A)") @@ -66,7 +66,7 @@ subsection \Number of nodes, with an example of tail-recursion\ -consts n_nodes :: "i => i" +consts n_nodes :: "i \ i" primrec "n_nodes(Lf) = 0" "n_nodes(Br(a, l, r)) = succ(n_nodes(l) #+ n_nodes(r))" @@ -74,7 +74,7 @@ lemma n_nodes_type [simp]: "t \ bt(A) \ n_nodes(t) \ nat" by (induct set: bt) auto -consts n_nodes_aux :: "i => i" +consts n_nodes_aux :: "i \ i" primrec "n_nodes_aux(Lf) = (\k \ nat. k)" "n_nodes_aux(Br(a, l, r)) = @@ -88,7 +88,7 @@ done definition - n_nodes_tail :: "i => i" where + n_nodes_tail :: "i \ i" where "n_nodes_tail(t) \ n_nodes_aux(t) ` 0" lemma "t \ bt(A) \ n_nodes_tail(t) = n_nodes(t)" @@ -98,7 +98,7 @@ subsection \Number of leaves\ consts - n_leaves :: "i => i" + n_leaves :: "i \ i" primrec "n_leaves(Lf) = 1" "n_leaves(Br(a, l, r)) = n_leaves(l) #+ n_leaves(r)" @@ -110,7 +110,7 @@ subsection \Reflecting trees\ consts - bt_reflect :: "i => i" + bt_reflect :: "i \ i" primrec "bt_reflect(Lf) = Lf" "bt_reflect(Br(a, l, r)) = Br(a, bt_reflect(r), bt_reflect(l))" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Brouwer.thy --- a/src/ZF/Induct/Brouwer.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Brouwer.thy Tue Sep 27 17:46:52 2022 +0100 @@ -42,7 +42,7 @@ subsection \The Martin-Löf wellordering type\ consts - Well :: "[i, i => i] => i" + Well :: "[i, i \ i] \ i" datatype \ "Vfrom(A \ (\x \ A. B(x)), csucc(nat \ |\x \ A. B(x)|))" \ \The union with \nat\ ensures that the cardinal is infinite.\ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Comb.thy --- a/src/ZF/Induct/Comb.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Comb.thy Tue Sep 27 17:46:52 2022 +0100 @@ -33,10 +33,10 @@ consts contract :: i abbreviation contract_syntax :: "[i,i] \ o" (infixl \\\<^sup>1\ 50) - where "p \\<^sup>1 q \ \ contract" + where "p \\<^sup>1 q \ \p,q\ \ contract" abbreviation contract_multi :: "[i,i] \ o" (infixl \\\ 50) - where "p \ q \ \ contract^*" + where "p \ q \ \p,q\ \ contract^*" inductive domains "contract" \ "comb \ comb" @@ -54,11 +54,11 @@ consts parcontract :: i -abbreviation parcontract_syntax :: "[i,i] => o" (infixl \\\<^sup>1\ 50) - where "p \\<^sup>1 q \ \ parcontract" +abbreviation parcontract_syntax :: "[i,i] \ o" (infixl \\\<^sup>1\ 50) + where "p \\<^sup>1 q \ \p,q\ \ parcontract" -abbreviation parcontract_multi :: "[i,i] => o" (infixl \\\ 50) - where "p \ q \ \ parcontract^+" +abbreviation parcontract_multi :: "[i,i] \ o" (infixl \\\ 50) + where "p \ q \ \p,q\ \ parcontract^+" inductive domains "parcontract" \ "comb \ comb" @@ -78,14 +78,14 @@ definition diamond :: "i \ o" where "diamond(r) \ - \x y. \r \ (\y'. \r \ (\z. \r \ \ r))" + \x y. \x,y\\r \ (\y'. \r \ (\z. \y,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)" + "\diamond(r); \x,y\:r^+\ \ + \y'. :r \ (\z. : r^+ \ \y,z\: r)" apply (unfold diamond_def) apply (erule trancl_induct) apply (blast intro: r_into_trancl) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Datatypes.thy --- a/src/ZF/Induct/Datatypes.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Datatypes.thy Tue Sep 27 17:46:52 2022 +0100 @@ -14,7 +14,7 @@ \ consts - data :: "[i, i] => i" + data :: "[i, i] \ i" datatype "data(A, B)" = Con0 diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/FoldSet.thy --- a/src/ZF/Induct/FoldSet.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/FoldSet.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,28 +8,28 @@ theory FoldSet imports ZF begin -consts fold_set :: "[i, i, [i,i]=>i, i] => i" +consts fold_set :: "[i, i, [i,i]\i, i] \ i" inductive domains "fold_set(A, B, f,e)" \ "Fin(A)*B" intros - emptyI: "e\B \ <0, e>\fold_set(A, B, f,e)" - consI: "\x\A; x \C; \ fold_set(A, B,f,e); f(x,y):B\ + emptyI: "e\B \ \0, e\\fold_set(A, B, f,e)" + consI: "\x\A; x \C; \C,y\ \ fold_set(A, B,f,e); f(x,y):B\ \ \fold_set(A, B, f, e)" type_intros Fin.intros definition - fold :: "[i, [i,i]=>i, i, i] => i" (\fold[_]'(_,_,_')\) where - "fold[B](f,e, A) \ THE x. \fold_set(A, B, f,e)" + fold :: "[i, [i,i]\i, i, i] \ i" (\fold[_]'(_,_,_')\) where + "fold[B](f,e, A) \ THE x. \A, x\\fold_set(A, B, f,e)" definition - setsum :: "[i=>i, i] => i" where + setsum :: "[i\i, i] \ i" where "setsum(g, C) \ if Finite(C) then - fold[int](%x y. g(x) $+ y, #0, C) else #0" + fold[int](\x y. g(x) $+ y, #0, C) else #0" (** foldSet **) -inductive_cases empty_fold_setE: "<0, x> \ fold_set(A, B, f,e)" +inductive_cases empty_fold_setE: "\0, x\ \ fold_set(A, B, f,e)" inductive_cases cons_fold_setE: " \ fold_set(A, B, f,e)" (* add-hoc lemmas *) @@ -44,8 +44,8 @@ (* fold_set monotonicity *) lemma fold_set_mono_lemma: - " \ fold_set(A, B, f, e) - \ \D. A<=D \ \ fold_set(D, B, f, e)" + "\C, x\ \ fold_set(A, B, f, e) + \ \D. A<=D \ \C, x\ \ fold_set(D, B, f, e)" apply (erule fold_set.induct) apply (auto intro: fold_set.intros) done @@ -57,7 +57,7 @@ done lemma fold_set_lemma: - "\fold_set(A, B, f, e) \ \fold_set(C, B, f, e) \ C<=A" + "\C, x\\fold_set(A, B, f, e) \ \C, x\\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 @@ -79,7 +79,7 @@ lemma (in fold_typing) Fin_imp_fold_set: - "C\Fin(A) \ (\x. \ fold_set(A, B, f,e))" + "C\Fin(A) \ (\x. \C, x\ \ fold_set(A, B, f,e))" apply (erule Fin_induct) apply (auto dest: fold_set.dom_subset [THEN subsetD] intro: fold_set.intros etype ftype) @@ -92,8 +92,8 @@ lemma (in fold_typing) fold_set_determ_lemma [rule_format]: "n\nat \ \C. |C| - (\x. \ fold_set(A, B, f,e)\ - (\y. \ fold_set(A, B, f,e) \ y=x))" + (\x. \C, x\ \ fold_set(A, B, f,e)\ + (\y. \C, y\ \ fold_set(A, B, f,e) \ y=x))" apply (erule nat_induct) apply (auto simp add: le_iff) apply (erule fold_set.cases) @@ -131,8 +131,8 @@ done lemma (in fold_typing) fold_set_determ: - "\\fold_set(A, B, f, e); - \fold_set(A, B, f, e)\ \ y=x" + "\\C, x\\fold_set(A, B, f, e); + \C, y\\fold_set(A, B, f, e)\ \ y=x" apply (frule fold_set.dom_subset [THEN subsetD], clarify) apply (drule Fin_into_Finite) apply (unfold Finite_def, clarify) @@ -143,7 +143,7 @@ (** The fold function **) lemma (in fold_typing) fold_equality: - " \ fold_set(A,B,f,e) \ fold[B](f,e,C) = y" + "\C,y\ \ fold_set(A,B,f,e) \ fold[B](f,e,C) = y" apply (unfold fold_def) apply (frule fold_set.dom_subset [THEN subsetD], clarify) apply (rule the_equality) @@ -161,7 +161,7 @@ text\This result is the right-to-left direction of the subsequent result\ lemma (in fold_typing) fold_set_imp_cons: - "\ \ fold_set(C, B, f, e); C \ Fin(A); c \ A; c\C\ + "\\C, y\ \ fold_set(C, B, f, e); C \ Fin(A); c \ A; c\C\ \ \ fold_set(cons(c, C), B, f, e)" apply (frule FinD [THEN fold_set_mono, THEN subsetD]) apply assumption @@ -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. \C, 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+) @@ -251,7 +251,7 @@ apply (auto intro: fold_typing.intro Finite_cons_lemma) done -lemma setsum_K0: "setsum((%i. #0), C) = #0" +lemma setsum_K0: "setsum((\i. #0), C) = #0" apply (case_tac "Finite (C) ") prefer 2 apply (simp add: setsum_def) apply (erule Finite_induct, auto) @@ -290,7 +290,7 @@ "Finite(I) \ (\i\I. Finite(C(i))) \ (\i\I. \j\I. i\j \ C(i) \ C(j) = 0) \ - setsum(f, \i\I. C(i)) = setsum (%i. setsum(f, C(i)), I)" + setsum(f, \i\I. C(i)) = setsum (\i. setsum(f, C(i)), I)" apply (erule Finite_induct, auto) apply (subgoal_tac "\i\B. x \ i") prefer 2 apply blast @@ -302,7 +302,7 @@ done -lemma setsum_addf: "setsum(%x. f(x) $+ g(x),C) = setsum(f, C) $+ setsum(g, C)" +lemma setsum_addf: "setsum(\x. f(x) $+ g(x),C) = setsum(f, C) $+ setsum(g, C)" apply (case_tac "Finite (C) ") prefer 2 apply (simp add: setsum_def) apply (erule Finite_induct, auto) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/ListN.thy --- a/src/ZF/Induct/ListN.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/ListN.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,19 +12,19 @@ @{cite "paulin-tlca"}. \ -consts listn :: "i=>i" +consts listn :: "i\i" inductive domains "listn(A)" \ "nat \ list(A)" intros - NilI: "<0,Nil> \ listn(A)" - ConsI: "\a \ A; \ listn(A)\ \ \ listn(A)" + NilI: "\0,Nil\ \ listn(A)" + ConsI: "\a \ A; \n,l\ \ listn(A)\ \ \ listn(A)" type_intros nat_typechecks list.intros lemma list_into_listn: "l \ list(A) \ \ listn(A)" by (induct set: list) (simp_all add: listn.intros) -lemma listn_iff: " \ listn(A) \ l \ list(A) \ length(l)=n" +lemma listn_iff: "\n,l\ \ listn(A) \ l \ list(A) \ length(l)=n" apply (rule iffI) apply (erule listn.induct) apply auto @@ -44,18 +44,18 @@ done lemma listn_append: - "\ \ listn(A); \ listn(A)\ \ \ listn(A)" + "\\n,l\ \ listn(A); \ listn(A)\ \ \ listn(A)" apply (erule listn.induct) apply (frule listn.dom_subset [THEN subsetD]) apply (simp_all add: listn.intros) done inductive_cases - Nil_listn_case: " \ listn(A)" + Nil_listn_case: "\i,Nil\ \ listn(A)" and Cons_listn_case: " \ listn(A)" inductive_cases - zero_listn_case: "<0,l> \ listn(A)" + zero_listn_case: "\0,l\ \ listn(A)" and succ_listn_case: " \ listn(A)" end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Multiset.thy Tue Sep 27 17:46:52 2022 +0100 @@ -13,68 +13,68 @@ abbreviation (input) \ \Short cut for multiset space\ - Mult :: "i=>i" where + Mult :: "i\i" where "Mult(A) \ A -||> nat-{0}" definition (* This is the original "restrict" from ZF.thy. Restricts the function f to the domain A FIXME: adapt Multiset to the new "restrict". *) - funrestrict :: "[i,i] => i" where + funrestrict :: "[i,i] \ i" where "funrestrict(f,A) \ \x \ A. f`x" definition (* M is a multiset *) - multiset :: "i => o" where + multiset :: "i \ o" where "multiset(M) \ \A. M \ A -> nat-{0} \ Finite(A)" definition - mset_of :: "i=>i" where + mset_of :: "i\i" where "mset_of(M) \ domain(M)" definition - munion :: "[i, i] => i" (infixl \+#\ 65) where + munion :: "[i, i] \ i" (infixl \+#\ 65) where "M +# N \ \x \ mset_of(M) \ mset_of(N). if x \ mset_of(M) \ mset_of(N) then (M`x) #+ (N`x) else (if x \ mset_of(M) then M`x else N`x)" definition (*convert a function to a multiset by eliminating 0*) - normalize :: "i => i" where + normalize :: "i \ i" where "normalize(f) \ if (\A. f \ A -> nat \ Finite(A)) then funrestrict(f, {x \ mset_of(f). 0 < f`x}) else 0" definition - mdiff :: "[i, i] => i" (infixl \-#\ 65) where + mdiff :: "[i, i] \ i" (infixl \-#\ 65) where "M -# N \ normalize(\x \ mset_of(M). if x \ mset_of(N) then M`x #- N`x else M`x)" definition (* set of elements of a multiset *) - msingle :: "i => i" (\{#_#}\) where - "{#a#} \ {}" + msingle :: "i \ i" (\{#_#}\) where + "{#a#} \ {\a, 1\}" definition - MCollect :: "[i, i=>o] => i" (*comprehension*) where + MCollect :: "[i, i\o] \ i" (*comprehension*) where "MCollect(M, P) \ funrestrict(M, {x \ mset_of(M). P(x)})" definition (* Counts the number of occurrences of an element in a multiset *) - mcount :: "[i, i] => i" where + mcount :: "[i, i] \ i" where "mcount(M, a) \ if a \ mset_of(M) then M`a else 0" definition - msize :: "i => i" where - "msize(M) \ setsum(%a. $# mcount(M,a), mset_of(M))" + msize :: "i \ i" where + "msize(M) \ setsum(\a. $# mcount(M,a), mset_of(M))" abbreviation - melem :: "[i,i] => o" (\(_/ :# _)\ [50, 51] 50) where + melem :: "[i,i] \ o" (\(_/ :# _)\ [50, 51] 50) where "a :# M \ a \ mset_of(M)" syntax - "_MColl" :: "[pttrn, i, o] => i" (\(1{# _ \ _./ _#})\) + "_MColl" :: "[pttrn, i, o] \ i" (\(1{# _ \ _./ _#})\) translations "{#x \ M. P#}" == "CONST MCollect(M, \x. P)" @@ -83,28 +83,28 @@ definition (* multirel1 has to be a set (not a predicate) so that we can form its transitive closure and reason about wf(.) and acc(.) *) - multirel1 :: "[i,i]=>i" where + multirel1 :: "[i,i]\i" where "multirel1(A, r) \ - { \ Mult(A)*Mult(A). + {\M, N\ \ 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). \b,a\ \ r)}" definition - multirel :: "[i, i] => i" where + multirel :: "[i, i] \ i" where "multirel(A, r) \ multirel1(A, r)^+" (* ordinal multiset orderings *) definition - omultiset :: "i => o" where + omultiset :: "i \ o" where "omultiset(M) \ \i. Ord(i) \ M \ Mult(field(Memrel(i)))" definition - mless :: "[i, i] => o" (infixl \<#\ 50) where - "M <# N \ \i. Ord(i) \ \ multirel(field(Memrel(i)), Memrel(i))" + mless :: "[i, i] \ o" (infixl \<#\ 50) where + "M <# N \ \i. Ord(i) \ \M, N\ \ multirel(field(Memrel(i)), Memrel(i))" definition - mle :: "[i, i] => o" (infixl \<#=\ 50) where + mle :: "[i, i] \ o" (infixl \<#=\ 50) where "M <#= N \ (omultiset(M) \ M = N) | M <# N" @@ -366,9 +366,9 @@ apply (drule not_empty_multiset_imp_exist, assumption, clarify) apply (subgoal_tac "Finite (mset_of (M) - {a}) ") prefer 2 apply (simp add: Finite_Diff) -apply (subgoal_tac "setsum (%x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0") +apply (subgoal_tac "setsum (\x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0") prefer 2 apply (simp add: cons_Diff, simp) -apply (subgoal_tac "#0 $\ setsum (%x. $# mcount (M, x), mset_of (M) - {a}) ") +apply (subgoal_tac "#0 $\ setsum (\x. $# mcount (M, x), mset_of (M) - {a}) ") apply (rule_tac [2] g_zpos_imp_setsum_zpos) apply (auto simp add: Finite_Diff not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym]) apply (rule not_zneg_int_of [THEN bexE]) @@ -376,8 +376,8 @@ done lemma setsum_mcount_Int: - "Finite(A) \ setsum(%a. $# mcount(N, a), A \ mset_of(N)) - = setsum(%a. $# mcount(N, a), A)" + "Finite(A) \ setsum(\a. $# mcount(N, a), A \ mset_of(N)) + = setsum(\a. $# mcount(N, a), A)" apply (induct rule: Finite_induct) apply auto apply (subgoal_tac "Finite (B \ mset_of (N))") @@ -477,9 +477,9 @@ lemma setsum_decr: "Finite(A) \ (\M. multiset(M) \ - (\a \ mset_of(M). setsum(%z. $# mcount(M(a:=M`a #- 1), z), A) = - (if a \ A then setsum(%z. $# mcount(M, z), A) $- #1 - else setsum(%z. $# mcount(M, z), A))))" + (\a \ mset_of(M). setsum(\z. $# mcount(M(a:=M`a #- 1), z), A) = + (if a \ A then setsum(\z. $# mcount(M, z), A) $- #1 + else setsum(\z. $# mcount(M, z), A))))" apply (unfold multiset_def) apply (erule Finite_induct) apply (auto simp add: multiset_fun_iff) @@ -493,19 +493,19 @@ lemma setsum_decr2: "Finite(A) \ \M. multiset(M) \ (\a \ mset_of(M). - setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) = - (if a \ A then setsum(%x. $# mcount(M, x), A) $- $# M`a - else setsum(%x. $# mcount(M, x), A)))" + setsum(\x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) = + (if a \ A then setsum(\x. $# mcount(M, x), A) $- $# M`a + else setsum(\x. $# mcount(M, x), A)))" apply (simp add: multiset_def) apply (erule Finite_induct) apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) done lemma setsum_decr3: "\Finite(A); multiset(M); a \ mset_of(M)\ - \ setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) = - (if a \ A then setsum(%x. $# mcount(M, x), A) $- $# M`a - else setsum(%x. $# mcount(M, x), A))" -apply (subgoal_tac "setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ") + \ setsum(\x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) = + (if a \ A then setsum(\x. $# mcount(M, x), A) $- $# M`a + else setsum(\x. $# mcount(M, x), A))" +apply (subgoal_tac "setsum (\x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (\x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ") apply (rule_tac [2] setsum_Diff [symmetric]) apply (rule sym, rule ssubst, blast) apply (rule sym, drule setsum_decr2, auto) @@ -530,22 +530,22 @@ done lemma multiset_induct_aux: - assumes prem1: "\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(, M))" + assumes prem1: "\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(\a, 1\, M))" and prem2: "\M b. \multiset(M); b \ mset_of(M); P(M)\ \ P(M(b:= M`b #+ 1))" shows "\n \ nat; P(0)\ \ (\M. multiset(M)\ - (setsum(%x. $# mcount(M, x), {x \ mset_of(M). 0 < M`x}) = $# n) \ P(M))" + (setsum(\x. $# mcount(M, x), {x \ mset_of(M). 0 < M`x}) = $# n) \ P(M))" apply (erule nat_induct, clarify) apply (frule msize_eq_0_iff) apply (auto simp add: mset_of_def multiset_def multiset_fun_iff msize_def) -apply (subgoal_tac "setsum (%x. $# mcount (M, x), A) =$# succ (x) ") +apply (subgoal_tac "setsum (\x. $# mcount (M, x), A) =$# succ (x) ") apply (drule setsum_succD, auto) apply (case_tac "1 multiset(M); P(0); - (\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(, M))); + (\M a. \multiset(M); a\mset_of(M); P(M)\ \ P(cons(\a, 1\, M))); (\M b. \multiset(M); b \ mset_of(M); P(M)\ \ P(M(b:= M`b #+ 1)))\ \ P(M)" apply (subgoal_tac "\n \ nat. setsum (\x. $# mcount (M, x), {x \ mset_of (M) . 0 < M ` x}) = $# n") @@ -614,7 +614,7 @@ done lemma munion_single_case1: - "\multiset(M); a \mset_of(M)\ \ M +# {#a#} = cons(, M)" + "\multiset(M); a \mset_of(M)\ \ M +# {#a#} = cons(\a, 1\, M)" apply (simp add: multiset_def msingle_def) apply (auto simp add: munion_def) apply (unfold mset_of_def, simp) @@ -726,10 +726,10 @@ by (auto simp add: multirel1_def) lemma multirel1_iff: -" \ multirel1(A, r) \ +" \N, M\ \ multirel1(A, r) \ (\a. a \ A \ (\M0. M0 \ Mult(A) \ (\K. K \ Mult(A) \ - M=M0 +# {#a#} \ N=M0 +# K \ (\b \ mset_of(K). \ r))))" + M=M0 +# {#a#} \ N=M0 +# K \ (\b \ mset_of(K). \b,a\ \ r))))" by (auto simp add: multirel1_def Mult_iff_multiset Bex_def) @@ -762,12 +762,12 @@ subsection\Toward the proof of well-foundedness of multirel1\ -lemma not_less_0 [iff]: " \ multirel1(A, r)" +lemma not_less_0 [iff]: "\M,0\ \ multirel1(A, r)" by (auto simp add: multirel1_def Mult_iff_multiset) lemma less_munion: "\ \ multirel1(A, r); M0 \ Mult(A)\ \ - (\M. \ multirel1(A, r) \ N = M +# {#a#}) | - (\K. K \ Mult(A) \ (\b \ mset_of(K). \ r) \ N = M0 +# K)" + (\M. \M, M0\ \ multirel1(A, r) \ N = M +# {#a#}) | + (\K. K \ Mult(A) \ (\b \ mset_of(K). \b, a\ \ r) \ N = M0 +# K)" apply (frule multirel1_type [THEN subsetD]) apply (simp add: multirel1_iff) apply (auto simp add: munion_eq_conv_exist) @@ -787,10 +787,10 @@ lemma acc_0: "acc(0)=0" by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD]) -lemma lemma1: "\\b \ A. \ r \ +lemma lemma1: "\\b \ A. \b,a\ \ r \ (\M \ acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r))); M0 \ acc(multirel1(A, r)); a \ A; - \M. \ multirel1(A, r) \ M +# {#a#} \ acc(multirel1(A, r))\ + \M. \M,M0\ \ multirel1(A, r) \ M +# {#a#} \ acc(multirel1(A, r))\ \ M0 +# {#a#} \ acc(multirel1(A, r))" apply (subgoal_tac "M0 \ Mult(A) ") prefer 2 @@ -801,7 +801,7 @@ apply (rename_tac "N") apply (drule less_munion, blast) apply (auto simp add: Mult_iff_multiset) -apply (erule_tac P = "\x \ mset_of (K) . \ r" in rev_mp) +apply (erule_tac P = "\x \ mset_of (K) . \x, a\ \ r" in rev_mp) apply (erule_tac P = "mset_of (K) \A" in rev_mp) apply (erule_tac M = K in multiset_induct) (* three subgoals *) @@ -813,13 +813,13 @@ apply (subgoal_tac "aa \ A") prefer 2 apply blast apply (drule_tac x = "M0 +# M" and P = - "%x. x \ acc(multirel1(A, r)) \ Q(x)" for Q in spec) + "\x. x \ acc(multirel1(A, r)) \ Q(x)" for Q in spec) apply (simp add: munion_assoc [symmetric]) (* subgoal 3 \ additional conditions *) apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset) done -lemma lemma2: "\\b \ A. \ r +lemma lemma2: "\\b \ A. \b,a\ \ r \ (\M \ acc(multirel1(A, r)). M +# {#b#} :acc(multirel1(A, r))); M \ acc(multirel1(A, r)); a \ A\ \ M +# {#a#} \ acc(multirel1(A, r))" apply (erule acc_induct) @@ -914,12 +914,12 @@ (* One direction *) lemma multirel_implies_one_step: -" \ multirel(A, r) \ +"\M,N\ \ multirel(A, r) \ trans[A](r) \ (\I J K. I \ Mult(A) \ J \ Mult(A) \ K \ Mult(A) \ N = I +# J \ M = I +# K \ J \ 0 \ - (\k \ mset_of(K). \j \ mset_of(J). \ r))" + (\k \ mset_of(K). \j \ mset_of(J). \k,j\ \ r))" apply (simp add: multirel_def Ball_def Bex_def) apply (erule converse_trancl_induct) apply (simp_all add: multirel1_iff Mult_iff_multiset) @@ -936,11 +936,11 @@ apply (rule_tac x = " (Ka -# {#a#}) +# K" in exI, simp (no_asm_simp)) apply (simp_all add: Un_subset_iff) apply (simp (no_asm_simp) add: munion_assoc [symmetric]) -apply (drule_tac t = "%M. M-#{#a#}" in subst_context) +apply (drule_tac t = "\M. M-#{#a#}" in subst_context) apply (simp add: mdiff_union_single_conv melem_diff_single, clarify) apply (erule disjE, simp) apply (erule disjE, simp) -apply (drule_tac x = a and P = "%x. x :# Ka \ Q(x)" for Q in spec) +apply (drule_tac x = a and P = "\x. x :# Ka \ Q(x)" for Q in spec) apply clarify apply (rule_tac x = xa in exI) apply (simp (no_asm_simp)) @@ -955,7 +955,7 @@ apply (rule conjI) apply (simp (no_asm_simp) add: multiset_equality mcount_elem [THEN succ_pred_eq_self]) apply (rule conjI) -apply (drule_tac t = "%M. M-#{#a#}" in subst_context) +apply (drule_tac t = "\M. M-#{#a#}" in subst_context) apply (simp add: mdiff_union_inverse2) apply (simp_all (no_asm_simp) add: multiset_equality) apply (rule diff_add_commute [symmetric]) @@ -984,7 +984,7 @@ "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)) + (msize(J) = $# n \ J \0 \ (\k \ mset_of(K). \j \ mset_of(J). \k, j\ \ r)) \ \ multirel(A, r))" apply (simp add: Mult_iff_multiset) apply (erule nat_induct, clarify) @@ -1002,19 +1002,19 @@ (*Now we know J' \ 0*) apply (drule sym, rotate_tac -1, simp) apply (erule_tac V = "$# x = msize (J') " in thin_rl) -apply (frule_tac M = K and P = "%x. \ r" in multiset_partition) +apply (frule_tac M = K and P = "\x. \x,a\ \ r" in multiset_partition) apply (erule_tac P = "\k \ mset_of (K) . P(k)" for P in rev_mp) apply (erule ssubst) apply (simp add: Ball_def, auto) -apply (subgoal_tac "< (I +# {# x \ K. \ r#}) +# {# x \ K. \ r#}, (I +# {# x \ K. \ r#}) +# J'> \ multirel(A, r) ") +apply (subgoal_tac "< (I +# {# x \ K. \x, a\ \ r#}) +# {# x \ K. \x, a\ \ r#}, (I +# {# x \ K. \x, a\ \ r#}) +# J'> \ multirel(A, r) ") prefer 2 - apply (drule_tac x = "I +# {# x \ K. \ r#}" in spec) + apply (drule_tac x = "I +# {# x \ K. \x, a\ \ r#}" in spec) apply (rotate_tac -1) apply (drule_tac x = "J'" in spec) apply (rotate_tac -1) - apply (drule_tac x = "{# x \ K. \ r#}" in spec, simp) apply blast + apply (drule_tac x = "{# x \ K. \x, a\ \ r#}" in spec, simp) apply blast apply (simp add: munion_assoc [symmetric] multirel_def) -apply (rule_tac b = "I +# {# x \ K. \ r#} +# J'" in trancl_trans, blast) +apply (rule_tac b = "I +# {# x \ K. \x, a\ \ r#} +# J'" in trancl_trans, blast) apply (rule r_into_trancl) apply (simp add: multirel1_iff Mult_iff_multiset) apply (rule_tac x = a in exI) @@ -1024,7 +1024,7 @@ done lemma one_step_implies_multirel: - "\J \ 0; \k \ mset_of(K). \j \ mset_of(J). \ r; + "\J \ 0; \k \ mset_of(K). \j \ mset_of(J). \k,j\ \ r; I \ Mult(A); J \ Mult(A); K \ Mult(A)\ \ \ multirel(A, r)" apply (subgoal_tac "multiset (J) ") @@ -1038,7 +1038,7 @@ (*irreflexivity*) lemma multirel_irrefl_lemma: - "Finite(A) \ part_ord(A, r) \ (\x \ A. \y \ A. \ r) \A=0" + "Finite(A) \ part_ord(A, r) \ (\x \ A. \y \ A. \x,y\ \ r) \A=0" apply (erule Finite_induct) apply (auto dest: subset_consI [THEN [2] part_ord_subset]) apply (auto simp add: part_ord_def irrefl_def) @@ -1066,7 +1066,7 @@ done lemma multirel_trans: - "\ \ multirel(A, r); \ multirel(A, r)\ \ \ multirel(A,r)" + "\\M, N\ \ multirel(A, r); \N, K\ \ multirel(A, r)\ \ \M, K\ \ multirel(A,r)" apply (simp add: multirel_def) apply (blast intro: trancl_trans) done @@ -1084,7 +1084,7 @@ (** Monotonicity of multiset union **) lemma munion_multirel1_mono: -"\ \ multirel1(A, r); K \ Mult(A)\ \ \ multirel1(A, r)" +"\\M,N\ \ multirel1(A, r); K \ Mult(A)\ \ \ multirel1(A, r)" apply (frule multirel1_type [THEN subsetD]) apply (auto simp add: multirel1_iff Mult_iff_multiset) apply (rule_tac x = a in exI) @@ -1096,11 +1096,11 @@ done lemma munion_multirel_mono2: - "\ \ multirel(A, r); K \ Mult(A)\\ \ multirel(A, r)" + "\\M, N\ \ multirel(A, r); K \ Mult(A)\\ \ multirel(A, r)" apply (frule multirel_type [THEN subsetD]) apply (simp (no_asm_use) add: multirel_def) apply clarify -apply (drule_tac psi = " \ multirel1 (A, r) ^+" in asm_rl) +apply (drule_tac psi = "\M,N\ \ multirel1 (A, r) ^+" in asm_rl) apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) @@ -1115,16 +1115,16 @@ done lemma munion_multirel_mono1: - "\ \ multirel(A, r); K \ Mult(A)\ \ \ multirel(A, r)" + "\\M, N\ \ multirel(A, r); K \ Mult(A)\ \ \ multirel(A, r)" apply (frule multirel_type [THEN subsetD]) -apply (rule_tac P = "%x. \ multirel(A, r)" for u in munion_commute [THEN subst]) +apply (rule_tac P = "\x. \x,u\ \ multirel(A, r)" for u in munion_commute [THEN subst]) apply (subst munion_commute [of N]) apply (rule munion_multirel_mono2) apply (auto simp add: Mult_iff_multiset) done lemma munion_multirel_mono: - "\ \ multirel(A, r); \ multirel(A, r)\ + "\\M,K\ \ multirel(A, r); \N,L\ \ multirel(A, r)\ \ \ multirel(A, r)" apply (subgoal_tac "M \ Mult(A) \ N \ Mult(A) \ K \ Mult(A) \ L \ Mult(A) ") prefer 2 apply (blast dest: multirel_type [THEN subsetD]) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Mutil.thy --- a/src/ZF/Induct/Mutil.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Mutil.thy Tue Sep 27 17:46:52 2022 +0100 @@ -14,13 +14,13 @@ consts domino :: i - tiling :: "i => i" + tiling :: "i \ i" inductive domains "domino" \ "Pow(nat \ nat)" intros - horiz: "\i \ nat; j \ nat\ \ {, } \ domino" - vertl: "\i \ nat; j \ nat\ \ {, } \ domino" + horiz: "\i \ nat; j \ nat\ \ {\i,j\, } \ domino" + vertl: "\i \ nat; j \ nat\ \ {\i,j\, } \ domino" type_intros empty_subsetI cons_subsetI PowI SigmaI nat_succI inductive @@ -32,13 +32,13 @@ type_elims PowD [elim_format] definition - evnodd :: "[i, i] => i" where - "evnodd(A,b) \ {z \ A. \i j. z = \ (i #+ j) mod 2 = b}" + evnodd :: "[i, i] \ i" where + "evnodd(A,b) \ {z \ A. \i j. z = \i,j\ \ (i #+ j) mod 2 = b}" subsection \Basic properties of evnodd\ -lemma evnodd_iff: ": evnodd(A,b) \ : A \ (i#+j) mod 2 = b" +lemma evnodd_iff: "\i,j\: evnodd(A,b) \ \i,j\: A \ (i#+j) mod 2 = b" by (unfold evnodd_def) blast lemma evnodd_subset: "evnodd(A, b) \ A" @@ -54,8 +54,8 @@ by (simp add: evnodd_def Collect_Diff) lemma evnodd_cons [simp]: - "evnodd(cons(,C), b) = - (if (i#+j) mod 2 = b then cons(, evnodd(C,b)) else evnodd(C,b))" + "evnodd(cons(\i,j\,C), b) = + (if (i#+j) mod 2 = b then cons(\i,j\, evnodd(C,b)) else evnodd(C,b))" by (simp add: evnodd_def Collect_cons) lemma evnodd_0 [simp]: "evnodd(0, b) = 0" @@ -143,7 +143,7 @@ theorem mutil_not_tiling: "\m \ nat; n \ nat; t = (succ(m)#+succ(m))*(succ(n)#+succ(n)); - t' = t - {<0,0>} - {}\ + t' = t - {\0,0\} - {}\ \ t' \ tiling(domino)" apply (rule notI) apply (drule tiling_domino_0_1) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Ntree.thy --- a/src/ZF/Induct/Ntree.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Ntree.thy Tue Sep 27 17:46:52 2022 +0100 @@ -13,9 +13,9 @@ \ consts - ntree :: "i => i" - maptree :: "i => i" - maptree2 :: "[i, i] => i" + ntree :: "i \ i" + maptree :: "i \ i" + maptree2 :: "[i, i] \ i" datatype "ntree(A)" = Branch ("a \ A", "h \ (\n \ nat. n -> ntree(A))") monos UN_mono [OF subset_refl Pi_mono] \ \MUST have this form\ @@ -31,12 +31,12 @@ type_intros FiniteFun_in_univ' definition - ntree_rec :: "[[i, i, i] => i, i] => i" where + ntree_rec :: "[[i, i, i] \ i, i] \ i" where "ntree_rec(b) \ Vrecursor(\pr. ntree_case(\x h. b(x, h, \i \ domain(h). pr`(h`i))))" definition - ntree_copy :: "i => i" where + ntree_copy :: "i \ i" where "ntree_copy(z) \ ntree_rec(\x h r. Branch(x,r), z)" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Primrec.thy --- a/src/ZF/Induct/Primrec.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Primrec.thy Tue Sep 27 17:46:52 2022 +0100 @@ -21,32 +21,32 @@ "SC \ \l \ list(nat). list_case(0, \x xs. succ(x), l)" definition - CONSTANT :: "i=>i" where + CONSTANT :: "i\i" where "CONSTANT(k) \ \l \ list(nat). k" definition - PROJ :: "i=>i" where + PROJ :: "i\i" where "PROJ(i) \ \l \ list(nat). list_case(0, \x xs. x, drop(i,l))" definition - COMP :: "[i,i]=>i" where + COMP :: "[i,i]\i" where "COMP(g,fs) \ \l \ list(nat). g ` map(\f. f`l, fs)" definition - PREC :: "[i,i]=>i" where + PREC :: "[i,i]\i" where "PREC(f,g) \ \l \ list(nat). list_case(0, \x xs. rec(x, f`xs, \y r. g ` Cons(r, Cons(y, xs))), l)" \ \Note that \g\ is applied first to \<^term>\PREC(f,g)`y\ and then to \y\!\ consts - ACK :: "i=>i" + ACK :: "i\i" primrec "ACK(0) = SC" "ACK(succ(i)) = PREC (CONSTANT (ACK(i) ` [1]), COMP(ACK(i), [PROJ(0)]))" abbreviation - ack :: "[i,i]=>i" where + ack :: "[i,i]\i" where "ack(x,y) \ ACK(x) ` [y]" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/PropLog.thy --- a/src/ZF/Induct/PropLog.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/PropLog.thy Tue Sep 27 17:46:52 2022 +0100 @@ -27,26 +27,26 @@ datatype propn = Fls | Var ("n \ nat") (\#_\ [100] 100) - | Imp ("p \ propn", "q \ propn") (infixr \=>\ 90) + | Imp ("p \ propn", "q \ propn") (infixr \\\ 90) subsection \The proof system\ -consts thms :: "i => i" +consts thms :: "i \ i" abbreviation - thms_syntax :: "[i,i] => o" (infixl \|-\ 50) + thms_syntax :: "[i,i] \ o" (infixl \|-\ 50) where "H |- p \ p \ thms(H)" inductive domains "thms(H)" \ "propn" intros H: "\p \ H; p \ propn\ \ H |- p" - K: "\p \ propn; q \ propn\ \ H |- p=>q=>p" + K: "\p \ propn; q \ propn\ \ H |- p\q\p" S: "\p \ propn; q \ propn; r \ propn\ - \ H |- (p=>q=>r) => (p=>q) => p=>r" - DN: "p \ propn \ H |- ((p=>Fls) => Fls) => p" - MP: "\H |- p=>q; H |- p; p \ propn; q \ propn\ \ H |- q" + \ H |- (p\q\r) \ (p\q) \ p\r" + DN: "p \ propn \ H |- ((p\Fls) \ Fls) \ p" + MP: "\H |- p\q; H |- p; p \ propn; q \ propn\ \ H |- q" type_intros "propn.intros" declare propn.intros [simp] @@ -57,14 +57,14 @@ subsubsection \Semantics of propositional logic.\ consts - is_true_fun :: "[i,i] => i" + is_true_fun :: "[i,i] \ i" primrec "is_true_fun(Fls, t) = 0" "is_true_fun(Var(v), t) = (if v \ t then 1 else 0)" - "is_true_fun(p=>q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)" + "is_true_fun(p\q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)" definition - is_true :: "[i,i] => o" where + is_true :: "[i,i] \ o" where "is_true(p,t) \ is_true_fun(p,t) = 1" \ \this definition is required since predicates can't be recursive\ @@ -74,7 +74,7 @@ lemma is_true_Var [simp]: "is_true(#v,t) \ v \ t" by (simp add: is_true_def) -lemma is_true_Imp [simp]: "is_true(p=>q,t) \ (is_true(p,t)\is_true(q,t))" +lemma is_true_Imp [simp]: "is_true(p\q,t) \ (is_true(p,t)\is_true(q,t))" by (simp add: is_true_def) @@ -86,7 +86,7 @@ \ definition - logcon :: "[i,i] => o" (infixl \|=\ 50) where + logcon :: "[i,i] \ o" (infixl \|=\ 50) where "H |= p \ \t. (\q \ H. is_true(q,t)) \ is_true(p,t)" @@ -96,11 +96,11 @@ \ consts - hyps :: "[i,i] => i" + hyps :: "[i,i] \ i" primrec "hyps(Fls, t) = 0" - "hyps(Var(v), t) = (if v \ t then {#v} else {#v=>Fls})" - "hyps(p=>q, t) = hyps(p,t) \ hyps(q,t)" + "hyps(Var(v), t) = (if v \ t then {#v} else {#v\Fls})" + "hyps(p\q, t) = hyps(p,t) \ hyps(q,t)" @@ -115,15 +115,15 @@ lemmas thms_in_pl = thms.dom_subset [THEN subsetD] -inductive_cases ImpE: "p=>q \ propn" +inductive_cases ImpE: "p\q \ propn" -lemma thms_MP: "\H |- p=>q; H |- p\ \ H |- q" +lemma thms_MP: "\H |- p\q; H |- p\ \ H |- q" \ \Stronger Modus Ponens rule: no typechecking!\ apply (rule thms.MP) apply (erule asm_rl thms_in_pl thms_in_pl [THEN ImpE])+ done -lemma thms_I: "p \ propn \ H |- p=>p" +lemma thms_I: "p \ propn \ H |- p\p" \ \Rule is called \I\ for Identity Combinator, not for Introduction.\ apply (rule thms.S [THEN thms_MP, THEN thms_MP]) apply (rule_tac [5] thms.K) @@ -144,13 +144,13 @@ lemmas weaken_left_Un1 = Un_upper1 [THEN weaken_left] lemmas weaken_left_Un2 = Un_upper2 [THEN weaken_left] -lemma weaken_right: "\H |- q; p \ propn\ \ H |- p=>q" +lemma weaken_right: "\H |- q; p \ propn\ \ H |- p\q" by (simp_all add: thms.K [THEN thms_MP] thms_in_pl) subsubsection \The deduction theorem\ -theorem deduction: "\cons(p,H) |- q; p \ propn\ \ H |- p=>q" +theorem deduction: "\cons(p,H) |- q; p \ propn\ \ H |- p\q" apply (erule thms.induct) apply (blast intro: thms_I thms.H [THEN weaken_right]) apply (blast intro: thms.K [THEN weaken_right]) @@ -173,7 +173,7 @@ apply (simp_all add: propn.intros) done -lemma thms_notE: "\H |- p=>Fls; H |- p; q \ propn\ \ H |- q" +lemma thms_notE: "\H |- p\Fls; H |- p; q \ propn\ \ H |- q" by (erule thms_MP [THEN thms_FlsE]) @@ -190,16 +190,16 @@ subsubsection \Towards the completeness proof\ -lemma Fls_Imp: "\H |- p=>Fls; q \ propn\ \ H |- p=>q" +lemma Fls_Imp: "\H |- p\Fls; q \ propn\ \ H |- p\q" apply (frule thms_in_pl) apply (rule deduction) apply (rule weaken_left_cons [THEN thms_notE]) apply (blast intro: thms.H elim: ImpE)+ done -lemma Imp_Fls: "\H |- p; H |- q=>Fls\ \ H |- (p=>q)=>Fls" +lemma Imp_Fls: "\H |- p; H |- q\Fls\ \ H |- (p\q)\Fls" apply (frule thms_in_pl) - apply (frule thms_in_pl [of concl: "q=>Fls"]) + apply (frule thms_in_pl [of concl: "q\Fls"]) apply (rule deduction) apply (erule weaken_left_cons [THEN thms_MP]) apply (rule consI1 [THEN thms.H, THEN thms_MP]) @@ -207,7 +207,7 @@ done lemma hyps_thms_if: - "p \ propn \ hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)" + "p \ propn \ hyps(p,t) |- (if is_true(p,t) then p else p\Fls)" \ \Typical example of strengthening the induction statement.\ apply simp apply (induct_tac p) @@ -234,14 +234,14 @@ \ lemma thms_excluded_middle: - "\p \ propn; q \ propn\ \ H |- (p=>q) => ((p=>Fls)=>q) => q" + "\p \ propn; q \ propn\ \ H |- (p\q) \ ((p\Fls)\q) \ q" apply (rule deduction [THEN deduction]) apply (rule thms.DN [THEN thms_MP]) apply (best intro!: propn_SIs intro: propn_Is)+ done lemma thms_excluded_middle_rule: - "\cons(p,H) |- q; cons(p=>Fls,H) |- q; p \ propn\ \ H |- q" + "\cons(p,H) |- q; cons(p\Fls,H) |- q; p \ propn\ \ H |- q" \ \Hard to prove directly because it requires cuts\ apply (rule thms_excluded_middle [THEN thms_MP, THEN thms_MP]) apply (blast intro!: propn_SIs intro: propn_Is)+ @@ -255,16 +255,16 @@ \ lemma hyps_Diff: - "p \ propn \ hyps(p, t-{v}) \ cons(#v=>Fls, hyps(p,t)-{#v})" + "p \ propn \ hyps(p, t-{v}) \ cons(#v\Fls, hyps(p,t)-{#v})" by (induct set: propn) auto text \ - For the case \<^prop>\hyps(p,t)-cons(#v => Fls,Y) |- p\ we also have - \<^prop>\hyps(p,t)-{#v=>Fls} \ hyps(p, cons(v,t))\. + For the case \<^prop>\hyps(p,t)-cons(#v \ Fls,Y) |- p\ we also have + \<^prop>\hyps(p,t)-{#v\Fls} \ hyps(p, cons(v,t))\. \ lemma hyps_cons: - "p \ propn \ hyps(p, cons(v,t)) \ cons(#v, hyps(p,t)-{#v=>Fls})" + "p \ propn \ hyps(p, cons(v,t)) \ cons(#v, hyps(p,t)-{#v\Fls})" by (induct set: propn) auto text \Two lemmas for use with \weaken_left\\ @@ -277,11 +277,11 @@ text \ The set \<^term>\hyps(p,t)\ is finite, and elements have the form - \<^term>\#v\ or \<^term>\#v=>Fls\; could probably prove the stronger + \<^term>\#v\ or \<^term>\#v\Fls\; could probably prove the stronger \<^prop>\hyps(p,t) \ Fin(hyps(p,0) \ hyps(p,nat))\. \ -lemma hyps_finite: "p \ propn \ hyps(p,t) \ Fin(\v \ nat. {#v, #v=>Fls})" +lemma hyps_finite: "p \ propn \ hyps(p,t) \ Fin(\v \ nat. {#v, #v\Fls})" by (induct set: propn) auto lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left] @@ -304,7 +304,7 @@ apply (blast intro: cons_Diff_same [THEN weaken_left]) apply (blast intro: cons_Diff_subset2 [THEN weaken_left] hyps_Diff [THEN Diff_weaken_left]) - txt \Case \<^prop>\hyps(p,t)-cons(#v => Fls,Y) |- p\\ + txt \Case \<^prop>\hyps(p,t)-cons(#v \ Fls,Y) |- p\\ apply (rule thms_excluded_middle_rule) apply (erule_tac [3] propn.intros) apply (blast intro: cons_Diff_subset2 [THEN weaken_left] @@ -321,7 +321,7 @@ apply (blast intro: completeness_0_lemma) done -lemma logcon_Imp: "\cons(p,H) |= q\ \ H |= p=>q" +lemma logcon_Imp: "\cons(p,H) |= q\ \ H |= p\q" \ \A semantic analogue of the Deduction Theorem\ by (simp add: logcon_def) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Rmap.thy --- a/src/ZF/Induct/Rmap.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Rmap.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,14 +8,14 @@ theory Rmap imports ZF begin consts - rmap :: "i=>i" + rmap :: "i\i" inductive domains "rmap(r)" \ "list(domain(r)) \ list(range(r))" intros - NilI: " \ rmap(r)" + NilI: "\Nil,Nil\ \ rmap(r)" - ConsI: "\: r; \ rmap(r)\ + ConsI: "\\x,y\: r; \xs,ys\ \ rmap(r)\ \ \ rmap(r)" type_intros domainI rangeI list.intros @@ -28,7 +28,7 @@ done inductive_cases - Nil_rmap_case [elim!]: " \ rmap(r)" + Nil_rmap_case [elim!]: "\Nil,zs\ \ rmap(r)" and Cons_rmap_case [elim!]: " \ rmap(r)" declare rmap.intros [intro] diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Term.thy --- a/src/ZF/Induct/Term.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Term.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,7 +12,7 @@ \ consts - "term" :: "i => i" + "term" :: "i \ i" datatype "term(A)" = Apply ("a \ A", "l \ list(term(A))") monos list_mono @@ -21,28 +21,28 @@ declare Apply [TC] definition - term_rec :: "[i, [i, i, i] => i] => i" where + term_rec :: "[i, [i, i, i] \ i] \ i" where "term_rec(t,d) \ Vrec(t, \t g. term_case(\x zs. d(x, zs, map(\z. g`z, zs)), t))" definition - term_map :: "[i => i, i] => i" where + term_map :: "[i \ i, i] \ i" where "term_map(f,t) \ term_rec(t, \x zs rs. Apply(f(x), rs))" definition - term_size :: "i => i" where + term_size :: "i \ i" where "term_size(t) \ term_rec(t, \x zs rs. succ(list_add(rs)))" definition - reflect :: "i => i" where + reflect :: "i \ i" where "reflect(t) \ term_rec(t, \x zs rs. Apply(x, rev(rs)))" definition - preorder :: "i => i" where + preorder :: "i \ i" where "preorder(t) \ term_rec(t, \x zs rs. Cons(x, flat(rs)))" definition - postorder :: "i => i" where + postorder :: "i \ i" where "postorder(t) \ term_rec(t, \x zs rs. flat(rs) @ [x])" lemma term_unfold: "term(A) = A * list(term(A))" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Induct/Tree_Forest.thy --- a/src/ZF/Induct/Tree_Forest.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Induct/Tree_Forest.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,9 +10,9 @@ subsection \Datatype definition\ consts - tree :: "i => i" - forest :: "i => i" - tree_forest :: "i => i" + tree :: "i \ i" + forest :: "i \ i" + tree_forest :: "i \ i" datatype "tree(A)" = Tcons ("a \ A", "f \ forest(A)") and "forest(A)" = Fnil | Fcons ("t \ tree(A)", "f \ forest(A)") @@ -118,12 +118,12 @@ subsection \Operations\ consts - map :: "[i => i, i] => i" - size :: "i => i" - preorder :: "i => i" - list_of_TF :: "i => i" - of_list :: "i => i" - reflect :: "i => i" + map :: "[i \ i, i] \ i" + size :: "i \ i" + preorder :: "i \ i" + list_of_TF :: "i \ i" + of_list :: "i \ i" + reflect :: "i \ i" primrec "list_of_TF (Tcons(x,f)) = [Tcons(x,f)]" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Int.thy --- a/src/ZF/Int.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Int.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,84 +10,84 @@ 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,y1\,\x2,y2\> \ x1#+y2 = x2#+y1}" definition int :: i where "int \ (nat*nat)//intrel" definition - int_of :: "i=>i" \ \coercion from nat to int\ (\$# _\ [80] 80) where + int_of :: "i\i" \ \coercion from nat to int\ (\$# _\ [80] 80) where "$# m \ intrel `` {}" definition - intify :: "i=>i" \ \coercion from ANYTHING to int\ where + intify :: "i\i" \ \coercion from ANYTHING to int\ where "intify(m) \ if m \ int then m else $#0" definition - raw_zminus :: "i=>i" where - "raw_zminus(z) \ \\z. intrel``{}" + raw_zminus :: "i\i" where + "raw_zminus(z) \ \\x,y\\z. intrel``{\y,x\}" definition - zminus :: "i=>i" (\$- _\ [80] 80) where + zminus :: "i\i" (\$- _\ [80] 80) where "$- z \ raw_zminus (intify(z))" definition - znegative :: "i=>o" where - "znegative(z) \ \x y. x y\nat \ \z" + znegative :: "i\o" where + "znegative(z) \ \x y. x y\nat \ \x,y\\z" definition - iszero :: "i=>o" where + iszero :: "i\o" where "iszero(z) \ z = $# 0" definition - raw_nat_of :: "i=>i" where - "raw_nat_of(z) \ natify (\\z. x#-y)" + raw_nat_of :: "i\i" where + "raw_nat_of(z) \ natify (\\x,y\\z. x#-y)" definition - nat_of :: "i=>i" where + nat_of :: "i\i" where "nat_of(z) \ raw_nat_of (intify(z))" definition - zmagnitude :: "i=>i" where + 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))" definition - raw_zmult :: "[i,i]=>i" where - (*Cannot use UN here or in zadd because of the form of congruent2. + raw_zmult :: "[i,i]\i" where + (*Cannot use UN\x1,y2\ here or in zadd because of the form of congruent2. Perhaps a "curried" or even polymorphic congruent predicate would be better.*) "raw_zmult(z1,z2) \ - \p1\z1. \p2\z2. split(%x1 y1. split(%x2 y2. + \p1\z1. \p2\z2. split(\x1 y1. split(\x2 y2. intrel``{}, p2), p1)" definition - zmult :: "[i,i]=>i" (infixl \$*\ 70) where + zmult :: "[i,i]\i" (infixl \$*\ 70) where "z1 $* z2 \ raw_zmult (intify(z1),intify(z2))" definition - raw_zadd :: "[i,i]=>i" where + raw_zadd :: "[i,i]\i" where "raw_zadd (z1, z2) \ - \z1\z1. \z2\z2. let =z1; =z2 + \z1\z1. \z2\z2. let \x1,y1\=z1; \x2,y2\=z2 in intrel``{}" definition - zadd :: "[i,i]=>i" (infixl \$+\ 65) where + zadd :: "[i,i]\i" (infixl \$+\ 65) where "z1 $+ z2 \ raw_zadd (intify(z1),intify(z2))" definition - zdiff :: "[i,i]=>i" (infixl \$-\ 65) where + zdiff :: "[i,i]\i" (infixl \$-\ 65) where "z1 $- z2 \ z1 $+ zminus(z2)" definition - zless :: "[i,i]=>o" (infixl \$<\ 50) where + zless :: "[i,i]\o" (infixl \$<\ 50) where "z1 $< z2 \ znegative(z1 $- z2)" definition - zle :: "[i,i]=>o" (infixl \$\\ 50) where + zle :: "[i,i]\o" (infixl \$\\ 50) where "z1 $\ z2 \ z1 $< z2 | intify(z1)=intify(z2)" @@ -98,18 +98,18 @@ (** Natural deduction for intrel **) lemma intrel_iff [simp]: - "<,>: intrel \ + "<\x1,y1\,\x2,y2\>: intrel \ x1\nat \ y1\nat \ x2\nat \ y2\nat \ x1#+y2 = x2#+y1" by (simp add: intrel_def) lemma intrelI [intro!]: "\x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat\ - \ <,>: intrel" + \ <\x1,y1\,\x2,y2\>: intrel" by (simp add: intrel_def) lemma intrelE [elim!]: "\p \ intrel; - \x1 y1 x2 y2. \p = <,>; x1#+y2 = x2#+y1; + \x1 y1 x2 y2. \p = <\x1,y1\,\x2,y2\>; x1#+y2 = x2#+y1; x1\nat; y1\nat; x2\nat; y2\nat\ \ Q\ \ Q" by (simp add: intrel_def, blast) @@ -126,7 +126,7 @@ apply (fast elim!: sym int_trans_lemma) done -lemma image_intrel_int: "\m\nat; n\nat\ \ intrel `` {} \ int" +lemma image_intrel_int: "\m\nat; n\nat\ \ intrel `` {\m,n\} \ int" by (simp add: int_def) declare equiv_intrel [THEN eq_equiv_class_iff, simp] @@ -208,7 +208,7 @@ subsection\\<^term>\zminus\: unary negation on \<^term>\int\\ -lemma zminus_congruent: "(%. intrel``{}) respects intrel" +lemma zminus_congruent: "(\\x,y\. intrel``{\y,x\}) respects intrel" by (auto simp add: congruent_def add_ac) lemma raw_zminus_type: "z \ int \ raw_zminus(z) \ int" @@ -235,13 +235,13 @@ by auto lemma raw_zminus: - "\x\nat; y\nat\ \ raw_zminus(intrel``{}) = intrel `` {}" + "\x\nat; y\nat\ \ raw_zminus(intrel``{\x,y\}) = intrel `` {\y,x\}" apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) done lemma zminus: "\x\nat; y\nat\ - \ $- (intrel``{}) = intrel `` {}" + \ $- (intrel``{\x,y\}) = intrel `` {\y,x\}" by (simp add: zminus_def raw_zminus image_intrel_int) lemma raw_zminus_zminus: "z \ int \ raw_zminus (raw_zminus(z)) = z" @@ -259,7 +259,7 @@ subsection\\<^term>\znegative\: the test for negative integers\ -lemma znegative: "\x\nat; y\nat\ \ znegative(intrel``{}) \ xx\nat; y\nat\ \ znegative(intrel``{\x,y\}) \ xx\nat; y\nat\ \ raw_nat_of(intrel``{}) = x#-y" + "\x\nat; y\nat\ \ raw_nat_of(intrel``{\x,y\}) = x#-y" by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)" @@ -375,7 +375,7 @@ text\Congruence Property for Addition\ lemma zadd_congruent2: - "(%z1 z2. let =z1; =z2 + "(\z1 z2. let \x1,y1\=z1; \x2,y2\=z2 in intrel``{}) respects2 intrel" apply (simp add: congruent2_def) @@ -400,7 +400,7 @@ lemma raw_zadd: "\x1\nat; y1\nat; x2\nat; y2\nat\ - \ raw_zadd (intrel``{}, intrel``{}) = + \ raw_zadd (intrel``{\x1,y1\}, intrel``{\x2,y2\}) = intrel `` {}" apply (simp add: raw_zadd_def UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) @@ -409,7 +409,7 @@ lemma zadd: "\x1\nat; y1\nat; x2\nat; y2\nat\ - \ (intrel``{}) $+ (intrel``{}) = + \ (intrel``{\x1,y1\}) $+ (intrel``{\x2,y2\}) = intrel `` {}" by (simp add: zadd_def raw_zadd image_intrel_int) @@ -489,14 +489,14 @@ text\Congruence property for multiplication\ lemma zmult_congruent2: - "(%p1 p2. split(%x1 y1. split(%x2 y2. + "(\p1 p2. split(\x1 y1. split(\x2 y2. intrel``{}, p2), p1)) respects2 intrel" apply (rule equiv_intrel [THEN congruent2_commuteI], auto) (*Proof that zmult is congruent in one argument*) apply (rename_tac x y) -apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context]) -apply (drule_tac t = "%u. y#*u" in subst_context) +apply (frule_tac t = "\u. x#*u" in sym [THEN subst_context]) +apply (drule_tac t = "\u. y#*u" in subst_context) apply (erule add_left_cancel)+ apply (simp_all add: add_mult_distrib_left) done @@ -513,14 +513,14 @@ lemma raw_zmult: "\x1\nat; y1\nat; x2\nat; y2\nat\ - \ raw_zmult(intrel``{}, intrel``{}) = + \ raw_zmult(intrel``{\x1,y1\}, intrel``{\x2,y2\}) = intrel `` {}" by (simp add: raw_zmult_def UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) lemma zmult: "\x1\nat; y1\nat; x2\nat; y2\nat\ - \ (intrel``{}) $* (intrel``{}) = + \ (intrel``{\x1,y1\}) $* (intrel``{\x2,y2\}) = intrel `` {}" by (simp add: zmult_def raw_zmult image_intrel_int) @@ -822,7 +822,7 @@ lemma zadd_left_cancel: "\w \ int; w': int\ \ (z $+ w' = z $+ w) \ (w' = w)" apply safe -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) +apply (drule_tac t = "\x. x $+ ($-z) " in subst_context) apply (simp add: zadd_ac) done @@ -835,7 +835,7 @@ lemma zadd_right_cancel: "\w \ int; w': int\ \ (w' $+ z = w $+ z) \ (w' = w)" apply safe -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) +apply (drule_tac t = "\x. x $+ ($-z) " in subst_context) apply (simp add: zadd_ac) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/IntDiv.thy --- a/src/ZF/IntDiv.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/IntDiv.thy Tue Sep 27 17:46:52 2022 +0100 @@ -34,65 +34,65 @@ begin definition - quorem :: "[i,i] => o" where - "quorem \ % . + quorem :: "[i,i] \ o" where + "quorem \ \\a,b\ \q,r\. a = b$*q $+ r \ (#0$ #0$\r \ r$(#0$ b$ r $\ #0)" definition - adjust :: "[i,i] => i" where - "adjust(b) \ %. if #0 $\ r$-b then <#2$*q $+ #1,r$-b> + adjust :: "[i,i] \ i" where + "adjust(b) \ \\q,r\. if #0 $\ r$-b then <#2$*q $+ #1,r$-b> else <#2$*q,r>" (** the division algorithm **) definition - posDivAlg :: "i => i" where + posDivAlg :: "i \ i" where (*for the case a>=0, b>0*) -(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) +(*recdef posDivAlg "inv_image less_than (\\a,b\. nat_of(a $- b $+ #1))"*) "posDivAlg(ab) \ - wfrec(measure(int*int, %. nat_of (a $- b $+ #1)), + wfrec(measure(int*int, \\a,b\. nat_of (a $- b $+ #1)), ab, - % f. if (a$#0) then <#0,a> + \\a,b\ f. if (a$#0) then <#0,a> else adjust(b, f ` ))" -(*for the case a<0, b>0*) +(*for the case a\0, b\0*) definition - negDivAlg :: "i => i" where -(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) + negDivAlg :: "i \ i" where +(*recdef negDivAlg "inv_image less_than (\\a,b\. nat_of(- a $- b))"*) "negDivAlg(ab) \ - wfrec(measure(int*int, %. nat_of ($- a $- b)), + wfrec(measure(int*int, \\a,b\. nat_of ($- a $- b)), ab, - % f. if (#0 $\ a$+b | b$\#0) then <#-1,a$+b> + \\a,b\ f. if (#0 $\ a$+b | b$\#0) then <#-1,a$+b> else adjust(b, f ` ))" (*for the general case @{term"b\0"}*) definition - negateSnd :: "i => i" where - "negateSnd \ %. " + negateSnd :: "i \ i" where + "negateSnd \ \\q,r\. " (*The full division algorithm considers all possible signs for a, b including the special case a=0, b<0, because negDivAlg requires a<0*) definition - divAlg :: "i => i" where + divAlg :: "i \ i" where "divAlg \ - %. if #0 $\ a then - if #0 $\ b then posDivAlg () + \\a,b\. if #0 $\ a then + if #0 $\ b then posDivAlg (\a,b\) else if a=#0 then <#0,#0> else negateSnd (negDivAlg (<$-a,$-b>)) else - if #0$) + if #0$a,b\) else negateSnd (posDivAlg (<$-a,$-b>))" definition - zdiv :: "[i,i]=>i" (infixl \zdiv\ 70) where + zdiv :: "[i,i]\i" (infixl \zdiv\ 70) where "a zdiv b \ fst (divAlg ())" definition - zmod :: "[i,i]=>i" (infixl \zmod\ 70) where + zmod :: "[i,i]\i" (infixl \zmod\ 70) where "a zmod b \ snd (divAlg ())" @@ -380,7 +380,7 @@ lemma unique_quotient: - "\quorem (, ); quorem (, ); b \ int; b \ #0; + "\quorem (\a,b\, \q,r\); quorem (\a,b\, ); b \ int; b \ #0; q \ int; q' \ int\ \ q = q'" apply (simp add: split_ifs quorem_def neq_iff_zless) apply safe @@ -391,7 +391,7 @@ done lemma unique_remainder: - "\quorem (, ); quorem (, ); b \ int; b \ #0; + "\quorem (\a,b\, \q,r\); quorem (\a,b\, ); b \ int; b \ #0; q \ int; q' \ int; r \ int; r' \ int\ \ r = r'" apply (subgoal_tac "q = q'") @@ -404,7 +404,7 @@ the Division Algorithm for \a\0\ and \b>0\\ lemma adjust_eq [simp]: - "adjust(b, ) = (let diff = r$-b in + "adjust(b, \q,r\) = (let diff = r$-b in if #0 $\ diff then <#2$*q $+ #1,diff> else <#2$*q,r>)" by (simp add: Let_def adjust_def) @@ -421,7 +421,7 @@ lemma posDivAlg_eqn: "\#0 $< b; a \ int; b \ int\ \ - posDivAlg() = + posDivAlg(\a,b\) = (if a$ else adjust(b, posDivAlg ()))" apply (rule posDivAlg_unfold [THEN trans]) apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym]) @@ -431,10 +431,10 @@ lemma posDivAlg_induct_lemma [rule_format]: assumes prem: "\a b. \a \ int; b \ int; - \ (a $< b | b $\ #0) \ P()\ \ P()" - shows " \ int*int \ P()" -using wf_measure [where A = "int*int" and f = "%.nat_of (a $- b $+ #1)"] -proof (induct "" arbitrary: u v rule: wf_induct) + \ (a $< b | b $\ #0) \ P()\ \ P(\a,b\)" + shows "\u,v\ \ int*int \ P(\u,v\)" +using wf_measure [where A = "int*int" and f = "\\a,b\.nat_of (a $- b $+ #1)"] +proof (induct "\u,v\" arbitrary: u v rule: wf_induct) case (step x) hence uv: "u \ int" "v \ int" by auto thus ?case @@ -452,7 +452,7 @@ and ih: "\a b. \a \ int; b \ int; \ (a $< b | b $\ #0) \ P(a, #2 $* b)\ \ P(a,b)" shows "P(u,v)" -apply (subgoal_tac "(%. P (x,y)) ()") +apply (subgoal_tac "(\\x,y\. P (x,y)) (\u,v\)") apply simp apply (rule posDivAlg_induct_lemma) apply (simp (no_asm_use)) @@ -526,7 +526,7 @@ (*Typechecking for posDivAlg*) lemma posDivAlg_type [rule_format]: - "\a \ int; b \ int\ \ posDivAlg() \ int * int" + "\a \ int; b \ int\ \ posDivAlg(\a,b\) \ int * int" apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) apply assumption+ apply (case_tac "#0 $< ba") @@ -542,7 +542,7 @@ (*Correctness of posDivAlg: it computes quotients correctly*) lemma posDivAlg_correct [rule_format]: "\a \ int; b \ int\ - \ #0 $\ a \ #0 $< b \ quorem (, posDivAlg())" + \ #0 $\ a \ #0 $< b \ quorem (\a,b\, posDivAlg(\a,b\))" apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) apply auto apply (simp_all add: quorem_def) @@ -576,7 +576,7 @@ lemma negDivAlg_eqn: "\#0 $< b; a \ int; b \ int\ \ - negDivAlg() = + negDivAlg(\a,b\) = (if #0 $\ a$+b then <#-1,a$+b> else adjust(b, negDivAlg ()))" apply (rule negDivAlg_unfold [THEN trans]) @@ -588,10 +588,10 @@ assumes prem: "\a b. \a \ int; b \ int; \ (#0 $\ a $+ b | b $\ #0) \ P()\ - \ P()" - shows " \ int*int \ P()" -using wf_measure [where A = "int*int" and f = "%.nat_of ($- a $- b)"] -proof (induct "" arbitrary: u v rule: wf_induct) + \ P(\a,b\)" + shows "\u,v\ \ int*int \ P(\u,v\)" +using wf_measure [where A = "int*int" and f = "\\a,b\.nat_of ($- a $- b)"] +proof (induct "\u,v\" arbitrary: u v rule: wf_induct) case (step x) hence uv: "u \ int" "v \ int" by auto thus ?case @@ -609,7 +609,7 @@ \ (#0 $\ a $+ b | b $\ #0) \ P(a, #2 $* b)\ \ P(a,b)" shows "P(u,v)" -apply (subgoal_tac " (%. P (x,y)) ()") +apply (subgoal_tac " (\\x,y\. P (x,y)) (\u,v\)") apply simp apply (rule negDivAlg_induct_lemma) apply (simp (no_asm_use)) @@ -620,7 +620,7 @@ (*Typechecking for negDivAlg*) lemma negDivAlg_type: - "\a \ int; b \ int\ \ negDivAlg() \ int * int" + "\a \ int; b \ int\ \ negDivAlg(\a,b\) \ int * int" apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) apply assumption+ apply (case_tac "#0 $< ba") @@ -638,7 +638,7 @@ It doesn't work if a=0 because the 0/b=0 rather than -1*) lemma negDivAlg_correct [rule_format]: "\a \ int; b \ int\ - \ a $< #0 \ #0 $< b \ quorem (, negDivAlg())" + \ a $< #0 \ #0 $< b \ quorem (\a,b\, negDivAlg(\a,b\))" apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) apply auto apply (simp_all add: quorem_def) @@ -687,7 +687,7 @@ apply (simp add: linear_arith_lemma integ_of_type vimage_iff) done -lemma negateSnd_eq [simp]: "negateSnd () = " +lemma negateSnd_eq [simp]: "negateSnd (\q,r\) = " apply (unfold negateSnd_def) apply auto done @@ -699,7 +699,7 @@ lemma quorem_neg: "\quorem (<$-a,$-b>, qr); a \ int; b \ int; qr \ int * int\ - \ quorem (, negateSnd(qr))" + \ quorem (\a,b\, negateSnd(qr))" apply clarify apply (auto elim: zless_asym simp add: quorem_def zless_zminus) txt\linear arithmetic from here on\ @@ -711,7 +711,7 @@ done lemma divAlg_correct: - "\b \ #0; a \ int; b \ int\ \ quorem (, divAlg())" + "\b \ #0; a \ int; b \ int\ \ quorem (\a,b\, divAlg(\a,b\))" apply (auto simp add: quorem_0 divAlg_def) apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct posDivAlg_type negDivAlg_type) @@ -720,7 +720,7 @@ apply (auto simp add: zle_def) done -lemma divAlg_type: "\a \ int; b \ int\ \ divAlg() \ int * int" +lemma divAlg_type: "\a \ int; b \ int\ \ divAlg(\a,b\) \ int * int" apply (auto simp add: divAlg_def) apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) done @@ -802,20 +802,20 @@ lemma quorem_div_mod: "\b \ #0; a \ int; b \ int\ - \ quorem (, )" + \ quorem (\a,b\, )" apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) done -(*Surely quorem(,) implies @{term"a \ int"}, but it doesn't matter*) +(*Surely quorem(\a,b\,\q,r\) implies @{term"a \ int"}, but it doesn't matter*) lemma quorem_div: - "\quorem(,); b \ #0; a \ int; b \ int; q \ int\ + "\quorem(\a,b\,\q,r\); b \ #0; a \ int; b \ int; q \ int\ \ a zdiv b = q" by (blast intro: quorem_div_mod [THEN unique_quotient]) lemma quorem_mod: - "\quorem(,); b \ #0; a \ int; b \ int; q \ int; r \ int\ + "\quorem(\a,b\,\q,r\); b \ #0; a \ int; b \ int; q \ int; r \ int\ \ a zmod b = r" by (blast intro: quorem_div_mod [THEN unique_remainder]) @@ -953,7 +953,7 @@ apply (simp add: int_0_less_mult_iff) apply (blast dest: zless_trans) (*linear arithmetic...*) -apply (drule_tac t = "%x. x $- r" in subst_context) +apply (drule_tac t = "\x. x $- r" in subst_context) apply (drule sym) apply (simp add: zcompare_rls) done @@ -963,12 +963,12 @@ apply (simp add: int_0_le_mult_iff zcompare_rls) apply (blast dest: zle_zless_trans) apply (simp add: zdiff_zmult_distrib2) -apply (drule_tac t = "%x. x $- a $* q" in subst_context) +apply (drule_tac t = "\x. x $- a $* q" in subst_context) apply (simp add: zcompare_rls) done lemma self_quotient: - "\quorem(,); a \ int; q \ int; a \ #0\ \ q = #1" + "\quorem(\a,a\,\q,r\); a \ int; q \ int; a \ #0\ \ q = #1" apply (simp add: split_ifs quorem_def neq_iff_zless) apply (rule zle_anti_sym) apply safe @@ -984,7 +984,7 @@ done lemma self_remainder: - "\quorem(,); a \ int; q \ int; r \ int; a \ #0\ \ r = #0" + "\quorem(\a,a\,\q,r\); a \ int; q \ int; r \ int; a \ #0\ \ r = #0" apply (frule self_quotient) apply (auto simp add: quorem_def) done @@ -1291,7 +1291,7 @@ (** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) lemma zmult1_lemma: - "\quorem(, ); c \ int; c \ #0\ + "\quorem(\b,c\, \q,r\); c \ int; c \ #0\ \ quorem (, )" apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) @@ -1354,7 +1354,7 @@ a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) lemma zadd1_lemma: - "\quorem(, ); quorem(, ); + "\quorem(\a,c\, \aq,ar\); quorem(\b,c\, \bq,br\); c \ int; c \ #0\ \ quorem (, )" apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 @@ -1511,7 +1511,7 @@ done lemma zdiv_zmult2_lemma: - "\quorem (, ); a \ int; b \ int; b \ #0; #0 $< c\ + "\quorem (\a,b\, \q,r\); a \ int; b \ int; b \ #0; #0 $< c\ \ quorem (, )" apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def neq_iff_zless int_0_less_mult_iff diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/List.thy --- a/src/ZF/List.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/List.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,7 +8,7 @@ theory List imports Datatype ArithSimp begin consts - list :: "i=>i" + list :: "i\i" datatype "list(A)" = Nil | Cons ("a \ A", "l \ list(A)") @@ -16,7 +16,7 @@ syntax "_Nil" :: i (\[]\) - "_List" :: "is => i" (\[(_)]\) + "_List" :: "is \ i" (\[(_)]\) translations "[x, xs]" == "CONST Cons(x, [xs])" @@ -25,9 +25,9 @@ consts - length :: "i=>i" - hd :: "i=>i" - tl :: "i=>i" + length :: "i\i" + hd :: "i\i" + tl :: "i\i" primrec "length([]) = 0" @@ -43,9 +43,9 @@ consts - map :: "[i=>i, i] => i" - set_of_list :: "i=>i" - app :: "[i,i]=>i" (infixr \@\ 60) + map :: "[i\i, i] \ i" + set_of_list :: "i\i" + app :: "[i,i]\i" (infixr \@\ 60) (*map is a binding operator -- it applies to meta-level functions, not object-level functions. This simplifies the final form of term_rec_conv, @@ -64,9 +64,9 @@ consts - rev :: "i=>i" - flat :: "i=>i" - list_add :: "i=>i" + rev :: "i\i" + flat :: "i\i" + list_add :: "i\i" primrec "rev([]) = []" @@ -81,7 +81,7 @@ "list_add(Cons(a,l)) = a #+ list_add(l)" consts - drop :: "[i,i]=>i" + drop :: "[i,i]\i" primrec drop_0: "drop(0,l) = l" @@ -92,25 +92,25 @@ definition (* Function `take' returns the first n elements of a list *) - take :: "[i,i]=>i" where + take :: "[i,i]\i" where "take(n, as) \ list_rec(\n\nat. [], - %a l r. \n\nat. nat_case([], %m. Cons(a, r`m), n), as)`n" + \a l r. \n\nat. nat_case([], \m. Cons(a, r`m), n), as)`n" definition - nth :: "[i, i]=>i" where + nth :: "[i, i]\i" where \ \returns the (n+1)th element of a list, or 0 if the list is too short.\ "nth(n, as) \ list_rec(\n\nat. 0, - %a l r. \n\nat. nat_case(a, %m. r`m, n), as) ` n" + \a l r. \n\nat. nat_case(a, \m. r`m, n), as) ` n" definition - list_update :: "[i, i, i]=>i" where + list_update :: "[i, i, i]\i" where "list_update(xs, i, v) \ list_rec(\n\nat. Nil, - %u us vs. \n\nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i" + \u us vs. \n\nat. nat_case(Cons(v, us), \m. Cons(u, vs`m), n), xs)`i" consts - filter :: "[i=>o, i] => i" - upt :: "[i, i] =>i" + filter :: "[i\o, i] \ i" + upt :: "[i, i] \i" primrec "filter(P, Nil) = Nil" @@ -122,11 +122,11 @@ "upt(i, succ(j)) = (if i \ j then upt(i, j)@[j] else Nil)" definition - min :: "[i,i] =>i" where + min :: "[i,i] \i" where "min(x, y) \ (if x \ y then x else y)" definition - max :: "[i, i] =>i" where + max :: "[i, i] \i" where "max(x, y) \ (if x \ y then y else x)" (*** Aspects of the datatype definition ***) @@ -286,12 +286,12 @@ (*** theorems about map ***) -lemma map_ident [simp]: "l \ list(A) \ map(%u. u, l) = l" +lemma map_ident [simp]: "l \ list(A) \ map(\u. u, l) = l" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done -lemma map_compose: "l \ list(A) \ map(h, map(j,l)) = map(%u. h(j(u)), l)" +lemma map_compose: "l \ list(A) \ map(h, map(j,l)) = map(\u. h(j(u)), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done @@ -309,7 +309,7 @@ lemma list_rec_map: "l \ list(A) \ list_rec(c, d, map(h,l)) = - list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)" + list_rec(c, \x xs r. d(h(x), map(h,xs), r), l)" apply (induct_tac "l") apply (simp_all (no_asm_simp)) done @@ -510,10 +510,10 @@ lemma filter_is_subset: "xs:list(A) \ set_of_list(filter(P,xs)) \ set_of_list(xs)" by (induct_tac "xs", auto) -lemma filter_False [simp]: "xs:list(A) \ filter(%p. False, xs) = Nil" +lemma filter_False [simp]: "xs:list(A) \ filter(\p. False, xs) = Nil" by (induct_tac "xs", auto) -lemma filter_True [simp]: "xs:list(A) \ filter(%p. True, xs) = xs" +lemma filter_True [simp]: "xs:list(A) \ filter(\p. True, xs) = xs" by (induct_tac "xs", auto) (** length **) @@ -853,18 +853,18 @@ text\Crafty definition to eliminate a type argument\ consts - zip_aux :: "[i,i]=>i" + zip_aux :: "[i,i]\i" primrec (*explicit lambda is required because both arguments of "un" vary*) "zip_aux(B,[]) = - (\ys \ list(B). list_case([], %y l. [], ys))" + (\ys \ list(B). list_case([], \y l. [], ys))" "zip_aux(B,Cons(x,l)) = (\ys \ list(B). - list_case(Nil, %y zs. Cons(, zip_aux(B,l)`zs), ys))" + list_case(Nil, \y zs. Cons(\x,y\, zip_aux(B,l)`zs), ys))" definition - zip :: "[i, i]=>i" where + zip :: "[i, i]\i" where "zip(xs, ys) \ zip_aux(set_of_list(ys),xs)`ys" @@ -897,7 +897,7 @@ lemma zip_Cons_Cons [simp]: "\xs:list(A); ys:list(B); x \ A; y \ B\ \ - zip(Cons(x,xs), Cons(y, ys)) = Cons(, zip(xs, ys))" + zip(Cons(x,xs), Cons(y, ys)) = Cons(\x,y\, zip(xs, ys))" apply (simp add: zip_def, auto) apply (rule zip_aux_unique, auto) apply (simp add: list_on_set_of_list [of _ B]) @@ -967,7 +967,7 @@ lemma set_of_list_zip [rule_format]: "\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, y\:A*B. \i\nat. i < min(length(xs), length(ys)) \ x = nth(i, xs) \ y = nth(i, ys)}" by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth) @@ -1179,9 +1179,9 @@ (** sublist (a generalization of nth to sets) **) definition - sublist :: "[i, i] => i" where + sublist :: "[i, i] \ i" where "sublist(xs, A)\ - map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))" + map(fst, (filter(\p. snd(p): A, zip(xs, upt(0,length(xs))))))" lemma sublist_0 [simp]: "xs:list(A) \sublist(xs, 0) =Nil" by (unfold sublist_def, auto) @@ -1191,8 +1191,8 @@ 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):A, zip(xs, upt(i,i #+ length(xs))))) = + map(fst, filter(\p. snd(p):nat \ snd(p) #+ i \ A, zip(xs,upt(0,length(xs)))))" apply (erule list_append_induct) apply (simp (no_asm_simp)) apply (auto simp add: add_commute length_app filter_append map_app_distrib) @@ -1250,7 +1250,7 @@ text\Repetition of a List Element\ -consts repeat :: "[i,i]=>i" +consts repeat :: "[i,i]\i" primrec "repeat(a,0) = []" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Nat.thy --- a/src/ZF/Nat.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Nat.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,49 +9,49 @@ definition nat :: i where - "nat \ lfp(Inf, %X. {0} \ {succ(i). i \ X})" + "nat \ lfp(Inf, \X. {0} \ {succ(i). i \ X})" definition - quasinat :: "i => o" where + quasinat :: "i \ o" where "quasinat(n) \ n=0 | (\m. n = succ(m))" definition (*Has an unconditional succ case, which is used in "recursor" below.*) - nat_case :: "[i, i=>i, i]=>i" where + nat_case :: "[i, i\i, i]\i" where "nat_case(a,b,k) \ THE y. k=0 \ y=a | (\x. k=succ(x) \ y=b(x))" definition - nat_rec :: "[i, i, [i,i]=>i]=>i" where + nat_rec :: "[i, i, [i,i]\i]\i" where "nat_rec(k,a,b) \ - wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))" + wfrec(Memrel(nat), k, \n f. nat_case(a, \m. b(m, f`m), n))" (*Internalized relations on the naturals*) definition Le :: i where - "Le \ {:nat*nat. x \ y}" + "Le \ {\x,y\:nat*nat. x \ y}" definition Lt :: i where - "Lt \ {:nat*nat. x < y}" + "Lt \ {\x, y\:nat*nat. x < y}" definition Ge :: i where - "Ge \ {:nat*nat. y \ x}" + "Ge \ {\x,y\:nat*nat. y \ x}" definition Gt :: i where - "Gt \ {:nat*nat. y < x}" + "Gt \ {\x,y\:nat*nat. y < x}" definition - greater_than :: "i=>i" where + greater_than :: "i\i" where "greater_than(n) \ {i \ nat. n < i}" text\No need for a less-than operator: a natural number is its list of predecessors!\ -lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \ {succ(i). i \ X})" +lemma nat_bnd_mono: "bnd_mono(Inf, \X. {0} \ {succ(i). i \ X})" apply (rule bnd_monoI) apply (cut_tac infinity, blast, blast) 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]: "\x,y\ \ Le \ x \ y \ x \ nat \ y \ nat" by (force simp add: Le_def) end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/OrdQuant.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,22 +10,22 @@ definition (* Ordinal Quantifiers *) - oall :: "[i, i => o] => o" where + oall :: "[i, i \ o] \ o" where "oall(A, P) \ \x. x P(x)" definition - oex :: "[i, i => o] => o" where + oex :: "[i, i \ o] \ o" where "oex(A, P) \ \x. x P(x)" definition (* Ordinal Union *) - OUnion :: "[i, i => i] => i" where + OUnion :: "[i, i \ i] \ i" where "OUnion(i,B) \ {z: \x\i. B(x). Ord(i)}" syntax - "_oall" :: "[idt, i, o] => o" (\(3\_<_./ _)\ 10) - "_oex" :: "[idt, i, o] => o" (\(3\_<_./ _)\ 10) - "_OUNION" :: "[idt, i, i] => i" (\(3\_<_./ _)\ 10) + "_oall" :: "[idt, i, o] \ o" (\(3\_<_./ _)\ 10) + "_oex" :: "[idt, i, o] \ o" (\(3\_<_./ _)\ 10) + "_OUNION" :: "[idt, i, i] \ i" (\(3\_<_./ _)\ 10) translations "\x "CONST oall(a, \x. P)" "\x "CONST oex(a, \x. P)" @@ -127,7 +127,7 @@ (*Congruence rule for rewriting*) lemma oall_cong [cong]: "\a=a'; \x. x P(x) <-> P'(x)\ - \ oall(a, %x. P(x)) <-> oall(a', %x. P'(x))" + \ oall(a, \x. P(x)) <-> oall(a', \x. P'(x))" by (simp add: oall_def) @@ -151,7 +151,7 @@ lemma oex_cong [cong]: "\a=a'; \x. x P(x) <-> P'(x)\ - \ oex(a, %x. P(x)) <-> oex(a', %x. P'(x))" + \ oex(a, \x. P(x)) <-> oex(a', \x. P'(x))" apply (simp add: oex_def cong add: conj_cong) done @@ -184,16 +184,16 @@ subsection \Quantification over a class\ definition - "rall" :: "[i=>o, i=>o] => o" where + "rall" :: "[i\o, i\o] \ o" where "rall(M, P) \ \x. M(x) \ P(x)" definition - "rex" :: "[i=>o, i=>o] => o" where + "rex" :: "[i\o, i\o] \ o" where "rex(M, P) \ \x. M(x) \ P(x)" syntax - "_rall" :: "[pttrn, i=>o, o] => o" (\(3\_[_]./ _)\ 10) - "_rex" :: "[pttrn, i=>o, o] => o" (\(3\_[_]./ _)\ 10) + "_rall" :: "[pttrn, i\o, o] \ o" (\(3\_[_]./ _)\ 10) + "_rex" :: "[pttrn, i\o, o] \ o" (\(3\_[_]./ _)\ 10) translations "\x[M]. P" \ "CONST rall(M, \x. P)" "\x[M]. P" \ "CONST rex(M, \x. P)" @@ -249,10 +249,10 @@ "(\x. M(x) \ P(x) <-> P'(x)) \ (\x[M]. P(x)) <-> (\x[M]. P'(x))" by (simp add: rex_def cong: conj_cong) -lemma rall_is_ball [simp]: "(\x[%z. z\A]. P(x)) <-> (\x\A. P(x))" +lemma rall_is_ball [simp]: "(\x[\z. z\A]. P(x)) <-> (\x\A. P(x))" by blast -lemma rex_is_bex [simp]: "(\x[%z. z\A]. P(x)) <-> (\x\A. P(x))" +lemma rex_is_bex [simp]: "(\x[\z. z\A]. P(x)) <-> (\x\A. P(x))" by blast lemma atomize_rall: "(\x. M(x) \ P(x)) \ Trueprop (\x[M]. P(x))" @@ -323,8 +323,8 @@ subsubsection\Sets as Classes\ definition - setclass :: "[i,i] => o" (\##_\ [40] 40) where - "setclass(A) \ %x. x \ A" + setclass :: "[i,i] \ o" (\##_\ [40] 40) where + "setclass(A) \ \x. x \ A" lemma setclass_iff [simp]: "setclass(A,x) <-> x \ A" by (simp add: setclass_def) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Order.thy --- a/src/ZF/Order.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Order.thy Tue Sep 27 17:46:52 2022 +0100 @@ -16,15 +16,15 @@ counterparts.\ definition - part_ord :: "[i,i]=>o" (*Strict partial ordering*) where + part_ord :: "[i,i]\o" (*Strict partial ordering*) where "part_ord(A,r) \ irrefl(A,r) \ trans[A](r)" definition - linear :: "[i,i]=>o" (*Strict total ordering*) where - "linear(A,r) \ (\x\A. \y\A. :r | x=y | :r)" + linear :: "[i,i]\o" (*Strict total ordering*) where + "linear(A,r) \ (\x\A. \y\A. \x,y\:r | x=y | \y,x\:r)" definition - tot_ord :: "[i,i]=>o" (*Strict total ordering*) where + tot_ord :: "[i,i]\o" (*Strict total ordering*) where "tot_ord(A,r) \ part_ord(A,r) \ linear(A,r)" definition @@ -40,31 +40,31 @@ "Partial_order(r) \ partial_order_on(field(r), r)" definition - well_ord :: "[i,i]=>o" (*Well-ordering*) where + well_ord :: "[i,i]\o" (*Well-ordering*) where "well_ord(A,r) \ tot_ord(A,r) \ wf[A](r)" definition - mono_map :: "[i,i,i,i]=>i" (*Order-preserving maps*) where + mono_map :: "[i,i,i,i]\i" (*Order-preserving maps*) where "mono_map(A,r,B,s) \ - {f \ A->B. \x\A. \y\A. :r \ :s}" + {f \ A->B. \x\A. \y\A. \x,y\:r \ :s}" definition - ord_iso :: "[i,i,i,i]=>i" (\(\_, _\ \/ \_, _\)\ 51) (*Order isomorphisms*) where + ord_iso :: "[i,i,i,i]\i" (\(\_, _\ \/ \_, _\)\ 51) (*Order isomorphisms*) where "\A,r\ \ \B,s\ \ - {f \ bij(A,B). \x\A. \y\A. :r \ :s}" + {f \ bij(A,B). \x\A. \y\A. \x,y\:r \ :s}" definition - pred :: "[i,i,i]=>i" (*Set of predecessors*) where - "pred(A,x,r) \ {y \ A. :r}" + pred :: "[i,i,i]\i" (*Set of predecessors*) where + "pred(A,x,r) \ {y \ A. \y,x\:r}" definition - ord_iso_map :: "[i,i,i,i]=>i" (*Construction for linearity theorem*) where + ord_iso_map :: "[i,i,i,i]\i" (*Construction for linearity theorem*) where "ord_iso_map(A,r,B,s) \ - \x\A. \y\B. \f \ ord_iso(pred(A,x,r), r, pred(B,y,s), s). {}" + \x\A. \y\B. \f \ ord_iso(pred(A,x,r), r, pred(B,y,s), s). {\x,y\}" definition - first :: "[i, i, i] => o" where - "first(u, X, R) \ u \ X \ (\v\X. v\u \ \ R)" + first :: "[i, i, i] \ o" where + "first(u, X, R) \ u \ X \ (\v\X. v\u \ \u,v\ \ R)" subsection\Immediate Consequences of the Definitions\ @@ -74,7 +74,7 @@ lemma linearE: "\linear(A,r); x \ A; y \ A; - :r \ P; x=y \ P; :r \ P\ + \x,y\:r \ P; x=y \ P; \y,x\:r \ P\ \ P" by (simp add: linear_def, blast) @@ -102,12 +102,12 @@ (** 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) \ \y,x\:r \ y \ A" by (unfold pred_def, blast) lemmas predI = conjI [THEN pred_iff [THEN iffD2]] -lemma predE: "\y \ pred(A,x,r); \y \ A; :r\ \ P\ \ P" +lemma predE: "\y \ pred(A,x,r); \y \ A; \y,x\:r\ \ P\ \ P" by (simp add: pred_def) lemma pred_subset_under: "pred(A,x,r) \ r -`` {x}" @@ -121,7 +121,7 @@ by (simp add: pred_def, blast) lemma trans_pred_pred_eq: - "\trans[A](r); :r; x \ A; y \ A\ + "\trans[A](r); \y,x\:r; x \ A; y \ A\ \ pred(pred(A,x,r), y, r) = pred(A,y,r)" by (unfold trans_on_def pred_def, blast) @@ -251,7 +251,7 @@ lemma ord_isoI: "\f \ bij(A, B); - \x y. \x \ A; y \ A\ \ \ r \ \ s\ + \x y. \x \ A; y \ A\ \ \x, y\ \ r \ \ s\ \ f \ ord_iso(A,r,B,s)" by (simp add: ord_iso_def) @@ -267,11 +267,11 @@ (*Needed? But ord_iso_converse is!*) lemma ord_iso_apply: - "\f \ ord_iso(A,r,B,s); : r; x \ A; y \ A\ \ \ s" + "\f \ ord_iso(A,r,B,s); \x,y\: r; x \ A; y \ A\ \ \ s" by (simp add: ord_iso_def) lemma ord_iso_converse: - "\f \ ord_iso(A,r,B,s); : s; x \ B; y \ B\ + "\f \ ord_iso(A,r,B,s); \x,y\: s; x \ B; y \ B\ \ \ r" apply (simp add: ord_iso_def, clarify) apply (erule bspec [THEN bspec, THEN iffD2]) @@ -437,10 +437,10 @@ (*Tricky; a lot of forward proof!*) lemma well_ord_iso_preserving: - "\well_ord(A,r); well_ord(B,s); : r; + "\well_ord(A,r); well_ord(B,s); \a,c\: r; f \ ord_iso(pred(A,a,r), r, pred(B,b,s), s); g \ ord_iso(pred(A,c,r), r, pred(B,d,s), s); - a \ A; c \ A; b \ B; d \ B\ \ : s" + a \ A; c \ A; b \ B; d \ B\ \ \b,d\: s" apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], (erule asm_rl predI predE)+) apply (subgoal_tac "b = g`a") apply (simp (no_asm_simp)) @@ -552,7 +552,7 @@ (*Trivial case: b=a*) apply clarify apply blast -(*Harder case: : r*) +(*Harder case: \a, xa\: r*) apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], (erule asm_rl predI predE)+) apply (frule ord_iso_restrict_pred) @@ -603,7 +603,7 @@ apply (rule wf_on_not_refl [THEN notE]) apply (erule well_ord_is_wf) apply assumption -apply (subgoal_tac ": ord_iso_map (A,r,B,s) ") +apply (subgoal_tac "\x,y\: ord_iso_map (A,r,B,s) ") apply (drule rangeI) apply (simp add: pred_def) apply (unfold ord_iso_map_def, blast) @@ -661,7 +661,7 @@ lemma subset_vimage_vimage_iff: "\Preorder(r); A \ field(r); B \ field(r)\ \ - r -`` A \ r -`` B \ (\a\A. \b\B. \ r)" + r -`` A \ r -`` B \ (\a\A. \b\B. \a, b\ \ r)" apply (auto simp: subset_def preorder_on_def refl_def vimage_def image_def) apply blast unfolding trans_on_def @@ -674,7 +674,7 @@ lemma subset_vimage1_vimage1_iff: "\Preorder(r); a \ field(r); b \ field(r)\ \ - r -`` {a} \ r -`` {b} \ \ r" + r -`` {a} \ r -`` {b} \ \a, b\ \ r" by (simp add: subset_vimage_vimage_iff) lemma Refl_antisym_eq_Image1_Image1_iff: diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/OrderArith.thy --- a/src/ZF/OrderArith.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/OrderArith.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,29 +9,29 @@ definition (*disjoint sum of two relations; underlies ordinal addition*) - radd :: "[i,i,i,i]=>i" where + radd :: "[i,i,i,i]\i" where "radd(A,r,B,s) \ {z: (A+B) * (A+B). - (\x y. z = ) | - (\x' x. z = \ :r) | - (\y' y. z = \ :s)}" + (\x y. z = \Inl(x), Inr(y)\) | + (\x' x. z = \Inl(x'), Inl(x)\ \ \x',x\:r) | + (\y' y. z = \Inr(y'), Inr(y)\ \ \y',y\:s)}" definition (*lexicographic product of two relations; underlies ordinal multiplication*) - rmult :: "[i,i,i,i]=>i" where + rmult :: "[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 = \\x',y'\, \x,y\\ \ + (\x',x\: r | (x'=x \ \y',y\: s))}" definition (*inverse image of a relation*) - rvimage :: "[i,i,i]=>i" where - "rvimage(A,f,r) \ {z \ A*A. \x y. z = \ : r}" + rvimage :: "[i,i,i]\i" where + "rvimage(A,f,r) \ {z \ A*A. \x y. z = \x,y\ \ \f`x,f`y\: r}" definition measure :: "[i, i\i] \ i" where - "measure(A,f) \ {: A*A. f(x) < f(y)}" + "measure(A,f) \ {\x,y\: A*A. f(x) < f(y)}" subsection\Addition of Relations -- Disjoint Sum\ @@ -39,19 +39,19 @@ 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" + "\Inl(a), Inr(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" + "\Inl(a'), Inl(a)\ \ radd(A,r,B,s) \ a':A \ a \ A \ \a',a\:r" by (unfold radd_def, blast) lemma radd_Inr_iff [iff]: - " \ radd(A,r,B,s) \ b':B \ b \ B \ :s" + "\Inr(b'), Inr(b)\ \ radd(A,r,B,s) \ b':B \ b \ B \ \b',b\:s" by (unfold radd_def, blast) lemma radd_Inr_Inl_iff [simp]: - " \ radd(A,r,B,s) \ False" + "\Inr(b), Inl(a)\ \ radd(A,r,B,s) \ False" by (unfold radd_def, blast) declare radd_Inr_Inl_iff [THEN iffD1, dest!] @@ -59,10 +59,10 @@ subsubsection\Elimination Rule\ lemma raddE: - "\ \ radd(A,r,B,s); + "\\p',p\ \ radd(A,r,B,s); \x y. \p'=Inl(x); x \ A; p=Inr(y); y \ B\ \ Q; - \x' x. \p'=Inl(x'); p=Inl(x); : r; x':A; x \ A\ \ Q; - \y' y. \p'=Inr(y'); p=Inr(y); : s; y':B; y \ B\ \ Q + \x' x. \p'=Inl(x'); p=Inl(x); \x',x\: r; x':A; x \ A\ \ Q; + \y' y. \p'=Inr(y'); p=Inr(y); \y',y\: s; y':B; y \ B\ \ Q \ \ Q" by (unfold radd_def, blast) @@ -116,8 +116,8 @@ lemma sum_bij: "\f \ bij(A,C); g \ bij(B,D)\ - \ (\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \ bij(A+B, C+D)" -apply (rule_tac d = "case (%x. Inl (converse(f)`x), %y. Inr(converse(g)`y))" + \ (\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z)) \ bij(A+B, C+D)" +apply (rule_tac d = "case (\x. Inl (converse(f)`x), \y. Inr(converse(g)`y))" in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) @@ -125,7 +125,7 @@ lemma sum_ord_iso_cong: "\f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s')\ \ - (\z\A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) + (\z\A+B. case(\x. Inl(f`x), \y. Inr(g`y), z)) \ ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))" apply (unfold ord_iso_def) apply (safe intro!: sum_bij) @@ -136,23 +136,23 @@ (*Could we prove an ord_iso result? Perhaps ord_iso(A+B, radd(A,r,B,s), A \ B, r \ s) *) lemma sum_disjoint_bij: "A \ B = 0 \ - (\z\A+B. case(%x. x, %y. y, z)) \ bij(A+B, A \ B)" -apply (rule_tac d = "%z. if z \ A then Inl (z) else Inr (z) " in lam_bijective) + (\z\A+B. case(\x. x, \y. y, z)) \ bij(A+B, A \ B)" +apply (rule_tac d = "\z. if z \ A then Inl (z) else Inr (z) " in lam_bijective) apply auto done subsubsection\Associativity\ lemma sum_assoc_bij: - "(\z\(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) + "(\z\(A+B)+C. case(case(Inl, \y. Inr(Inl(y))), \y. Inr(Inr(y)), z)) \ bij((A+B)+C, A+(B+C))" -apply (rule_tac d = "case (%x. Inl (Inl (x)), case (%x. Inl (Inr (x)), Inr))" +apply (rule_tac d = "case (\x. Inl (Inl (x)), case (\x. Inl (Inr (x)), Inr))" in lam_bijective) apply auto done lemma sum_assoc_ord_iso: - "(\z\(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) + "(\z\(A+B)+C. case(case(Inl, \y. Inr(Inl(y))), \y. Inr(Inr(y)), z)) \ ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))" by (rule sum_assoc_bij [THEN ord_isoI], auto) @@ -163,16 +163,16 @@ subsubsection\Rewrite rule. Can be used to obtain introduction rules\ 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)" + "\\a',b'\, \a,b\\ \ rmult(A,r,B,s) \ + (\a',a\: r \ a':A \ a \ A \ b': B \ b \ B) | + (\b',b\: s \ a'=a \ a \ A \ b': B \ b \ B)" by (unfold rmult_def, blast) lemma rmultE: - "\<, > \ rmult(A,r,B,s); - \: r; a':A; a \ A; b':B; b \ B\ \ Q; - \: s; a \ A; a'=a; b':B; b \ B\ \ Q + "\\\a',b'\, \a,b\\ \ rmult(A,r,B,s); + \\a',a\: r; a':A; a \ A; b':B; b \ B\ \ Q; + \\b',b\: s; a \ A; a'=a; b':B; b \ B\ \ Q \ \ Q" by blast @@ -195,7 +195,7 @@ apply (rule wf_onI2) apply (erule SigmaE) apply (erule ssubst) -apply (subgoal_tac "\b\B. : Ba", blast) +apply (subgoal_tac "\b\B. \x,b\: Ba", blast) apply (erule_tac a = x in wf_on_induct, assumption) apply (rule ballI) apply (erule_tac a = b in wf_on_induct, assumption) @@ -221,8 +221,8 @@ lemma prod_bij: "\f \ bij(A,C); g \ bij(B,D)\ - \ (lam :A*B. ) \ bij(A*B, C*D)" -apply (rule_tac d = "%. " + \ (lam \x,y\:A*B. \f`x, g`y\) \ bij(A*B, C*D)" +apply (rule_tac d = "\\x,y\. \converse (f) `x, converse (g) `y\" in lam_bijective) apply (typecheck add: bij_is_inj inj_is_fun) apply (auto simp add: left_inverse_bij right_inverse_bij) @@ -230,7 +230,7 @@ lemma prod_ord_iso_cong: "\f \ ord_iso(A,r,A',r'); g \ ord_iso(B,s,B',s')\ - \ (lam :A*B. ) + \ (lam \x,y\:A*B. \f`x, g`y\) \ ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))" apply (unfold ord_iso_def) apply (safe intro!: prod_bij) @@ -238,13 +238,13 @@ apply (blast intro: bij_is_inj [THEN inj_apply_equality]) done -lemma singleton_prod_bij: "(\z\A. ) \ bij(A, {x}*A)" +lemma singleton_prod_bij: "(\z\A. \x,z\) \ bij(A, {x}*A)" by (rule_tac d = snd in lam_bijective, auto) (*Used??*) lemma singleton_prod_ord_iso: "well_ord({x},xr) \ - (\z\A. ) \ ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))" + (\z\A. \x,z\) \ ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))" apply (rule singleton_prod_bij [THEN ord_isoI]) apply (simp (no_asm_simp)) apply (blast dest: well_ord_is_wf [THEN wf_on_not_refl]) @@ -254,7 +254,7 @@ case_cong, id_conv, comp_lam, case_case.*) lemma prod_sum_singleton_bij: "a\C \ - (\x\C*B + D. case(%x. x, %y., x)) + (\x\C*B + D. case(\x. x, \y.\a,y\, x)) \ bij(C*B + D, C*B \ {a}*D)" apply (rule subst_elem) apply (rule id_bij [THEN sum_bij, THEN comp_bij]) @@ -268,7 +268,7 @@ lemma prod_sum_singleton_ord_iso: "\a \ A; well_ord(A,r)\ \ - (\x\pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y., x)) + (\x\pred(A,a,r)*B + pred(B,b,s). case(\x. x, \y.\a,y\, x)) \ ord_iso(pred(A,a,r)*B + pred(B,b,s), radd(A*B, rmult(A,r,B,s), B, s), pred(A,a,r)*B \ {a}*pred(B,b,s), rmult(A,r,B,s))" @@ -280,13 +280,13 @@ subsubsection\Distributive law\ lemma sum_prod_distrib_bij: - "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) + "(lam \x,z\:(A+B)*C. case(\y. Inl(\y,z\), \y. Inr(\y,z\), x)) \ bij((A+B)*C, (A*C)+(B*C))" -by (rule_tac d = "case (%., %.) " +by (rule_tac d = "case (\\x,y\.\Inl (x),y\, \\x,y\.\Inr (x),y\) " in lam_bijective, auto) lemma sum_prod_distrib_ord_iso: - "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) + "(lam \x,z\:(A+B)*C. case(\y. Inl(\y,z\), \y. Inr(\y,z\), x)) \ ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t), (A*C)+(B*C), radd(A*C, rmult(A,r,C,t), B*C, rmult(B,s,C,t)))" by (rule sum_prod_distrib_bij [THEN ord_isoI], auto) @@ -294,11 +294,11 @@ subsubsection\Associativity\ lemma prod_assoc_bij: - "(lam <, z>:(A*B)*C. >) \ bij((A*B)*C, A*(B*C))" -by (rule_tac d = "%>. <, z>" in lam_bijective, auto) + "(lam \\x,y\, z\:(A*B)*C. \x,\y,z\\) \ bij((A*B)*C, A*(B*C))" +by (rule_tac d = "\\x, \y,z\\. \\x,y\, z\" in lam_bijective, auto) lemma prod_assoc_ord_iso: - "(lam <, z>:(A*B)*C. >) + "(lam \\x,y\, z\:(A*B)*C. \x,\y,z\\) \ ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))" by (rule prod_assoc_bij [THEN ord_isoI], auto) @@ -307,7 +307,7 @@ subsubsection\Rewrite rule\ -lemma rvimage_iff: " \ rvimage(A,f,r) \ : r \ a \ A \ b \ A" +lemma rvimage_iff: "\a,b\ \ rvimage(A,f,r) \ \f`a,f`b\: r \ a \ A \ b \ A" by (unfold rvimage_def, blast) subsubsection\Type checking\ @@ -371,7 +371,7 @@ text\But note that the combination of \wf_imp_wf_on\ and \wf_rvimage\ gives \<^prop>\wf(r) \ wf[C](rvimage(A,f,r))\\ -lemma wf_on_rvimage: "\f \ A->B; wf[B](r)\ \ wf[A](rvimage(A,f,r))" +lemma wf_on_rvimage: "\f \ A\B; wf[B](r)\ \ wf[A](rvimage(A,f,r))" apply (rule wf_onI2) apply (subgoal_tac "\z\A. f`z=f`y \ z \ Ba") apply blast @@ -408,11 +408,11 @@ definition - wfrank :: "[i,i]=>i" where - "wfrank(r,a) \ wfrec(r, a, %x f. \y \ r-``{x}. succ(f`y))" + wfrank :: "[i,i]\i" where + "wfrank(r,a) \ wfrec(r, a, \x f. \y \ r-``{x}. succ(f`y))" definition - wftype :: "i=>i" where + wftype :: "i\i" where "wftype(r) \ \y \ range(r). succ(wfrank(r,y))" lemma wfrank: "wf(r) \ wfrank(r,a) = (\y \ r-``{a}. succ(wfrank(r,y)))" @@ -424,7 +424,7 @@ apply (rule Ord_succ [THEN Ord_UN], blast) done -lemma wfrank_lt: "\wf(r); \ r\ \ wfrank(r,a) < wfrank(r,b)" +lemma wfrank_lt: "\wf(r); \a,b\ \ r\ \ wfrank(r,a) < wfrank(r,b)" apply (rule_tac a1 = b in wfrank [THEN ssubst], assumption) apply (rule UN_I [THEN ltI]) apply (simp add: Ord_wfrank vimage_iff)+ @@ -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)x,y\ \ 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)" + and ok: "\a u v. \\u,v\ \ s; v \ B(a); a \ A\ + \ (\a'\A. \a',a\ \ r \ u \ B(a')) | u \ B(a)" shows "wf[\a\A. B(a)](s)" apply (rule wf_onI2) apply (erule UN_E) @@ -547,18 +547,18 @@ subsubsection\Bijections involving Powersets\ lemma Pow_sum_bij: - "(\Z \ Pow(A+B). <{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}>) + "(\Z \ Pow(A+B). \{x \ A. Inl(x) \ Z}, {y \ B. Inr(y) \ Z}\) \ bij(Pow(A+B), Pow(A)*Pow(B))" -apply (rule_tac d = "%. {Inl (x). x \ X} \ {Inr (y). y \ Y}" +apply (rule_tac d = "\\X,Y\. {Inl (x). x \ X} \ {Inr (y). y \ Y}" in lam_bijective) apply force+ done -text\As a special case, we have \<^term>\bij(Pow(A*B), A -> Pow(B))\\ +text\As a special case, we have \<^term>\bij(Pow(A*B), A \ Pow(B))\\ lemma Pow_Sigma_bij: "(\r \ Pow(Sigma(A,B)). \x \ A. r``{x}) \ bij(Pow(Sigma(A,B)), \x \ A. Pow(B(x)))" -apply (rule_tac d = "%f. \x \ A. \y \ f`x. {}" in lam_bijective) +apply (rule_tac d = "\f. \x \ A. \y \ f`x. {\x,y\}" in lam_bijective) apply (blast intro: lam_type) apply (blast dest: apply_type, simp_all) apply fast (*strange, but blast can't do it*) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/OrderType.thy --- a/src/ZF/OrderType.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/OrderType.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,40 +12,40 @@ here. But a definition by transfinite recursion would be much simpler!\ definition - ordermap :: "[i,i]=>i" where - "ordermap(A,r) \ \x\A. wfrec[A](r, x, %x f. f `` pred(A,x,r))" + ordermap :: "[i,i]\i" where + "ordermap(A,r) \ \x\A. wfrec[A](r, x, \x f. f `` pred(A,x,r))" definition - ordertype :: "[i,i]=>i" where + ordertype :: "[i,i]\i" where "ordertype(A,r) \ ordermap(A,r)``A" definition (*alternative definition of ordinal numbers*) - Ord_alt :: "i => o" where + Ord_alt :: "i \ o" where "Ord_alt(X) \ well_ord(X, Memrel(X)) \ (\u\X. u=pred(X, u, Memrel(X)))" definition (*coercion to ordinal: if not, just 0*) - ordify :: "i=>i" where + ordify :: "i\i" where "ordify(x) \ if Ord(x) then x else 0" definition (*ordinal multiplication*) - omult :: "[i,i]=>i" (infixl \**\ 70) where + omult :: "[i,i]\i" (infixl \**\ 70) where "i ** j \ ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))" definition (*ordinal addition*) - raw_oadd :: "[i,i]=>i" where + raw_oadd :: "[i,i]\i" where "raw_oadd(i,j) \ ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))" definition - oadd :: "[i,i]=>i" (infixl \++\ 65) where + oadd :: "[i,i]\i" (infixl \++\ 65) where "i ++ j \ raw_oadd(ordify(i),ordify(j))" definition (*ordinal subtraction*) - odiff :: "[i,i]=>i" (infixl \--\ 65) where + odiff :: "[i,i]\i" (infixl \--\ 65) where "i -- j \ ordertype(i-j, Memrel(i))" @@ -128,14 +128,14 @@ (*The theorem above is \wf[A](r); x \ A\ -\ ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \ A . \ r}} +\ ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \ A . \y,x\ \ r}} NOTE: the definition of ordermap used here delivers ordinals only if r is transitive. If r is the predecessor relation on the naturals then ordermap(nat,predr) ` n equals {n-1} and not n. A more complicated definition, like - ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y \ A . \ r}}, + ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y \ A . \y,x\ \ r}}, might eliminate the need for r to be transitive. *) @@ -169,7 +169,7 @@ subsubsection\ordermap preserves the orderings in both directions\ lemma ordermap_mono: - "\: r; wf[A](r); w \ A; x \ A\ + "\\w,x\: r; wf[A](r); w \ A; x \ A\ \ ordermap(A,r)`w \ ordermap(A,r)`x" apply (erule_tac x1 = x in ordermap_unfold [THEN ssubst], assumption, blast) done @@ -177,7 +177,7 @@ (*linearity of r is crucial here*) lemma converse_ordermap_mono: "\ordermap(A,r)`w \ ordermap(A,r)`x; well_ord(A,r); w \ A; x \ A\ - \ : r" + \ \w,x\: r" apply (unfold well_ord_def tot_ord_def, safe) apply (erule_tac x=w and y=x in linearE, assumption+) apply (blast elim!: mem_not_refl [THEN notE]) @@ -333,7 +333,7 @@ text\Addition with 0\ -lemma bij_sum_0: "(\z\A+0. case(%x. x, %y. y, z)) \ bij(A+0, A)" +lemma bij_sum_0: "(\z\A+0. case(\x. x, \y. y, z)) \ bij(A+0, A)" apply (rule_tac d = Inl in lam_bijective, safe) apply (simp_all (no_asm_simp)) done @@ -345,7 +345,7 @@ apply force done -lemma bij_0_sum: "(\z\0+A. case(%x. x, %y. y, z)) \ bij(0+A, A)" +lemma bij_0_sum: "(\z\0+A. case(\x. x, \y. y, z)) \ bij(0+A, A)" apply (rule_tac d = Inr in lam_bijective, safe) apply (simp_all (no_asm_simp)) done @@ -364,7 +364,7 @@ "a \ A \ (\x\pred(A,a,r). Inl(x)) \ bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))" apply (unfold pred_def) -apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective) +apply (rule_tac d = "case (\x. x, \y. y) " in lam_bijective) apply auto done @@ -382,7 +382,7 @@ id(A+pred(B,b,s)) \ bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))" apply (unfold pred_def id_def) -apply (rule_tac d = "%z. z" in lam_bijective, auto) +apply (rule_tac d = "\z. z" in lam_bijective, auto) done lemma ordertype_pred_Inr_eq: @@ -686,7 +686,7 @@ lemma bij_sum_Diff: "A<=B \ (\y\B. if(y \ A, Inl(y), Inr(y))) \ bij(B, A+(B-A))" -apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective) +apply (rule_tac d = "case (\x. x, \y. y) " in lam_bijective) apply (blast intro!: if_type) apply (fast intro!: case_type) apply (erule_tac [2] sumE) @@ -757,14 +757,14 @@ subsubsection\A useful unfolding law\ lemma pred_Pair_eq: - "\a \ A; b \ B\ \ pred(A*B, , rmult(A,r,B,s)) = + "\a \ A; b \ B\ \ pred(A*B, \a,b\, rmult(A,r,B,s)) = pred(A,a,r)*B \ ({a} * pred(B,b,s))" apply (unfold pred_def, blast) done lemma ordertype_pred_Pair_eq: "\a \ A; b \ B; well_ord(A,r); well_ord(B,s)\ \ - ordertype(pred(A*B, , rmult(A,r,B,s)), rmult(A,r,B,s)) = + ordertype(pred(A*B, \a,b\, rmult(A,r,B,s)), rmult(A,r,B,s)) = ordertype(pred(A,a,r)*B + pred(B,b,s), radd(A*B, rmult(A,r,B,s), B, s))" apply (simp (no_asm_simp) add: pred_Pair_eq) @@ -849,7 +849,7 @@ apply (unfold omult_def) apply (rule_tac s1="Memrel(i)" in ord_isoI [THEN ordertype_eq, THEN trans]) -apply (rule_tac c = snd and d = "%z.<0,z>" in lam_bijective) +apply (rule_tac c = snd and d = "\z.\0,z\" in lam_bijective) apply (auto elim!: snd_type well_ord_Memrel ordertype_Memrel) done @@ -857,7 +857,7 @@ apply (unfold omult_def) apply (rule_tac s1="Memrel(i)" in ord_isoI [THEN ordertype_eq, THEN trans]) -apply (rule_tac c = fst and d = "%z." in lam_bijective) +apply (rule_tac c = fst and d = "\z.\z,0\" in lam_bijective) apply (auto elim!: fst_type well_ord_Memrel ordertype_Memrel) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Ordinal.thy --- a/src/ZF/Ordinal.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Ordinal.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,23 +8,23 @@ theory Ordinal imports WF Bool equalities begin definition - Memrel :: "i=>i" where - "Memrel(A) \ {z\A*A . \x y. z= \ x\y }" + Memrel :: "i\i" where + "Memrel(A) \ {z\A*A . \x y. z=\x,y\ \ x\y }" definition - Transset :: "i=>o" where + Transset :: "i\o" where "Transset(i) \ \x\i. x<=i" definition - Ord :: "i=>o" where + Ord :: "i\o" where "Ord(i) \ Transset(i) \ (\x\i. Transset(x))" definition - lt :: "[i,i] => o" (infixl \<\ 50) (*less-than on ordinals*) where + lt :: "[i,i] \ o" (infixl \<\ 50) (*less-than on ordinals*) where "i i\j \ Ord(j)" definition - Limit :: "i=>o" where + Limit :: "i\o" where "Limit(i) \ Ord(i) \ 0 (\y. y succ(y)Transset(C); \C\ \ a\C \ b\C" + "\Transset(C); \a,b\\C\ \ a\C \ b\C" apply (simp add: Pair_def) apply (blast dest: Transset_doubleton_D) done @@ -230,7 +230,7 @@ done -text\Recall that \<^term>\i \ j\ abbreviates \<^term>\i \\ +text\Recall that \<^term>\i \ j\ abbreviates \<^term>\i!\ lemma le_iff: "i \ j <-> i Ord(j))" by (unfold lt_def, blast) @@ -267,14 +267,14 @@ 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]: "\a,b\ \ Memrel(A) <-> a\b \ a\A \ b\A" by (unfold Memrel_def, blast) -lemma MemrelI [intro!]: "\a \ b; a \ A; b \ A\ \ \ Memrel(A)" +lemma MemrelI [intro!]: "\a \ b; a \ A; b \ A\ \ \a,b\ \ Memrel(A)" by auto lemma MemrelE [elim!]: - "\ \ Memrel(A); + "\\a,b\ \ Memrel(A); \a \ A; b \ A; a\b\ \ P\ \ P" by auto @@ -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) \ \a,b\ \ Memrel(A) <-> a\b \ b\A" by (unfold Transset_def, blast) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Perm.thy --- a/src/ZF/Perm.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Perm.thy Tue Sep 27 17:46:52 2022 +0100 @@ -14,28 +14,28 @@ definition (*composition of relations and functions; NOT Suppes's relative product*) - comp :: "[i,i]=>i" (infixr \O\ 60) where + 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=\x,z\ \ \x,y\:s \ \y,z\:r}" definition (*the identity function for A*) - id :: "i=>i" where + id :: "i\i" where "id(A) \ (\x\A. x)" definition (*one-to-one functions from A to B*) - inj :: "[i,i]=>i" where + inj :: "[i,i]\i" where "inj(A,B) \ { f \ A->B. \w\A. \x\A. f`w=f`x \ w=x}" definition (*onto functions from A to B*) - surj :: "[i,i]=>i" where + surj :: "[i,i]\i" where "surj(A,B) \ { f \ A->B . \y\B. \x\A. f`x=y}" definition (*one-to-one and onto functions*) - bij :: "[i,i]=>i" where + bij :: "[i,i]\i" where "bij(A,B) \ inj(A,B) \ surj(A,B)" @@ -89,7 +89,7 @@ text\Good for dealing with sets of pairs, but a bit ugly in use [used in AC]\ lemma inj_equality: - "\:f; :f; f \ inj(A,B)\ \ a=c" + "\\a,b\:f; \c,b\:f; f \ inj(A,B)\ \ a=c" apply (unfold inj_def) apply (blast dest: Pair_mem_PiD) done @@ -146,12 +146,12 @@ subsection\Identity Function\ -lemma idI [intro!]: "a \ A \ \ id(A)" +lemma idI [intro!]: "a \ A \ \a,a\ \ id(A)" apply (unfold id_def) apply (erule lamI) done -lemma idE [elim!]: "\p \ id(A); \x.\x \ A; p=\ \ P\ \ P" +lemma idE [elim!]: "\p \ id(A); \x.\x \ A; p=\x,x\\ \ P\ \ P" by (simp add: id_def lam_def, blast) lemma id_type: "id(A) \ A->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]: "\x,y\ \ id(A) \ x=y \ y \ A" by auto @@ -261,18 +261,18 @@ text\The inductive definition package could derive these theorems for \<^term>\r O s\\ -lemma compI [intro]: "\:s; :r\ \ \ r O s" +lemma compI [intro]: "\\a,b\:s; \b,c\:r\ \ \a,c\ \ r O s" by (unfold comp_def, blast) lemma compE [elim!]: "\xz \ r O s; - \x y z. \xz=; :s; :r\ \ P\ + \x y z. \xz=\x,z\; \x,y\:s; \y,z\:r\ \ P\ \ P" by (unfold comp_def, blast) lemma compEpair: - "\ \ r O s; - \y. \:s; :r\ \ P\ + "\\a,c\ \ r O s; + \y. \\a,y\:s; \y,c\:r\ \ P\ \ P" by (erule compE, simp) @@ -366,7 +366,7 @@ "\g \ inj(A,B); f \ inj(B,C)\ \ (f O g) \ inj(A,C)" apply (frule inj_is_fun [of g]) apply (frule inj_is_fun [of f]) -apply (rule_tac d = "%y. converse (g) ` (converse (f) ` y)" in f_imp_injective) +apply (rule_tac d = "\y. converse (g) ` (converse (f) ` y)" in f_imp_injective) apply (blast intro: comp_fun, simp) done @@ -443,7 +443,7 @@ lemma comp_eq_id_iff: "\f \ A->B; g \ B->A\ \ f O g = id(B) \ (\y\B. f`(g`y)=y)" apply (unfold id_def, safe) - apply (drule_tac t = "%h. h`y " in subst_context) + apply (drule_tac t = "\h. h`y " in subst_context) apply simp apply (rule fun_extension) apply (blast intro: comp_fun lam_type) @@ -473,7 +473,7 @@ lemma inj_disjoint_Un: "\f \ inj(A,B); g \ inj(C,D); B \ D = 0\ \ (\a\A \ C. if a \ A then f`a else g`a) \ inj(A \ C, B \ D)" -apply (rule_tac d = "%z. if z \ B then converse (f) `z else converse (g) `z" +apply (rule_tac d = "\z. if z \ B then converse (f) `z else converse (g) `z" in lam_injective) apply (auto simp add: inj_is_fun [THEN apply_type]) done @@ -546,7 +546,7 @@ lemma inj_extend: "\f \ inj(A,B); a\A; b\B\ - \ cons(,f) \ inj(cons(a,A), cons(b,B))" + \ cons(\a,b\,f) \ inj(cons(a,A), cons(b,B))" apply (unfold inj_def) apply (force intro: apply_type simp add: fun_extend) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/QPair.thy --- a/src/ZF/QPair.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/QPair.thy Tue Sep 27 17:46:52 2022 +0100 @@ -21,53 +21,53 @@ \ definition - QPair :: "[i, i] => i" (\<(_;/ _)>\) where + QPair :: "[i, i] \ i" (\<(_;/ _)>\) where " \ a+b" definition - qfst :: "i => i" where + qfst :: "i \ i" where "qfst(p) \ THE a. \b. p=" definition - qsnd :: "i => i" where + qsnd :: "i \ i" where "qsnd(p) \ THE b. \a. p=" definition - qsplit :: "[[i, i] => 'a, i] => 'a::{}" (*for pattern-matching*) where + qsplit :: "[[i, i] \ 'a, i] \ 'a::{}" (*for pattern-matching*) where "qsplit(c,p) \ c(qfst(p), qsnd(p))" definition - qconverse :: "i => i" where + qconverse :: "i \ i" where "qconverse(r) \ {z. w \ r, \x y. w= \ z=}" definition - QSigma :: "[i, i => i] => i" where + QSigma :: "[i, i \ i] \ i" where "QSigma(A,B) \ \x\A. \y\B(x). {}" syntax - "_QSUM" :: "[idt, i, i] => i" (\(3QSUM _ \ _./ _)\ 10) + "_QSUM" :: "[idt, i, i] \ i" (\(3QSUM _ \ _./ _)\ 10) translations - "QSUM x \ A. B" => "CONST QSigma(A, %x. B)" + "QSUM x \ A. B" => "CONST QSigma(A, \x. B)" abbreviation qprod (infixr \<*>\ 80) where - "A <*> B \ QSigma(A, %_. B)" + "A <*> B \ QSigma(A, \_. B)" definition - qsum :: "[i,i]=>i" (infixr \<+>\ 65) where + qsum :: "[i,i]\i" (infixr \<+>\ 65) where "A <+> B \ ({0} <*> A) \ ({1} <*> B)" definition - QInl :: "i=>i" where + QInl :: "i\i" where "QInl(a) \ <0;a>" definition - QInr :: "i=>i" where + QInr :: "i\i" where "QInr(b) \ <1;b>" definition - qcase :: "[i=>i, i=>i, i]=>i" where - "qcase(c,d) \ qsplit(%y z. cond(y, d(z), c(z)))" + qcase :: "[i\i, i\i, i]\i" where + "qcase(c,d) \ qsplit(\y z. cond(y, d(z), c(z)))" subsection\Quine ordered pairing\ @@ -149,14 +149,14 @@ subsubsection\Eliminator: qsplit\ (*A META-equality, so that it applies to higher types as well...*) -lemma qsplit [simp]: "qsplit(%x y. c(x,y), ) \ c(a,b)" +lemma qsplit [simp]: "qsplit(\x y. c(x,y), ) \ c(a,b)" by (simp add: qsplit_def) lemma qsplit_type [elim!]: "\p \ QSigma(A,B); \x y.\x \ A; y \ B(x)\ \ c(x,y):C() -\ \ qsplit(%x y. c(x,y), p) \ C(p)" +\ \ qsplit(\x y. c(x,y), p) \ C(p)" by auto lemma expand_qsplit: @@ -298,7 +298,7 @@ lemma Part_QInr: "Part(A <+> B,QInr) = {QInr(y). y \ B}" by blast -lemma Part_QInr2: "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y \ Part(B,h)}" +lemma Part_QInr2: "Part(A <+> B, \x. QInr(h(x))) = {QInr(y). y \ Part(B,h)}" by blast lemma Part_qsum_equality: "C \ A <+> B \ Part(C,QInl) \ Part(C,QInr) = C" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/QUniv.thy --- a/src/ZF/QUniv.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/QUniv.thy Tue Sep 27 17:46:52 2022 +0100 @@ -20,7 +20,7 @@ case_eqns qcase_QInl qcase_QInr definition - quniv :: "i => i" where + quniv :: "i \ i" where "quniv(A) \ Pow(univ(eclose(A)))" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Resid/Confluence.thy --- a/src/ZF/Resid/Confluence.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Resid/Confluence.thy Tue Sep 27 17:46:52 2022 +0100 @@ -6,14 +6,14 @@ theory Confluence imports Reduction begin definition - confluence :: "i=>o" where + confluence :: "i\o" where "confluence(R) \ - \x y. \ R \ (\z. \ R \ (\u. \ R \ \ R))" + \x y. \x,y\ \ R \ (\z.\x,z\ \ R \ (\u.\y,u\ \ R \ \z,u\ \ R))" definition strip :: "o" where "strip \ \x y. (x =\ y) \ - (\z.(x =1=> z) \ (\u.(y =1=> u) \ (z=\u)))" + (\z.(x =1\ z) \ (\u.(y =1\ u) \ (z=\u)))" (* ------------------------------------------------------------------------- *) @@ -69,11 +69,11 @@ abbreviation Sconv1_rel (infixl \<-1->\ 50) where - "a<-1->b \ \ Sconv1" + "a<-1->b \ \a,b\ \ Sconv1" abbreviation Sconv_rel (infixl \<-\\ 50) where - "a<-\b \ \ Sconv" + "a<-\b \ \a,b\ \ Sconv" inductive domains "Sconv1" \ "lambda*lambda" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Resid/Redex.thy --- a/src/ZF/Resid/Redex.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Resid/Redex.thy Tue Sep 27 17:46:52 2022 +0100 @@ -19,28 +19,28 @@ abbreviation Ssub_rel (infixl \\\ 70) where - "a \ b \ \ Ssub" + "a \ b \ \a,b\ \ Ssub" abbreviation Scomp_rel (infixl \\\ 70) where - "a \ b \ \ Scomp" + "a \ b \ \a,b\ \ Scomp" abbreviation "regular(a) \ a \ Sreg" -consts union_aux :: "i=>i" +consts union_aux :: "i\i" primrec (*explicit lambda is required because both arguments of "\" vary*) "union_aux(Var(n)) = - (\t \ redexes. redexes_case(%j. Var(n), %x. 0, %b x y.0, t))" + (\t \ redexes. redexes_case(\j. Var(n), \x. 0, \b x y.0, t))" "union_aux(Fun(u)) = - (\t \ redexes. redexes_case(%j. 0, %y. Fun(union_aux(u)`y), - %b y z. 0, t))" + (\t \ redexes. redexes_case(\j. 0, \y. Fun(union_aux(u)`y), + \b y z. 0, t))" "union_aux(App(b,f,a)) = (\t \ redexes. - redexes_case(%j. 0, %y. 0, - %c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))" + redexes_case(\j. 0, \y. 0, + \c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))" definition union (infixl \\\ 70) where diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Resid/Reduction.thy --- a/src/ZF/Resid/Reduction.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Resid/Reduction.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,10 +9,10 @@ consts lambda :: "i" - unmark :: "i=>i" + unmark :: "i\i" abbreviation - Apl :: "[i,i]=>i" where + Apl :: "[i,i]\i" where "Apl(n,m) \ App(0,n,m)" inductive @@ -75,19 +75,19 @@ abbreviation Sred1_rel (infixl \-1->\ 50) where - "a -1-> b \ \ Sred1" + "a -1-> b \ \a,b\ \ Sred1" abbreviation Sred_rel (infixl \-\\ 50) where - "a -\ b \ \ Sred" + "a -\ b \ \a,b\ \ Sred" abbreviation - Spar_red1_rel (infixl \=1=>\ 50) where - "a =1=> b \ \ Spar_red1" + Spar_red1_rel (infixl \=1\\ 50) where + "a =1\ b \ \a,b\ \ Spar_red1" abbreviation Spar_red_rel (infixl \=\\ 50) where - "a =\ b \ \ Spar_red" + "a =\ b \ \a,b\ \ Spar_red" inductive @@ -115,10 +115,10 @@ inductive domains "Spar_red1" \ "lambda*lambda" intros - beta: "\m =1=> m'; n =1=> n'\ \ Apl(Fun(m),n) =1=> n'/m'" - rvar: "n \ nat \ Var(n) =1=> Var(n)" - rfun: "m =1=> m' \ Fun(m) =1=> Fun(m')" - rapl: "\m =1=> m'; n =1=> n'\ \ Apl(m,n) =1=> Apl(m',n')" + beta: "\m =1\ m'; n =1\ n'\ \ Apl(Fun(m),n) =1\ n'/m'" + rvar: "n \ nat \ Var(n) =1\ Var(n)" + rfun: "m =1\ m' \ Fun(m) =1\ Fun(m')" + rapl: "\m =1\ m'; n =1\ n'\ \ Apl(m,n) =1\ Apl(m',n')" type_intros red_typechecks declare Spar_red1.intros [intro, simp] @@ -126,7 +126,7 @@ inductive domains "Spar_red" \ "lambda*lambda" intros - one_step: "m =1=> n \ m =\ n" + one_step: "m =1\ n \ m =\ n" trans: "\m=\n; n=\p\ \ m=\p" type_intros Spar_red1.dom_subset [THEN subsetD] red_typechecks @@ -150,7 +150,7 @@ declare bool_typechecks [intro] -inductive_cases [elim!]: "Fun(t) =1=> Fun(u)" +inductive_cases [elim!]: "Fun(t) =1\ Fun(u)" @@ -190,10 +190,10 @@ (* ------------------------------------------------------------------------- *) -lemma refl_par_red1: "m \ lambda\ m =1=> m" +lemma refl_par_red1: "m \ lambda\ m =1\ m" by (erule lambda.induct, simp_all) -lemma red1_par_red1: "m-1->n \ m=1=>n" +lemma red1_par_red1: "m-1->n \ m=1\n" by (erule Sred1.induct, simp_all add: refl_par_red1) lemma red_par_red: "m-\n \ m=\n" @@ -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+) @@ -237,12 +237,12 @@ (* ------------------------------------------------------------------------- *) lemma completeness_l [rule_format]: - "u \ v \ regular(v) \ unmark(u) =1=> unmark(u|>v)" + "u \ v \ regular(v) \ unmark(u) =1\ unmark(u|>v)" apply (erule Scomp.induct) apply (auto simp add: unmmark_subst_rec) done -lemma completeness: "\u \ lambda; u \ v; regular(v)\ \ u =1=> unmark(u|>v)" +lemma completeness: "\u \ lambda; u \ v; regular(v)\ \ u =1\ unmark(u|>v)" by (drule completeness_l, simp_all add: lambda_unmark) end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Resid/Residuals.thy --- a/src/ZF/Resid/Residuals.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Resid/Residuals.thy Tue Sep 27 17:46:52 2022 +0100 @@ -26,7 +26,7 @@ type_intros subst_type nat_typechecks redexes.intros bool_typechecks definition - res_func :: "[i,i]=>i" (infixl \|>\ 70) where + res_func :: "[i,i]\i" (infixl \|>\ 70) where "u |> v \ THE w. residuals(u,v,w)" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Resid/Substitution.thy --- a/src/ZF/Resid/Substitution.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Resid/Substitution.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,7 +8,7 @@ in the recursive calls ***) consts - lift_aux :: "i=>i" + lift_aux :: "i\i" primrec "lift_aux(Var(i)) = (\k \ nat. if ik \ nat. App(b, lift_aux(f)`k, lift_aux(a)`k))" definition - lift_rec :: "[i,i]=> i" where + lift_rec :: "[i,i]\ i" where "lift_rec(r,k) \ lift_aux(r)`k" abbreviation - lift :: "i=>i" where + lift :: "i\i" where "lift(r) \ lift_rec(r,0)" consts - subst_aux :: "i=>i" + subst_aux :: "i\i" primrec "subst_aux(Var(i)) = (\r \ redexes. \k \ nat. if kr \ redexes. \k \ nat. App(b, subst_aux(f)`r`k, subst_aux(a)`r`k))" definition - subst_rec :: "[i,i,i]=> i" (**NOTE THE ARGUMENT ORDER BELOW**) where + subst_rec :: "[i,i,i]\ i" (**NOTE THE ARGUMENT ORDER BELOW**) where "subst_rec(u,r,k) \ subst_aux(r)`u`k" abbreviation - subst :: "[i,i]=>i" (infixl \'/\ 70) where + subst :: "[i,i]\i" (infixl \'/\ 70) where "u/v \ subst_rec(u,v,0)" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Sum.thy --- a/src/ZF/Sum.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Sum.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,20 +9,20 @@ text\And the "Part" primitive for simultaneous recursive type definitions\ -definition sum :: "[i,i]=>i" (infixr \+\ 65) where +definition sum :: "[i,i]\i" (infixr \+\ 65) where "A+B \ {0}*A \ {1}*B" -definition Inl :: "i=>i" where - "Inl(a) \ <0,a>" +definition Inl :: "i\i" where + "Inl(a) \ \0,a\" -definition Inr :: "i=>i" where - "Inr(b) \ <1,b>" +definition Inr :: "i\i" where + "Inr(b) \ \1,b\" -definition "case" :: "[i=>i, i=>i, i]=>i" where - "case(c,d) \ (%. cond(y, d(z), c(z)))" +definition "case" :: "[i\i, i\i, i]\i" where + "case(c,d) \ (\\y,z\. cond(y, d(z), c(z)))" (*operator for selecting out the various summands*) -definition Part :: "[i,i=>i] => i" where +definition Part :: "[i,i\i] \ i" where "Part(A,h) \ {x \ A. \z. x = h(z)}" subsection\Rules for the \<^term>\Part\ Primitive\ @@ -154,8 +154,8 @@ by auto lemma case_case: "z \ A+B \ - case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = - case(%x. c(c'(x)), %y. d(d'(y)), z)" + case(c, d, case(\x. Inl(c'(x)), \y. Inr(d'(y)), z)) = + case(\x. c(c'(x)), \y. d(d'(y)), z)" by auto @@ -179,10 +179,10 @@ lemma PartD1: "a \ Part(A,h) \ a \ A" by (simp add: Part_def) -lemma Part_id: "Part(A,%x. x) = A" +lemma Part_id: "Part(A,\x. x) = A" by blast -lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y \ Part(B,h)}" +lemma Part_Inr2: "Part(A+B, \x. Inr(h(x))) = {Inr(y). y \ Part(B,h)}" by blast lemma Part_sum_equality: "C \ A+B \ Part(C,Inl) \ Part(C,Inr) = C" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Trancl.thy --- a/src/ZF/Trancl.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Trancl.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,44 +8,44 @@ theory Trancl imports Fixedpt Perm begin definition - refl :: "[i,i]=>o" where - "refl(A,r) \ (\x\A. \ r)" + refl :: "[i,i]\o" where + "refl(A,r) \ (\x\A. \x,x\ \ r)" definition - irrefl :: "[i,i]=>o" where - "irrefl(A,r) \ \x\A. \ r" + irrefl :: "[i,i]\o" where + "irrefl(A,r) \ \x\A. \x,x\ \ r" definition - sym :: "i=>o" where - "sym(r) \ \x y. : r \ : r" + sym :: "i\o" where + "sym(r) \ \x y. \x,y\: r \ \y,x\: r" definition - asym :: "i=>o" where - "asym(r) \ \x y. :r \ \ :r" + asym :: "i\o" where + "asym(r) \ \x y. \x,y\:r \ \ \y,x\:r" definition - antisym :: "i=>o" where - "antisym(r) \ \x y.:r \ :r \ x=y" + antisym :: "i\o" where + "antisym(r) \ \x y.\x,y\:r \ \y,x\:r \ x=y" definition - trans :: "i=>o" where - "trans(r) \ \x y z. : r \ : r \ : r" + trans :: "i\o" where + "trans(r) \ \x y z. \x,y\: r \ \y,z\: r \ \x,z\: r" definition - trans_on :: "[i,i]=>o" (\trans[_]'(_')\) where + trans_on :: "[i,i]\o" (\trans[_]'(_')\) where "trans[A](r) \ \x\A. \y\A. \z\A. - : r \ : r \ : r" + \x,y\: r \ \y,z\: r \ \x,z\: r" definition - rtrancl :: "i=>i" (\(_^*)\ [100] 100) (*refl/transitive closure*) where - "r^* \ lfp(field(r)*field(r), %s. id(field(r)) \ (r O s))" + rtrancl :: "i\i" (\(_^*)\ [100] 100) (*refl/transitive closure*) where + "r^* \ lfp(field(r)*field(r), \s. id(field(r)) \ (r O s))" definition - trancl :: "i=>i" (\(_^+)\ [100] 100) (*transitive closure*) where + trancl :: "i\i" (\(_^+)\ [100] 100) (*transitive closure*) where "r^+ \ r O r^*" definition - equiv :: "[i,i]=>o" where + equiv :: "[i,i]\o" where "equiv(A,r) \ r \ A*A \ refl(A,r) \ sym(r) \ trans(r)" @@ -54,37 +54,37 @@ subsubsection\irreflexivity\ lemma irreflI: - "\\x. x \ A \ \ r\ \ irrefl(A,r)" + "\\x. x \ A \ \x,x\ \ r\ \ irrefl(A,r)" by (simp add: irrefl_def) -lemma irreflE: "\irrefl(A,r); x \ A\ \ \ r" +lemma irreflE: "\irrefl(A,r); x \ A\ \ \x,x\ \ r" by (simp add: irrefl_def) subsubsection\symmetry\ lemma symI: - "\\x y.: r \ : r\ \ sym(r)" + "\\x y.\x,y\: r \ \y,x\: r\ \ sym(r)" by (unfold sym_def, blast) -lemma symE: "\sym(r); : r\ \ : r" +lemma symE: "\sym(r); \x,y\: r\ \ \y,x\: r" by (unfold sym_def, blast) subsubsection\antisymmetry\ lemma antisymI: - "\\x y.\: r; : r\ \ x=y\ \ antisym(r)" + "\\x y.\\x,y\: r; \y,x\: r\ \ x=y\ \ antisym(r)" by (simp add: antisym_def, blast) -lemma antisymE: "\antisym(r); : r; : r\ \ x=y" +lemma antisymE: "\antisym(r); \x,y\: r; \y,x\: r\ \ x=y" by (simp add: antisym_def, blast) subsubsection\transitivity\ -lemma transD: "\trans(r); :r; :r\ \ :r" +lemma transD: "\trans(r); \a,b\:r; \b,c\:r\ \ \a,c\:r" by (unfold trans_def, blast) lemma trans_onD: - "\trans[A](r); :r; :r; a \ A; b \ A; c \ A\ \ :r" + "\trans[A](r); \a,b\:r; \b,c\:r; a \ A; b \ A; c \ A\ \ \a,c\:r" by (unfold trans_on_def, blast) lemma trans_imp_trans_on: "trans(r) \ trans[A](r)" @@ -97,7 +97,7 @@ subsection\Transitive closure of a relation\ lemma rtrancl_bnd_mono: - "bnd_mono(field(r)*field(r), %s. id(field(r)) \ (r O s))" + "bnd_mono(field(r)*field(r), \s. id(field(r)) \ (r O s))" by (rule bnd_monoI, blast+) lemma rtrancl_mono: "r<=s \ r^* \ s^*" @@ -122,19 +122,19 @@ done (*Reflexivity of rtrancl*) -lemma rtrancl_refl: "\a \ field(r)\ \ \ r^*" +lemma rtrancl_refl: "\a \ field(r)\ \ \a,a\ \ r^*" apply (rule rtrancl_unfold [THEN ssubst]) apply (erule idI [THEN UnI1]) done (*Closure under composition with r *) -lemma rtrancl_into_rtrancl: "\ \ r^*; \ r\ \ \ r^*" +lemma rtrancl_into_rtrancl: "\\a,b\ \ r^*; \b,c\ \ r\ \ \a,c\ \ r^*" apply (rule rtrancl_unfold [THEN ssubst]) apply (rule compI [THEN UnI2], assumption, assumption) done (*rtrancl of r contains all pairs in r *) -lemma r_into_rtrancl: " \ r \ \ r^*" +lemma r_into_rtrancl: "\a,b\ \ r \ \a,b\ \ r^*" by (rule rtrancl_refl [THEN rtrancl_into_rtrancl], blast+) (*The premise ensures that r consists entirely of pairs*) @@ -148,22 +148,22 @@ (** standard induction rule **) lemma rtrancl_full_induct [case_names initial step, consumes 1]: - "\ \ r^*; - \x. x \ field(r) \ P(); - \x y z.\P(); : r^*; : r\ \ P()\ - \ P()" + "\\a,b\ \ r^*; + \x. x \ field(r) \ P(\x,x\); + \x y z.\P(\x,y\); \x,y\: r^*; \y,z\: r\ \ P(\x,z\)\ + \ P(\a,b\)" by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast) (*nice induction rule. Tried adding the typing hypotheses y,z \ field(r), but these caused expensive case splits!*) lemma rtrancl_induct [case_names initial step, induct set: rtrancl]: - "\ \ r^*; + "\\a,b\ \ r^*; P(a); - \y z.\ \ r^*; \ r; P(y)\ \ P(z) + \y z.\\a,y\ \ r^*; \y,z\ \ r; P(y)\ \ P(z) \ \ P(b)" (*by induction on this formula*) -apply (subgoal_tac "\y. = \ P (y) ") +apply (subgoal_tac "\y. \a,b\ = \a,y\ \ P (y) ") (*now solve first subgoal: this formula is sufficient*) apply (erule spec [THEN mp], rule refl) (*now do the induction*) @@ -182,10 +182,10 @@ (*elimination of rtrancl -- by induction on a special formula*) lemma rtranclE: - "\ \ r^*; (a=b) \ P; - \y.\ \ r^*; \ r\ \ P\ + "\\a,b\ \ r^*; (a=b) \ P; + \y.\\a,y\ \ r^*; \y,b\ \ r\ \ P\ \ P" -apply (subgoal_tac "a = b | (\y. \ r^* \ \ r) ") +apply (subgoal_tac "a = b | (\y. \a,y\ \ r^* \ \y,b\ \ r) ") (*see HOL/trancl*) apply blast apply (erule rtrancl_induct, blast+) @@ -207,13 +207,13 @@ (** Conversions between trancl and rtrancl **) -lemma trancl_into_rtrancl: " \ r^+ \ \ r^*" +lemma trancl_into_rtrancl: "\a,b\ \ r^+ \ \a,b\ \ r^*" apply (unfold trancl_def) apply (blast intro: rtrancl_into_rtrancl) done (*r^+ contains all pairs in r *) -lemma r_into_trancl: " \ r \ \ r^+" +lemma r_into_trancl: "\a,b\ \ r \ \a,b\ \ r^+" apply (unfold trancl_def) apply (blast intro!: rtrancl_refl) done @@ -224,12 +224,12 @@ (*intro rule by definition: from r^* and r *) -lemma rtrancl_into_trancl1: "\ \ r^*; \ r\ \ \ r^+" +lemma rtrancl_into_trancl1: "\\a,b\ \ r^*; \b,c\ \ r\ \ \a,c\ \ r^+" by (unfold trancl_def, blast) (*intro rule from r and r^* *) lemma rtrancl_into_trancl2: - "\ \ r; \ r^*\ \ \ r^+" + "\\a,b\ \ r; \b,c\ \ r^*\ \ \a,c\ \ r^+" apply (erule rtrancl_induct) apply (erule r_into_trancl) apply (blast intro: r_into_trancl trancl_trans) @@ -237,14 +237,14 @@ (*Nice induction rule for trancl*) lemma trancl_induct [case_names initial step, induct set: trancl]: - "\ \ r^+; - \y. \ \ r\ \ P(y); - \y z.\ \ r^+; \ r; P(y)\ \ P(z) + "\\a,b\ \ r^+; + \y. \\a,y\ \ r\ \ P(y); + \y z.\\a,y\ \ r^+; \y,z\ \ r; P(y)\ \ P(z) \ \ P(b)" apply (rule compEpair) apply (unfold trancl_def, assumption) (*by induction on this formula*) -apply (subgoal_tac "\z. \ r \ P (z) ") +apply (subgoal_tac "\z. \y,z\ \ r \ P (z) ") (*now solve first subgoal: this formula is sufficient*) apply blast apply (erule rtrancl_induct) @@ -253,11 +253,11 @@ (*elimination of r^+ -- NOT an induction rule*) lemma tranclE: - "\ \ r^+; - \ r \ P; - \y.\ \ r^+; \ r\ \ P + "\\a,b\ \ r^+; + \a,b\ \ r \ P; + \y.\\a,y\ \ r^+; \y,b\ \ r\ \ P \ \ P" -apply (subgoal_tac " \ r | (\y. \ r^+ \ \ r) ") +apply (subgoal_tac "\a,b\ \ r | (\y. \a,y\ \ r^+ \ \y,b\ \ r) ") apply blast apply (rule compEpair) apply (unfold trancl_def, assumption) @@ -320,7 +320,7 @@ (** rtrancl **) -lemma rtrancl_converseD: ":converse(r)^* \ :converse(r^*)" +lemma rtrancl_converseD: "\x,y\:converse(r)^* \ \x,y\:converse(r^*)" apply (rule converseI) apply (frule rtrancl_type [THEN subsetD]) apply (erule rtrancl_induct) @@ -328,7 +328,7 @@ apply (blast intro: r_into_rtrancl rtrancl_trans) done -lemma rtrancl_converseI: ":converse(r^*) \ :converse(r)^*" +lemma rtrancl_converseI: "\x,y\:converse(r^*) \ \x,y\:converse(r)^*" apply (drule converseD) apply (frule rtrancl_type [THEN subsetD]) apply (erule rtrancl_induct) @@ -344,12 +344,12 @@ (** trancl **) -lemma trancl_converseD: ":converse(r)^+ \ :converse(r^+)" +lemma trancl_converseD: "\a, b\:converse(r)^+ \ \a, b\:converse(r^+)" apply (erule trancl_induct) apply (auto intro: r_into_trancl trancl_trans) done -lemma trancl_converseI: ":converse(r^+) \ :converse(r)^+" +lemma trancl_converseI: "\x,y\:converse(r^+) \ \x,y\:converse(r)^+" apply (drule converseD) apply (erule trancl_induct) apply (auto intro: r_into_trancl trancl_trans) @@ -362,8 +362,8 @@ done lemma converse_trancl_induct [case_names initial step, consumes 1]: -"\:r^+; \y. :r \ P(y); - \y z. \ \ r; \ r^+; P(z)\ \ P(y)\ +"\\a, b\:r^+; \y. \y, b\ :r \ P(y); + \y z. \\y, z\ \ r; \z, b\ \ r^+; P(z)\ \ P(y)\ \ P(a)" apply (drule converseI) apply (simp (no_asm_use) add: trancl_converse [symmetric]) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/AllocBase.thy --- a/src/ZF/UNITY/AllocBase.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/AllocBase.thy Tue Sep 27 17:46:52 2022 +0100 @@ -20,13 +20,13 @@ Nclients_pos: "Nclients \ nat-{0}" text\This function merely sums the elements of a list\ -consts tokens :: "i =>i" +consts tokens :: "i \i" item :: i (* Items to be merged/distributed *) primrec "tokens(Nil) = 0" "tokens (Cons(x,xs)) = x #+ tokens(xs)" -consts bag_of :: "i => i" +consts bag_of :: "i \ i" primrec "bag_of(Nil) = 0" "bag_of(Cons(x,xs)) = {#x#} +# bag_of(xs)" @@ -34,28 +34,28 @@ text\Definitions needed in Client.thy. We define a recursive predicate using 0 and 1 to code the truth values.\ -consts all_distinct0 :: "i=>i" +consts all_distinct0 :: "i\i" primrec "all_distinct0(Nil) = 1" "all_distinct0(Cons(a, l)) = (if a \ set_of_list(l) then 0 else all_distinct0(l))" definition - all_distinct :: "i=>o" where + all_distinct :: "i\o" where "all_distinct(l) \ all_distinct0(l)=1" definition - state_of :: "i =>i" \ \coersion from anyting to state\ where + state_of :: "i \i" \ \coersion from anyting to state\ where "state_of(s) \ if s \ state then s else st0" definition - lift :: "i =>(i=>i)" \ \simplifies the expression of programs\ where - "lift(x) \ %s. s`x" + lift :: "i \(i\i)" \ \simplifies the expression of programs\ where + "lift(x) \ \s. s`x" text\function to show that the set of variables is infinite\ consts - nat_list_inj :: "i=>i" - var_inj :: "i=>i" + nat_list_inj :: "i\i" + var_inj :: "i\i" primrec "nat_list_inj(0) = Nil" @@ -65,7 +65,7 @@ "var_inj(Var(l)) = length(l)" definition - nat_var_inj :: "i=>i" where + nat_var_inj :: "i\i" where "nat_var_inj(n) \ Var(nat_list_inj(n))" @@ -106,13 +106,13 @@ by (erule list.induct, auto) lemma tokens_mono_aux [rule_format]: - "xs\list(A) \ \ys\list(A). \prefix(A) + "xs\list(A) \ \ys\list(A). \xs, ys\\prefix(A) \ tokens(xs) \ tokens(ys)" apply (induct_tac "xs") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: prefix_def) done -lemma tokens_mono: "\prefix(A) \ tokens(xs) \ tokens(ys)" +lemma tokens_mono: "\xs, ys\\prefix(A) \ tokens(xs) \ tokens(ys)" apply (cut_tac prefix_type) apply (blast intro: tokens_mono_aux) done @@ -147,7 +147,7 @@ done lemma bag_of_mono_aux [rule_format]: - "xs\list(A) \ \ys\list(A). \prefix(A) + "xs\list(A) \ \ys\list(A). \xs, ys\\prefix(A) \ \MultLe(A, r)" apply (induct_tac "xs", simp_all, clarify) apply (frule_tac l = ys in bag_of_multiset) @@ -158,7 +158,7 @@ done lemma bag_of_mono [intro]: - "\\prefix(A); xs\list(A); ys\list(A)\ + "\\xs, ys\\prefix(A); xs\list(A); ys\list(A)\ \ \MultLe(A, r)" apply (blast intro: bag_of_mono_aux) done @@ -212,7 +212,7 @@ "l\list(A) \ C \ nat \ bag_of(sublist(l, C)) = - msetsum(%i. {#nth(i, l)#}, C \ length(l), A)" + msetsum(\i. {#nth(i, l)#}, C \ length(l), A)" apply (erule list_append_induct) apply (simp (no_asm)) apply (simp (no_asm_simp) add: sublist_append nth_append bag_of_sublist_lemma munion_commute bag_of_sublist_lemma msetsum_multiset munion_0) @@ -227,8 +227,8 @@ (*eliminating the assumption C<=nat*) lemma bag_of_sublist: "l\list(A) \ - bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C \ length(l), A)" -apply (subgoal_tac " bag_of (sublist (l, C \ nat)) = msetsum (%i. {#nth (i, l) #}, C \ length (l), A) ") + bag_of(sublist(l, C)) = msetsum(\i. {#nth(i, l)#}, C \ length(l), A)" +apply (subgoal_tac " bag_of (sublist (l, C \ nat)) = msetsum (\i. {#nth (i, l) #}, C \ length (l), A) ") apply (simp add: sublist_Int_eq) apply (simp add: bag_of_sublist_lemma2 Int_lower2 Int_assoc nat_Int_length_eq) done @@ -257,7 +257,7 @@ "\Finite(I); \i\I. \j\I. i\j \ A(i) \ A(j) = 0; l\list(B)\ \ bag_of(sublist(l, \i\I. A(i))) = - (msetsum(%i. bag_of(sublist(l, A(i))), I, B)) " + (msetsum(\i. bag_of(sublist(l, A(i))), I, B)) " apply (simp (no_asm_simp) del: UN_simps add: UN_simps [symmetric] bag_of_sublist) apply (subst msetsum_UN_disjoint [of _ _ _ "length (l)"]) @@ -308,7 +308,7 @@ (** Used in ClientImp **) lemma gen_Increains_state_of_eq: - "Increasing(A, r, %s. f(state_of(s))) = Increasing(A, r, f)" + "Increasing(A, r, \s. f(state_of(s))) = Increasing(A, r, f)" apply (unfold Increasing_def, auto) done @@ -318,7 +318,7 @@ gen_Increains_state_of_eq [THEN equalityD2, THEN subsetD] lemma Follows_state_of_eq: - "Follows(A, r, %s. f(state_of(s)), %s. g(state_of(s))) = + "Follows(A, r, \s. f(state_of(s)), \s. g(state_of(s))) = Follows(A, r, f, g)" apply (unfold Follows_def Increasing_def, auto) done @@ -383,7 +383,7 @@ lemma setsum_nsetsum_eq: "\Finite(A); \x\A. g(x)\nat\ - \ setsum(%x. $#(g(x)), A) = $# nsetsum(%x. g(x), A)" + \ setsum(\x. $#(g(x)), A) = $# nsetsum(\x. g(x), A)" apply (erule Finite_induct) apply (auto simp add: int_of_add) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/AllocImpl.thy --- a/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/AllocImpl.thy Tue Sep 27 17:46:52 2022 +0100 @@ -25,7 +25,7 @@ definition "alloc_giv_act \ - { \ state*state. + {\s, t\ \ state*state. \k. k = length(s`giv) \ t = s(giv := s`giv @ [nth(k, s`ask)], available_tok := s`available_tok #- nth(k, s`ask)) \ @@ -33,7 +33,7 @@ definition "alloc_rel_act \ - { \ state*state. + {\s, t\ \ state*state. t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel), NbR := succ(s`NbR)) \ s`NbR < length(s`rel)}" @@ -202,7 +202,7 @@ apply (auto intro!: AlwaysI simp add: guar_def) apply (subgoal_tac "G \ preserves (lift (giv))") prefer 2 apply (simp add: alloc_prog_ok_iff) -apply (rule_tac P = "%x y. \ prefix(tokbag)" and A = "list(nat)" +apply (rule_tac P = "\x y. \x,y\ \ prefix(tokbag)" and A = "list(nat)" in stable_Join_Stable) apply safety prefer 2 apply (simp add: lift_def, clarify) @@ -573,7 +573,7 @@ apply (rule LeadsTo_weaken_R) apply (rule Always_LeadsToD [OF safety length_ask_giv], assumption+, clarify) apply (simp add: INT_iff) -apply (drule_tac x = "length(x ` giv)" and P = "%x. f (x) \ NbT" for f in bspec) +apply (drule_tac x = "length(x ` giv)" and P = "\x. f (x) \ NbT" for f in bspec) apply simp apply (blast intro: le_trans) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/ClientImpl.thy --- a/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/ClientImpl.thy Tue Sep 27 17:46:52 2022 +0100 @@ -26,7 +26,7 @@ definition (** Release some client_tokens **) "client_rel_act \ - { \ state*state. + {\s,t\ \ state*state. \nrel \ nat. nrel = length(s`rel) \ t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) \ nrel < length(s`giv) \ @@ -36,11 +36,11 @@ (** Including t=s suppresses fairness, allowing the non-trivial part of the action to be ignored **) definition - "client_tok_act \ { \ state*state. t=s | + "client_tok_act \ {\s,t\ \ state*state. t=s | t = s(tok:=succ(s`tok mod NbT))}" definition - "client_ask_act \ { \ state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" + "client_ask_act \ {\s,t\ \ state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}" definition "client_prog \ @@ -178,15 +178,15 @@ by (force intro!: AlwaysI client_prog_Join_Stable_rel_le_giv) lemma def_act_eq: - "A \ { \ state*state. P(s, t)} \ A={ \ state*state. P(s, t)}" + "A \ {\s, t\ \ state*state. P(s, t)} \ A={\s, t\ \ state*state. P(s, t)}" by auto -lemma act_subset: "A={ \ state*state. P(s, t)} \ A<=state*state" +lemma act_subset: "A={\s,t\ \ state*state. P(s, t)} \ A<=state*state" by auto lemma transient_lemma: "client_prog \ - transient({s \ state. s`rel = k \ \ strict_prefix(nat) + transient({s \ state. s`rel = k \ \k, h\ \ strict_prefix(nat) \ \ prefix(nat) \ h pfixGe s`ask})" apply (rule_tac act = client_rel_act in transientI) apply (simp (no_asm) add: client_prog_def [THEN def_prg_Acts]) @@ -209,7 +209,7 @@ done lemma strict_prefix_is_prefix: - " \ strict_prefix(A) \ \ prefix(A) \ xs\ys" + "\xs, ys\ \ strict_prefix(A) \ \xs, ys\ \ prefix(A) \ xs\ys" apply (unfold strict_prefix_def id_def lam_def) apply (auto dest: prefix_type [THEN subsetD]) done @@ -217,7 +217,7 @@ 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) + {s \ state. s`rel = k \ \k,h\ \ strict_prefix(nat) \ \ prefix(nat) \ h pfixGe s`ask} \w {s \ state. \ strict_prefix(nat) \ \ prefix(nat) \ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Comp.thy --- a/src/ZF/UNITY/Comp.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Comp.thy Tue Sep 27 17:46:52 2022 +0100 @@ -18,33 +18,33 @@ theory Comp imports Union Increasing begin definition - component :: "[i,i]=>o" (infixl \component\ 65) where + component :: "[i,i]\o" (infixl \component\ 65) where "F component H \ (\G. F \ G = H)" definition - strict_component :: "[i,i]=>o" (infixl \strict'_component\ 65) where + strict_component :: "[i,i]\o" (infixl \strict'_component\ 65) where "F strict_component H \ F component H \ F\H" definition (* A stronger form of the component relation *) - component_of :: "[i,i]=>o" (infixl \component'_of\ 65) where + component_of :: "[i,i]\o" (infixl \component'_of\ 65) where "F component_of H \ (\G. F ok G \ F \ G = H)" definition - strict_component_of :: "[i,i]=>o" (infixl \strict'_component'_of\ 65) where + strict_component_of :: "[i,i]\o" (infixl \strict'_component'_of\ 65) where "F strict_component_of H \ F component_of H \ F\H" definition (* Preserves a state function f, in particular a variable *) - preserves :: "(i=>i)=>i" where + preserves :: "(i\i)\i" where "preserves(f) \ {F:program. \z. F: stable({s:state. f(s) = z})}" definition - fun_pair :: "[i=>i, i =>i] =>(i=>i)" where - "fun_pair(f, g) \ %x. " + fun_pair :: "[i\i, i \i] \(i\i)" where + "fun_pair(f, g) \ \x. " definition - localize :: "[i=>i, i] => i" where + localize :: "[i\i, i] \ i" where "localize(f,F) \ mk_program(Init(F), Acts(F), AllowedActs(F) \ (\G\preserves(f). Acts(G)))" @@ -156,7 +156,7 @@ done lemma preserves_imp_eq: - "\F \ preserves(f); act \ Acts(F); \ act\ \ f(s) = f(t)" + "\F \ preserves(f); act \ Acts(F); \s,t\ \ act\ \ f(s) = f(t)" apply (unfold preserves_def stable_def constrains_def) apply (subgoal_tac "s \ state \ t \ state") prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) @@ -226,11 +226,11 @@ "\F \ preserves(f); \x \ state. f(x):A\ \ F \ Increasing.increasing(A, r, f)" apply (unfold Increasing.increasing_def) apply (auto intro: preserves_into_program) -apply (rule_tac P = "%x. :r" in preserves_imp_stable, auto) +apply (rule_tac P = "\x. \k, x\:r" in preserves_imp_stable, auto) done lemma preserves_id_subset_stable: - "st_set(A) \ preserves(%x. x) \ stable(A)" + "st_set(A) \ preserves(\x. x) \ stable(A)" apply (unfold preserves_def stable_def constrains_def, auto) apply (drule_tac x = xb in spec) apply (drule_tac x = act in bspec) @@ -238,7 +238,7 @@ done lemma preserves_id_imp_stable: - "\F \ preserves(%x. x); st_set(A)\ \ F \ stable(A)" + "\F \ preserves(\x. x); st_set(A)\ \ F \ stable(A)" by (blast intro: preserves_id_subset_stable [THEN subsetD]) (** Added by Sidi **) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Constrains.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,8 +9,8 @@ imports UNITY begin -consts traces :: "[i, i] => i" - (* Initial states and program => (final state, reversed trace to it)... +consts traces :: "[i, i] \ i" + (* Initial states and program \ (final state, reversed trace to it)... the domain may also be state*list(state) *) inductive domains @@ -20,13 +20,13 @@ (*Initial trace is empty*) Init: "s: init \ \ traces(init,acts)" - Acts: "\act:acts; \ traces(init,acts); : act\ + Acts: "\act:acts; \s,evs\ \ traces(init,acts); : act\ \ \ traces(init, acts)" type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1 -consts reachable :: "i=>i" +consts reachable :: "i\i" inductive domains "reachable(F)" \ "Init(F) \ (\act\Acts(F). field(act))" @@ -40,20 +40,20 @@ definition - Constrains :: "[i,i] => i" (infixl \Co\ 60) where + Constrains :: "[i,i] \ i" (infixl \Co\ 60) where "A Co B \ {F:program. F:(reachable(F) \ A) co B}" definition - op_Unless :: "[i, i] => i" (infixl \Unless\ 60) where + op_Unless :: "[i, i] \ i" (infixl \Unless\ 60) where "A Unless B \ (A-B) Co (A \ B)" definition - Stable :: "i => i" where + Stable :: "i \ i" where "Stable(A) \ A Co A" definition (*Always is the weak form of "invariant"*) - Always :: "i => i" where + Always :: "i \ i" where "Always(A) \ initially(A) \ Stable(A)" @@ -80,7 +80,7 @@ declare state_Int_reachable [iff] lemma reachable_equiv_traces: -"F \ program \ reachable(F)={s \ state. \evs. :traces(Init(F), Acts(F))}" +"F \ program \ reachable(F)={s \ state. \evs. \s,evs\:traces(Init(F), Acts(F))}" apply (rule equalityI, safe) apply (blast dest: reachable_type [THEN subsetD]) apply (erule_tac [2] traces.induct) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Distributor.thy --- a/src/ZF/UNITY/Distributor.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Distributor.thy Tue Sep 27 17:46:52 2022 +0100 @@ -13,7 +13,7 @@ text\spec (14)\ definition - distr_follows :: "[i, i, i, i =>i] =>i" where + distr_follows :: "[i, i, i, i \i] \i" where "distr_follows(A, In, iIn, Out) \ (lift(In) IncreasingWrt prefix(A)/list(A)) \ (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \ @@ -22,17 +22,17 @@ (\n \ Nclients. lift(Out(n)) Fols - (%s. sublist(s`In, {k \ nat. k nth(k, s`iIn) = n})) + (\s. sublist(s`In, {k \ nat. k nth(k, s`iIn) = n})) Wrt prefix(A)/list(A))" definition - distr_allowed_acts :: "[i=>i]=>i" where + distr_allowed_acts :: "[i\i]\i" where "distr_allowed_acts(Out) \ {D \ program. AllowedActs(D) = cons(id(state), \G \ (\n\nat. preserves(lift(Out(n)))). Acts(G))}" definition - distr_spec :: "[i, i, i, i =>i]=>i" where + distr_spec :: "[i, i, i, i \i]\i" where "distr_spec(A, In, iIn, Out) \ distr_follows(A, In, iIn, Out) \ distr_allowed_acts(Out)" @@ -105,7 +105,7 @@ D \ G \ Always({s \ state. \elt \ set_of_list(s`iIn). elt < Nclients})\ \ D \ G \ Always - ({s \ state. msetsum(%n. bag_of (sublist(s`In, + ({s \ state. msetsum(\n. bag_of (sublist(s`In, {k \ nat. k < length(s`iIn) \ nth(k, s`iIn)= n})), Nclients, A) = @@ -131,7 +131,7 @@ apply (subgoal_tac "length (s ` iIn) \ nat") apply typecheck apply (subgoal_tac "m \ nat") -apply (drule_tac x = "nth(m, s`iIn) " and P = "%elt. X (elt) \ elt state. \elt \ set_of_list(s`iIn). elt < Nclients})) guarantees (\n \ Nclients. - (%s. msetsum(%i. bag_of(s`Out(i)), Nclients, A)) + (\s. msetsum(\i. bag_of(s`Out(i)), Nclients, A)) Fols - (%s. bag_of(sublist(s`In, length(s`iIn)))) + (\s. bag_of(sublist(s`In, length(s`iIn)))) Wrt MultLe(A, r)/Mult(A))" apply (cut_tac distr_spec) apply (rule guaranteesI, clarify) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/FP.thy --- a/src/ZF/UNITY/FP.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/FP.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,11 +12,11 @@ theory FP imports UNITY begin definition - FP_Orig :: "i=>i" where + FP_Orig :: "i\i" where "FP_Orig(F) \ \({A \ Pow(state). \B. F \ stable(A \ B)})" definition - FP :: "i=>i" where + FP :: "i\i" where "FP(F) \ {s\state. F \ stable({s})}" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Follows.thy --- a/src/ZF/UNITY/Follows.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Follows.thy Tue Sep 27 17:46:52 2022 +0100 @@ -10,7 +10,7 @@ theory Follows imports SubstAx Increasing begin definition - Follows :: "[i, i, i=>i, i=>i] => i" where + Follows :: "[i, i, i\i, i\i] \ i" where "Follows(A, r, f, g) \ Increasing(A, r, g) Int Increasing(A, r,f) Int @@ -18,27 +18,27 @@ (\k \ A. {s \ state. :r} \w {s \ state. :r})" abbreviation - Incr :: "[i=>i]=>i" where + Incr :: "[i\i]\i" where "Incr(f) \ Increasing(list(nat), prefix(nat), f)" abbreviation - n_Incr :: "[i=>i]=>i" where + n_Incr :: "[i\i]\i" where "n_Incr(f) \ Increasing(nat, Le, f)" abbreviation - s_Incr :: "[i=>i]=>i" where + s_Incr :: "[i\i]\i" where "s_Incr(f) \ Increasing(Pow(nat), SetLe(nat), f)" abbreviation - m_Incr :: "[i=>i]=>i" where + m_Incr :: "[i\i]\i" where "m_Incr(f) \ Increasing(Mult(nat), MultLe(nat, Le), f)" abbreviation - n_Fols :: "[i=>i, i=>i]=>i" (infixl \n'_Fols\ 65) where + n_Fols :: "[i\i, i\i]\i" (infixl \n'_Fols\ 65) where "f n_Fols g \ Follows(nat, Le, f, g)" abbreviation - Follows' :: "[i=>i, i=>i, i, i] => i" + Follows' :: "[i\i, i\i, i, i] \ i" (\(_ /Fols _ /Wrt (_ /'/ _))\ [60, 0, 0, 60] 60) where "f Fols g Wrt r/A \ Follows(A,r,f,g)" @@ -114,14 +114,14 @@ apply (rule single_LeadsTo_I, auto) apply (drule_tac x = "g (sa) " and A = B in bspec) apply auto -apply (drule_tac x = "f (sa) " and P = "%j. F \ X(j) \w Y(j)" for X Y in bspec) +apply (drule_tac x = "f (sa) " and P = "\j. F \ X(j) \w Y(j)" for X Y in bspec) apply auto apply (rule PSP_Stable [THEN LeadsTo_weaken], blast, blast) apply auto apply (force simp add: part_order_def refl_def) apply (force simp add: part_order_def refl_def) -apply (drule_tac x = "f1 (x)" and x1 = "f (sa) " and P2 = "%x y. \u\B. P (x,y,u)" for P in bspec [THEN bspec]) -apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "%x y. P (x,y) \ d (x,y) \ t" for P d in bspec [THEN bspec]) +apply (drule_tac x = "f1 (x)" and x1 = "f (sa) " and P2 = "\x y. \u\B. P (x,y,u)" for P in bspec [THEN bspec]) +apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "\x y. P (x,y) \ d (x,y) \ t" for P d in bspec [THEN bspec]) apply auto apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD) apply (auto simp add: part_order_def) @@ -135,7 +135,7 @@ F:{x \ state. \ t} \w {x \ state. \ t}" apply (unfold mono2_def Increasing_def) apply (rule single_LeadsTo_I, auto) -apply (drule_tac x = "f (sa) " and P = "%k. F \ Stable (X (k))" for X in bspec) +apply (drule_tac x = "f (sa) " and P = "\k. F \ Stable (X (k))" for X in bspec) apply auto apply (drule_tac x = "g (sa) " in bspec) apply auto @@ -144,7 +144,7 @@ apply (force simp add: part_order_def refl_def) apply (force simp add: part_order_def refl_def) apply (drule_tac x = "f (x) " and x1 = "f (sa) " in bspec [THEN bspec]) -apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "%x y. P (x,y) \ d (x,y) \ t" for P d in bspec [THEN bspec]) +apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "\x y. P (x,y) \ d (x,y) \ t" for P d in bspec [THEN bspec]) apply auto apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD) apply (auto simp add: part_order_def) @@ -180,7 +180,7 @@ done lemma Follows_constantI: - "\F \ program; c \ A; refl(A, r)\ \ F \ Follows(A, r, %x. c, %x. c)" + "\F \ program; c \ A; refl(A, r)\ \ F \ Follows(A, r, \x. c, \x. c)" apply (unfold Follows_def, auto) apply (auto simp add: refl_def) done @@ -210,7 +210,7 @@ lemma imp_Follows_comp2: "\F \ Follows(A, r, f1, f); F \ Follows(B, s, g1, g); mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t)\ - \ F \ Follows(C, t, %x. h(f1(x), g1(x)), %x. h(f(x), g(x)))" + \ F \ Follows(C, t, \x. h(f1(x), g1(x)), \x. h(f(x), g(x)))" apply (unfold Follows_def, clarify) apply (frule_tac f = g in IncreasingD) apply (frule_tac f = f in IncreasingD) @@ -324,13 +324,13 @@ lemma increasing_Un: "\F \ Increasing.increasing(Pow(A), SetLe(A), f); F \ Increasing.increasing(Pow(A), SetLe(A), g)\ - \ F \ Increasing.increasing(Pow(A), SetLe(A), %x. f(x) \ g(x))" + \ F \ Increasing.increasing(Pow(A), SetLe(A), \x. f(x) \ g(x))" by (rule_tac h = "(Un)" in imp_increasing_comp2, auto) lemma Increasing_Un: "\F \ Increasing(Pow(A), SetLe(A), f); F \ Increasing(Pow(A), SetLe(A), g)\ - \ F \ Increasing(Pow(A), SetLe(A), %x. f(x) \ g(x))" + \ F \ Increasing(Pow(A), SetLe(A), \x. f(x) \ g(x))" by (rule_tac h = "(Un)" in imp_Increasing_comp2, auto) lemma Always_Un: @@ -342,7 +342,7 @@ lemma Follows_Un: "\F \ Follows(Pow(A), SetLe(A), f1, f); F \ Follows(Pow(A), SetLe(A), g1, g)\ - \ F \ Follows(Pow(A), SetLe(A), %s. f1(s) \ g1(s), %s. f(s) \ g(s))" + \ F \ Follows(Pow(A), SetLe(A), \s. f1(s) \ g1(s), \s. f(s) \ g(s))" by (rule_tac h = "(Un)" in imp_Follows_comp2, auto) (** Multiset union properties (with the MultLe ordering) **) @@ -351,12 +351,12 @@ by (unfold MultLe_def refl_def, auto) lemma MultLe_refl1 [simp]: - "\multiset(M); mset_of(M)<=A\ \ \ MultLe(A, r)" + "\multiset(M); mset_of(M)<=A\ \ \M, M\ \ MultLe(A, r)" apply (unfold MultLe_def id_def lam_def) apply (auto simp add: Mult_iff_multiset) done -lemma MultLe_refl2 [simp]: "M \ Mult(A) \ \ MultLe(A, r)" +lemma MultLe_refl2 [simp]: "M \ Mult(A) \ \M, M\ \ MultLe(A, r)" by (unfold MultLe_def id_def lam_def, auto) @@ -371,7 +371,7 @@ done lemma MultLe_trans: - "\ \ MultLe(A,r); \ MultLe(A,r)\ \ \ MultLe(A,r)" + "\\M,K\ \ MultLe(A,r); \K,N\ \ MultLe(A,r)\ \ \M,N\ \ MultLe(A,r)" apply (cut_tac A=A in trans_on_MultLe) apply (drule trans_onD, assumption) apply (auto dest: MultLe_type [THEN subsetD]) @@ -405,7 +405,7 @@ done lemma empty_le_MultLe [simp]: -"\multiset(M); mset_of(M)<= A\ \ <0, M> \ MultLe(A, r)" +"\multiset(M); mset_of(M)<= A\ \ \0, M\ \ MultLe(A, r)" apply (unfold MultLe_def) apply (case_tac "M=0") apply (auto simp add: FiniteFun.intros) @@ -414,11 +414,11 @@ apply (auto simp add: Mult_iff_multiset) done -lemma empty_le_MultLe2 [simp]: "M \ Mult(A) \ <0, M> \ MultLe(A, r)" +lemma empty_le_MultLe2 [simp]: "M \ Mult(A) \ \0, M\ \ MultLe(A, r)" by (simp add: Mult_iff_multiset) lemma munion_mono: -"\ \ MultLe(A, r); \ MultLe(A, r)\ \ +"\\M, N\ \ MultLe(A, r); \K, L\ \ MultLe(A, r)\ \ \ MultLe(A, r)" apply (unfold MultLe_def) apply (auto intro: munion_multirel_mono1 munion_multirel_mono2 @@ -428,13 +428,13 @@ lemma increasing_munion: "\F \ Increasing.increasing(Mult(A), MultLe(A,r), f); F \ Increasing.increasing(Mult(A), MultLe(A,r), g)\ - \ F \ Increasing.increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))" + \ F \ Increasing.increasing(Mult(A),MultLe(A,r), \x. f(x) +# g(x))" by (rule_tac h = munion in imp_increasing_comp2, auto) lemma Increasing_munion: "\F \ Increasing(Mult(A), MultLe(A,r), f); F \ Increasing(Mult(A), MultLe(A,r), g)\ - \ F \ Increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))" + \ F \ Increasing(Mult(A),MultLe(A,r), \x. f(x) +# g(x))" by (rule_tac h = munion in imp_Increasing_comp2, auto) lemma Always_munion: @@ -449,7 +449,7 @@ lemma Follows_munion: "\F \ Follows(Mult(A), MultLe(A, r), f1, f); F \ Follows(Mult(A), MultLe(A, r), g1, g)\ - \ F \ Follows(Mult(A), MultLe(A, r), %s. f1(s) +# g1(s), %s. f(s) +# g(s))" + \ F \ Follows(Mult(A), MultLe(A, r), \s. f1(s) +# g1(s), \s. f(s) +# g(s))" by (rule_tac h = munion in imp_Follows_comp2, auto) (** Used in ClientImp **) @@ -460,8 +460,8 @@ multiset(f(i, s)) \ mset_of(f(i, s))<=A ; Finite(I); F \ program\ \ F \ Follows(Mult(A), - MultLe(A, r), %x. msetsum(%i. f'(i, x), I, A), - %x. msetsum(%i. f(i, x), I, A))" + MultLe(A, r), \x. msetsum(\i. f'(i, x), I, A), + \x. msetsum(\i. f(i, x), I, A))" apply (erule rev_mp) apply (drule Finite_into_Fin) apply (erule Fin_induct) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/GenPrefix.thy --- a/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/GenPrefix.thy Tue Sep 27 17:46:52 2022 +0100 @@ -2,7 +2,7 @@ Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2001 University of Cambridge -:gen_prefix(r) +\xs,ys\:gen_prefix(r) if ys = xs' @ zs where length(xs) = length(xs') and corresponding elements of xs, xs' are pairwise related by r @@ -16,11 +16,11 @@ begin definition (*really belongs in ZF/Trancl*) - part_order :: "[i, i] => o" where + part_order :: "[i, i] \ o" where "part_order(A, r) \ refl(A,r) \ trans[A](r) \ antisym(r)" consts - gen_prefix :: "[i, i] => i" + gen_prefix :: "[i, i] \ i" inductive (* Parameter A is the domain of zs's elements *) @@ -30,35 +30,35 @@ intros Nil: "<[],[]>:gen_prefix(A, r)" - prepend: "\:gen_prefix(A, r); :r; x \ A; y \ A\ + prepend: "\\xs,ys\:gen_prefix(A, r); \x,y\:r; x \ A; y \ A\ \ : gen_prefix(A, r)" - append: "\:gen_prefix(A, r); zs:list(A)\ + append: "\\xs,ys\:gen_prefix(A, r); zs:list(A)\ \ :gen_prefix(A, r)" type_intros app_type list.Nil list.Cons definition - prefix :: "i=>i" where + prefix :: "i\i" where "prefix(A) \ gen_prefix(A, id(A))" definition - strict_prefix :: "i=>i" where + strict_prefix :: "i\i" where "strict_prefix(A) \ prefix(A) - id(list(A))" (* less or equal and greater or equal over prefixes *) abbreviation - pfixLe :: "[i, i] => o" (infixl \pfixLe\ 50) where - "xs pfixLe ys \ :gen_prefix(nat, Le)" + pfixLe :: "[i, i] \ o" (infixl \pfixLe\ 50) where + "xs pfixLe ys \ \xs, ys\:gen_prefix(nat, Le)" abbreviation - pfixGe :: "[i, i] => o" (infixl \pfixGe\ 50) where - "xs pfixGe ys \ :gen_prefix(nat, Ge)" + pfixGe :: "[i, i] \ o" (infixl \pfixGe\ 50) where + "xs pfixGe ys \ \xs, ys\:gen_prefix(nat, Ge)" lemma reflD: - "\refl(A, r); x \ A\ \ :r" + "\refl(A, r); x \ A\ \ \x,x\:r" apply (unfold refl_def, auto) done @@ -69,7 +69,7 @@ declare Nil_gen_prefix [simp] -lemma gen_prefix_length_le: " \ gen_prefix(A, r) \ length(xs) \ length(ys)" +lemma gen_prefix_length_le: "\xs,ys\ \ gen_prefix(A, r) \ length(xs) \ length(ys)" apply (erule gen_prefix.induct) apply (subgoal_tac [3] "ys \ list (A) ") apply (auto dest: gen_prefix.dom_subset [THEN subsetD] @@ -81,15 +81,15 @@ "\ \ 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)))" + \x,y\:r \ \xs, ys\ \ gen_prefix(A, r)))" apply (erule gen_prefix.induct) prefer 3 apply (force intro: gen_prefix.append, auto) done lemma Cons_gen_prefixE: "\ \ gen_prefix(A, r); - \y ys. \zs = Cons(y, ys); y \ A; x \ A; :r; - \ gen_prefix(A, r)\ \ P\ \ P" + \y ys. \zs = Cons(y, ys); y \ A; x \ A; \x,y\:r; + \xs,ys\ \ gen_prefix(A, r)\ \ P\ \ P" apply (frule gen_prefix.dom_subset [THEN subsetD], auto) apply (blast dest: Cons_gen_prefix_aux) done @@ -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 \ \x,y\:r \ \xs,ys\ \ gen_prefix(A, r))" apply (auto intro: gen_prefix.prepend) done declare Cons_gen_prefix_Cons [iff] @@ -145,7 +145,7 @@ (* A lemma for proving gen_prefix_trans_comp *) lemma append_gen_prefix [rule_format (no_asm)]: "xs \ list(A) \ - \zs. \ gen_prefix(A, r) \ : gen_prefix(A, r)" + \zs. \ gen_prefix(A, r) \ \xs, zs\: gen_prefix(A, r)" apply (erule list.induct) apply (auto dest: gen_prefix.dom_subset [THEN subsetD]) done @@ -153,8 +153,8 @@ (* Lemma proving transitivity and more*) lemma gen_prefix_trans_comp [rule_format (no_asm)]: - ": gen_prefix(A, r) \ - (\z \ list(A). \ gen_prefix(A, s)\ \ gen_prefix(A, s O r))" + "\x, y\: gen_prefix(A, r) \ + (\z \ list(A). \y,z\ \ gen_prefix(A, s)\\x, z\ \ gen_prefix(A, s O r))" apply (erule gen_prefix.induct) apply (auto elim: ConsE simp add: Nil_gen_prefix) apply (subgoal_tac "ys \ list (A) ") @@ -180,19 +180,19 @@ done lemma prefix_gen_prefix_trans: - "\ \ prefix(A); \ gen_prefix(A, r); r<=A*A\ - \ \ gen_prefix(A, r)" + "\\x,y\ \ prefix(A); \y, z\ \ gen_prefix(A, r); r<=A*A\ + \ \x, z\ \ gen_prefix(A, r)" apply (unfold prefix_def) -apply (rule_tac P = "%r. \ gen_prefix (A, r) " in right_comp_id [THEN subst]) +apply (rule_tac P = "\r. \x,z\ \ gen_prefix (A, r) " in right_comp_id [THEN subst]) apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ done lemma gen_prefix_prefix_trans: -"\ \ gen_prefix(A,r); \ prefix(A); r<=A*A\ - \ \ gen_prefix(A, r)" +"\\x,y\ \ gen_prefix(A,r); \y, z\ \ prefix(A); r<=A*A\ + \ \x, z\ \ gen_prefix(A, r)" apply (unfold prefix_def) -apply (rule_tac P = "%r. \ gen_prefix (A, r) " in left_comp_id [THEN subst]) +apply (rule_tac P = "\r. \x,z\ \ gen_prefix (A, r) " in left_comp_id [THEN subst]) apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+ done @@ -231,7 +231,7 @@ lemma same_gen_prefix_gen_prefix: "\refl(A, r); xs \ list(A)\ \ - : gen_prefix(A, r) \ \ gen_prefix(A, r)" + : gen_prefix(A, r) \ \ys,zs\ \ gen_prefix(A, r)" apply (unfold refl_def) apply (induct_tac "xs") apply (simp_all (no_asm_simp)) @@ -240,11 +240,11 @@ 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 \ \z,y\:r \ \zs,ys\ \ gen_prefix(A,r)))" apply (induct_tac "xs", auto) done -lemma gen_prefix_take_append: "\refl(A,r); \ gen_prefix(A, r); zs \ list(A)\ +lemma gen_prefix_take_append: "\refl(A,r); \xs,ys\ \ gen_prefix(A, r); zs \ list(A)\ \ \ gen_prefix(A, r)" apply (erule gen_prefix.induct) apply (simp (no_asm_simp)) @@ -254,7 +254,7 @@ apply (simp_all (no_asm_simp) add: diff_is_0_iff [THEN iffD2] take_type) done -lemma gen_prefix_append_both: "\refl(A, r); \ gen_prefix(A,r); +lemma gen_prefix_append_both: "\refl(A, r); \xs,ys\ \ gen_prefix(A,r); length(xs) = length(ys); zs \ list(A)\ \ \ gen_prefix(A, r)" apply (drule_tac zs = zs in gen_prefix_take_append, assumption+) @@ -267,7 +267,7 @@ by (auto simp add: app_assoc) lemma append_one_gen_prefix_lemma [rule_format]: - "\ \ gen_prefix(A, r); refl(A, r)\ + "\\xs,ys\ \ gen_prefix(A, r); refl(A, r)\ \ length(xs) < length(ys) \ \ gen_prefix(A, r)" apply (erule gen_prefix.induct, blast) @@ -283,14 +283,14 @@ apply (blast intro: gen_prefix.append) apply (erule_tac V = "length (xs) < length (ys) \u" for u in thin_rl) apply (erule_tac a = zs in list.cases, auto) -apply (rule_tac P1 = "%x. :w" for u v w in nat_diff_split [THEN iffD2]) +apply (rule_tac P1 = "\x. :w" for u v w in nat_diff_split [THEN iffD2]) apply auto apply (simplesubst append_cons_conv) apply (rule_tac [2] gen_prefix.append) apply (auto elim: ConsE simp add: gen_prefix_append_both) done -lemma append_one_gen_prefix: "\: gen_prefix(A, r); length(xs) < length(ys); refl(A, r)\ +lemma append_one_gen_prefix: "\\xs,ys\: gen_prefix(A, r); length(xs) < length(ys); refl(A, r)\ \ \ gen_prefix(A, r)" apply (blast intro: append_one_gen_prefix_lemma) done @@ -300,13 +300,13 @@ lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \ list(A) \ \ys \ list(A). \i \ nat. i < length(xs) - \ : gen_prefix(A, r) \ :r" + \ \xs, ys\: gen_prefix(A, r) \ :r" apply (induct_tac "xs", simp, clarify) apply simp apply (erule natE, auto) done -lemma gen_prefix_imp_nth: "\ \ gen_prefix(A,r); i < length(xs)\ +lemma gen_prefix_imp_nth: "\\xs,ys\ \ gen_prefix(A,r); i < length(xs)\ \ :r" apply (cut_tac A = A in gen_prefix.dom_subset) apply (rule gen_prefix_imp_nth_lemma) @@ -316,7 +316,7 @@ lemma nth_imp_gen_prefix [rule_format]: "xs \ list(A) \ \ys \ list(A). length(xs) \ length(ys) \ (\i. i < length(xs) \ :r) - \ \ gen_prefix(A, r)" + \ \xs, ys\ \ gen_prefix(A, r)" apply (induct_tac "xs") apply (simp_all (no_asm_simp)) apply clarify @@ -324,7 +324,7 @@ apply (force intro!: nat_0_le simp add: lt_nat_in_nat) done -lemma gen_prefix_iff_nth: "( \ gen_prefix(A,r)) \ +lemma gen_prefix_iff_nth: "(\xs,ys\ \ gen_prefix(A,r)) \ (xs \ list(A) \ ys \ list(A) \ length(xs) \ length(ys) \ (\i. i < length(xs) \ : r))" apply (rule iffI) @@ -363,7 +363,7 @@ (* Monotonicity of "set" operator WRT prefix *) lemma set_of_list_prefix_mono: -" \ prefix(A) \ set_of_list(xs) \ set_of_list(ys)" +"\xs,ys\ \ prefix(A) \ set_of_list(xs) \ set_of_list(ys)" apply (unfold prefix_def) apply (erule gen_prefix.induct) @@ -392,13 +392,13 @@ declare prefix_Nil [iff] lemma Cons_prefix_Cons: -" \ prefix(A) \ (x=y \ \ prefix(A) \ y \ A)" +" \ prefix(A) \ (x=y \ \xs,ys\ \ prefix(A) \ y \ A)" apply (unfold prefix_def, auto) done declare Cons_prefix_Cons [iff] lemma same_prefix_prefix: -"xs \ list(A)\ \ prefix(A) \ ( \ prefix(A))" +"xs \ list(A)\ \ prefix(A) \ (\ys,zs\ \ prefix(A))" apply (unfold prefix_def) apply (subgoal_tac "refl (A,id (A))") apply (simp (no_asm_simp)) @@ -407,13 +407,13 @@ declare same_prefix_prefix [simp] lemma same_prefix_prefix_Nil: "xs \ list(A) \ \ prefix(A) \ ( \ prefix(A))" -apply (rule_tac P = "%x. :v \ w(x)" for u v w in app_right_Nil [THEN subst]) +apply (rule_tac P = "\x. \u, x\:v \ w(x)" for u v w in app_right_Nil [THEN subst]) apply (rule_tac [2] same_prefix_prefix, auto) done declare same_prefix_prefix_Nil [simp] lemma prefix_appendI: -"\ \ prefix(A); zs \ list(A)\ \ \ prefix(A)" +"\\xs,ys\ \ prefix(A); zs \ list(A)\ \ \ prefix(A)" apply (unfold prefix_def) apply (erule gen_prefix.append, assumption) done @@ -422,13 +422,13 @@ 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) \ \zs,ys\ \ prefix(A)))" apply (unfold prefix_def) apply (auto simp add: gen_prefix_Cons) done lemma append_one_prefix: - "\ \ prefix(A); length(xs) < length(ys)\ + "\\xs,ys\ \ prefix(A); length(xs) < length(ys)\ \ \ prefix(A)" apply (unfold prefix_def) apply (subgoal_tac "refl (A, id (A))") @@ -437,7 +437,7 @@ done lemma prefix_length_le: -" \ prefix(A) \ length(xs) \ length(ys)" +"\xs,ys\ \ prefix(A) \ length(xs) \ length(ys)" apply (unfold prefix_def) apply (blast dest: gen_prefix_length_le) done @@ -454,7 +454,7 @@ done lemma strict_prefix_length_lt_aux: - " \ prefix(A) \ xs\ys \ length(xs) < length(ys)" + "\xs,ys\ \ prefix(A) \ xs\ys \ length(xs) < length(ys)" apply (unfold prefix_def) apply (erule gen_prefix.induct, clarify) apply (subgoal_tac [!] "ys \ list(A) \ xs \ list(A)") @@ -467,7 +467,7 @@ done lemma strict_prefix_length_lt: - ":strict_prefix(A) \ length(xs) < length(ys)" + "\xs,ys\:strict_prefix(A) \ length(xs) < length(ys)" apply (unfold strict_prefix_def) apply (rule strict_prefix_length_lt_aux [THEN mp]) apply (auto dest: prefix_type [THEN subsetD]) @@ -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)" + "\xs,zs\ \ 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) ") @@ -494,7 +494,7 @@ lemma prefix_snoc: "\xs \ list(A); ys \ list(A); y \ A\ \ - \ prefix(A) \ (xs = ys@[y] | \ prefix(A))" + \ prefix(A) \ (xs = ys@[y] | \xs,ys\ \ prefix(A))" apply (simp (no_asm) add: prefix_iff) apply (rule iffI, clarify) apply (erule_tac xs = ysa in rev_list_elim, simp) @@ -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)))" + (\xs,ys\ \ prefix(A) | (\us. xs = ys@us \ \us,zs\ \ prefix(A)))" apply (erule list_append_induct, force, clarify) apply (rule iffI) apply (simp add: add: app_assoc [symmetric]) @@ -519,13 +519,13 @@ (*Although the prefix ordering is not linear, the prefixes of a list are linearly ordered.*) lemma common_prefix_linear_lemma [rule_format]: "\zs \ list(A); xs \ list(A); ys \ list(A)\ - \ \ prefix(A) \ \ prefix(A) - \ \ prefix(A) | \ prefix(A)" + \ \xs, zs\ \ prefix(A) \ \ys,zs\ \ prefix(A) + \\xs,ys\ \ prefix(A) | \ys,xs\ \ prefix(A)" apply (erule list_append_induct, auto) done -lemma common_prefix_linear: "\ \ prefix(A); \ prefix(A)\ - \ \ prefix(A) | \ prefix(A)" +lemma common_prefix_linear: "\\xs, zs\ \ prefix(A); \ys,zs\ \ prefix(A)\ + \ \xs,ys\ \ prefix(A) | \ys,xs\ \ prefix(A)" apply (cut_tac prefix_type) apply (blast del: disjCI intro: common_prefix_linear_lemma) done @@ -577,7 +577,7 @@ lemma prefix_imp_pfixLe: -":prefix(nat)\ xs pfixLe ys" +"\xs,ys\:prefix(nat)\ xs pfixLe ys" apply (unfold prefix_def) apply (rule gen_prefix_mono [THEN subsetD], auto) @@ -610,14 +610,14 @@ by (blast intro: antisym_gen_prefix [THEN antisymE]) lemma prefix_imp_pfixGe: - ":prefix(nat) \ xs pfixGe ys" + "\xs,ys\:prefix(nat) \ xs pfixGe ys" apply (unfold prefix_def Ge_def) apply (rule gen_prefix_mono [THEN subsetD], auto) done (* Added by Sidi \ prefix and take *) lemma prefix_imp_take: -" \ prefix(A) \ xs = take(length(xs), ys)" +"\xs, ys\ \ prefix(A) \ xs = take(length(xs), ys)" apply (unfold prefix_def) apply (erule gen_prefix.induct) @@ -630,14 +630,14 @@ apply (simp_all (no_asm_simp) add: diff_is_0_iff) done -lemma prefix_length_equal: "\ \ prefix(A); length(xs)=length(ys)\ \ xs = ys" +lemma prefix_length_equal: "\\xs,ys\ \ prefix(A); length(xs)=length(ys)\ \ xs = ys" apply (cut_tac A = A in prefix_type) apply (drule subsetD, auto) apply (drule prefix_imp_take) apply (erule trans, simp) done -lemma prefix_length_le_equal: "\ \ prefix(A); length(ys) \ length(xs)\ \ xs = ys" +lemma prefix_length_le_equal: "\\xs,ys\ \ prefix(A); length(ys) \ length(xs)\ \ xs = ys" by (blast intro: prefix_length_equal le_anti_sym prefix_length_le) lemma take_prefix [rule_format]: "xs \ list(A) \ \n \ nat. \ prefix(A)" @@ -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: "\xs,ys\ \ prefix(A) \ (xs=take(length(xs), ys) \ xs \ list(A) \ ys \ list(A))" apply (rule iffI) apply (frule prefix_type [THEN subsetD]) apply (blast intro: prefix_imp_take, clarify) @@ -654,13 +654,13 @@ apply (blast intro: take_prefix length_type) done -lemma prefix_imp_nth: "\ \ prefix(A); i < length(xs)\ \ nth(i,xs) = nth(i,ys)" +lemma prefix_imp_nth: "\\xs,ys\ \ prefix(A); i < length(xs)\ \ nth(i,xs) = nth(i,ys)" by (auto dest!: gen_prefix_imp_nth simp add: prefix_def) lemma nth_imp_prefix: "\xs \ list(A); ys \ list(A); length(xs) \ length(ys); \i. i < length(xs) \ nth(i, xs) = nth(i,ys)\ - \ \ prefix(A)" + \ \xs,ys\ \ prefix(A)" apply (auto simp add: prefix_def nth_imp_gen_prefix) apply (auto intro!: nth_imp_gen_prefix simp add: prefix_def) apply (blast intro: nth_type lt_trans2) @@ -668,7 +668,7 @@ lemma length_le_prefix_imp_prefix: "\length(xs) \ length(ys); - \ prefix(A); \ prefix(A)\ \ \ prefix(A)" + \xs,zs\ \ prefix(A); \ys,zs\ \ prefix(A)\ \ \xs,ys\ \ prefix(A)" apply (cut_tac A = A in prefix_type) apply (rule nth_imp_prefix, blast, blast) apply assumption diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Guar.thy --- a/src/ZF/UNITY/Guar.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Guar.thy Tue Sep 27 17:46:52 2022 +0100 @@ -37,38 +37,38 @@ case, proving equivalence with Chandy and Sanders's n-ary definitions*) definition - ex_prop :: "i => o" where + ex_prop :: "i \ o" where "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 :: "i \ o" where "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 :: "i \ o" where "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 :: "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)))" definition - guar :: "[i, i] => i" (infixl \guarantees\ 55) where + guar :: "[i, i] \ i" (infixl \guarantees\ 55) where (*higher than membership, lower than Co*) "X guarantees Y \ {F \ program. \G \ program. F ok G \ F \ G \ X \ F \ G \ Y}" definition (* Weakest guarantees *) - wg :: "[i,i] => i" where + wg :: "[i,i] \ i" where "wg(F,Y) \ \({X \ Pow(program). F:(X guarantees Y)})" definition (* Weakest existential property stronger than X *) - wx :: "i =>i" where + wx :: "i \i" where "wx(X) \ \({Y \ Pow(program). Y<=X \ ex_prop(Y)})" definition @@ -77,13 +77,13 @@ "welldef \ {F \ program. Init(F) \ 0}" definition - refines :: "[i, i, i] => o" (\(3_ refines _ wrt _)\ [10,10,10] 10) where + 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) \ (G \ H \ welldef \ X)" definition - iso_refines :: "[i,i, i] => o" (\(3_ iso'_refines _ wrt _)\ [10,10,10] 10) where + iso_refines :: "[i,i, i] \ o" (\(3_ iso'_refines _ wrt _)\ [10,10,10] 10) where "G iso_refines F wrt X \ F \ welldef \ X \ G \ welldef \ X" @@ -94,7 +94,7 @@ lemma ex1 [rule_format]: "GG \ Fin(program) - \ ex_prop(X) \ GG \ X\0 \ OK(GG, (%G. G)) \(\G \ GG. G) \ X" + \ ex_prop(X) \ GG \ X\0 \ OK(GG, (\G. G)) \(\G \ GG. G) \ X" apply (unfold ex_prop_def) apply (erule Fin_induct) apply (simp_all add: OK_cons_iff) @@ -103,7 +103,7 @@ lemma ex2 [rule_format]: "X \ program \ - (\GG \ Fin(program). GG \ X \ 0 \ OK(GG,(%G. G))\(\G \ GG. G) \ X) + (\GG \ Fin(program). GG \ X \ 0 \ OK(GG,(\G. G))\(\G \ GG. G) \ X) \ ex_prop(X)" apply (unfold ex_prop_def, clarify) apply (drule_tac x = "{F,G}" in bspec) @@ -115,7 +115,7 @@ lemma ex_prop_finite: "ex_prop(X) \ (X\program \ - (\GG \ Fin(program). GG \ X \ 0 \ OK(GG,(%G. G))\(\G \ GG. G) \ X))" + (\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 @@ -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+) @@ -157,7 +157,7 @@ (*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)" + (\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) @@ -427,7 +427,7 @@ by (simp (no_asm_simp) add: wg_equiv) lemma wg_finite [rule_format]: -"\FF \ Fin(program). FF \ X \ 0 \ OK(FF, %F. F) +"\FF \ Fin(program). FF \ X \ 0 \ OK(FF, \F. F) \ (\F \ FF. ((\F \ FF. F) \ wg(F,X)) \ ((\F \ FF. F) \ X))" apply clarify apply (subgoal_tac "F component_of (\F \ FF. F) ") diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Increasing.thy --- a/src/ZF/UNITY/Increasing.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Increasing.thy Tue Sep 27 17:46:52 2022 +0100 @@ -11,19 +11,19 @@ theory Increasing imports Constrains Monotonicity begin definition - increasing :: "[i, i, i=>i] => i" (\increasing[_]'(_, _')\) where + increasing :: "[i, i, i\i] \ i" (\increasing[_]'(_, _')\) where "increasing[A](r, f) \ {F \ program. (\k \ A. F \ stable({s \ state. \ r})) \ (\x \ state. f(x):A)}" definition - Increasing :: "[i, i, i=>i] => i" (\Increasing[_]'(_, _')\) where + Increasing :: "[i, i, i\i] \ i" (\Increasing[_]'(_, _')\) where "Increasing[A](r, f) \ {F \ program. (\k \ A. F \ Stable({s \ state. \ r})) \ (\x \ state. f(x):A)}" abbreviation (input) - IncWrt :: "[i=>i, i, i] => i" (\(_ IncreasingWrt _ '/ _)\ [60, 0, 60] 60) where + IncWrt :: "[i\i, i, i] \ i" (\(_ IncreasingWrt _ '/ _)\ [60, 0, 60] 60) where "f IncreasingWrt r/A \ Increasing[A](r,f)" @@ -47,7 +47,7 @@ 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) @@ -116,7 +116,7 @@ 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 @@ -163,7 +163,7 @@ lemma imp_increasing_comp2: "\F \ increasing[A](r, f); F \ increasing[B](s, g); mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\ - \ F \ increasing[C](t, %x. h(f(x), g(x)))" + \ F \ increasing[C](t, \x. h(f(x), g(x)))" apply (unfold increasing_def stable_def part_order_def constrains_def mono2_def, clarify, simp) apply clarify @@ -195,7 +195,7 @@ lemma imp_Increasing_comp2: "\F \ Increasing[A](r, f); F \ Increasing[B](s, g); mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\ \ - F \ Increasing[C](t, %x. h(f(x), g(x)))" + F \ Increasing[C](t, \x. h(f(x), g(x)))" apply (unfold Increasing_def stable_def part_order_def constrains_def mono2_def Stable_def Constrains_def, safe) apply (simp_all add: ActsD) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Merge.thy --- a/src/ZF/UNITY/Merge.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Merge.thy Tue Sep 27 17:46:52 2022 +0100 @@ -13,50 +13,50 @@ definition (*spec (10)*) - merge_increasing :: "[i, i, i] =>i" where + merge_increasing :: "[i, i, i] \i" where "merge_increasing(A, Out, iOut) \ program guarantees (lift(Out) IncreasingWrt prefix(A)/list(A)) Int (lift(iOut) IncreasingWrt prefix(nat)/list(nat))" definition (*spec (11)*) - merge_eq_Out :: "[i, i] =>i" where + merge_eq_Out :: "[i, i] \i" where "merge_eq_Out(Out, iOut) \ program guarantees Always({s \ state. length(s`Out) = length(s`iOut)})" definition (*spec (12)*) - merge_bounded :: "i=>i" where + merge_bounded :: "i\i" where "merge_bounded(iOut) \ program guarantees Always({s \ state. \elt \ set_of_list(s`iOut). elti, i, i] =>i" where + merge_follows :: "[i, i\i, i, i] \i" where "merge_follows(A, In, Out, iOut) \ (\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))" definition (*spec: preserves part*) - merge_preserves :: "[i=>i] =>i" where + merge_preserves :: "[i\i] \i" where "merge_preserves(In) \ \n \ nat. preserves(lift(In(n)))" definition (* environmental constraints*) - merge_allowed_acts :: "[i, i] =>i" where + merge_allowed_acts :: "[i, i] \i" where "merge_allowed_acts(Out, iOut) \ {F \ program. AllowedActs(F) = cons(id(state), (\G \ preserves(lift(Out)) \ preserves(lift(iOut)). Acts(G)))}" definition - merge_spec :: "[i, i =>i, i, i]=>i" where + merge_spec :: "[i, i \i, i, i]\i" where "merge_spec(A, In, Out, iOut) \ merge_increasing(A, Out, iOut) \ merge_eq_Out(Out, iOut) \ merge_bounded(iOut) \ merge_follows(A, In, Out, iOut) @@ -146,7 +146,7 @@ "\G \ preserves(lift(iOut)); G: preserves(lift(Out)); M \ Allowed(G)\ \ M \ G \ Always - ({s \ state. msetsum(%i. bag_of(sublist(s`Out, + ({s \ state. msetsum(\i. bag_of(sublist(s`Out, {k \ nat. k < length(s`iOut) \ nth(k, s`iOut)=i})), Nclients, A) = bag_of(s`Out)})" apply (rule Always_Diff_Un_eq [THEN iffD1]) @@ -166,7 +166,7 @@ apply (subgoal_tac "xa \ nat") apply (simp_all add: Ord_mem_iff_lt) prefer 2 apply (blast intro: lt_trans) -apply (drule_tac x = "nth (xa, x`iOut)" and P = "%elt. X (elt) \ elt (\n \ Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A)) guarantees - (%s. bag_of(s`Out)) Fols - (%s. msetsum(%i. bag_of(s`In(i)),Nclients, A)) Wrt MultLe(A, r)/Mult(A)" + (\s. bag_of(s`Out)) Fols + (\s. msetsum(\i. bag_of(s`In(i)),Nclients, A)) Wrt MultLe(A, r)/Mult(A)" apply (cut_tac merge_spec) apply (rule merge_bag_Follows_lemma [THEN Always_Follows1, THEN guaranteesI]) apply (simp_all add: M_ok_iff, clarify) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Monotonicity.thy --- a/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Monotonicity.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,37 +12,37 @@ begin definition - mono1 :: "[i, i, i, i, i=>i] => o" where + 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. \x,y\ \ r \ \ s) \ (\x \ A. f(x) \ B)" (* monotonicity of a 2-place meta-function f *) definition - mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o" where + mono2 :: "[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) \ + \x,y\ \ r \ \u,v\ \ s \ \ t) \ (\x \ A. \y \ B. f(x,y) \ C)" (* Internalized relations on sets and multisets *) definition - SetLe :: "i =>i" where - "SetLe(A) \ { \ Pow(A)*Pow(A). x \ y}" + SetLe :: "i \i" where + "SetLe(A) \ {\x,y\ \ Pow(A)*Pow(A). x \ y}" definition - MultLe :: "[i,i] =>i" where + MultLe :: "[i,i] \i" where "MultLe(A, r) \ multirel(A, r - id(A)) \ id(Mult(A))" lemma mono1D: - "\mono1(A, r, B, s, f); \ r; x \ A; y \ A\ \ \ s" + "\mono1(A, r, B, s, f); \x, y\ \ r; x \ A; y \ A\ \ \ s" by (unfold mono1_def, auto) lemma mono2D: "\mono2(A, r, B, s, C, t, f); - \ r; \ s; x \ A; y \ A; u \ B; v \ B\ + \x, y\ \ r; \u,v\ \ s; x \ A; y \ A; u \ B; v \ B\ \ \ t" by (unfold mono2_def, auto) @@ -72,12 +72,12 @@ by (blast intro: le_in_nat take_mono_left_lemma) lemma take_mono_right: - "\ \ prefix(A); i \ nat\ + "\\xs,ys\ \ prefix(A); i \ nat\ \ \ prefix(A)" by (auto simp add: prefix_iff) lemma take_mono: - "\i \ j; \ prefix(A); j \ nat\ + "\i \ j; \xs, ys\ \ prefix(A); j \ nat\ \ \ prefix(A)" apply (rule_tac b = "take (j, xs) " in prefix_trans) apply (auto dest: prefix_type [THEN subsetD] intro: take_mono_left take_mono_right) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/MultisetSum.thy --- a/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/MultisetSum.thy Tue Sep 27 17:46:52 2022 +0100 @@ -9,23 +9,23 @@ begin definition - lcomm :: "[i, i, [i,i]=>i]=>o" where + 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 \ B. f(x, y) \ B)" definition - general_setsum :: "[i,i,i, [i,i]=>i, i=>i] =>i" where + general_setsum :: "[i,i,i, [i,i]\i, i\i] \i" where "general_setsum(C, B, e, f, g) \ - if Finite(C) then fold[cons(e, B)](%x y. f(g(x), y), e, C) else e" + if Finite(C) then fold[cons(e, B)](\x y. f(g(x), y), e, C) else e" definition - msetsum :: "[i=>i, i, i]=>i" where + msetsum :: "[i\i, i, i]\i" where "msetsum(g, C, B) \ normalize(general_setsum(C, Mult(B), 0, (+#), g))" definition (* sum for naturals *) - nsetsum :: "[i=>i, i] =>i" where + nsetsum :: "[i\i, i] \i" where "nsetsum(g, C) \ general_setsum(C, nat, 0, (#+), g)" @@ -121,7 +121,7 @@ "\I \ Fin(K); \i \ K. C(i) \ Fin(A)\ \ (\x \ A. multiset(f(x)) \ mset_of(f(x))\B) \ (\i \ I. \j \ I. i\j \ C(i) \ C(j) = 0) \ - msetsum(f, \i \ I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)" + msetsum(f, \i \ I. C(i), B) = msetsum (\i. msetsum(f, C(i),B), I, B)" apply (erule Fin_induct, auto) apply (subgoal_tac "\i \ y. x \ i") prefer 2 apply blast @@ -141,7 +141,7 @@ "\C \ Fin(A); \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)" + msetsum(\x. f(x) +# g(x), C, B) = msetsum(f, C, B) +# msetsum(g, C, B)" apply (subgoal_tac "\x \ A. multiset (f(x) +# g(x)) \ mset_of (f(x) +# g(x)) \ B") apply (erule Fin_induct) apply (auto simp add: munion_ac) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Mutex.thy --- a/src/ZF/UNITY/Mutex.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Mutex.thy Tue Sep 27 17:46:52 2022 +0100 @@ -36,37 +36,37 @@ definition (** The program for process U **) - "U0 \ {:state*state. t = s(u:=1, m:=#1) \ s`m = #0}" + "U0 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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 \ {\s,t\: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}, diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/State.thy --- a/src/ZF/UNITY/State.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/State.thy Tue Sep 27 17:46:52 2022 +0100 @@ -17,8 +17,8 @@ type_intros nat_subset_univ [THEN list_subset_univ, THEN subsetD] consts - type_of :: "i=>i" - default_val :: "i=>i" + type_of :: "i\i" + default_val :: "i\i" definition "state \ \x \ var. cons(default_val(x), type_of(x))" @@ -27,13 +27,13 @@ "st0 \ \x \ var. default_val(x)" definition - st_set :: "i=>o" where + st_set :: "i\o" where (* To prevent typing conditions like `A<=state' from being used in combination with the rules `constrains_weaken', etc. *) "st_set(A) \ A<=state" definition - st_compl :: "i=>i" where + st_compl :: "i\i" where "st_compl(A) \ state-A" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/SubstAx.thy Tue Sep 27 17:46:52 2022 +0100 @@ -13,11 +13,11 @@ definition (* The definitions below are not `conventional', but yield simpler rules *) - Ensures :: "[i,i] => i" (infixl \Ensures\ 60) where + Ensures :: "[i,i] \ i" (infixl \Ensures\ 60) where "A Ensures B \ {F \ program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" definition - LeadsTo :: "[i, i] => i" (infixl \\w\ 60) where + LeadsTo :: "[i, i] \ i" (infixl \\w\ 60) where "A \w B \ {F \ program. F:(reachable(F) \ A) \ (reachable(F) \ B)}" @@ -280,7 +280,7 @@ lemma LessThan_induct: "\\m \ nat. F:(A \ f-``{m}) \w ((A \ f-``m) \ B); A<=f-``nat; F \ program\ \ F \ A \w B" -apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN LeadsTo_wf_induct]) +apply (rule_tac A1 = nat and f1 = "\x. x" in wf_measure [THEN LeadsTo_wf_induct]) apply (simp_all add: nat_measure_field) apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric]) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/UNITY.thy --- a/src/ZF/UNITY/UNITY.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/UNITY.thy Tue Sep 27 17:46:52 2022 +0100 @@ -19,7 +19,7 @@ id(state) \ acts \ id(state) \ allowed}" definition - mk_program :: "[i,i,i]=>i" where + mk_program :: "[i,i,i]\i" where \ \The definition yields a program thanks to the coercions init \ state, acts \ Pow(state*state), etc.\ "mk_program(init, acts, allowed) \ @@ -32,69 +32,69 @@ (* Coercion from anything to program *) definition - programify :: "i=>i" where + programify :: "i\i" where "programify(F) \ if F \ program then F else SKIP" definition - RawInit :: "i=>i" where + RawInit :: "i\i" where "RawInit(F) \ fst(F)" definition - Init :: "i=>i" where + Init :: "i\i" where "Init(F) \ RawInit(programify(F))" definition - RawActs :: "i=>i" where + RawActs :: "i\i" where "RawActs(F) \ cons(id(state), fst(snd(F)))" definition - Acts :: "i=>i" where + Acts :: "i\i" where "Acts(F) \ RawActs(programify(F))" definition - RawAllowedActs :: "i=>i" where + RawAllowedActs :: "i\i" where "RawAllowedActs(F) \ cons(id(state), snd(snd(F)))" definition - AllowedActs :: "i=>i" where + AllowedActs :: "i\i" where "AllowedActs(F) \ RawAllowedActs(programify(F))" definition - Allowed :: "i =>i" where + Allowed :: "i \i" where "Allowed(F) \ {G \ program. Acts(G) \ AllowedActs(F)}" definition - initially :: "i=>i" where + initially :: "i\i" where "initially(A) \ {F \ program. Init(F)\A}" -definition "constrains" :: "[i, i] => i" (infixl \co\ 60) where +definition "constrains" :: "[i, i] \ i" (infixl \co\ 60) where "A co B \ {F \ program. (\act \ Acts(F). act``A\B) \ st_set(A)}" \ \the condition \<^term>\st_set(A)\ makes the definition slightly stronger than the HOL one\ -definition unless :: "[i, i] => i" (infixl \unless\ 60) where +definition unless :: "[i, i] \ i" (infixl \unless\ 60) where "A unless B \ (A - B) co (A \ B)" definition - stable :: "i=>i" where + stable :: "i\i" where "stable(A) \ A co A" definition - strongest_rhs :: "[i, i] => i" where + strongest_rhs :: "[i, i] \ i" where "strongest_rhs(F, A) \ \({B \ Pow(state). F \ A co B})" definition - invariant :: "i => i" where + invariant :: "i \ i" where "invariant(A) \ initially(A) \ stable(A)" (* meta-function composition *) definition - metacomp :: "[i=>i, i=>i] => (i=>i)" (infixl \comp\ 65) where - "f comp g \ %x. f(g(x))" + metacomp :: "[i\i, i\i] \ (i\i)" (infixl \comp\ 65) where + "f comp g \ \x. f(g(x))" definition - pg_compl :: "i=>i" where + pg_compl :: "i\i" where "pg_compl(X)\ program - X" text\SKIP\ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/Union.thy --- a/src/ZF/UNITY/Union.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/Union.thy Tue Sep 27 17:46:52 2022 +0100 @@ -15,34 +15,34 @@ definition (*FIXME: conjoin Init(F) \ Init(G) \ 0 *) - ok :: "[i, i] => o" (infixl \ok\ 65) where + ok :: "[i, i] \ o" (infixl \ok\ 65) where "F ok G \ Acts(F) \ AllowedActs(G) \ Acts(G) \ AllowedActs(F)" definition (*FIXME: conjoin (\i \ I. Init(F(i))) \ 0 *) - OK :: "[i, i=>i] => o" where + OK :: "[i, i\i] \ o" where "OK(I,F) \ (\i \ I. \j \ I-{i}. Acts(F(i)) \ AllowedActs(F(j)))" definition - JOIN :: "[i, i=>i] => i" where + JOIN :: "[i, i\i] \ i" where "JOIN(I,F) \ if I = 0 then SKIP else mk_program(\i \ I. Init(F(i)), \i \ I. Acts(F(i)), \i \ I. AllowedActs(F(i)))" definition - Join :: "[i, i] => i" (infixl \\\ 65) where + Join :: "[i, i] \ i" (infixl \\\ 65) where "F \ G \ mk_program (Init(F) \ Init(G), Acts(F) \ Acts(G), AllowedActs(F) \ AllowedActs(G))" definition (*Characterizes safety properties. Used with specifying AllowedActs*) - safety_prop :: "i => o" where + safety_prop :: "i \ o" where "safety_prop(X) \ X\program \ SKIP \ X \ (\G \ program. Acts(G) \ (\F \ X. Acts(F)) \ G \ X)" syntax - "_JOIN1" :: "[pttrns, i] => i" (\(3\_./ _)\ 10) - "_JOIN" :: "[pttrn, i, i] => i" (\(3\_ \ _./ _)\ 10) + "_JOIN1" :: "[pttrns, i] \ i" (\(3\_./ _)\ 10) + "_JOIN" :: "[pttrn, i, i] \ i" (\(3\_ \ _./ _)\ 10) translations "\x \ A. B" == "CONST JOIN(A, (\x. B))" @@ -141,10 +141,10 @@ subsection\Eliminating programify form JOIN and OK expressions\ -lemma OK_programify [iff]: "OK(I, %x. programify(F(x))) \ OK(I, F)" +lemma OK_programify [iff]: "OK(I, \x. programify(F(x))) \ OK(I, F)" by (simp add: OK_def) -lemma JOIN_programify [iff]: "JOIN(I, %x. programify(F(x))) = JOIN(I, F)" +lemma JOIN_programify [iff]: "JOIN(I, \x. programify(F(x))) = JOIN(I, F)" by (simp add: JOIN_def) @@ -422,7 +422,7 @@ 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+) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/UNITY/WFair.thy --- a/src/ZF/UNITY/WFair.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/UNITY/WFair.thy Tue Sep 27 17:46:52 2022 +0100 @@ -16,26 +16,26 @@ definition (* This definition specifies weak fairness. The rest of the theory is generic to all forms of fairness.*) - transient :: "i=>i" where + transient :: "i\i" where "transient(A) \{F \ program. (\act\Acts(F). A<=domain(act) \ act``A \ state-A) \ st_set(A)}" definition - ensures :: "[i,i] => i" (infixl \ensures\ 60) where + ensures :: "[i,i] \ i" (infixl \ensures\ 60) where "A ensures B \ ((A-B) co (A \ B)) \ transient(A-B)" consts (*LEADS-TO constant for the inductive definition*) - leads :: "[i, i]=>i" + leads :: "[i, i]\i" inductive domains "leads(D, F)" \ "Pow(D)*Pow(D)" intros - Basis: "\F \ A ensures B; A \ Pow(D); B \ Pow(D)\ \ :leads(D, F)" - Trans: "\ \ leads(D, F); \ leads(D, F)\ \ :leads(D, F)" - Union: "\S \ Pow({A \ S. :leads(D, F)}); B \ Pow(D); S \ Pow(Pow(D))\ \ + Basis: "\F \ A ensures B; A \ Pow(D); B \ Pow(D)\ \ \A,B\:leads(D, F)" + Trans: "\\A,B\ \ leads(D, F); \B,C\ \ leads(D, F)\ \ \A,C\:leads(D, F)" + Union: "\S \ Pow({A \ S. \A, B\:leads(D, F)}); B \ Pow(D); S \ Pow(Pow(D))\ \ <\(S),B>:leads(D, F)" monos Pow_mono @@ -43,12 +43,12 @@ definition (* The Visible version of the LEADS-TO relation*) - leadsTo :: "[i, i] => i" (infixl \\\ 60) where - "A \ B \ {F \ program. :leads(state, F)}" + leadsTo :: "[i, i] \ i" (infixl \\\ 60) where + "A \ B \ {F \ program. \A,B\:leads(state, F)}" definition (* wlt(F, B) is the largest set that leads to B*) - wlt :: "[i, i] => i" where + wlt :: "[i, i] \ i" where "wlt(F, B) \ \({A \ Pow(state). F \ A \ B})" (** Ad-hoc set-theory rules **) @@ -510,7 +510,7 @@ apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best) done -lemma nat_measure_field: "field(measure(nat, %x. x)) = nat" +lemma nat_measure_field: "field(measure(nat, \x. x)) = nat" apply (unfold field_def) apply (simp add: measure_def) apply (rule equalityI, force, clarify) @@ -522,7 +522,7 @@ done -lemma Image_inverse_lessThan: "k measure(A, %x. x) -`` {k} = k" +lemma Image_inverse_lessThan: "k measure(A, \x. x) -`` {k} = k" apply (rule equalityI) apply (auto simp add: measure_def) apply (blast intro: ltD) @@ -538,7 +538,7 @@ F \ program; st_set(A); st_set(B); \m \ nat. F:(A \ f-``{m}) \ ((A \ f -`` m) \ B)\ \ F \ A \ B" -apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct]) +apply (rule_tac A1 = nat and f1 = "\x. x" in wf_measure [THEN leadsTo_wf_induct]) apply (simp_all add: nat_measure_field) apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric]) done @@ -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) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Univ.thy --- a/src/ZF/Univ.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Univ.thy Tue Sep 27 17:46:52 2022 +0100 @@ -14,26 +14,26 @@ theory Univ imports Epsilon Cardinal begin definition - Vfrom :: "[i,i]=>i" where - "Vfrom(A,i) \ transrec(i, %x f. A \ (\y\x. Pow(f`y)))" + Vfrom :: "[i,i]\i" where + "Vfrom(A,i) \ transrec(i, \x f. A \ (\y\x. Pow(f`y)))" abbreviation - Vset :: "i=>i" where + Vset :: "i\i" where "Vset(x) \ Vfrom(0,x)" definition - Vrec :: "[i, [i,i]=>i] =>i" where - "Vrec(a,H) \ transrec(rank(a), %x g. \z\Vset(succ(x)). + Vrec :: "[i, [i,i]\i] \i" where + "Vrec(a,H) \ transrec(rank(a), \x g. \z\Vset(succ(x)). H(z, \w\Vset(x). g`rank(w)`w)) ` a" definition - Vrecursor :: "[[i,i]=>i, i] =>i" where - "Vrecursor(H,a) \ transrec(rank(a), %x g. \z\Vset(succ(x)). + Vrecursor :: "[[i,i]\i, i] \i" where + "Vrecursor(H,a) \ transrec(rank(a), \x g. \z\Vset(succ(x)). H(\w\Vset(x). g`rank(w)`w, z)) ` a" definition - univ :: "i=>i" where + univ :: "i\i" where "univ(A) \ Vfrom(A,nat)" @@ -124,7 +124,7 @@ by (rule subset_mem_Vfrom, safe) lemma Pair_in_Vfrom: - "\a \ Vfrom(A,i); b \ Vfrom(A,i)\ \ \ Vfrom(A,succ(succ(i)))" + "\a \ Vfrom(A,i); b \ Vfrom(A,i)\ \ \a,b\ \ Vfrom(A,succ(succ(i)))" apply (unfold Pair_def) apply (blast intro: doubleton_in_Vfrom) done @@ -219,7 +219,7 @@ done lemma Pair_in_VLimit: - "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i)\ \ \ Vfrom(A,i)" + "\a \ Vfrom(A,i); b \ Vfrom(A,i); Limit(i)\ \ \a,b\ \ Vfrom(A,i)" txt\Infer that a, b occur at ordinals x,xa < i.\ apply (erule Limit_VfromE, assumption) apply (erule Limit_VfromE, assumption) @@ -283,12 +283,12 @@ 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: "\\a,b\ \ C; Transset(C)\ \ a: C \ b: C" by (unfold Pair_def Transset_def, blast) lemma Transset_Pair_subset_VLimit: - "\ \ Vfrom(A,i); Transset(A); Limit(i)\ - \ \ Vfrom(A,i)" + "\\a,b\ \ Vfrom(A,i); Transset(A); Limit(i)\ + \ \a,b\ \ Vfrom(A,i)" apply (erule Transset_Pair_subset [THEN conjE]) apply (erule Transset_Vfrom) apply (blast intro: Pair_in_VLimit) @@ -496,7 +496,7 @@ apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def) done -text\This form avoids giant explosions in proofs. NOTE USE OF \\ +text\This form avoids giant explosions in proofs. NOTE the form of the premise!\ lemma def_Vrec: "\\x. h(x)\Vrec(x,H)\ \ h(a) = H(a, \x\Vset(rank(a)). h(x))" @@ -512,7 +512,7 @@ apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def) done -text\This form avoids giant explosions in proofs. NOTE USE OF \\ +text\This form avoids giant explosions in proofs. NOTE the form of the premise!\ lemma def_Vrecursor: "h \ Vrecursor(H) \ h(a) = H(\x\Vset(rank(a)). h(x), a)" apply simp @@ -595,7 +595,7 @@ done lemma Pair_in_univ: - "\a: univ(A); b: univ(A)\ \ \ univ(A)" + "\a: univ(A); b: univ(A)\ \ \a,b\ \ univ(A)" apply (unfold univ_def) apply (blast intro: Pair_in_VLimit Limit_nat) done @@ -619,8 +619,8 @@ apply (rule i_subset_Vfrom) done -text\n:nat \ n:univ(A)\ -lemmas nat_into_univ = nat_subset_univ [THEN subsetD] +lemma nat_into_univ: "n \ nat \ n \ univ(A)" + by (rule nat_subset_univ [THEN subsetD]) subsubsection\Instances for 1 and 2\ @@ -766,16 +766,16 @@ text\This weaker version says a, b exist at the same level\ lemmas Vfrom_doubleton_D = Transset_Vfrom [THEN Transset_doubleton_D] -(** Using only the weaker theorem would prove \ Vfrom(X,i) +(** Using only the weaker theorem would prove \a,b\ \ Vfrom(X,i) implies a, b \ Vfrom(X,i), which is useless for induction. - Using only the stronger theorem would prove \ Vfrom(X,succ(succ(i))) + Using only the stronger theorem would prove \a,b\ \ Vfrom(X,succ(succ(i))) implies a, b \ Vfrom(X,i), leaving the succ(i) case untreated. The combination gives a reduction by precisely one level, which is most convenient for proofs. **) lemma Pair_in_Vfrom_D: - "\ \ Vfrom(X,succ(i)); Transset(X)\ + "\\a,b\ \ Vfrom(X,succ(i)); Transset(X)\ \ a \ Vfrom(X,i) \ b \ Vfrom(X,i)" apply (unfold Pair_def) apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/WF.thy --- a/src/ZF/WF.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/WF.thy Tue Sep 27 17:46:52 2022 +0100 @@ -19,34 +19,34 @@ theory WF imports Trancl begin definition - wf :: "i=>o" where + wf :: "i\o" where (*r is a well-founded relation*) - "wf(r) \ \Z. Z=0 | (\x\Z. \y. :r \ \ y \ Z)" + "wf(r) \ \Z. Z=0 | (\x\Z. \y. \y,x\:r \ \ y \ Z)" definition - wf_on :: "[i,i]=>o" (\wf[_]'(_')\) where + wf_on :: "[i,i]\o" (\wf[_]'(_')\) where (*r is well-founded on A*) "wf_on(A,r) \ wf(r \ A*A)" definition - is_recfun :: "[i, i, [i,i]=>i, i] =>o" where + is_recfun :: "[i, i, [i,i]\i, i] \o" where "is_recfun(r,a,H,f) \ (f = (\x\r-``{a}. H(x, restrict(f, r-``{x}))))" definition - the_recfun :: "[i, i, [i,i]=>i] =>i" where + the_recfun :: "[i, i, [i,i]\i] \i" where "the_recfun(r,a,H) \ (THE f. is_recfun(r,a,H,f))" definition - wftrec :: "[i, i, [i,i]=>i] =>i" where + wftrec :: "[i, i, [i,i]\i] \i" where "wftrec(r,a,H) \ H(a, the_recfun(r,a,H))" definition - wfrec :: "[i, i, [i,i]=>i] =>i" where + wfrec :: "[i, i, [i,i]\i] \i" where (*public version. Does not require r to be transitive*) - "wfrec(r,a,H) \ wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))" + "wfrec(r,a,H) \ wftrec(r^+, a, \x f. H(x, restrict(f,r-``{x})))" definition - wfrec_on :: "[i, i, i, [i,i]=>i] =>i" (\wfrec[_]'(_,_,_')\) where + wfrec_on :: "[i, i, i, [i,i]\i] \i" (\wfrec[_]'(_,_,_')\) where "wfrec[A](r,a,H) \ wfrec(r \ A*A, a, H)" @@ -80,7 +80,7 @@ text\If every non-empty subset of \<^term>\A\ has an \<^term>\r\-minimal element then we have \<^term>\wf[A](r)\.\ lemma wf_onI: - assumes prem: "\Z u. \Z<=A; u \ Z; \x\Z. \y\Z. :r\ \ False" + assumes prem: "\Z u. \Z<=A; u \ Z; \x\Z. \y\Z. \y,x\:r\ \ False" shows "wf[A](r)" apply (unfold wf_on_def wf_def) apply (rule equals0I [THEN disjCI, THEN allI]) @@ -89,9 +89,9 @@ text\If \<^term>\r\ allows well-founded induction over \<^term>\A\ then we have \<^term>\wf[A](r)\. Premise is equivalent to - \<^prop>\\B. \x\A. (\y. : r \ y \ B) \ x \ B \ A<=B\\ + \<^prop>\\B. \x\A. (\y. \y,x\: r \ y \ B) \ x \ B \ A<=B\\ lemma wf_onI2: - assumes prem: "\y B. \\x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A\ + assumes prem: "\y B. \\x\A. (\y\A. \y,x\:r \ y \ B) \ x \ B; y \ A\ \ y \ B" shows "wf[A](r)" apply (rule wf_onI) @@ -107,7 +107,7 @@ \<^term>\P(z)\ does not hold...\ lemma wf_induct_raw: "\wf(r); - \x.\\y. : r \ P(y)\ \ P(x)\ + \x.\\y. \y,x\: r \ P(y)\ \ P(x)\ \ P(a)" apply (unfold wf_def) apply (erule_tac x = "{z \ domain(r). \ P(z)}" in allE) @@ -119,7 +119,7 @@ text\The form of this rule is designed to match \wfI\\ lemma wf_induct2: "\wf(r); a \ A; field(r)<=A; - \x.\x \ A; \y. : r \ P(y)\ \ P(x)\ + \x.\x \ A; \y. \y,x\: r \ P(y)\ \ P(x)\ \ P(a)" apply (erule_tac P="a \ A" in rev_mp) apply (erule_tac a=a in wf_induct, blast) @@ -130,7 +130,7 @@ lemma wf_on_induct_raw [consumes 2, induct set: wf_on]: "\wf[A](r); a \ A; - \x.\x \ A; \y\A. : r \ P(y)\ \ P(x) + \x.\x \ A; \y\A. \y,x\: r \ P(y)\ \ P(x) \ \ P(a)" apply (unfold wf_on_def) apply (erule wf_induct2, assumption) @@ -145,7 +145,7 @@ then we have \<^term>\wf(r)\.\ lemma wfI: "\field(r)<=A; - \y B. \\x\A. (\y\A. :r \ y \ B) \ x \ B; y \ A\ + \y B. \\x\A. (\y\A. \y,x\:r \ y \ B) \ x \ B; y \ A\ \ y \ B\ \ wf(r)" apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf]) @@ -157,35 +157,35 @@ subsection\Basic Properties of Well-Founded Relations\ -lemma wf_not_refl: "wf(r) \ \ r" +lemma wf_not_refl: "wf(r) \ \a,a\ \ r" by (erule_tac a=a in wf_induct, blast) -lemma wf_not_sym [rule_format]: "wf(r) \ \x. :r \ \ r" +lemma wf_not_sym [rule_format]: "wf(r) \ \x. \a,x\:r \ \x,a\ \ r" by (erule_tac a=a in wf_induct, blast) -(* @{term"\wf(r); \ r; \P \ \ r\ \ P"} *) +(* @{term"\wf(r); \a,x\ \ r; \P \ \x,a\ \ r\ \ P"} *) lemmas wf_asym = wf_not_sym [THEN swap] -lemma wf_on_not_refl: "\wf[A](r); a \ A\ \ \ r" +lemma wf_on_not_refl: "\wf[A](r); a \ A\ \ \a,a\ \ r" by (erule_tac a=a in wf_on_induct, assumption, blast) lemma wf_on_not_sym: - "\wf[A](r); a \ A\ \ (\b. b\A \ :r \ \r)" + "\wf[A](r); a \ A\ \ (\b. b\A \ \a,b\:r \ \b,a\\r)" apply (atomize (full), intro impI) apply (erule_tac a=a in wf_on_induct, assumption, blast) done lemma wf_on_asym: - "\wf[A](r); \Z \ \ r; - \ r \ Z; \Z \ a \ A; \Z \ b \ A\ \ Z" + "\wf[A](r); \Z \ \a,b\ \ r; + \b,a\ \ r \ Z; \Z \ a \ A; \Z \ b \ A\ \ Z" by (blast dest: wf_on_not_sym) (*Needed to prove well_ordI. Could also reason that wf[A](r) means wf(r \ A*A); thus wf( (r \ A*A)^+ ) and use wf_not_refl *) lemma wf_on_chain3: - "\wf[A](r); :r; :r; :r; a \ A; b \ A; c \ A\ \ P" -apply (subgoal_tac "\y\A. \z\A. :r \ :r \ :r \ P", + "\wf[A](r); \a,b\:r; \b,c\:r; \c,a\:r; a \ A; b \ A; c \ A\ \ P" +apply (subgoal_tac "\y\A. \z\A. \a,y\:r \ \y,z\:r \ \z,a\:r \ P", blast) apply (erule_tac a=a in wf_on_induct, assumption, blast) done @@ -228,16 +228,16 @@ lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function] lemma apply_recfun: - "\is_recfun(r,a,H,f); :r\ \ f`x = H(x, restrict(f,r-``{x}))" + "\is_recfun(r,a,H,f); \x,a\:r\ \ f`x = H(x, restrict(f,r-``{x}))" apply (unfold is_recfun_def) txt\replace f only on the left-hand side\ -apply (erule_tac P = "%x. t(x) = u" for t u in ssubst) +apply (erule_tac P = "\x. t(x) = u" for t u in ssubst) apply (simp add: underI) done lemma is_recfun_equal [rule_format]: "\wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g)\ - \ :r \ :r \ f`x=g`x" + \ \x,a\:r \ \x,b\:r \ f`x=g`x" apply (frule_tac f = f in is_recfun_type) apply (frule_tac f = g in is_recfun_type) apply (simp add: is_recfun_def) @@ -245,8 +245,8 @@ apply (intro impI) apply (elim ssubst) apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def) -apply (rule_tac t = "%z. H (x, z)" for x in subst_context) -apply (subgoal_tac "\y\r-``{x}. \z. :f \ :g") +apply (rule_tac t = "\z. H (x, z)" for x in subst_context) +apply (subgoal_tac "\y\r-``{x}. \z. \y,z\:f \ \y,z\:g") apply (blast dest: transD) apply (simp add: apply_iff) apply (blast dest: transD intro: sym) @@ -254,7 +254,7 @@ lemma is_recfun_cut: "\wf(r); trans(r); - is_recfun(r,a,H,f); is_recfun(r,b,H,g); :r\ + is_recfun(r,a,H,f); is_recfun(r,b,H,g); \b,a\:r\ \ restrict(f, r-``{b}) = g" apply (frule_tac f = f in is_recfun_type) apply (rule fun_extension) @@ -292,7 +292,7 @@ apply (rule lam_cong [OF refl]) apply (drule underD) apply (fold is_recfun_def) -apply (rule_tac t = "%z. H(x, z)" for x in subst_context) +apply (rule_tac t = "\z. H(x, z)" for x in subst_context) apply (rule fun_extension) apply (blast intro: is_recfun_type) apply (rule lam_type [THEN restrict_type2]) @@ -300,7 +300,7 @@ apply (blast dest: transD) apply atomize apply (frule spec [THEN mp], assumption) -apply (subgoal_tac " \ r") +apply (subgoal_tac "\xa,a1\ \ r") apply (drule_tac x1 = xa in spec [THEN mp], assumption) apply (simp add: vimage_singleton_iff apply_recfun is_recfun_cut) @@ -311,7 +311,7 @@ subsection\Unfolding \<^term>\wftrec(r,a,H)\\ lemma the_recfun_cut: - "\wf(r); trans(r); :r\ + "\wf(r); trans(r); \b,a\:r\ \ restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)" by (blast intro: is_recfun_cut unfold_the_recfun) @@ -366,7 +366,7 @@ text\Minimal-element characterization of well-foundedness\ lemma wf_eq_minimal: - "wf(r) \ (\Q x. x \ Q \ (\z\Q. \y. :r \ y\Q))" + "wf(r) \ (\Q x. x \ Q \ (\z\Q. \y. \y,z\:r \ y\Q))" by (unfold wf_def, blast) end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ZF.thy --- a/src/ZF/ZF.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ZF.thy Tue Sep 27 17:46:52 2022 +0100 @@ -7,14 +7,14 @@ subsection\Iteration of the function \<^term>\F\\ -consts iterates :: "[i=>i,i,i] => i" (\(_^_ '(_'))\ [60,1000,1000] 60) +consts iterates :: "[i\i,i,i] \ i" (\(_^_ '(_'))\ [60,1000,1000] 60) primrec "F^0 (x) = x" "F^(succ(n)) (x) = F(F^n (x))" definition - iterates_omega :: "[i=>i,i] => i" (\(_^\ '(_'))\ [60,1000] 60) where + iterates_omega :: "[i\i,i] \ i" (\(_^\ '(_'))\ [60,1000] 60) where "F^\ (x) \ \n\nat. F^n (x)" lemma iterates_triv: @@ -45,7 +45,7 @@ three cases of ordinals\ definition - transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where + transrec3 :: "[i, i, [i,i]\i, [i,i]\i] \i" where "transrec3(k, a, b, c) \ transrec(k, \x r. if x=0 then a diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ZF_Base.thy --- a/src/ZF/ZF_Base.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ZF_Base.thy Tue Sep 27 17:46:52 2022 +0100 @@ -49,10 +49,10 @@ 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{_ ./ _ \ _, _})\) + "_Replace" :: "[pttrn, pttrn, i, o] \ i" (\(1{_ ./ _ \ _, _})\) translations "{y. x\A, Q}" \ "CONST Replace(A, \x y. Q)" @@ -63,7 +63,7 @@ where "RepFun(A,f) \ {y . x\A, y=f(x)}" syntax - "_RepFun" :: "[i, pttrn, i] => i" (\(1{_ ./ _ \ _})\ [51,0,51]) + "_RepFun" :: "[i, pttrn, i] \ i" (\(1{_ ./ _ \ _})\ [51,0,51]) translations "{b. x\A}" \ "CONST RepFun(A, \x. b)" @@ -81,12 +81,12 @@ subsection \General union and intersection\ -definition Inter :: "i => i" (\\_\ [90] 90) +definition Inter :: "i \ i" (\\_\ [90] 90) where "\(A) \ { x\\(A) . \y\A. x\y}" syntax - "_UNION" :: "[pttrn, i, i] => i" (\(3\_\_./ _)\ 10) - "_INTER" :: "[pttrn, i, i] => i" (\(3\_\_./ _)\ 10) + "_UNION" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) + "_INTER" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) translations "\x\A. B" == "CONST Union({B. x\A})" "\x\A. B" == "CONST Inter({B. x\A})" @@ -97,7 +97,7 @@ (*Unordered pairs (Upair) express binary union/intersection and cons; set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*) -definition Upair :: "[i, i] => i" +definition Upair :: "[i, i] \ i" where "Upair(a,b) \ {y. x\Pow(Pow(0)), (x=0 \ y=a) | (x=Pow(0) \ y=b)}" definition Subset :: "[i, i] \ o" (infixl \\\ 50) \ \subset relation\ @@ -112,10 +112,10 @@ definition Int :: "[i, i] \ i" (infixl \\\ 70) \ \binary intersection\ where "A \ B \ \(Upair(A,B))" -definition cons :: "[i, i] => i" +definition cons :: "[i, i] \ i" where "cons(a,A) \ Upair(a,a) \ A" -definition succ :: "i => i" +definition succ :: "i \ i" where "succ(i) \ cons(i, i)" nonterminal "is" @@ -160,14 +160,14 @@ 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 '(_,_,_')\) + old_if :: "[o, i, i] \ i" (\if '(_,_,_')\) where "if(P,a,b) \ If(P,a,b)" subsection \Ordered Pairing\ (* this "symmetric" definition works better than {{a}, {a,b}} *) -definition Pair :: "[i, i] => i" +definition Pair :: "[i, i] \ i" where "Pair(a,b) \ {{a,a}, {a,b}}" definition fst :: "i \ i" @@ -182,10 +182,10 @@ (* Patterns -- extends pre-defined type "pttrn" used in abstractions *) nonterminal patterns syntax - "_pattern" :: "patterns => pttrn" (\\_\\) - "" :: "pttrn => patterns" (\_\) - "_patterns" :: "[pttrn, patterns] => patterns" (\_,/_\) - "_Tuple" :: "[i, is] => i" (\\(_,/ _)\\) + "_pattern" :: "patterns \ pttrn" (\\_\\) + "" :: "pttrn \ patterns" (\_\) + "_patterns" :: "[pttrn, patterns] \ patterns" (\_,/_\) + "_Tuple" :: "[i, is] \ i" (\\(_,/ _)\\) translations "\x, y, z\" == "\x, \y, z\\" "\x, y\" == "CONST Pair(x, y)" @@ -195,7 +195,7 @@ definition Sigma :: "[i, i \ i] \ i" where "Sigma(A,B) \ \x\A. \y\B(x). {\x,y\}" -abbreviation cart_prod :: "[i, i] => i" (infixr \\\ 80) \ \Cartesian product\ +abbreviation cart_prod :: "[i, i] \ i" (infixr \\\ 80) \ \Cartesian product\ where "A \ B \ Sigma(A, \_. B)" @@ -249,9 +249,9 @@ (* binder syntax *) syntax - "_PROD" :: "[pttrn, i, i] => i" (\(3\_\_./ _)\ 10) - "_SUM" :: "[pttrn, i, i] => i" (\(3\_\_./ _)\ 10) - "_lam" :: "[pttrn, i, i] => i" (\(3\_\_./ _)\ 10) + "_PROD" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) + "_SUM" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) + "_lam" :: "[pttrn, i, i] \ i" (\(3\_\_./ _)\ 10) translations "\x\A. B" == "CONST Pi(A, \x. B)" "\x\A. B" == "CONST Sigma(A, \x. B)" @@ -270,18 +270,18 @@ not_mem (infixl \\:\ 50) syntax (ASCII) - "_Ball" :: "[pttrn, i, o] => o" (\(3ALL _:_./ _)\ 10) - "_Bex" :: "[pttrn, i, o] => o" (\(3EX _:_./ _)\ 10) - "_Collect" :: "[pttrn, i, o] => i" (\(1{_: _ ./ _})\) - "_Replace" :: "[pttrn, pttrn, i, o] => i" (\(1{_ ./ _: _, _})\) - "_RepFun" :: "[i, pttrn, i] => i" (\(1{_ ./ _: _})\ [51,0,51]) - "_UNION" :: "[pttrn, i, i] => i" (\(3UN _:_./ _)\ 10) - "_INTER" :: "[pttrn, i, i] => i" (\(3INT _:_./ _)\ 10) - "_PROD" :: "[pttrn, i, i] => i" (\(3PROD _:_./ _)\ 10) - "_SUM" :: "[pttrn, i, i] => i" (\(3SUM _:_./ _)\ 10) - "_lam" :: "[pttrn, i, i] => i" (\(3lam _:_./ _)\ 10) - "_Tuple" :: "[i, is] => i" (\<(_,/ _)>\) - "_pattern" :: "patterns => pttrn" (\<_>\) + "_Ball" :: "[pttrn, i, o] \ o" (\(3ALL _:_./ _)\ 10) + "_Bex" :: "[pttrn, i, o] \ o" (\(3EX _:_./ _)\ 10) + "_Collect" :: "[pttrn, i, o] \ i" (\(1{_: _ ./ _})\) + "_Replace" :: "[pttrn, pttrn, i, o] \ i" (\(1{_ ./ _: _, _})\) + "_RepFun" :: "[i, pttrn, i] \ i" (\(1{_ ./ _: _})\ [51,0,51]) + "_UNION" :: "[pttrn, i, i] \ i" (\(3UN _:_./ _)\ 10) + "_INTER" :: "[pttrn, i, i] \ i" (\(3INT _:_./ _)\ 10) + "_PROD" :: "[pttrn, i, i] \ i" (\(3PROD _:_./ _)\ 10) + "_SUM" :: "[pttrn, i, i] \ i" (\(3SUM _:_./ _)\ 10) + "_lam" :: "[pttrn, i, i] \ i" (\(3lam _:_./ _)\ 10) + "_Tuple" :: "[i, is] \ i" (\<(_,/ _)>\) + "_pattern" :: "patterns \ pttrn" (\<_>\) subsection \Substitution\ @@ -510,7 +510,7 @@ lemma Collect_cong [cong]: "\A=B; \x. x\B \ P(x) <-> Q(x)\ - \ Collect(A, %x. P(x)) = Collect(B, %x. Q(x))" + \ Collect(A, \x. P(x)) = Collect(B, \x. Q(x))" by (simp add: Collect_def) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/Zorn.thy --- a/src/ZF/Zorn.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/Zorn.thy Tue Sep 27 17:46:52 2022 +0100 @@ -11,23 +11,23 @@ Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.\ definition - Subset_rel :: "i=>i" where - "Subset_rel(A) \ {z \ A*A . \x y. z= \ x<=y \ x\y}" + Subset_rel :: "i\i" where + "Subset_rel(A) \ {z \ A*A . \x y. z=\x,y\ \ x<=y \ x\y}" definition - chain :: "i=>i" where + chain :: "i\i" where "chain(A) \ {F \ Pow(A). \X\F. \Y\F. X<=Y | Y<=X}" definition - super :: "[i,i]=>i" where + super :: "[i,i]\i" where "super(A,c) \ {d \ chain(A). c<=d \ c\d}" definition - maxchain :: "i=>i" where + maxchain :: "i\i" where "maxchain(A) \ {c \ chain(A). super(A,c)=0}" definition - increasing :: "i=>i" where + increasing :: "i\i" where "increasing(A) \ {f \ Pow(A)->Pow(A). \x. x<=A \ x<=f`x}" @@ -42,7 +42,7 @@ the induction rule lets us assume that condition! Many inductive proofs are therefore unconditional.\ consts - "TFin" :: "[i,i]=>i" + "TFin" :: "[i,i]\i" inductive domains "TFin(S,next)" \ "Pow(S)" @@ -396,7 +396,7 @@ \X \ Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X))\ \ (\ x \ S. \({y \ TFin(S,next). x \ y})) \ inj(S, TFin(S,next) - {S})" -apply (rule_tac d = "%y. ch` (S-y) " in lam_injective) +apply (rule_tac d = "\y. ch` (S-y) " in lam_injective) apply (rule DiffI) apply (rule Collect_subset [THEN TFin_UnionI]) apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE) @@ -435,8 +435,8 @@ text \Reimported from HOL by Clemens Ballarin.\ -definition Chain :: "i => i" where - "Chain(r) = {A \ Pow(field(r)). \a\A. \b\A. \ r | \ r}" +definition Chain :: "i \ i" where + "Chain(r) = {A \ Pow(field(r)). \a\A. \b\A. \a, b\ \ r | \b, a\ \ r}" lemma mono_Chain: "r \ s \ Chain(r) \ Chain(s)" @@ -445,8 +445,8 @@ theorem Zorn_po: assumes po: "Partial_order(r)" - and u: "\C\Chain(r). \u\field(r). \a\C. \ r" - shows "\m\field(r). \a\field(r). \ r \ a = m" + and u: "\C\Chain(r). \u\field(r). \a\C. \a, u\ \ r" + shows "\m\field(r). \a\field(r). \m, a\ \ r \ a = m" proof - have "Preorder(r)" using po by (simp add: partial_order_on_def) \ \Mirror r in the set of subsets below (wrt r) elements of A (?).\ @@ -477,11 +477,11 @@ fix a b assume "a \ field(r)" "r -`` {a} \ C" "b \ field(r)" "r -`` {b} \ C" hence "r -`` {a} \ r -`` {b} | r -`` {b} \ r -`` {a}" using 2 by auto - then show " \ r | \ r" + then show "\a, b\ \ r | \b, a\ \ r" using \Preorder(r)\ \a \ field(r)\ \b \ field(r)\ by (simp add: subset_vimage1_vimage1_iff) qed - then obtain u where uA: "u \ field(r)" "\a\?A. \ r" + then obtain u where uA: "u \ field(r)" "\a\?A. \a, u\ \ r" using u apply auto apply (drule bspec) apply assumption @@ -498,7 +498,7 @@ apply (erule lamE) apply simp done - then show " \ r" using uA aB \Preorder(r)\ + then show "\a, u\ \ r" using uA aB \Preorder(r)\ by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+ qed then show "\U\field(r). \A\C. A \ r -`` {U}" @@ -508,7 +508,7 @@ obtain m B where "m \ field(r)" "B = r -`` {m}" "\x\field(r). B \ r -`` {x} \ B = r -`` {x}" by (auto elim!: lamE simp: ball_image_simp) - then have "\a\field(r). \ r \ a = m" + then have "\a\field(r). \m, a\ \ r \ a = m" using po \Preorder(r)\ \m \ field(r)\ by (auto simp: subset_vimage1_vimage1_iff Partial_order_eq_vimage1_vimage1_iff) then show ?thesis using \m \ field(r)\ by blast diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/equalities.thy --- a/src/ZF/equalities.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/equalities.thy Tue Sep 27 17:46:52 2022 +0100 @@ -36,18 +36,18 @@ subsection\Converse of a Relation\ -lemma converse_iff [simp]: "\ converse(r) \ \r" +lemma converse_iff [simp]: "\a,b\\ converse(r) \ \b,a\\r" by (unfold converse_def, blast) -lemma converseI [intro!]: "\r \ \converse(r)" +lemma converseI [intro!]: "\a,b\\r \ \b,a\\converse(r)" by (unfold converse_def, blast) -lemma converseD: " \ converse(r) \ \ r" +lemma converseD: "\a,b\ \ converse(r) \ \b,a\ \ r" by (unfold converse_def, blast) lemma converseE [elim!]: "\yx \ converse(r); - \x y. \yx=; \r\ \ P\ + \x y. \yx=\y,x\; \x,y\\r\ \ P\ \ P" by (unfold converse_def, blast) @@ -482,7 +482,7 @@ lemma INT_Union_eq: "0 \ A \ (\x\ \(A). B(x)) = (\y\A. \x\y. B(x))" apply (subgoal_tac "\x\A. x\0") - prefer 2 apply blast + prefer 2 apply blast apply (force simp add: Inter_def ball_conj_distrib) done @@ -580,14 +580,14 @@ (** Domain **) -lemma domain_iff: "a: domain(r) \ (\y. \ r)" +lemma domain_iff: "a: domain(r) \ (\y. \a,y\\ r)" by (unfold domain_def, blast) -lemma domainI [intro]: "\ r \ a: domain(r)" +lemma domainI [intro]: "\a,b\\ r \ a: domain(r)" by (unfold domain_def, blast) lemma domainE [elim!]: - "\a \ domain(r); \y. \ r \ P\ \ P" + "\a \ domain(r); \y. \a,y\\ r \ P\ \ P" by (unfold domain_def, blast) lemma domain_subset: "domain(Sigma(A,B)) \ A" @@ -599,7 +599,7 @@ lemma domain_0 [simp]: "domain(0) = 0" by blast -lemma domain_cons [simp]: "domain(cons(,r)) = cons(a, domain(r))" +lemma domain_cons [simp]: "domain(cons(\a,b\,r)) = cons(a, domain(r))" by blast lemma domain_Un_eq [simp]: "domain(A \ B) = domain(A) \ domain(B)" @@ -620,12 +620,12 @@ (** Range **) -lemma rangeI [intro]: "\ r \ b \ range(r)" +lemma rangeI [intro]: "\a,b\\ r \ b \ range(r)" apply (unfold range_def) apply (erule converseI [THEN domainI]) done -lemma rangeE [elim!]: "\b \ range(r); \x. \ r \ P\ \ P" +lemma rangeE [elim!]: "\b \ range(r); \x. \x,b\\ r \ P\ \ P" by (unfold range_def, blast) lemma range_subset: "range(A*B) \ B" @@ -640,7 +640,7 @@ lemma range_0 [simp]: "range(0) = 0" by blast -lemma range_cons [simp]: "range(cons(,r)) = cons(b, range(r))" +lemma range_cons [simp]: "range(cons(\a,b\,r)) = cons(b, range(r))" by blast lemma range_Un_eq [simp]: "range(A \ B) = range(A) \ range(B)" @@ -661,21 +661,21 @@ (** Field **) -lemma fieldI1: "\ r \ a \ field(r)" +lemma fieldI1: "\a,b\\ r \ a \ field(r)" by (unfold field_def, blast) -lemma fieldI2: "\ r \ b \ field(r)" +lemma fieldI2: "\a,b\\ r \ b \ field(r)" by (unfold field_def, blast) lemma fieldCI [intro]: - "(\ \r \ \ r) \ a \ field(r)" + "(\ \c,a\\r \ \a,b\\ r) \ a \ field(r)" apply (unfold field_def, blast) done lemma fieldE [elim!]: "\a \ field(r); - \x. \ r \ P; - \x. \ r \ P\ \ P" + \x. \a,x\\ r \ P; + \x. \x,a\\ r \ P\ \ P" by (unfold field_def, blast) lemma field_subset: "field(A*B) \ A \ B" @@ -706,7 +706,7 @@ lemma field_0 [simp]: "field(0) = 0" by blast -lemma field_cons [simp]: "field(cons(,r)) = cons(a, cons(b, field(r)))" +lemma field_cons [simp]: "field(cons(\a,b\,r)) = cons(a, cons(b, field(r)))" by blast lemma field_Un_eq [simp]: "field(A \ B) = field(A) \ field(B)" @@ -730,26 +730,26 @@ lemma rel_Un: "\r \ A*B; s \ C*D\ \ (r \ s) \ (A \ C) * (B \ D)" by blast -lemma domain_Diff_eq: "\ \ r; c\b\ \ domain(r-{}) = domain(r)" +lemma domain_Diff_eq: "\\a,c\ \ r; c\b\ \ domain(r-{\a,b\}) = domain(r)" by blast -lemma range_Diff_eq: "\ \ r; c\a\ \ range(r-{}) = range(r)" +lemma range_Diff_eq: "\\c,b\ \ r; c\a\ \ range(r-{\a,b\}) = range(r)" by blast subsection\Image of a Set under a Function or Relation\ -lemma image_iff: "b \ r``A \ (\x\A. \r)" +lemma image_iff: "b \ r``A \ (\x\A. \x,b\\r)" by (unfold image_def, blast) -lemma image_singleton_iff: "b \ r``{a} \ \r" +lemma image_singleton_iff: "b \ r``{a} \ \a,b\\r" by (rule image_iff [THEN iff_trans], blast) -lemma imageI [intro]: "\\ r; a\A\ \ b \ r``A" +lemma imageI [intro]: "\\a,b\\ r; a\A\ \ b \ r``A" by (unfold image_def, blast) lemma imageE [elim!]: - "\b: r``A; \x.\\ r; x\A\ \ P\ \ P" + "\b: r``A; \x.\\x,b\\ r; x\A\ \ P\ \ P" by (unfold image_def, blast) lemma image_subset: "r \ A*B \ r``C \ B" @@ -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(\x,y\)})" by blast lemma image_Int_subset: "r``(A \ B) \ (r``A) \ (r``B)" @@ -792,17 +792,17 @@ subsection\Inverse Image of a Set under a Function or Relation\ lemma vimage_iff: - "a \ r-``B \ (\y\B. \r)" + "a \ r-``B \ (\y\B. \a,y\\r)" by (unfold vimage_def image_def converse_def, blast) -lemma vimage_singleton_iff: "a \ r-``{b} \ \r" +lemma vimage_singleton_iff: "a \ r-``{b} \ \a,b\\r" by (rule vimage_iff [THEN iff_trans], blast) -lemma vimageI [intro]: "\\ r; b\B\ \ a \ r-``B" +lemma vimageI [intro]: "\\a,b\\ r; b\B\ \ a \ r-``B" by (unfold vimage_def, blast) lemma vimageE [elim!]: - "\a: r-``B; \x.\\ r; x\B\ \ P\ \ P" + "\a: r-``B; \x.\\a,x\\ r; x\B\ \ P\ \ P" apply (unfold vimage_def, blast) done @@ -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]: diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/CoUnit.thy --- a/src/ZF/ex/CoUnit.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/CoUnit.thy Tue Sep 27 17:46:52 2022 +0100 @@ -59,13 +59,13 @@ \ \Proving freeness results.\ by (fast elim!: counit2.free_elims) -lemma Con2_bnd_mono: "bnd_mono(univ(0), %x. Con2(x, x))" +lemma Con2_bnd_mono: "bnd_mono(univ(0), \x. Con2(x, x))" apply (unfold counit2.con_defs) apply (rule bnd_monoI) apply (assumption | rule subset_refl QPair_subset_univ QPair_mono)+ done -lemma lfp_Con2_in_counit2: "lfp(univ(0), %x. Con2(x,x)) \ counit2" +lemma lfp_Con2_in_counit2: "lfp(univ(0), \x. Con2(x,x)) \ counit2" apply (rule singletonI [THEN counit2.coinduct]) apply (rule qunivI [THEN singleton_subsetI]) apply (rule subset_trans [OF lfp_subset empty_subsetI [THEN univ_mono]]) @@ -90,7 +90,7 @@ apply (assumption | rule conjI counit2_Int_Vset_subset [THEN Int_Vset_subset])+ done -lemma counit2_eq_univ: "counit2 = {lfp(univ(0), %x. Con2(x,x))}" +lemma counit2_eq_univ: "counit2 = {lfp(univ(0), \x. Con2(x,x))}" apply (rule equalityI) apply (rule_tac [2] lfp_Con2_in_counit2 [THEN singleton_subsetI]) apply (rule subsetI) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/Commutation.thy --- a/src/ZF/ex/Commutation.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/Commutation.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,29 +8,29 @@ theory Commutation imports ZF begin definition - square :: "[i, i, i, i] => o" where + square :: "[i, i, i, i] \ o" where "square(r,s,t,u) \ - (\a b. \ r \ (\c. \ s \ (\x. \ t \ \ u)))" + (\a b. \a,b\ \ r \ (\c. \a, c\ \ s \ (\x. \b,x\ \ t \ \c,x\ \ u)))" definition - commute :: "[i, i] => o" where + commute :: "[i, i] \ o" where "commute(r,s) \ square(r,s,s,r)" definition - diamond :: "i=>o" where + diamond :: "i\o" where "diamond(r) \ commute(r, r)" definition - strip :: "i=>o" where + strip :: "i\o" where "strip(r) \ commute(r^*, r)" definition - Church_Rosser :: "i => o" where - "Church_Rosser(r) \ (\x y. \ (r \ converse(r))^* \ - (\z. \ r^* \ \ r^*))" + Church_Rosser :: "i \ o" where + "Church_Rosser(r) \ (\x y. \x,y\ \ (r \ converse(r))^* \ + (\z. \x,z\ \ r^* \ \y,z\ \ r^*))" definition - confluent :: "i=>o" where + confluent :: "i\o" where "confluent(r) \ diamond(r^*)" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/Group.thy --- a/src/ZF/ex/Group.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/Group.thy Tue Sep 27 17:46:52 2022 +0100 @@ -13,28 +13,28 @@ (*First, we must simulate a record declaration: record monoid = carrier :: i - mult :: "[i,i] => i" (infixl "\\" 70) + mult :: "[i,i] \ i" (infixl "\\" 70) one :: i ("\\") *) definition - carrier :: "i => i" where + carrier :: "i \ i" where "carrier(M) \ fst(M)" definition - mmult :: "[i, i, i] => i" (infixl \\\\ 70) where - "mmult(M,x,y) \ fst(snd(M)) ` " + mmult :: "[i, i, i] \ i" (infixl \\\\ 70) where + "mmult(M,x,y) \ fst(snd(M)) ` \x,y\" definition - one :: "i => i" (\\\\) where + one :: "i \ i" (\\\\) where "one(M) \ fst(snd(snd(M)))" definition - update_carrier :: "[i,i] => i" where + update_carrier :: "[i,i] \ i" where "update_carrier(M,A) \ " definition - m_inv :: "i => i => i" (\inv\ _\ [81] 80) where + m_inv :: "i \ i \ i" (\inv\ _\ [81] 80) where "inv\<^bsub>G\<^esub> x \ (THE y. y \ carrier(G) \ y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub> \ x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub>)" locale monoid = fixes G (structure) @@ -48,16 +48,16 @@ and r_one [simp]: "x \ carrier(G) \ x \ \ = x" text\Simulating the record\ -lemma carrier_eq [simp]: "carrier() = A" +lemma carrier_eq [simp]: "carrier(\A,Z\) = A" by (simp add: carrier_def) -lemma mult_eq [simp]: "mmult(, x, y) = M ` " +lemma mult_eq [simp]: "mmult(, x, y) = M ` \x,y\" by (simp add: mmult_def) lemma one_eq [simp]: "one() = I" by (simp add: one_def) -lemma update_carrier_eq [simp]: "update_carrier(,B) = " +lemma update_carrier_eq [simp]: "update_carrier(\A,Z\,B) = \B,Z\" by (simp add: update_carrier_def) lemma carrier_update_carrier [simp]: "carrier(update_carrier(M,B)) = B" @@ -302,9 +302,9 @@ subsection \Direct Products\ definition - DirProdGroup :: "[i,i] => i" (infixr \\\ 80) where + DirProdGroup :: "[i,i] \ i" (infixr \\\ 80) where "G \ H \ carrier(H), - (\<, > + (\<\g,h\, > \ (carrier(G) \ carrier(H)) \ (carrier(G) \ carrier(H)). \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>), <\\<^bsub>G\<^esub>, \\<^bsub>H\<^esub>>, 0>" @@ -329,14 +329,14 @@ lemma mult_DirProdGroup [simp]: "\g \ carrier(G); h \ carrier(H); g' \ carrier(G); h' \ carrier(H)\ - \ \\<^bsub>G \ H\<^esub> = \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>" + \ \g, h\ \\<^bsub>G \ H\<^esub> = \<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h'>" by (simp add: DirProdGroup_def) lemma inv_DirProdGroup [simp]: assumes "group(G)" and "group(H)" assumes g: "g \ carrier(G)" and h: "h \ carrier(H)" - shows "inv \<^bsub>G \ H\<^esub> = G\<^esub> g, inv\<^bsub>H\<^esub> h>" + shows "inv \<^bsub>G \ H\<^esub> \g, h\ = G\<^esub> g, inv\<^bsub>H\<^esub> h>" apply (rule group.inv_equality [OF DirProdGroup_group]) apply (simp_all add: assms group.l_inv) done @@ -344,7 +344,7 @@ subsection \Isomorphisms\ definition - hom :: "[i,i] => i" where + hom :: "[i,i] \ i" where "hom(G,H) \ {h \ carrier(G) -> carrier(H). (\x \ carrier(G). \y \ carrier(G). h ` (x \\<^bsub>G\<^esub> y) = (h ` x) \\<^bsub>H\<^esub> (h ` y))}" @@ -370,7 +370,7 @@ subsection \Isomorphisms\ definition - iso :: "[i,i] => i" (infixr \\\ 60) where + iso :: "[i,i] \ i" (infixr \\\ 60) where "G \ H \ hom(G,H) \ bij(carrier(G), carrier(H))" lemma (in group) iso_refl: "id(carrier(G)) \ G \ G" @@ -392,7 +392,7 @@ lemma DirProdGroup_commute_iso: assumes "group(G)" and "group(H)" - shows "(\ \ carrier(G \ H). ) \ (G \ H) \ (H \ G)" + shows "(\\x,y\ \ carrier(G \ H). \y,x\) \ (G \ H) \ (H \ G)" proof - interpret group G by fact interpret group H by fact @@ -401,7 +401,7 @@ lemma DirProdGroup_assoc_iso: assumes "group(G)" and "group(H)" and "group(I)" - shows "(\<,z> \ carrier((G \ H) \ I). >) + shows "(\<\x,y\,z> \ carrier((G \ H) \ I). y,z\>) \ ((G \ H) \ I) \ (G \ (H \ I))" proof - interpret group G by fact @@ -502,10 +502,10 @@ subsection \Bijections of a Set, Permutation Groups, Automorphism Groups\ definition - BijGroup :: "i=>i" where + BijGroup :: "i\i" where "BijGroup(S) \ \ bij(S,S) \ bij(S,S). g O f, + \\g,f\ \ bij(S,S) \ bij(S,S). g O f, id(S), 0>" @@ -537,11 +537,11 @@ definition - auto :: "i=>i" where + auto :: "i\i" where "auto(G) \ iso(G,G)" definition - AutoGroup :: "i=>i" where + AutoGroup :: "i\i" where "AutoGroup(G) \ update_carrier(BijGroup(carrier(G)), auto(G))" @@ -576,23 +576,23 @@ subsection\Cosets and Quotient Groups\ definition - r_coset :: "[i,i,i] => i" (infixl \#>\\ 60) where + r_coset :: "[i,i,i] \ i" (infixl \#>\\ 60) where "H #>\<^bsub>G\<^esub> a \ \h\H. {h \\<^bsub>G\<^esub> a}" definition - l_coset :: "[i,i,i] => i" (infixl \<#\\ 60) where + l_coset :: "[i,i,i] \ i" (infixl \<#\\ 60) where "a <#\<^bsub>G\<^esub> H \ \h\H. {a \\<^bsub>G\<^esub> h}" definition - RCOSETS :: "[i,i] => i" (\rcosets\ _\ [81] 80) where + RCOSETS :: "[i,i] \ i" (\rcosets\ _\ [81] 80) where "rcosets\<^bsub>G\<^esub> H \ \a\carrier(G). {H #>\<^bsub>G\<^esub> a}" definition - set_mult :: "[i,i,i] => i" (infixl \<#>\\ 60) where + set_mult :: "[i,i,i] \ i" (infixl \<#>\\ 60) where "H <#>\<^bsub>G\<^esub> K \ \h\H. \k\K. {h \\<^bsub>G\<^esub> k}" definition - SET_INV :: "[i,i] => i" (\set'_inv\ _\ [81] 80) where + SET_INV :: "[i,i] \ i" (\set'_inv\ _\ [81] 80) where "set_inv\<^bsub>G\<^esub> H \ \h\H. {inv\<^bsub>G\<^esub> h}" @@ -860,8 +860,8 @@ subsubsection\Two distinct right cosets are disjoint\ definition - r_congruent :: "[i,i] => i" (\rcong\ _\ [60] 60) where - "rcong\<^bsub>G\<^esub> H \ { \ carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \\<^bsub>G\<^esub> y \ H}" + r_congruent :: "[i,i] \ i" (\rcong\ _\ [60] 60) where + "rcong\<^bsub>G\<^esub> H \ {\x,y\ \ carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \\<^bsub>G\<^esub> y \ H}" lemma (in subgroup) equiv_rcong: @@ -940,7 +940,7 @@ subsection \Order of a Group and Lagrange's Theorem\ definition - order :: "i => i" where + order :: "i \ i" where "order(S) \ |carrier(S)|" lemma (in group) rcos_self: @@ -1020,10 +1020,10 @@ subsection \Quotient Groups: Factorization of a Group\ definition - FactGroup :: "[i,i] => i" (infixl \Mod\ 65) where + FactGroup :: "[i,i] \ i" (infixl \Mod\ 65) where \ \Actually defined for groups rather than monoids\ "G Mod H \ - G\<^esub> H, \ \ (rcosets\<^bsub>G\<^esub> H) \ (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>" + G\<^esub> H, \\K1,K2\ \ (rcosets\<^bsub>G\<^esub> H) \ (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>" lemma (in normal) setmult_closed: "\K1 \ rcosets H; K2 \ rcosets H\ \ K1 <#> K2 \ rcosets H" @@ -1084,7 +1084,7 @@ range of that homomorphism.\ definition - kernel :: "[i,i,i] => i" where + kernel :: "[i,i,i] \ i" where \ \the kernel of a homomorphism\ "kernel(G,H,h) \ {x \ carrier(G). h ` x = \\<^bsub>H\<^esub>}" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/LList.thy --- a/src/ZF/ex/LList.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/LList.thy Tue Sep 27 17:46:52 2022 +0100 @@ -16,7 +16,7 @@ theory LList imports ZF begin consts - llist :: "i=>i" + llist :: "i\i" codatatype "llist(A)" = LNil | LCons ("a \ A", "l \ llist(A)") @@ -24,7 +24,7 @@ (*Coinductive definition of equality*) consts - lleq :: "i=>i" + lleq :: "i\i" (*Previously used <*> in the domain and variant pairs as elements. But standard pairs work just as well. To use variant pairs, must change prefix @@ -32,7 +32,7 @@ coinductive domains "lleq(A)" \ "llist(A) * llist(A)" intros - LNil: " \ lleq(A)" + LNil: "\LNil, LNil\ \ lleq(A)" LCons: "\a \ A; \ lleq(A)\ \ \ lleq(A)" type_intros llist.intros @@ -40,10 +40,10 @@ (*Lazy list functions; flip is not definitional!*) definition - lconst :: "i => i" where - "lconst(a) \ lfp(univ(a), %l. LCons(a,l))" + lconst :: "i \ i" where + "lconst(a) \ lfp(univ(a), \l. LCons(a,l))" -axiomatization flip :: "i => i" +axiomatization flip :: "i \ i" where flip_LNil: "flip(LNil) = LNil" and flip_LCons: "\x \ bool; l \ llist(bool)\ @@ -170,7 +170,7 @@ lemma equal_llist_implies_leq: "\l=l'; l \ llist(A)\ \ \ lleq(A)" -apply (rule_tac X = "{. l \ llist (A) }" in lleq.coinduct) +apply (rule_tac X = "{\l,l\. l \ llist (A) }" in lleq.coinduct) apply blast apply safe apply (erule_tac a=la in llist.cases, fast+) @@ -183,7 +183,7 @@ (*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***) -lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), %l. LCons(a,l))" +lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), \l. LCons(a,l))" apply (unfold llist.con_defs ) apply (rule bnd_monoI) apply (blast intro: A_subset_univ QInr_subset_univ) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/Limit.thy --- a/src/ZF/ex/Limit.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/Limit.thy Tue Sep 27 17:46:52 2022 +0100 @@ -20,15 +20,15 @@ theory Limit imports ZF begin definition - rel :: "[i,i,i]=>o" where - "rel(D,x,y) \ :snd(D)" + rel :: "[i,i,i]\o" where + "rel(D,x,y) \ \x,y\:snd(D)" definition - set :: "i=>i" where + set :: "i\i" where "set(D) \ fst(D)" definition - po :: "i=>o" where + po :: "i\o" where "po(D) \ (\x \ set(D). rel(D,x,x)) \ (\x \ set(D). \y \ set(D). \z \ set(D). @@ -36,66 +36,66 @@ (\x \ set(D). \y \ set(D). rel(D,x,y) \ rel(D,y,x) \ x = y)" definition - chain :: "[i,i]=>o" where + 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))))" definition - isub :: "[i,i,i]=>o" where + isub :: "[i,i,i]\o" where "isub(D,X,x) \ x \ set(D) \ (\n \ nat. rel(D,X`n,x))" definition - islub :: "[i,i,i]=>o" where + islub :: "[i,i,i]\o" where "islub(D,X,x) \ isub(D,X,x) \ (\y. isub(D,X,y) \ rel(D,x,y))" definition - lub :: "[i,i]=>i" where + lub :: "[i,i]\i" where "lub(D,X) \ THE x. islub(D,X,x)" definition - cpo :: "i=>o" where + cpo :: "i\o" where "cpo(D) \ po(D) \ (\X. chain(D,X) \ (\x. islub(D,X,x)))" definition - pcpo :: "i=>o" where + pcpo :: "i\o" where "pcpo(D) \ cpo(D) \ (\x \ set(D). \y \ set(D). rel(D,x,y))" definition - bot :: "i=>i" where + bot :: "i\i" where "bot(D) \ THE x. x \ set(D) \ (\y \ set(D). rel(D,x,y))" definition - mono :: "[i,i]=>i" where + mono :: "[i,i]\i" where "mono(D,E) \ {f \ set(D)->set(E). \x \ set(D). \y \ set(D). rel(D,x,y) \ rel(E,f`x,f`y)}" definition - cont :: "[i,i]=>i" where + cont :: "[i,i]\i" where "cont(D,E) \ {f \ mono(D,E). \X. chain(D,X) \ f`(lub(D,X)) = lub(E,\n \ nat. f`(X`n))}" definition - cf :: "[i,i]=>i" where + cf :: "[i,i]\i" where "cf(D,E) \ cont(D,E)*cont(D,E). \x \ set(D). rel(E,(fst(y))`x,(snd(y))`x)}>" definition - suffix :: "[i,i]=>i" where + suffix :: "[i,i]\i" where "suffix(X,n) \ \m \ nat. X`(n #+ m)" definition - subchain :: "[i,i]=>o" where + subchain :: "[i,i]\o" where "subchain(X,Y) \ \m \ nat. \n \ nat. X`m = Y`(m #+ n)" definition - dominate :: "[i,i,i]=>o" where + dominate :: "[i,i,i]\o" where "dominate(D,X,Y) \ \m \ nat. \n \ nat. rel(D,X`m,Y`n)" definition - matrix :: "[i,i]=>o" where + 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)) \ @@ -103,89 +103,89 @@ (\n \ nat. \m \ nat. rel(D,M`n`m,M`succ(n)`succ(m)))" definition - projpair :: "[i,i,i,i]=>o" where + 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)))" definition - emb :: "[i,i,i]=>o" where + emb :: "[i,i,i]\o" where "emb(D,E,e) \ \p. projpair(D,E,e,p)" definition - Rp :: "[i,i,i]=>i" where + Rp :: "[i,i,i]\i" where "Rp(D,E,e) \ THE p. projpair(D,E,e,p)" definition (* Twice, constructions on cpos are more difficult. *) - iprod :: "i=>i" where + iprod :: "i\i" where "iprod(DD) \ <(\n \ nat. set(DD`n)), {x:(\n \ nat. set(DD`n))*(\n \ nat. set(DD`n)). \n \ nat. rel(DD`n,fst(x)`n,snd(x)`n)}>" definition - mkcpo :: "[i,i=>o]=>i" where + mkcpo :: "[i,i\o]\i" where (* Cannot use rel(D), is meta fun, need two more args *) "mkcpo(D,P) \ <{x \ set(D). P(x)},{x \ set(D)*set(D). rel(D,fst(x),snd(x))}>" definition - subcpo :: "[i,i]=>o" where + subcpo :: "[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)) \ (\X. chain(D,X) \ lub(E,X):set(D))" definition - subpcpo :: "[i,i]=>o" where + subpcpo :: "[i,i]\o" where "subpcpo(D,E) \ subcpo(D,E) \ bot(E):set(D)" definition - emb_chain :: "[i,i]=>o" where + 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))" definition - Dinf :: "[i,i]=>i" where + Dinf :: "[i,i]\i" where "Dinf(DD,ee) \ mkcpo(iprod(DD)) - (%x. \n \ nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)" + (\x. \n \ nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)" definition - e_less :: "[i,i,i,i]=>i" where + e_less :: "[i,i,i,i]\i" where (* Valid for m \ n only. *) - "e_less(DD,ee,m,n) \ rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)" + "e_less(DD,ee,m,n) \ rec(n#-m,id(set(DD`m)),\x y. ee`(m#+x) O y)" definition - e_gr :: "[i,i,i,i]=>i" where + e_gr :: "[i,i,i,i]\i" where (* Valid for n \ m only. *) "e_gr(DD,ee,m,n) \ rec(m#-n,id(set(DD`n)), - %x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))" + \x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))" definition - eps :: "[i,i,i,i]=>i" where + eps :: "[i,i,i,i]\i" where "eps(DD,ee,m,n) \ if(m \ n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))" definition - rho_emb :: "[i,i,i]=>i" where + rho_emb :: "[i,i,i]\i" where "rho_emb(DD,ee,n) \ \x \ set(DD`n). \m \ nat. eps(DD,ee,n,m)`x" definition - rho_proj :: "[i,i,i]=>i" where + rho_proj :: "[i,i,i]\i" where "rho_proj(DD,ee,n) \ \x \ set(Dinf(DD,ee)). x`n" definition - commute :: "[i,i,i,i=>i]=>o" where + commute :: "[i,i,i,i\i]\o" where "commute(DD,ee,E,r) \ (\n \ nat. emb(DD`n,E,r(n))) \ (\m \ nat. \n \ nat. m \ n \ r(n) O eps(DD,ee,m,n) = r(m))" definition - mediating :: "[i,i,i=>i,i=>i,i]=>o" where + mediating :: "[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))" @@ -199,10 +199,10 @@ lemma set_I: "x \ fst(D) \ x \ set(D)" by (simp add: set_def) -lemma rel_I: ":snd(D) \ rel(D,x,y)" +lemma rel_I: "\x,y\:snd(D) \ rel(D,x,y)" by (simp add: rel_def) -lemma rel_E: "rel(D,x,y) \ :snd(D)" +lemma rel_E: "rel(D,x,y) \ \x,y\:snd(D)" by (simp add: rel_def) (*----------------------------------------------------------------------*) @@ -936,7 +936,7 @@ apply (subst comp_assoc) apply (blast intro: cpo_cf cont_cf comp_mono comp_pres_cont dest: projpair_ep_cont) -apply (rule_tac P = "%x. rel (cf (E,D),p O e' O p',x)" +apply (rule_tac P = "\x. rel (cf (E,D),p O e' O p',x)" in projpair_p_cont [THEN cont_fun, THEN comp_id, THEN subst], assumption) apply (rule comp_mono) @@ -959,7 +959,7 @@ apply (subst comp_assoc) apply (blast intro: cpo_cf cont_cf comp_mono comp_pres_cont dest: projpair_ep_cont) -apply (rule_tac P = "%x. rel (cf (D,E), (e O p) O e',x)" +apply (rule_tac P = "\x. rel (cf (D,E), (e O p) O e',x)" in projpair_e_cont [THEN cont_fun, THEN id_comp, THEN subst], assumption) apply (blast intro: cpo_cf cont_cf comp_pres_cont projpair_rel comp_mono @@ -1108,8 +1108,8 @@ apply (blast intro: lam_type chain_iprod [THEN cpo_lub, THEN islub_in]) apply (rule rel_iprodI, simp) (*looks like something should be inserted into the assumptions!*) -apply (rule_tac P = "%t. rel (DD`na,t,lub (DD`na,\x \ nat. X`x`na))" - and b1 = "%n. X`n`na" in beta [THEN subst]) +apply (rule_tac P = "\t. rel (DD`na,t,lub (DD`na,\x \ nat. X`x`na))" + and b1 = "\n. X`n`na" in beta [THEN subst]) apply (simp del: beta_if add: chain_iprod [THEN cpo_lub, THEN islub_ub] iprodE chain_in)+ @@ -1832,12 +1832,12 @@ apply (rule_tac [3] comp_fun_apply [THEN subst]) apply (rename_tac [5] y) apply (rule_tac [5] P = - "%z. rel(DD`succ(y), + "\z. rel(DD`succ(y), (ee`y O Rp(DD'(y)`y,DD'(y)`succ(y),ee'(y)`y)) ` (x`succ(y)), z)" for DD' ee' in id_conv [THEN subst]) apply (rule_tac [6] rel_cf) -(* Dinf and cont_fun doesn't go well together, both Pi(_,%x._). *) +(* Dinf and cont_fun doesn't go well together, both Pi(_,\x._). *) (* solves 10 of 11 subgoals *) apply (assumption | rule Dinf_prod [THEN apply_type] cont_fun Rp_cont e_less_cont diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/NatSum.thy --- a/src/ZF/ex/NatSum.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/NatSum.thy Tue Sep 27 17:46:52 2022 +0100 @@ -18,7 +18,7 @@ theory NatSum imports ZF begin -consts sum :: "[i=>i, i] => i" +consts sum :: "[i\i, i] \ i" primrec "sum (f,0) = #0" "sum (f, succ(n)) = f($#n) $+ sum(f,n)" @@ -27,41 +27,41 @@ declare zdiff_zmult_distrib [simp] zdiff_zmult_distrib2 [simp] (*The sum of the first n odd numbers equals n squared.*) -lemma sum_of_odds: "n \ nat \ sum (%i. i $+ i $+ #1, n) = $#n $* $#n" +lemma sum_of_odds: "n \ nat \ sum (\i. i $+ i $+ #1, n) = $#n $* $#n" by (induct_tac "n", auto) (*The sum of the first n odd squares*) lemma sum_of_odd_squares: - "n \ nat \ #3 $* sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) = + "n \ nat \ #3 $* sum (\i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) = $#n $* (#4 $* $#n $* $#n $- #1)" by (induct_tac "n", auto) (*The sum of the first n odd cubes*) lemma sum_of_odd_cubes: "n \ nat - \ sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) = + \ sum (\i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) = $#n $* $#n $* (#2 $* $#n $* $#n $- #1)" by (induct_tac "n", auto) (*The sum of the first n positive integers equals n(n+1)/2.*) lemma sum_of_naturals: - "n \ nat \ #2 $* sum(%i. i, succ(n)) = $#n $* $#succ(n)" + "n \ nat \ #2 $* sum(\i. i, succ(n)) = $#n $* $#succ(n)" by (induct_tac "n", auto) lemma sum_of_squares: - "n \ nat \ #6 $* sum (%i. i$*i, succ(n)) = + "n \ nat \ #6 $* sum (\i. i$*i, succ(n)) = $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1)" by (induct_tac "n", auto) lemma sum_of_cubes: - "n \ nat \ #4 $* sum (%i. i$*i$*i, succ(n)) = + "n \ nat \ #4 $* sum (\i. i$*i$*i, succ(n)) = $#n $* $#n $* ($#n $+ #1) $* ($#n $+ #1)" by (induct_tac "n", auto) (** Sum of fourth powers **) lemma sum_of_fourth_powers: - "n \ nat \ #30 $* sum (%i. i$*i$*i$*i, succ(n)) = + "n \ nat \ #30 $* sum (\i. i$*i$*i$*i, succ(n)) = $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1) $* (#3 $* $#n $* $#n $+ #3 $* $#n $- #1)" by (induct_tac "n", auto) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/Primes.thy --- a/src/ZF/ex/Primes.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/Primes.thy Tue Sep 27 17:46:52 2022 +0100 @@ -8,22 +8,22 @@ theory Primes imports ZF begin definition - divides :: "[i,i]=>o" (infixl \dvd\ 50) where + divides :: "[i,i]\o" (infixl \dvd\ 50) where "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 :: "[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)" definition - gcd :: "[i,i]=>i" \ \Euclid's algorithm for the gcd\ where + gcd :: "[i,i]\i" \ \Euclid's algorithm for the gcd\ where "gcd(m,n) \ transrec(natify(n), - %n f. \m \ nat. + \n f. \m \ nat. if n=0 then m else f`(m mod n)`n) ` natify(m)" definition - coprime :: "[i,i]=>o" \ \the coprime relation\ where + coprime :: "[i,i]\o" \ \the coprime relation\ where "coprime(m,n) \ gcd(m,n) = 1" definition @@ -93,7 +93,7 @@ lemma gcd_non_0_raw: "\0 nat\ \ gcd(m,n) = gcd(n, m mod n)" apply (simp add: gcd_def) -apply (rule_tac P = "%z. left (z) = right" for left right in transrec [THEN ssubst]) +apply (rule_tac P = "\z. left (z) = right" for left right in transrec [THEN ssubst]) apply (simp add: ltD [THEN mem_imp_not_eq, THEN not_sym] mod_less_divisor [THEN ltD]) done diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/Ramsey.thy --- a/src/ZF/ex/Ramsey.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/Ramsey.thy Tue Sep 27 17:46:52 2022 +0100 @@ -27,23 +27,23 @@ theory Ramsey imports ZF begin definition - Symmetric :: "i=>o" where - "Symmetric(E) \ (\x y. :E \ :E)" + Symmetric :: "i\o" where + "Symmetric(E) \ (\x y. \x,y\:E \ \y,x\:E)" definition - Atleast :: "[i,i]=>o" where \ \not really necessary: ZF defines cardinality\ + Atleast :: "[i,i]\o" where \ \not really necessary: ZF defines cardinality\ "Atleast(n,S) \ (\f. f \ inj(n,S))" definition - Clique :: "[i,i,i]=>o" where - "Clique(C,V,E) \ (C \ V) \ (\x \ C. \y \ C. x\y \ \ E)" + Clique :: "[i,i,i]\o" where + "Clique(C,V,E) \ (C \ V) \ (\x \ C. \y \ C. x\y \ \x,y\ \ E)" definition - Indept :: "[i,i,i]=>o" where - "Indept(I,V,E) \ (I \ V) \ (\x \ I. \y \ I. x\y \ \ E)" + Indept :: "[i,i,i]\o" where + "Indept(I,V,E) \ (I \ V) \ (\x \ I. \y \ I. x\y \ \x,y\ \ E)" definition - Ramsey :: "[i,i,i]=>o" where + Ramsey :: "[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))" @@ -145,7 +145,7 @@ (*For the Atleast part, proves \(a \ I) from the second premise!*) lemma Indept_succ: - "\Indept(I, {z \ V-{a}. \ E}, E); Symmetric(E); a \ V; + "\Indept(I, {z \ V-{a}. \a,z\ \ E}, E); Symmetric(E); a \ V; Atleast(j,I)\ \ Indept(cons(a,I), V, E) \ Atleast(succ(j), cons(a,I))" apply (unfold Symmetric_def Indept_def) @@ -154,7 +154,7 @@ lemma Clique_succ: - "\Clique(C, {z \ V-{a}. :E}, E); Symmetric(E); a \ V; + "\Clique(C, {z \ V-{a}. \a,z\:E}, E); Symmetric(E); a \ V; Atleast(j,C)\ \ Clique(cons(a,C), V, E) \ Atleast(succ(j), cons(a,C))" apply (unfold Symmetric_def Clique_def) @@ -169,7 +169,7 @@ m \ nat; n \ nat\ \ Ramsey(succ(m#+n), succ(i), succ(j))" apply (unfold Ramsey_def, clarify) apply (erule Atleast_succD [THEN bexE]) -apply (erule_tac P1 = "%z.:E" in Atleast_partition [THEN disjE], +apply (erule_tac P1 = "\z.\x,z\:E" in Atleast_partition [THEN disjE], assumption+) (*case m*) apply (fast dest!: Indept_succ elim: Clique_superset) diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/Ring.thy --- a/src/ZF/ex/Ring.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/Ring.thy Tue Sep 27 17:46:52 2022 +0100 @@ -12,27 +12,27 @@ (*First, we must simulate a record declaration: record ring = monoid + - add :: "[i, i] => i" (infixl "\\" 65) + add :: "[i, i] \ i" (infixl "\\" 65) zero :: i ("\\") *) definition - add_field :: "i => i" where + add_field :: "i \ i" where "add_field(M) = fst(snd(snd(snd(M))))" definition - ring_add :: "[i, i, i] => i" (infixl \\\\ 65) where - "ring_add(M,x,y) = add_field(M) ` " + ring_add :: "[i, i, i] \ i" (infixl \\\\ 65) where + "ring_add(M,x,y) = add_field(M) ` \x,y\" definition - zero :: "i => i" (\\\\) where + zero :: "i \ i" (\\\\) where "zero(M) = fst(snd(snd(snd(snd(M)))))" lemma add_field_eq [simp]: "add_field() = A" by (simp add: add_field_def) -lemma add_eq [simp]: "ring_add(, x, y) = A ` " +lemma add_eq [simp]: "ring_add(, x, y) = A ` \x,y\" by (simp add: ring_add_def) lemma zero_eq [simp]: "zero() = Z" @@ -42,11 +42,11 @@ text \Derived operations.\ definition - a_inv :: "[i,i] => i" (\\\ _\ [81] 80) where + a_inv :: "[i,i] \ i" (\\\ _\ [81] 80) where "a_inv(R) \ m_inv ()" definition - minus :: "[i,i,i] => i" (\(_ \\ _)\ [65,66] 65) where + minus :: "[i,i,i] \ i" (\(_ \\ _)\ [65,66] 65) where "\x \ carrier(R); y \ carrier(R)\ \ x \\<^bsub>R\<^esub> y = x \\<^bsub>R\<^esub> (\\<^bsub>R\<^esub> y)" locale abelian_monoid = fixes G (structure) @@ -231,7 +231,7 @@ subsection \Morphisms\ definition - ring_hom :: "[i,i] => i" where + ring_hom :: "[i,i] \ i" where "ring_hom(R,S) \ {h \ carrier(R) -> carrier(S). (\x y. x \ carrier(R) \ y \ carrier(R) \ diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/ex/misc.thy --- a/src/ZF/ex/misc.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/ex/misc.thy Tue Sep 27 17:46:52 2022 +0100 @@ -31,7 +31,7 @@ text\A weird property of ordered pairs.\ -lemma "b\c \ \ = " +lemma "b\c \ \a,b\ \ \a,c\ = \a,a\" by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast) text\These two are cited in Benzmueller and Kohlhase's system description of @@ -81,7 +81,7 @@ 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 \ - (\x \ A. \y \ A. H`(f`) = g`)}) \ + (\x \ A. \y \ A. H`(f`\x,y\) = g`)}) \ J \ hom(A,f,B,g) \ K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" by force @@ -89,7 +89,7 @@ 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 \ - (\x \ A. \y \ A. H`(f`) = g`)}) + (\x \ A. \y \ A. H`(f`\x,y\) = g`)}) \ J \ hom(A,f,B,g) \ K \ hom(B,g,C,h) \ (K O J) \ hom(A,f,C,h)" by force diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/func.thy --- a/src/ZF/func.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/func.thy Tue Sep 27 17:46:52 2022 +0100 @@ -25,7 +25,7 @@ (*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. \x,y\: f)" by (unfold Pi_def function_def, blast) lemma fun_is_function: "f \ Pi(A,B) \ function(f)" @@ -36,7 +36,7 @@ by (simp add: Pi_iff relation_def, blast) lemma functionI: - "\\x y y'. \:r; :r\ \ y=y'\ \ function(r)" + "\\x y y'. \\x,y\:r; :r\ \ y=y'\ \ function(r)" by (simp add: function_def, blast) (*Functions are relations*) @@ -57,13 +57,13 @@ subsection\Function Application\ -lemma apply_equality2: "\: f; : f; f \ Pi(A,B)\ \ b=c" +lemma apply_equality2: "\\a,b\: f; \a,c\: f; f \ Pi(A,B)\ \ b=c" by (unfold Pi_def function_def, blast) -lemma function_apply_equality: "\: f; function(f)\ \ f`a = b" +lemma function_apply_equality: "\\a,b\: f; function(f)\ \ f`a = b" by (unfold apply_def function_def, blast) -lemma apply_equality: "\: f; f \ Pi(A,B)\ \ f`a = b" +lemma apply_equality: "\\a,b\: f; f \ Pi(A,B)\ \ f`a = b" apply (unfold Pi_def) apply (blast intro: function_apply_equality) done @@ -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) \ \a,b\: f \ a \ A \ f`a = b" apply (frule fun_is_rel) apply (blast intro!: apply_Pair apply_equality) done @@ -109,7 +109,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, \x. {y \ B(x). P(x,y)})) \ f \ Pi(A,B) \ (\x\A. P(x, f`x))" by (blast intro: Pi_type dest: apply_type) @@ -120,13 +120,13 @@ (** Elimination of membership in a function **) -lemma domain_type: "\ \ f; f \ Pi(A,B)\ \ a \ A" +lemma domain_type: "\\a,b\ \ f; f \ Pi(A,B)\ \ a \ A" by (blast dest: fun_is_rel) -lemma range_type: "\ \ f; f \ Pi(A,B)\ \ b \ B(a)" +lemma range_type: "\\a,b\ \ 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: "\\a,b\: f; f \ Pi(A,B)\ \ a \ A \ b \ B(a) \ f`a = b" by (blast intro: domain_type range_type apply_equality) subsection\Lambda Abstraction\ @@ -141,7 +141,7 @@ \ \ P" by (simp add: lam_def, blast) -lemma lamD: "\: (\x\A. b(x))\ \ c = b(a)" +lemma lamD: "\\a,c\: (\x\A. b(x))\ \ c = b(a)" by (simp add: lam_def) lemma lam_type [TC]: @@ -190,7 +190,7 @@ by (unfold Pi_def function_def, blast) (*The singleton function*) -lemma singleton_fun [simp]: "{} \ {a} -> {b}" +lemma singleton_fun [simp]: "{\a,b\} \ {a} -> {b}" by (unfold Pi_def function_def, blast) lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)" @@ -395,21 +395,21 @@ subsection\Extensions of Functions\ lemma fun_extend: - "\f \ A->B; c\A\ \ cons(,f) \ cons(c,A) -> cons(b,B)" + "\f \ A->B; c\A\ \ cons(\c,b\,f) \ cons(c,A) -> cons(b,B)" apply (frule singleton_fun [THEN fun_disjoint_Un], blast) apply (simp add: cons_eq) done lemma fun_extend3: - "\f \ A->B; c\A; b \ B\ \ cons(,f) \ cons(c,A) -> B" + "\f \ A->B; c\A; b \ B\ \ cons(\c,b\,f) \ cons(c,A) -> B" by (blast intro: fun_extend [THEN fun_weaken_type]) lemma extend_apply: - "c \ domain(f) \ cons(,f)`a = (if a=c then b else f`a)" + "c \ domain(f) \ cons(\c,b\,f)`a = (if a=c then b else f`a)" by (auto simp add: apply_def) lemma fun_extend_apply [simp]: - "\f \ A->B; c\A\ \ cons(,f)`a = (if a=c then b else f`a)" + "\f \ A->B; c\A\ \ cons(\c,b\,f)`a = (if a=c then b else f`a)" apply (rule extend_apply) apply (simp add: Pi_def, blast) done @@ -418,7 +418,7 @@ (*For Finite.ML. Inclusion of right into left is easy*) lemma cons_fun_eq: - "c \ A \ cons(c,A) -> B = (\f \ A->B. \b\B. {cons(, f)})" + "c \ A \ cons(c,A) -> B = (\f \ A->B. \b\B. {cons(\c,b\, f)})" apply (rule equalityI) apply (safe elim!: fun_extend3) (*Inclusion of left into right*) @@ -435,14 +435,14 @@ apply (erule consE, simp_all) done -lemma succ_fun_eq: "succ(n) -> B = (\f \ n->B. \b\B. {cons(, f)})" +lemma succ_fun_eq: "succ(n) -> B = (\f \ n->B. \b\B. {cons(\n,b\, f)})" by (simp add: succ_def mem_not_refl cons_fun_eq) subsection\Function Updates\ definition - update :: "[i,i,i] => i" where + update :: "[i,i,i] \ i" where "update(f,a,b) \ \x\cons(a, domain(f)). if(x=a, b, f`x)" nonterminal updbinds and updbind @@ -451,10 +451,10 @@ (* Let expressions *) - "_updbind" :: "[i, i] => updbind" (\(2_ :=/ _)\) - "" :: "updbind => updbinds" (\_\) - "_updbinds" :: "[updbind, updbinds] => updbinds" (\_,/ _\) - "_Update" :: "[i, updbinds] => i" (\_/'((_)')\ [900,0] 900) + "_updbind" :: "[i, i] \ updbind" (\(2_ :=/ _)\) + "" :: "updbind \ updbinds" (\_\) + "_updbinds" :: "[updbind, updbinds] \ updbinds" (\_,/ _\) + "_Update" :: "[i, updbinds] \ i" (\_/'((_)')\ [900,0] 900) translations "_Update (f, _updbinds(b,bs))" == "_Update (_Update(f,b), bs)" @@ -568,11 +568,11 @@ subsubsection\Images\ lemma image_pair_mono: - "\\x y. :r \ :s; A<=B\ \ r``A \ s``B" + "\\x y. \x,y\:r \ \x,y\:s; A<=B\ \ r``A \ s``B" by blast lemma vimage_pair_mono: - "\\x y. :r \ :s; A<=B\ \ r-``A \ s-``B" + "\\x y. \x,y\:r \ \x,y\:s; A<=B\ \ r-``A \ s-``B" by blast lemma image_mono: "\r<=s; A<=B\ \ r``A \ s``B" diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/pair.thy --- a/src/ZF/pair.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/pair.thy Tue Sep 27 17:46:52 2022 +0100 @@ -29,7 +29,7 @@ \ -(** Lemmas for showing that uniquely determines a and b **) +(** Lemmas for showing that \a,b\ uniquely determines a and b **) lemma singleton_eq_iff [iff]: "{a} = {b} \ a=b" by (rule extension [THEN iff_trans], blast) @@ -37,7 +37,7 @@ 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,b\ = \c,d\ \ a=c \ b=d" by (simp add: Pair_def doubleton_eq_iff, blast) lemmas Pair_inject = Pair_iff [THEN iffD1, THEN conjE, elim!] @@ -45,7 +45,7 @@ lemmas Pair_inject1 = Pair_iff [THEN iffD1, THEN conjunct1] lemmas Pair_inject2 = Pair_iff [THEN iffD1, THEN conjunct2] -lemma Pair_not_0: " \ 0" +lemma Pair_not_0: "\a,b\ \ 0" apply (unfold Pair_def) apply (blast elim: equalityE) done @@ -54,7 +54,7 @@ declare sym [THEN Pair_neq_0, elim!] -lemma Pair_neq_fst: "=a \ P" +lemma Pair_neq_fst: "\a,b\=a \ P" proof (unfold Pair_def) assume eq: "{{a, a}, {a, b}} = a" have "{a, a} \ {{a, a}, {a, b}}" by (rule consI1) @@ -63,7 +63,7 @@ ultimately show "P" by (rule mem_asym) qed -lemma Pair_neq_snd: "=b \ P" +lemma Pair_neq_snd: "\a,b\=b \ P" proof (unfold Pair_def) assume eq: "{{a, a}, {a, b}} = b" have "{a, b} \ {{a, a}, {a, b}}" by blast @@ -77,10 +77,10 @@ text\Generalizes Cartesian product\ -lemma Sigma_iff [simp]: ": Sigma(A,B) \ a \ A \ b \ B(a)" +lemma Sigma_iff [simp]: "\a,b\: Sigma(A,B) \ a \ A \ b \ B(a)" by (simp add: Sigma_def) -lemma SigmaI [TC,intro!]: "\a \ A; b \ B(a)\ \ \ Sigma(A,B)" +lemma SigmaI [TC,intro!]: "\a \ A; b \ B(a)\ \ \a,b\ \ Sigma(A,B)" by simp lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1] @@ -89,12 +89,12 @@ (*The general elimination rule*) lemma SigmaE [elim!]: "\c \ Sigma(A,B); - \x y.\x \ A; y \ B(x); c=\ \ P + \x y.\x \ A; y \ B(x); c=\x,y\\ \ P \ \ P" by (unfold Sigma_def, blast) lemma SigmaE2 [elim!]: - "\ \ Sigma(A,B); + "\\a,b\ \ Sigma(A,B); \a \ A; b \ B(a)\ \ P \ \ P" by (unfold Sigma_def, blast) @@ -120,10 +120,10 @@ subsection\Projections \<^term>\fst\ and \<^term>\snd\\ -lemma fst_conv [simp]: "fst() = a" +lemma fst_conv [simp]: "fst(\a,b\) = a" by (simp add: fst_def) -lemma snd_conv [simp]: "snd() = b" +lemma snd_conv [simp]: "snd(\a,b\) = b" by (simp add: snd_def) lemma fst_type [TC]: "p \ Sigma(A,B) \ fst(p) \ A" @@ -139,33 +139,33 @@ subsection\The Eliminator, \<^term>\split\\ (*A META-equality, so that it applies to higher types as well...*) -lemma split [simp]: "split(%x y. c(x,y), ) \ c(a,b)" +lemma split [simp]: "split(\x y. c(x,y), \a,b\) \ c(a,b)" by (simp add: split_def) lemma split_type [TC]: "\p \ Sigma(A,B); - \x y.\x \ A; y \ B(x)\ \ c(x,y):C() -\ \ split(%x y. c(x,y), p) \ C(p)" + \x y.\x \ A; y \ B(x)\ \ c(x,y):C(\x,y\) +\ \ split(\x y. c(x,y), p) \ C(p)" by (erule SigmaE, auto) lemma expand_split: "u \ A*B \ - R(split(c,u)) \ (\x\A. \y\B. u = \ R(c(x,y)))" + R(split(c,u)) \ (\x\A. \y\B. u = \x,y\ \ R(c(x,y)))" by (auto simp add: split_def) subsection\A version of \<^term>\split\ for Formulae: Result Type \<^typ>\o\\ -lemma splitI: "R(a,b) \ split(R, )" +lemma splitI: "R(a,b) \ split(R, \a,b\)" by (simp add: split_def) lemma splitE: "\split(R,z); z \ Sigma(A,B); - \x y. \z = ; R(x,y)\ \ P + \x y. \z = \x,y\; R(x,y)\ \ P \ \ P" by (auto simp add: split_def) -lemma splitD: "split(R,) \ R(a,b)" +lemma splitD: "split(R,\a,b\) \ R(a,b)" by (simp add: split_def) text \ @@ -173,11 +173,11 @@ \ lemma split_paired_Bex_Sigma [simp]: - "(\z \ Sigma(A,B). P(z)) \ (\x \ A. \y \ B(x). P())" + "(\z \ Sigma(A,B). P(z)) \ (\x \ A. \y \ B(x). P(\x,y\))" by blast lemma split_paired_Ball_Sigma [simp]: - "(\z \ Sigma(A,B). P(z)) \ (\x \ A. \y \ B(x). P())" + "(\z \ Sigma(A,B). P(z)) \ (\x \ A. \y \ B(x). P(\x,y\))" by blast end diff -r 0c18df79b1c8 -r a642599ffdea src/ZF/upair.thy --- a/src/ZF/upair.thy Tue Sep 27 17:03:23 2022 +0100 +++ b/src/ZF/upair.thy Tue Sep 27 17:46:52 2022 +0100 @@ -229,11 +229,11 @@ (** Rewrite rules for boolean case-splitting: faster than split_if [split] **) -lemmas split_if_eq1 = split_if [of "%x. x = b"] for b -lemmas split_if_eq2 = split_if [of "%x. a = x"] for a +lemmas split_if_eq1 = split_if [of "\x. x = b"] for b +lemmas split_if_eq2 = split_if [of "\x. a = x"] for a -lemmas split_if_mem1 = split_if [of "%x. x \ b"] for b -lemmas split_if_mem2 = split_if [of "%x. a \ x"] for a +lemmas split_if_mem1 = split_if [of "\x. x \ b"] for b +lemmas split_if_mem2 = split_if [of "\x. a \ x"] for a lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2