converted to Isar script
authorpaulson
Sat, 21 Sep 2002 21:10:34 +0200
changeset 13574 f9796358e66f
parent 13573 37b22343c35a
child 13575 ecb6ecd9af13
converted to Isar script
src/ZF/Integ/EquivClass.ML
src/ZF/Integ/EquivClass.thy
src/ZF/IsaMakefile
--- a/src/ZF/Integ/EquivClass.ML	Fri Sep 20 11:49:38 2002 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,246 +0,0 @@
-(*  Title:      ZF/EquivClass.ML
-    ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1994  University of Cambridge
-
-Equivalence relations in Zermelo-Fraenkel Set Theory 
-*)
-
-(*** Suppes, Theorem 70: r is an equiv relation iff converse(r) O r = r ***)
-
-(** first half: equiv(A,r) ==> converse(r) O r = r **)
-
-Goalw [trans_def,sym_def]
-    "[| sym(r); trans(r) |] ==> converse(r) O r <= r";
-by (Blast_tac 1);
-qed "sym_trans_comp_subset";
-
-Goalw [refl_def]
-    "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r";
-by (Blast_tac 1);
-qed "refl_comp_subset";
-
-Goalw [equiv_def]
-    "equiv(A,r) ==> converse(r) O r = r";
-by (blast_tac (subset_cs addSIs [sym_trans_comp_subset, refl_comp_subset]) 1);
-qed "equiv_comp_eq";
-
-(*second half*)
-Goalw [equiv_def,refl_def,sym_def,trans_def]
-    "[| converse(r) O r = r;  domain(r) = A |] ==> equiv(A,r)";
-by (etac equalityE 1);
-by (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r" 1);
-by (ALLGOALS Fast_tac);
-qed "comp_equivI";
-
-(** Equivalence classes **)
-
-(*Lemma for the next result*)
-Goalw [trans_def,sym_def]
-    "[| sym(r);  trans(r);  <a,b>: r |] ==> r``{a} <= r``{b}";
-by (Blast_tac 1);
-qed "equiv_class_subset";
-
-Goalw [equiv_def]
-    "[| equiv(A,r);  <a,b>: r |] ==> r``{a} = r``{b}";
-by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset]));
-by (rewtac sym_def);
-by (Blast_tac 1);
-qed "equiv_class_eq";
-
-Goalw [equiv_def,refl_def]
-    "[| equiv(A,r);  a: A |] ==> a: r``{a}";
-by (Blast_tac 1);
-qed "equiv_class_self";
-
-(*Lemma for the next result*)
-Goalw [equiv_def,refl_def]
-    "[| equiv(A,r);  r``{b} <= r``{a};  b: A |] ==> <a,b>: r";
-by (Blast_tac 1);
-qed "subset_equiv_class";
-
-Goal "[| r``{a} = r``{b};  equiv(A,r);  b: A |] ==> <a,b>: r";
-by (REPEAT (ares_tac[equalityD2, subset_equiv_class] 1));
-qed "eq_equiv_class";
-
-(*thus r``{a} = r``{b} as well*)
-Goalw [equiv_def,trans_def,sym_def]
-    "[| equiv(A,r);  x: (r``{a} Int r``{b}) |] ==> <a,b>: r";
-by (Blast_tac 1);
-qed "equiv_class_nondisjoint";
-
-Goalw [equiv_def] "equiv(A,r) ==> r <= A*A";
-by (safe_tac subset_cs);
-qed "equiv_type";
-
-Goal "equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A";
-by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq]
-                      addDs [equiv_type]) 1);
-qed "equiv_class_eq_iff";
-
-Goal "[| equiv(A,r);  x: A;  y: A |] ==> r``{x} = r``{y} <-> <x,y>: r";
-by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq]
-                      addDs [equiv_type]) 1);
-qed "eq_equiv_class_iff";
-
-(*** Quotients ***)
-
-(** Introduction/elimination rules -- needed? **)
-
-Goalw [quotient_def] "x:A ==> r``{x}: A//r";
-by (etac RepFunI 1);
-qed "quotientI";
-AddTCs [quotientI];
-
-val major::prems = Goalw [quotient_def]
-    "[| X: A//r;  !!x. [| X = r``{x};  x:A |] ==> P |]   \
-\    ==> P";
-by (rtac (major RS RepFunE) 1);
-by (eresolve_tac prems 1);
-by (assume_tac 1);
-qed "quotientE";
-
-Goalw [equiv_def,refl_def,quotient_def]
-    "equiv(A,r) ==> Union(A//r) = A";
-by (Blast_tac 1);
-qed "Union_quotient";
-
-Goalw [quotient_def]
-    "[| equiv(A,r);  X: A//r;  Y: A//r |] ==> X=Y | (X Int Y <= 0)";
-by (safe_tac (claset() addSIs [equiv_class_eq]));
-by (assume_tac 1);
-by (rewrite_goals_tac [equiv_def,trans_def,sym_def]);
-by (Blast_tac 1);
-qed "quotient_disj";
-
-(**** Defining unary operations upon equivalence classes ****)
-
-(** These proofs really require as local premises
-     equiv(A,r);  congruent(r,b)
-**)
-
-(*Conversion rule*)
-val prems as [equivA,bcong,_] = goal (the_context ())
-    "[| equiv(A,r);  congruent(r,b);  a: A |] ==> (UN x:r``{a}. b(x)) = b(a)";
-by (cut_facts_tac prems 1);
-by (rtac ([refl RS UN_cong, UN_constant] MRS trans) 1);
-by (etac equiv_class_self 2);
-by (assume_tac 2);
-by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]);
-by (Blast_tac 1);
-qed "UN_equiv_class";
-
-(*type checking of  UN x:r``{a}. b(x) *)
-val prems = Goalw [quotient_def]
-    "[| equiv(A,r);  congruent(r,b);  X: A//r;   \
-\       !!x.  x : A ==> b(x) : B |]     \
-\    ==> (UN x:X. b(x)) : B";
-by (cut_facts_tac prems 1);
-by Safe_tac;
-by (asm_simp_tac (simpset() addsimps UN_equiv_class::prems) 1);
-qed "UN_equiv_class_type";
-
-(*Sufficient conditions for injectiveness.  Could weaken premises!
-  major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B
-*)
-val prems = Goalw [quotient_def]
-    "[| equiv(A,r);   congruent(r,b);  \
-\       (UN x:X. b(x))=(UN y:Y. b(y));  X: A//r;  Y: A//r;  \
-\       !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |]         \
-\    ==> X=Y";
-by (cut_facts_tac prems 1);
-by Safe_tac;
-by (rtac equiv_class_eq 1);
-by (REPEAT (ares_tac prems 1));
-by (etac box_equals 1);
-by (REPEAT (ares_tac [UN_equiv_class] 1));
-qed "UN_equiv_class_inject";
-
-
-(**** Defining binary operations upon equivalence classes ****)
-
-
-Goalw [congruent_def,congruent2_def,equiv_def,refl_def]
-    "[| equiv(A,r);  congruent2(r,b);  a: A |] ==> congruent(r,b(a))";
-by (Blast_tac 1);
-qed "congruent2_implies_congruent";
-
-val equivA::prems = goalw (the_context ()) [congruent_def]
-    "[| equiv(A,r);  congruent2(r,b);  a: A |] ==> \
-\    congruent(r, %x1. UN x2:r``{a}. b(x1,x2))";
-by (cut_facts_tac (equivA::prems) 1);
-by Safe_tac;
-by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1);
-by (assume_tac 1);
-by (asm_simp_tac (simpset() addsimps [equivA RS UN_equiv_class,
-                                     congruent2_implies_congruent]) 1);
-by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
-by (Blast_tac 1);
-qed "congruent2_implies_congruent_UN";
-
-val prems as equivA::_ = goal (the_context ())
-    "[| equiv(A,r);  congruent2(r,b);  a1: A;  a2: A |]  \
-\    ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)";
-by (cut_facts_tac prems 1);
-by (asm_simp_tac (simpset() addsimps [equivA RS UN_equiv_class,
-                                     congruent2_implies_congruent,
-                                     congruent2_implies_congruent_UN]) 1);
-qed "UN_equiv_class2";
-
-(*type checking*)
-val prems = Goalw [quotient_def]
-    "[| equiv(A,r);  congruent2(r,b);                   \
-\       X1: A//r;  X2: A//r;                            \
-\       !!x1 x2.  [| x1: A; x2: A |] ==> b(x1,x2) : B   \
-\    |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B";
-by (cut_facts_tac prems 1);
-by Safe_tac;
-by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
-                             congruent2_implies_congruent_UN,
-                             congruent2_implies_congruent, quotientI]) 1));
-qed "UN_equiv_class_type2";
-
-
-(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
-  than the direct proof*)
-val prems = Goalw [congruent2_def,equiv_def,refl_def]
-    "[| equiv(A,r);     \
-\       !! y z w. [| w: A;  <y,z> : r |] ==> b(y,w) = b(z,w);      \
-\       !! y z w. [| w: A;  <y,z> : r |] ==> b(w,y) = b(w,z)       \
-\    |] ==> congruent2(r,b)";
-by (cut_facts_tac prems 1);
-by Safe_tac;
-by (rtac trans 1);
-by (REPEAT (ares_tac prems 1
-     ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
-qed "congruent2I";
-
-val [equivA,commute,congt] = Goal
-    "[| equiv(A,r);     \
-\       !! y z. [| y: A;  z: A |] ==> b(y,z) = b(z,y);        \
-\       !! y z w. [| w: A;  <y,z>: r |] ==> b(w,y) = b(w,z)     \
-\    |] ==> congruent2(r,b)";
-by (resolve_tac [equivA RS congruent2I] 1);
-by (rtac (commute RS trans) 1);
-by (rtac (commute RS trans RS sym) 3);
-by (rtac sym 5);
-by (REPEAT (ares_tac [congt] 1
-     ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1));
-qed "congruent2_commuteI";
-
-(*Obsolete?*)
-val [equivA,ZinA,congt,commute] = Goalw [quotient_def]
-    "[| equiv(A,r);  Z: A//r;  \
-\       !!w. [| w: A |] ==> congruent(r, %z. b(w,z));    \
-\       !!x y. [| x: A;  y: A |] ==> b(y,x) = b(x,y)    \
-\    |] ==> congruent(r, %w. UN z: Z. b(w,z))";
-val congt' = rewrite_rule [congruent_def] congt;
-by (cut_facts_tac [ZinA] 1);
-by (rewtac congruent_def);
-by Safe_tac;
-by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1);
-by (assume_tac 1);
-by (asm_simp_tac (simpset() addsimps [commute,
-                                     [equivA, congt] MRS UN_equiv_class]) 1);
-by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1));
-qed "congruent_commuteI";
--- a/src/ZF/Integ/EquivClass.thy	Fri Sep 20 11:49:38 2002 +0200
+++ b/src/ZF/Integ/EquivClass.thy	Sat Sep 21 21:10:34 2002 +0200
@@ -3,21 +3,260 @@
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1994  University of Cambridge
 
