src/ZF/CardinalArith.ML
changeset 5067 62b6288e6005
parent 4477 b3e5857d8d99
child 5116 8eb343ab5748
--- a/src/ZF/CardinalArith.ML	Mon Jun 22 15:53:24 1998 +0200
+++ b/src/ZF/CardinalArith.ML	Mon Jun 22 17:12:27 1998 +0200
@@ -17,7 +17,7 @@
 
 (** Cardinal addition is commutative **)
 
-goalw CardinalArith.thy [eqpoll_def] "A+B eqpoll B+A";
+Goalw [eqpoll_def] "A+B eqpoll B+A";
 by (rtac exI 1);
 by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")] 
     lam_bijective 1);
@@ -25,19 +25,19 @@
 by (ALLGOALS (Asm_simp_tac));
 qed "sum_commute_eqpoll";
 
-goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i";
+Goalw [cadd_def] "i |+| j = j |+| i";
 by (rtac (sum_commute_eqpoll RS cardinal_cong) 1);
 qed "cadd_commute";
 
 (** Cardinal addition is associative **)
 
-goalw CardinalArith.thy [eqpoll_def] "(A+B)+C eqpoll A+(B+C)";
+Goalw [eqpoll_def] "(A+B)+C eqpoll A+(B+C)";
 by (rtac exI 1);
 by (rtac sum_assoc_bij 1);
 qed "sum_assoc_eqpoll";
 
 (*Unconditional version requires AC*)
-goalw CardinalArith.thy [cadd_def]
+Goalw [cadd_def]
     "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==>  \
 \             (i |+| j) |+| k = i |+| (j |+| k)";
 by (rtac cardinal_cong 1);
@@ -51,25 +51,25 @@
 
 (** 0 is the identity for addition **)
 
-goalw CardinalArith.thy [eqpoll_def] "0+A eqpoll A";
+Goalw [eqpoll_def] "0+A eqpoll A";
 by (rtac exI 1);
 by (rtac bij_0_sum 1);
 qed "sum_0_eqpoll";
 
-goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
+Goalw [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
 by (asm_simp_tac (simpset() addsimps [sum_0_eqpoll RS cardinal_cong, 
 				      Card_cardinal_eq]) 1);
 qed "cadd_0";
 
 (** Addition by another cardinal **)
 
-goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A+B";
+Goalw [lepoll_def, inj_def] "A lepoll A+B";
 by (res_inst_tac [("x", "lam x:A. Inl(x)")] exI 1);
 by (asm_simp_tac (simpset() addsimps [lam_type]) 1);
 qed "sum_lepoll_self";
 
 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
-goalw CardinalArith.thy [cadd_def]
+Goalw [cadd_def]
     "!!K. [| Card(K);  Ord(L) |] ==> K le (K |+| L)";
 by (rtac ([Card_cardinal_le, well_ord_lepoll_imp_Card_le] MRS le_trans) 1);
 by (rtac sum_lepoll_self 3);
@@ -78,7 +78,7 @@
 
 (** Monotonicity of addition **)
 
-goalw CardinalArith.thy [lepoll_def]
+Goalw [lepoll_def]
      "!!A B C D. [| A lepoll C;  B lepoll D |] ==> A + B  lepoll  C + D";
 by (REPEAT (etac exE 1));
 by (res_inst_tac [("x", "lam z:A+B. case(%w. Inl(f`w), %y. Inr(fa`y), z)")] 
@@ -91,7 +91,7 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [left_inverse])));
 qed "sum_lepoll_mono";
 
-goalw CardinalArith.thy [cadd_def]
+Goalw [cadd_def]
     "!!K. [| K' le K;  L' le L |] ==> (K' |+| L') le (K |+| L)";
 by (safe_tac (claset() addSDs [le_subset_iff RS iffD1]));
 by (rtac well_ord_lepoll_imp_Card_le 1);
@@ -101,7 +101,7 @@
 
 (** Addition of finite cardinals is "ordinary" addition **)
 
