--- a/src/ZF/List.thy Thu Jul 18 15:21:42 2002 +0200
+++ b/src/ZF/List.thy Fri Jul 19 13:28:19 2002 +0200
@@ -28,14 +28,15 @@
consts
length :: "i=>i"
- hd :: "i=>i"
- tl :: "i=>i"
+ hd :: "i=>i"
+ tl :: "i=>i"
primrec
"length([]) = 0"
"length(Cons(a,l)) = succ(length(l))"
primrec
+ "hd([]) = 0"
"hd(Cons(a,l)) = a"
primrec
--- a/src/ZF/Main.thy Thu Jul 18 15:21:42 2002 +0200
+++ b/src/ZF/Main.thy Fri Jul 19 13:28:19 2002 +0200
@@ -40,6 +40,9 @@
==> Ord(F^n (x))"
by (induct n rule: nat_induct, simp_all)
+lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))"
+by (induct_tac n, simp_all)
+
(* belongs to theory IntDiv *)
lemmas posDivAlg_induct = posDivAlg_induct [consumes 2]
--- a/src/ZF/Ordinal.thy Thu Jul 18 15:21:42 2002 +0200
+++ b/src/ZF/Ordinal.thy Fri Jul 19 13:28:19 2002 +0200
@@ -300,11 +300,16 @@
apply (rule foundation [THEN disjE, THEN allI], erule disjI1, blast)
done
-(*Transset(i) does not suffice, though ALL j:i.Transset(j) does*)
+text{*The premise @{term "Ord(i)"} does not suffice.*}
lemma trans_Memrel:
"Ord(i) ==> trans(Memrel(i))"
by (unfold Ord_def Transset_def trans_def, blast)
+text{*However, the following premise is strong enough.*}
+lemma Transset_trans_Memrel:
+ "\<forall>j\<in>i. Transset(j) ==> trans(Memrel(i))"
+by (unfold Transset_def trans_def, blast)
+
(*If Transset(A) then Memrel(A) internalizes the membership relation below A*)
lemma Transset_Memrel_iff:
"Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A"