-Equivalence relations in Zermelo-Fraenkel Set Theory 
 *)
 
-EquivClass = Trancl + Perm + 
+header{*Equivalence Relations*}
+
+theory EquivClass = Trancl + Perm:
 
 constdefs
 
-  quotient    :: [i,i]=>i    (infixl "'/'/" 90)  (*set of equiv classes*)
+  quotient   :: "[i,i]=>i"    (infixl "'/'/" 90)  (*set of equiv classes*)
       "A//r == {r``{x} . x:A}"
 
-  congruent   :: [i,i=>i]=>o
+  congruent  :: "[i,i=>i]=>o"
       "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)"
 
-  congruent2  :: [i,[i,i]=>i]=>o
-      "congruent2(r,b) == ALL y1 z1 y2 z2. 
+  congruent2 :: "[i,[i,i]=>i]=>o"
+      "congruent2(r,b) == ALL y1 z1 y2 z2.
            <y1,z1>:r --> <y2,z2>:r --> b(y1,y2) = b(z1,z2)"
 
+subsection{*Suppes, Theorem 70:
+    @{term r} is an equiv relation iff @{term "converse(r) O r = r"}*}
+
+(** first half: equiv(A,r) ==> converse(r) O r = r **)
+
+lemma sym_trans_comp_subset:
+    "[| sym(r); trans(r) |] ==> converse(r) O r <= r"
+apply (unfold trans_def sym_def, blast)
+done
+
+lemma refl_comp_subset:
+    "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r"
+apply (unfold refl_def, blast)
+done
+
+lemma equiv_comp_eq:
+    "equiv(A,r) ==> converse(r) O r = r"
+apply (unfold equiv_def)
+apply (blast del: subsetI
+             intro!: sym_trans_comp_subset refl_comp_subset)
+done
+
+(*second half*)
+lemma comp_equivI:
+    "[| converse(r) O r = r;  domain(r) = A |] ==> equiv(A,r)"
+apply (unfold equiv_def refl_def sym_def trans_def)
+apply (erule equalityE)
+apply (subgoal_tac "ALL 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);  <a,b>: r |] ==> r``{a} <= r``{b}"
+by (unfold trans_def sym_def, blast)
+
+lemma equiv_class_eq:
+    "[| 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)
+done
+
+lemma equiv_class_self:
+    "[| equiv(A,r);  a: A |] ==> a: r``{a}"
+by (unfold equiv_def refl_def, blast)
+
+(*Lemma for the next result*)
+lemma subset_equiv_class:
+    "[| equiv(A,r);  r``{b} <= r``{a};  b: A |] ==> <a,b>: r"
+by (unfold equiv_def refl_def, blast)
+
+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} Int 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) ==> <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} <-> <x,y>: r"
+by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
+
+(*** Quotients ***)
+
+(** Introduction/elimination rules -- needed? **)
+
+lemma quotientI [TC]: "x:A ==> r``{x}: A//r"
+apply (unfold quotient_def)
+apply (erule RepFunI)
+done
+
+lemma quotientE:
+    "[| X: A//r;  !!x. [| X = r``{x};  x:A |] ==> P |] ==> P"
+by (unfold quotient_def, blast)
+
+lemma Union_quotient:
+    "equiv(A,r) ==> Union(A//r) = A"
+by (unfold equiv_def refl_def quotient_def, blast)
+
+lemma quotient_disj:
+    "[| equiv(A,r);  X: A//r;  Y: A//r |] ==> X=Y | (X Int Y <= 0)"
+apply (unfold quotient_def)
+apply (safe intro!: equiv_class_eq, assumption)
+apply (unfold equiv_def trans_def sym_def, blast)
+done
+
+subsection{*Defining Unary Operations upon Equivalence Classes*}
+
+(** Could have a locale with the premises equiv(A,r)  and  congruent(r,b)
+**)
+
+(*Conversion rule*)
+lemma UN_equiv_class:
+    "[| equiv(A,r);  congruent(r,b);  a: A |] ==> (UN x:r``{a}. b(x)) = b(a)"
+apply (rule trans [OF refl [THEN UN_cong] UN_constant])
+apply (erule_tac [2] equiv_class_self)
+prefer 2 apply assumption
+apply (unfold equiv_def sym_def congruent_def, blast)
+done
+
+(*type checking of  UN x:r``{a}. b(x) *)
+lemma UN_equiv_class_type:
+    "[| equiv(A,r);  congruent(r,b);  X: A//r;  !!x.  x : A ==> b(x) : B |]
+     ==> (UN x:X. b(x)) : B"
+apply (unfold quotient_def, safe)
+apply (simp (no_asm_simp) add: UN_equiv_class)
+done
+
+(*Sufficient conditions for injectiveness.  Could weaken premises!
+  major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B
+*)
+lemma UN_equiv_class_inject:
+    "[| equiv(A,r);   congruent(r,b);
+        (UN x:X. b(x))=(UN y:Y. b(y));  X: A//r;  Y: A//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)
+apply (rotate_tac -2) 
+apply (simp add: UN_equiv_class [of A r b])  
+done
+
+
+subsection{*Defining Binary Operations upon Equivalence Classes*}
+
+
+lemma congruent2_implies_congruent:
+    "[| equiv(A,r);  congruent2(r,b);  a: A |] ==> congruent(r,b(a))"
+apply (unfold congruent_def congruent2_def equiv_def refl_def, blast)
+done
+
+lemma congruent2_implies_congruent_UN:
+    "[| equiv(A,r);  congruent2(r,b);  a: A |] ==>
+     congruent(r, %x1. UN x2:r``{a}. b(x1,x2))"
+apply (unfold congruent_def, safe)
+apply (frule equiv_type [THEN subsetD], assumption)
+apply clarify 
+apply (simp add: UN_equiv_class [of A r] congruent2_implies_congruent)
+apply (unfold congruent2_def equiv_def refl_def, blast)
+done
+
+lemma UN_equiv_class2:
+    "[| equiv(A,r);  congruent2(r,b);  a1: A;  a2: A |]
+     ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)"
+by (simp add: UN_equiv_class [of A r] congruent2_implies_congruent
+              congruent2_implies_congruent_UN)
+
+(*type checking*)
+lemma UN_equiv_class_type2:
+    "[| equiv(A,r);  congruent2(r,b);
+        X1: A//r;  X2: A//r;
+        !!x1 x2.  [| x1: A; x2: A |] ==> b(x1,x2) : B
+     |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"
+apply (unfold quotient_def, safe)
+apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN 
+                    congruent2_implies_congruent quotientI)
+done
+
+
+(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
+  than the direct proof*)
+lemma congruent2I:
+    "[| equiv(A,r);
+        !! y z w. [| w: A;  <y,z> : r |] ==> b(y,w) = b(z,w);
+        !! y z w. [| w: A;  <y,z> : r |] ==> b(w,y) = b(w,z)
+     |] ==> congruent2(r,b)"
+apply (unfold congruent2_def equiv_def refl_def, safe)
+apply (blast intro: trans) 
+done
+
+lemma congruent2_commuteI:
+ assumes equivA: "equiv(A,r)"
+     and commute: "!! y z. [| y: A;  z: A |] ==> b(y,z) = b(z,y)"
+     and congt:   "!! y z w. [| w: A;  <y,z>: r |] ==> b(w,y) = b(w,z)"
+ shows "congruent2(r,b)"
+apply (insert equivA [THEN equiv_type, THEN subsetD]) 
+apply (rule congruent2I [OF equivA])
+apply (rule commute [THEN trans])
+apply (rule_tac [3] commute [THEN trans, symmetric])
+apply (rule_tac [5] sym) 
+apply (blast intro: congt)+
+done
+
+(*Obsolete?*)
+lemma congruent_commuteI:
+    "[| equiv(A,r);  Z: A//r;
+        !!w. [| w: A |] ==> congruent(r, %z. b(w,z));
+        !!x y. [| x: A;  y: A |] ==> b(y,x) = b(x,y)
+     |] ==> congruent(r, %w. UN z: Z. b(w,z))"
+apply (simp (no_asm) add: congruent_def)
+apply (safe elim!: quotientE)
+apply (frule equiv_type [THEN subsetD], assumption)
+apply (simp add: UN_equiv_class [of A r]) 
+apply (simp add: congruent_def) 
+done
+
+ML
+{*
+val sym_trans_comp_subset = thm "sym_trans_comp_subset";
+val refl_comp_subset = thm "refl_comp_subset";
+val equiv_comp_eq = thm "equiv_comp_eq";
+val comp_equivI = thm "comp_equivI";
+val equiv_class_subset = thm "equiv_class_subset";
+val equiv_class_eq = thm "equiv_class_eq";
+val equiv_class_self = thm "equiv_class_self";
+val subset_equiv_class = thm "subset_equiv_class";
+val eq_equiv_class = thm "eq_equiv_class";
+val equiv_class_nondisjoint = thm "equiv_class_nondisjoint";
+val equiv_type = thm "equiv_type";
+val equiv_class_eq_iff = thm "equiv_class_eq_iff";
+val eq_equiv_class_iff = thm "eq_equiv_class_iff";
+val quotientI = thm "quotientI";
+val quotientE = thm "quotientE";
+val Union_quotient = thm "Union_quotient";
+val quotient_disj = thm "quotient_disj";
+val UN_equiv_class = thm "UN_equiv_class";
+val UN_equiv_class_type = thm "UN_equiv_class_type";
+val UN_equiv_class_inject = thm "UN_equiv_class_inject";
+val congruent2_implies_congruent = thm "congruent2_implies_congruent";
+val congruent2_implies_congruent_UN = thm "congruent2_implies_congruent_UN";
+val congruent_commuteI = thm "congruent_commuteI";
+val UN_equiv_class2 = thm "UN_equiv_class2";
+val UN_equiv_class_type2 = thm "UN_equiv_class_type2";
+val congruent2I = thm "congruent2I";
+val congruent2_commuteI = thm "congruent2_commuteI";
+val congruent_commuteI = thm "congruent_commuteI";
+*}
+
 end
--- a/src/ZF/IsaMakefile	Fri Sep 20 11:49:38 2002 +0200
+++ b/src/ZF/IsaMakefile	Sat Sep 21 21:10:34 2002 +0200
@@ -33,7 +33,7 @@
   CardinalArith.thy Cardinal_AC.thy \
   Datatype.ML Datatype.thy Epsilon.thy Finite.thy	\
   Fixedpt.thy Inductive.ML Inductive.thy 	\
-  InfDatatype.thy Integ/Bin.thy Integ/EquivClass.ML	\
+  InfDatatype.thy Integ/Bin.thy \
   Integ/EquivClass.thy Integ/Int.thy Integ/IntArith.thy	\
   Integ/IntDiv.thy Integ/int_arith.ML			\
   Let.ML Let.thy List.thy Main.ML Main.thy	\