-goalw CardinalArith.thy [eqpoll_def] "succ(A)+B eqpoll succ(A+B)";
+Goalw [eqpoll_def] "succ(A)+B eqpoll succ(A+B)";
 by (rtac exI 1);
 by (res_inst_tac [("c", "%z. if(z=Inl(A),A+B,z)"), 
                   ("d", "%z. if(z=A+B,Inl(A),z)")] 
@@ -113,7 +113,7 @@
 
 (*Pulling the  succ(...)  outside the |...| requires m, n: nat  *)
 (*Unconditional version requires AC*)
-goalw CardinalArith.thy [cadd_def]
+Goalw [cadd_def]
     "!!m n. [| Ord(m);  Ord(n) |] ==> succ(m) |+| n = |succ(m |+| n)|";
 by (rtac (sum_succ_eqpoll RS cardinal_cong RS trans) 1);
 by (rtac (succ_eqpoll_cong RS cardinal_cong) 1);
@@ -136,7 +136,7 @@
 (** Cardinal multiplication is commutative **)
 
 (*Easier to prove the two directions separately*)
-goalw CardinalArith.thy [eqpoll_def] "A*B eqpoll B*A";
+Goalw [eqpoll_def] "A*B eqpoll B*A";
 by (rtac exI 1);
 by (res_inst_tac [("c", "%<x,y>.<y,x>"), ("d", "%<x,y>.<y,x>")] 
     lam_bijective 1);
@@ -144,19 +144,19 @@
 by (ALLGOALS (Asm_simp_tac));
 qed "prod_commute_eqpoll";
 
-goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i";
+Goalw [cmult_def] "i |*| j = j |*| i";
 by (rtac (prod_commute_eqpoll RS cardinal_cong) 1);
 qed "cmult_commute";
 
 (** Cardinal multiplication is associative **)
 
-goalw CardinalArith.thy [eqpoll_def] "(A*B)*C eqpoll A*(B*C)";
+Goalw [eqpoll_def] "(A*B)*C eqpoll A*(B*C)";
 by (rtac exI 1);
 by (rtac prod_assoc_bij 1);
 qed "prod_assoc_eqpoll";
 
 (*Unconditional version requires AC*)
-goalw CardinalArith.thy [cmult_def]
+Goalw [cmult_def]
     "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==>  \
 \             (i |*| j) |*| k = i |*| (j |*| k)";
 by (rtac cardinal_cong 1);
@@ -170,12 +170,12 @@
 
 (** Cardinal multiplication distributes over addition **)
 
-goalw CardinalArith.thy [eqpoll_def] "(A+B)*C eqpoll (A*C)+(B*C)";
+Goalw [eqpoll_def] "(A+B)*C eqpoll (A*C)+(B*C)";
 by (rtac exI 1);
 by (rtac sum_prod_distrib_bij 1);
 qed "sum_prod_distrib_eqpoll";
 
-goalw CardinalArith.thy [cadd_def, cmult_def]
+Goalw [cadd_def, cmult_def]
     "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==>  \
 \             (i |+| j) |*| k = (i |*| k) |+| (j |*| k)";
 by (rtac cardinal_cong 1);
@@ -189,38 +189,38 @@
 
 (** Multiplication by 0 yields 0 **)
 
-goalw CardinalArith.thy [eqpoll_def] "0*A eqpoll 0";
+Goalw [eqpoll_def] "0*A eqpoll 0";
 by (rtac exI 1);
 by (rtac lam_bijective 1);
 by Safe_tac;
 qed "prod_0_eqpoll";
 
-goalw CardinalArith.thy [cmult_def] "0 |*| i = 0";
+Goalw [cmult_def] "0 |*| i = 0";
 by (asm_simp_tac (simpset() addsimps [prod_0_eqpoll RS cardinal_cong, 
 				      Card_0 RS Card_cardinal_eq]) 1);
 qed "cmult_0";
 
 (** 1 is the identity for multiplication **)
 
-goalw CardinalArith.thy [eqpoll_def] "{x}*A eqpoll A";
+Goalw [eqpoll_def] "{x}*A eqpoll A";
 by (rtac exI 1);
 by (resolve_tac [singleton_prod_bij RS bij_converse_bij] 1);
 qed "prod_singleton_eqpoll";
 
-goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
+Goalw [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
 by (asm_simp_tac (simpset() addsimps [prod_singleton_eqpoll RS cardinal_cong, 
 				      Card_cardinal_eq]) 1);
 qed "cmult_1";
 
 (*** Some inequalities for multiplication ***)
 
-goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A";
+Goalw [lepoll_def, inj_def] "A lepoll A*A";
 by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1);
 by (simp_tac (simpset() addsimps [lam_type]) 1);
 qed "prod_square_lepoll";
 
 (*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
-goalw CardinalArith.thy [cmult_def] "!!K. Card(K) ==> K le K |*| K";
+Goalw [cmult_def] "!!K. Card(K) ==> K le K |*| K";
 by (rtac le_trans 1);
 by (rtac well_ord_lepoll_imp_Card_le 2);
 by (rtac prod_square_lepoll 3);
@@ -231,13 +231,13 @@
 
 (** Multiplication by a non-zero cardinal **)
 
-goalw CardinalArith.thy [lepoll_def, inj_def] "!!b. b: B ==> A lepoll A*B";
+Goalw [lepoll_def, inj_def] "!!b. b: B ==> A lepoll A*B";
 by (res_inst_tac [("x", "lam x:A. <x,b>")] exI 1);
 by (asm_simp_tac (simpset() addsimps [lam_type]) 1);
 qed "prod_lepoll_self";
 
 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
-goalw CardinalArith.thy [cmult_def]
+Goalw [cmult_def]
     "!!K. [| Card(K);  Ord(L);  0<L |] ==> K le (K |*| L)";
 by (rtac ([Card_cardinal_le, well_ord_lepoll_imp_Card_le] MRS le_trans) 1);
 by (rtac prod_lepoll_self 3);
@@ -246,7 +246,7 @@
 
 (** Monotonicity of multiplication **)
 
-goalw CardinalArith.thy [lepoll_def]
+Goalw [lepoll_def]
      "!!A B C D. [| A lepoll C;  B lepoll D |] ==> A * B  lepoll  C * D";
 by (REPEAT (etac exE 1));
 by (res_inst_tac [("x", "lam <w,y>:A*B. <f`w, fa`y>")] exI 1);
@@ -257,7 +257,7 @@
 by (asm_simp_tac (simpset() addsimps [left_inverse]) 1);
 qed "prod_lepoll_mono";
 
-goalw CardinalArith.thy [cmult_def]
+Goalw [cmult_def]
     "!!K. [| K' le K;  L' le L |] ==> (K' |*| L') le (K |*| L)";
 by (safe_tac (claset() addSDs [le_subset_iff RS iffD1]));
 by (rtac well_ord_lepoll_imp_Card_le 1);
@@ -267,7 +267,7 @@
 
 (*** Multiplication of finite cardinals is "ordinary" multiplication ***)
 
-goalw CardinalArith.thy [eqpoll_def] "succ(A)*B eqpoll B + A*B";
+Goalw [eqpoll_def] "succ(A)*B eqpoll B + A*B";
 by (rtac exI 1);
 by (res_inst_tac [("c", "%<x,y>. if(x=A, Inl(y), Inr(<x,y>))"), 
                   ("d", "case(%y. <A,y>, %z. z)")] 
@@ -278,7 +278,7 @@
 qed "prod_succ_eqpoll";
 
 (*Unconditional version requires AC*)
-goalw CardinalArith.thy [cmult_def, cadd_def]
+Goalw [cmult_def, cadd_def]
     "!!m n. [| Ord(m);  Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)";
 by (rtac (prod_succ_eqpoll RS cardinal_cong RS trans) 1);
 by (rtac (cardinal_cong RS sym) 1);
@@ -295,7 +295,7 @@
 				      nat_cadd_eq_add]) 1);
 qed "nat_cmult_eq_mult";
 
