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