-goal CardinalArith.thy "!!m n. Card(n) ==> 2 |*| n = n |+| n";
+Goal "!!m n. Card(n) ==> 2 |*| n = n |+| n";
 by (asm_simp_tac 
     (simpset() addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, 
 			 Card_is_Ord, cadd_0,
@@ -309,7 +309,7 @@
   lam z:cons(u,A). if(z=u, 0, if(z : nat, succ(z), z))  and inverse
   %y. if(y:nat, nat_case(u,%z.z,y), y).  If f: inj(nat,A) then
   range(f) behaves like the natural numbers.*)
-goalw CardinalArith.thy [lepoll_def]
+Goalw [lepoll_def]
     "!!i. nat lepoll A ==> cons(u,A) lepoll A";
 by (etac exE 1);
 by (res_inst_tac [("x",
@@ -329,32 +329,32 @@
                setloop split_tac [expand_if]) 1);
 qed "nat_cons_lepoll";
 
-goal CardinalArith.thy "!!i. nat lepoll A ==> cons(u,A) eqpoll A";
+Goal "!!i. nat lepoll A ==> cons(u,A) eqpoll A";
 by (etac (nat_cons_lepoll RS eqpollI) 1);
 by (rtac (subset_consI RS subset_imp_lepoll) 1);
 qed "nat_cons_eqpoll";
 
 (*Specialized version required below*)
-goalw CardinalArith.thy [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A";
+Goalw [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A";
 by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1);
 qed "nat_succ_eqpoll";
 
-goalw CardinalArith.thy [InfCard_def] "InfCard(nat)";
+Goalw [InfCard_def] "InfCard(nat)";
 by (blast_tac (claset() addIs [Card_nat, le_refl, Card_is_Ord]) 1);
 qed "InfCard_nat";
 
-goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
+Goalw [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
 by (etac conjunct1 1);
 qed "InfCard_is_Card";
 
-goalw CardinalArith.thy [InfCard_def]
+Goalw [InfCard_def]
     "!!K L. [| InfCard(K);  Card(L) |] ==> InfCard(K Un L)";
 by (asm_simp_tac (simpset() addsimps [Card_Un, Un_upper1_le RSN (2,le_trans), 
 				      Card_is_Ord]) 1);
 qed "InfCard_Un";
 
 (*Kunen's Lemma 10.11*)
-goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
+Goalw [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
 by (etac conjE 1);
 by (forward_tac [Card_is_Ord] 1);
 by (rtac (ltI RS non_succ_LimitI) 1);
@@ -381,13 +381,13 @@
 
 (** Establishing the well-ordering **)
 
-goalw CardinalArith.thy [inj_def]
+Goalw [inj_def]
  "!!K. Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)";
 by (fast_tac (claset() addss (simpset())
                        addIs [lam_type, Un_least_lt RS ltD, ltI]) 1);
 qed "csquare_lam_inj";
 
-goalw CardinalArith.thy [csquare_rel_def]
+Goalw [csquare_rel_def]
  "!!K. Ord(K) ==> well_ord(K*K, csquare_rel(K))";
 by (rtac (csquare_lam_inj RS well_ord_rvimage) 1);
 by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
@@ -395,7 +395,7 @@
 
 (** Characterising initial segments of the well-ordering **)
 
-goalw CardinalArith.thy [csquare_rel_def]
+Goalw [csquare_rel_def]
  "!!K. [| x<K;  y<K;  z<K |] ==> \
 \      <<x,y>, <z,z>> : csquare_rel(K) --> x le z & y le z";
 by (REPEAT (etac ltE 1));
@@ -406,7 +406,7 @@
 by (ALLGOALS (asm_simp_tac (simpset() addsimps [lt_def, succI2, Ord_succ])));
 qed_spec_mp "csquareD";
 
-goalw CardinalArith.thy [pred_def]
+Goalw [pred_def]
  "!!K. z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)";
 by (safe_tac (claset_of ZF.thy addSEs [SigmaE]));  (*avoids using succCI,...*)
 by (rtac (csquareD RS conjE) 1);
@@ -415,7 +415,7 @@
 by (ALLGOALS Blast_tac);
 qed "pred_csquare_subset";
 
-goalw CardinalArith.thy [csquare_rel_def]
+Goalw [csquare_rel_def]
  "!!K. [| x<z;  y<z;  z<K |] ==>  <<x,y>, <z,z>> : csquare_rel(K)";
 by (subgoals_tac ["x<K", "y<K"] 1);
 by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2));
@@ -425,7 +425,7 @@
 qed "csquare_ltI";
 
 (*Part of the traditional proof.  UNUSED since it's harder to prove & apply *)
-goalw CardinalArith.thy [csquare_rel_def]
+Goalw [csquare_rel_def]
  "!!K. [| x le z;  y le z;  z<K |] ==> \
 \      <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z";
 by (subgoals_tac ["x<K", "y<K"] 1);
@@ -441,7 +441,7 @@
 
 (** The cardinality of initial segments **)
 
-goal CardinalArith.thy
+Goal
     "!!K. [| Limit(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
 \         ordermap(K*K, csquare_rel(K)) ` <x,y> <               \
 \         ordermap(K*K, csquare_rel(K)) ` <z,z>";
@@ -456,7 +456,7 @@
 qed "ordermap_z_lt";
 
 (*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *)
-goalw CardinalArith.thy [cmult_def]
+Goalw [cmult_def]
   "!!K. [| Limit(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
 \       | ordermap(K*K, csquare_rel(K)) ` <x,y> | le  |succ(z)| |*| |succ(z)|";
 by (rtac (well_ord_rmult RS well_ord_lepoll_imp_Card_le) 1);
@@ -476,7 +476,7 @@
 qed "ordermap_csquare_le";
 
 (*Kunen: "... so the order type <= K" *)
-goal CardinalArith.thy
+Goal
     "!!K. [| InfCard(K);  ALL y:K. InfCard(y) --> y |*| y = y |]  ==>  \
 \         ordertype(K*K, csquare_rel(K)) le K";
 by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1);
@@ -509,7 +509,7 @@
 qed "ordertype_csquare_le";
 
 (*Main result: Kunen's Theorem 10.12*)
-goal CardinalArith.thy "!!K. InfCard(K) ==> K |*| K = K";
+Goal "!!K. InfCard(K) ==> K |*| K = K";
 by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1);
 by (etac rev_mp 1);
 by (trans_ind_tac "K" [] 1);
@@ -527,7 +527,7 @@
 qed "InfCard_csquare_eq";
 
 (*Corollary for arbitrary well-ordered sets (all sets, assuming AC)*)
-goal CardinalArith.thy
+Goal
     "!!A. [| well_ord(A,r);  InfCard(|A|) |] ==> A*A eqpoll A";
 by (resolve_tac [prod_eqpoll_cong RS eqpoll_trans] 1);
 by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1));
@@ -539,7 +539,7 @@
 
 (** Toward's Kunen's Corollary 10.13 (1) **)
 
-goal CardinalArith.thy "!!K. [| InfCard(K);  L le K;  0<L |] ==> K |*| L = K";
+Goal "!!K. [| InfCard(K);  L le K;  0<L |] ==> K |*| L = K";
 by (rtac le_anti_sym 1);
 by (etac ltE 2 THEN
     REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2));
@@ -549,7 +549,7 @@
 qed "InfCard_le_cmult_eq";
 
 (*Corollary 10.13 (1), for cardinal multiplication*)
-goal CardinalArith.thy
+Goal
     "!!K. [| InfCard(K);  InfCard(L) |] ==> K |*| L = K Un L";
 by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1);
 by (typechk_tac [InfCard_is_Card, Card_is_Ord]);
@@ -562,7 +562,7 @@
 qed "InfCard_cmult_eq";
 
 (*This proof appear to be the simplest!*)
-goal CardinalArith.thy "!!K. InfCard(K) ==> K |+| K = K";
+Goal "!!K. InfCard(K) ==> K |+| K = K";
 by (asm_simp_tac
     (simpset() addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1);
 by (rtac InfCard_le_cmult_eq 1);
@@ -571,7 +571,7 @@
 qed "InfCard_cdouble_eq";
 
 (*Corollary 10.13 (1), for cardinal addition*)
-goal CardinalArith.thy "!!K. [| InfCard(K);  L le K |] ==> K |+| L = K";
+Goal "!!K. [| InfCard(K);  L le K |] ==> K |+| L = K";
 by (rtac le_anti_sym 1);
 by (etac ltE 2 THEN
     REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2));
@@ -580,7 +580,7 @@
 by (asm_simp_tac (simpset() addsimps [InfCard_cdouble_eq]) 1);
 qed "InfCard_le_cadd_eq";
 
-goal CardinalArith.thy
+Goal
     "!!K. [| InfCard(K);  InfCard(L) |] ==> K |+| L = K Un L";
 by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1);
 by (typechk_tac [InfCard_is_Card, Card_is_Ord]);
@@ -600,7 +600,7 @@
 (*** For every cardinal number there exists a greater one
      [Kunen's Theorem 10.16, which would be trivial using AC] ***)
 
-goalw CardinalArith.thy [jump_cardinal_def] "Ord(jump_cardinal(K))";
+Goalw [jump_cardinal_def] "Ord(jump_cardinal(K))";
 by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
 by (blast_tac (claset() addSIs [Ord_ordertype]) 2);
 by (rewtac Transset_def);
@@ -613,14 +613,14 @@
 qed "Ord_jump_cardinal";
 
 (*Allows selective unfolding.  Less work than deriving intro/elim rules*)
-goalw CardinalArith.thy [jump_cardinal_def]
+Goalw [jump_cardinal_def]
      "i : jump_cardinal(K) <->   \
 \         (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))";
 by (fast_tac subset_cs 1);      (*It's vital to avoid reasoning about <=*)
 qed "jump_cardinal_iff";
 
 (*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*)
-goal CardinalArith.thy "!!K. Ord(K) ==> K < jump_cardinal(K)";
+Goal "!!K. Ord(K) ==> K < jump_cardinal(K)";
 by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1);
 by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
 by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel]));
@@ -631,7 +631,7 @@
 
 (*The proof by contradiction: the bijection f yields a wellordering of X
   whose ordertype is jump_cardinal(K).  *)
-goal CardinalArith.thy
+Goal
     "!!K. [| well_ord(X,r);  r <= K * K;  X <= K;       \
 \            f : bij(ordertype(X,r), jump_cardinal(K))  \
 \         |] ==> jump_cardinal(K) : jump_cardinal(K)";
@@ -649,7 +649,7 @@
 qed "Card_jump_cardinal_lemma";
 
 (*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*)
-goal CardinalArith.thy "Card(jump_cardinal(K))";
+Goal "Card(jump_cardinal(K))";
 by (rtac (Ord_jump_cardinal RS CardI) 1);
 by (rewtac eqpoll_def);
 by (safe_tac (claset() addSDs [ltD, jump_cardinal_iff RS iffD1]));
@@ -658,7 +658,7 @@
 
 (*** Basic properties of successor cardinals ***)
 
-goalw CardinalArith.thy [csucc_def]
+Goalw [csucc_def]
     "!!K. Ord(K) ==> Card(csucc(K)) & K < csucc(K)";
 by (rtac LeastI 1);
 by (REPEAT (ares_tac [conjI, Card_jump_cardinal, K_lt_jump_cardinal,
@@ -669,18 +669,18 @@
 
 bind_thm ("lt_csucc", csucc_basic RS conjunct2);
 
-goal CardinalArith.thy "!!K. Ord(K) ==> 0 < csucc(K)";
+Goal "!!K. Ord(K) ==> 0 < csucc(K)";
 by (resolve_tac [[Ord_0_le, lt_csucc] MRS lt_trans1] 1);
 by (REPEAT (assume_tac 1));
 qed "Ord_0_lt_csucc";
 
-goalw CardinalArith.thy [csucc_def]
+Goalw [csucc_def]
     "!!K L. [| Card(L);  K<L |] ==> csucc(K) le L";
 by (rtac Least_le 1);
 by (REPEAT (ares_tac [conjI, Card_is_Ord] 1));
 qed "csucc_le";
 
-goal CardinalArith.thy
+Goal
     "!!K. [| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K";
 by (rtac iffI 1);
 by (rtac Card_lt_imp_lt 2);
@@ -694,13 +694,13 @@
      ORELSE eresolve_tac [ltE, Card_is_Ord] 1));
 qed "lt_csucc_iff";
 
-goal CardinalArith.thy
+Goal
     "!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K";
 by (asm_simp_tac 
     (simpset() addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
 qed "Card_lt_csucc_iff";
 
-goalw CardinalArith.thy [InfCard_def]
+Goalw [InfCard_def]
     "!!K. InfCard(K) ==> InfCard(csucc(K))";
 by (asm_simp_tac (simpset() addsimps [Card_csucc, Card_is_Ord, 
 				      lt_csucc RS leI RSN (2,le_trans)]) 1);
@@ -709,7 +709,7 @@
 
 (*** Finite sets ***)
 
-goal CardinalArith.thy
+Goal
     "!!n. n: nat ==> ALL A. A eqpoll n --> A : Fin(A)";
 by (etac nat_induct 1);
 by (simp_tac (simpset() addsimps (eqpoll_0_iff::Fin.intrs)) 1);
@@ -729,19 +729,19 @@
 by (blast_tac (claset() addIs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
 val lemma = result();
 
-goalw CardinalArith.thy [Finite_def] "!!A. Finite(A) ==> A : Fin(A)";
+Goalw [Finite_def] "!!A. Finite(A) ==> A : Fin(A)";
 by (blast_tac (claset() addIs [lemma RS spec RS mp]) 1);
 qed "Finite_into_Fin";
 
-goal CardinalArith.thy "!!A. A : Fin(U) ==> Finite(A)";
+Goal "!!A. A : Fin(U) ==> Finite(A)";
 by (fast_tac (claset() addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1);
 qed "Fin_into_Finite";
 
-goal CardinalArith.thy "Finite(A) <-> A : Fin(A)";
+Goal "Finite(A) <-> A : Fin(A)";
 by (blast_tac (claset() addIs [Finite_into_Fin, Fin_into_Finite]) 1);
 qed "Finite_Fin_iff";
 
-goal CardinalArith.thy
+Goal
     "!!A. [| Finite(A); Finite(B) |] ==> Finite(A Un B)";
 by (blast_tac (claset() addSIs [Fin_into_Finite, Fin_UnI] 
                         addSDs [Finite_into_Fin]
@@ -752,7 +752,7 @@
 
 (** Removing elements from a finite set decreases its cardinality **)
 
-goal CardinalArith.thy
+Goal
     "!!A. A: Fin(U) ==> x~:A --> ~ cons(x,A) lepoll A";
 by (etac Fin_induct 1);
 by (simp_tac (simpset() addsimps [lepoll_0_iff]) 1);
@@ -762,7 +762,7 @@
 by (Blast_tac 1);
 qed "Fin_imp_not_cons_lepoll";
 
-goal CardinalArith.thy
+Goal
     "!!a A. [| Finite(A);  a~:A |] ==> |cons(a,A)| = succ(|A|)";
 by (rewtac cardinal_def);
 by (rtac Least_equality 1);
@@ -783,7 +783,7 @@
 qed "Finite_imp_cardinal_cons";
 
 
-goal CardinalArith.thy "!!a A. [| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|";
+Goal "!!a A. [| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|";
 by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
 by (assume_tac 1);
 by (asm_simp_tac (simpset() addsimps [Finite_imp_cardinal_cons,
@@ -791,7 +791,7 @@
 by (asm_simp_tac (simpset() addsimps [cons_Diff]) 1);
 qed "Finite_imp_succ_cardinal_Diff";
 
-goal CardinalArith.thy "!!a A. [| Finite(A);  a:A |] ==> |A-{a}| < |A|";
+Goal "!!a A. [| Finite(A);  a:A |] ==> |A-{a}| < |A|";
 by (rtac succ_leE 1);
 by (asm_simp_tac (simpset() addsimps [Finite_imp_succ_cardinal_Diff, 
 				      Ord_cardinal RS le_refl]) 1);
@@ -803,7 +803,7 @@
 val nat_implies_well_ord =
   (transfer CardinalArith.thy nat_into_Ord) RS well_ord_Memrel;
 
-goal CardinalArith.thy "!!m n. [| m:nat; n:nat |] ==> m + n eqpoll m #+ n";
+Goal "!!m n. [| m:nat; n:nat |] ==> m + n eqpoll m #+ n";
 by (rtac eqpoll_trans 1);
 by (resolve_tac [well_ord_radd RS well_ord_cardinal_eqpoll RS eqpoll_sym] 1);
 by (REPEAT (etac nat_implies_well_ord 1));