# HG changeset patch # User wenzelm # Date 898528347 -7200 # Node ID 62b6288e6005486b7e848be70bec439d7623d15f # Parent 30271d90644f30826220f9d9ae0ae2c0ba8965e8 isatool fixgoal; diff -r 30271d90644f -r 62b6288e6005 src/ZF/AC.ML --- a/src/ZF/AC.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/AC.ML Mon Jun 22 17:12:27 1998 +0200 @@ -18,13 +18,13 @@ qed "AC_Pi"; (*Using dtac, this has the advantage of DELETING the universal quantifier*) -goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)"; +Goal "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)"; by (rtac AC_Pi 1); by (etac bspec 1); by (assume_tac 1); qed "AC_ball_Pi"; -goal AC.thy "EX f. f: (PROD X: Pow(C)-{0}. X)"; +Goal "EX f. f: (PROD X: Pow(C)-{0}. X)"; by (res_inst_tac [("B1", "%x. x")] (AC_Pi RS exE) 1); by (etac exI 2); by (Blast_tac 1); @@ -43,12 +43,12 @@ by (ALLGOALS Blast_tac); qed "non_empty_family"; -goal AC.thy "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x"; +Goal "!!A. 0 ~: A ==> EX f: A->Union(A). ALL x:A. f`x : x"; by (rtac AC_func 1); by (REPEAT (ares_tac [non_empty_family] 1)); qed "AC_func0"; -goal AC.thy "EX f: (Pow(C)-{0}) -> C. ALL x:(Pow(C)-{0}). f`x : x"; +Goal "EX f: (Pow(C)-{0}) -> C. ALL x:(Pow(C)-{0}). f`x : x"; by (resolve_tac [AC_func0 RS bexE] 1); by (rtac bexI 2); by (assume_tac 2); @@ -56,7 +56,7 @@ by (ALLGOALS Blast_tac); qed "AC_func_Pow"; -goal AC.thy "!!A. 0 ~: A ==> EX f. f: (PROD x:A. x)"; +Goal "!!A. 0 ~: A ==> EX f. f: (PROD x:A. x)"; by (rtac AC_Pi 1); by (REPEAT (ares_tac [non_empty_family] 1)); qed "AC_Pi0"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/AC/AC0_AC1.ML --- a/src/ZF/AC/AC0_AC1.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/AC/AC0_AC1.ML Mon Jun 22 17:12:27 1998 +0200 @@ -6,19 +6,19 @@ AC0 comes from Suppes, AC1 from Rubin & Rubin *) -goal thy "!!A. 0~:A ==> A <= Pow(Union(A))-{0}"; +Goal "!!A. 0~:A ==> A <= Pow(Union(A))-{0}"; by (Fast_tac 1); qed "subset_Pow_Union"; -goal thy "!!f. [| f:(PROD X:A. X); D<=A |] ==> EX g. g:(PROD X:D. X)"; +Goal "!!f. [| f:(PROD X:A. X); D<=A |] ==> EX g. g:(PROD X:D. X)"; by (fast_tac (claset() addSIs [restrict_type, apply_type]) 1); val lemma1 = result(); -goalw thy AC_defs "!!Z. AC0 ==> AC1"; +Goalw AC_defs "!!Z. AC0 ==> AC1"; by (fast_tac (FOL_cs addSEs [lemma1, subset_Pow_Union]) 1); qed "AC0_AC1"; -goalw thy AC_defs "!!Z. AC1 ==> AC0"; +Goalw AC_defs "!!Z. AC1 ==> AC0"; by (Deepen_tac 0 1); (*Large search space. Faster proof by by (fast_tac (claset() addSIs [notI, singletonI] addSEs [notE, DiffE]) 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/AC/AC10_AC15.ML --- a/src/ZF/AC/AC10_AC15.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/AC/AC10_AC15.ML Mon Jun 22 17:12:27 1998 +0200 @@ -27,13 +27,13 @@ (* - ex_fun_AC13_AC15 *) (* ********************************************************************** *) -goalw thy [lepoll_def] "!!A. A~=0 ==> B lepoll A*B"; +Goalw [lepoll_def] "!!A. A~=0 ==> B lepoll A*B"; by (etac not_emptyE 1); by (res_inst_tac [("x","lam z:B. ")] exI 1); by (fast_tac (claset() addSIs [snd_conv, lam_injective]) 1); qed "lepoll_Sigma"; -goal thy "!!A. 0~:A ==> ALL B:{cons(0,x*nat). x:A}. ~Finite(B)"; +Goal "!!A. 0~:A ==> ALL B:{cons(0,x*nat). x:A}. ~Finite(B)"; by (rtac ballI 1); by (etac RepFunE 1); by (hyp_subst_tac 1); @@ -44,18 +44,18 @@ by (Fast_tac 1); qed "cons_times_nat_not_Finite"; -goal thy "!!A. [| Union(C)=A; a:A |] ==> EX B:C. a:B & B <= A"; +Goal "!!A. [| Union(C)=A; a:A |] ==> EX B:C. a:B & B <= A"; by (Fast_tac 1); val lemma1 = result(); -goalw thy [pairwise_disjoint_def] +Goalw [pairwise_disjoint_def] "!!A. [| pairwise_disjoint(A); B:A; C:A; a:B; a:C |] ==> B=C"; by (dtac IntI 1 THEN (assume_tac 1)); by (dres_inst_tac [("A","B Int C")] not_emptyI 1); by (Fast_tac 1); val lemma2 = result(); -goalw thy [sets_of_size_between_def] +Goalw [sets_of_size_between_def] "!!A. ALL B:{cons(0, x*nat). x:A}. pairwise_disjoint(f`B) & \ \ sets_of_size_between(f`B, 2, n) & Union(f`B)=B \ \ ==> ALL B:A. EX! u. u:f`cons(0, B*nat) & u <= cons(0, B*nat) & \ @@ -72,7 +72,7 @@ by (rtac lemma2 1 THEN (REPEAT (assume_tac 1))); val lemma3 = result(); -goalw thy [lepoll_def] "!!A. [| A lepoll i; Ord(i) |] ==> {P(a). a:A} lepoll i"; +Goalw [lepoll_def] "!!A. [| A lepoll i; Ord(i) |] ==> {P(a). a:A} lepoll i"; by (etac exE 1); by (res_inst_tac [("x", "lam x:RepFun(A, P). LEAST j. EX a:A. x=P(a) & f`a=j")] exI 1); @@ -88,7 +88,7 @@ by (fast_tac (claset() addEs [sym, left_inverse RS ssubst]) 1); val lemma4 = result(); -goal thy "!!A. [| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B); \ +Goal "!!A. [| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B); \ \ u(B) lepoll succ(n) |] \ \ ==> (lam x:A. {fst(x). x:u(x)-{0}})`B ~= 0 & \ \ (lam x:A. {fst(x). x:u(x)-{0}})`B <= B & \ @@ -105,7 +105,7 @@ Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1); val lemma5 = result(); -goal thy "!!A. [| EX f. ALL B:{cons(0, x*nat). x:A}. \ +Goal "!!A. [| EX f. ALL B:{cons(0, x*nat). x:A}. \ \ pairwise_disjoint(f`B) & \ \ sets_of_size_between(f`B, 2, succ(n)) & \ \ Union(f`B)=B; n:nat |] \ @@ -123,7 +123,7 @@ (* AC10(n) ==> AC11 *) (* ********************************************************************** *) -goalw thy AC_defs "!!Z. [| n:nat; 1 le n; AC10(n) |] ==> AC11"; +Goalw AC_defs "!!Z. [| n:nat; 1 le n; AC10(n) |] ==> AC11"; by (rtac bexI 1 THEN (assume_tac 2)); by (Fast_tac 1); qed "AC10_AC11"; @@ -132,7 +132,7 @@ (* AC11 ==> AC12 *) (* ********************************************************************** *) -goalw thy AC_defs "!! Z. AC11 ==> AC12"; +Goalw AC_defs "!! Z. AC11 ==> AC12"; by (fast_tac (FOL_cs addSEs [bexE] addIs [bexI]) 1); qed "AC11_AC12"; @@ -140,7 +140,7 @@ (* AC12 ==> AC15 *) (* ********************************************************************** *) -goalw thy AC_defs "!!Z. AC12 ==> AC15"; +Goalw AC_defs "!!Z. AC12 ==> AC15"; by Safe_tac; by (etac allE 1); by (etac impE 1); @@ -166,7 +166,7 @@ (* AC10(n) implies AC13(n) for (1 le n) *) (* ********************************************************************** *) -goalw thy AC_defs "!!n. [| n:nat; 1 le n; AC10(n) |] ==> AC13(n)"; +Goalw AC_defs "!!n. [| n:nat; 1 le n; AC10(n) |] ==> AC13(n)"; by Safe_tac; by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE), ex_fun_AC13_AC15]) 1); @@ -180,7 +180,7 @@ (* AC1 ==> AC13(1) *) (* ********************************************************************** *) -goalw thy AC_defs "!!Z. AC1 ==> AC13(1)"; +Goalw AC_defs "!!Z. AC1 ==> AC13(1)"; by (rtac allI 1); by (etac allE 1); by (rtac impI 1); @@ -197,7 +197,7 @@ (* AC13(m) ==> AC13(n) for m <= n *) (* ********************************************************************** *) -goalw thy AC_defs "!!m n. [| m:nat; n:nat; m le n; AC13(m) |] ==> AC13(n)"; +Goalw AC_defs "!!m n. [| m:nat; n:nat; m le n; AC13(m) |] ==> AC13(n)"; by (dtac nat_le_imp_lepoll 1 THEN REPEAT (assume_tac 1)); by (fast_tac (claset() addSEs [lepoll_trans]) 1); qed "AC13_mono"; @@ -210,7 +210,7 @@ (* AC13(n) ==> AC14 if 1 <= n *) (* ********************************************************************** *) -goalw thy AC_defs "!!n. [| n:nat; 1 le n; AC13(n) |] ==> AC14"; +Goalw AC_defs "!!n. [| n:nat; 1 le n; AC13(n) |] ==> AC14"; by (fast_tac (FOL_cs addIs [bexI]) 1); qed "AC13_AC14"; @@ -218,7 +218,7 @@ (* AC14 ==> AC15 *) (* ********************************************************************** *) -goalw thy AC_defs "!!Z. AC14 ==> AC15"; +Goalw AC_defs "!!Z. AC14 ==> AC15"; by (Fast_tac 1); qed "AC14_AC15"; @@ -230,11 +230,11 @@ (* AC13(1) ==> AC1 *) (* ********************************************************************** *) -goal thy "!!A. [| A~=0; A lepoll 1 |] ==> EX a. A={a}"; +Goal "!!A. [| A~=0; A lepoll 1 |] ==> EX a. A={a}"; by (fast_tac (claset() addSEs [not_emptyE, lepoll_1_is_sing]) 1); qed "lemma_aux"; -goal thy "!!f. ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1 \ +Goal "!!f. ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1 \ \ ==> (lam x:A. THE y. f(x)={y}) : (PROD X:A. X)"; by (rtac lam_type 1); by (dtac bspec 1 THEN (assume_tac 1)); @@ -244,7 +244,7 @@ by (fast_tac (claset() addEs [ssubst]) 1); val lemma = result(); -goalw thy AC_defs "!!Z. AC13(1) ==> AC1"; +Goalw AC_defs "!!Z. AC13(1) ==> AC1"; by (fast_tac (claset() addSEs [lemma]) 1); qed "AC13_AC1"; @@ -252,6 +252,6 @@ (* AC11 ==> AC14 *) (* ********************************************************************** *) -goalw thy [AC11_def, AC14_def] "!!Z. AC11 ==> AC14"; +Goalw [AC11_def, AC14_def] "!!Z. AC11 ==> AC14"; by (fast_tac (claset() addSIs [AC10_AC13]) 1); qed "AC11_AC14"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Arith.ML --- a/src/ZF/Arith.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Arith.ML Mon Jun 22 17:12:27 1998 +0200 @@ -22,12 +22,12 @@ val rec_trans = rec_def RS def_transrec RS trans; -goal Arith.thy "rec(0,a,b) = a"; +Goal "rec(0,a,b) = a"; by (rtac rec_trans 1); by (rtac nat_case_0 1); qed "rec_0"; -goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))"; +Goal "rec(succ(m),a,b) = b(m, rec(m,a,b))"; by (rtac rec_trans 1); by (Simp_tac 1); qed "rec_succ"; @@ -47,7 +47,7 @@ Addsimps [rec_type, nat_0_le, nat_le_refl]; val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]; -goal Arith.thy "!!k. [| 0 EX j: nat. k = succ(j)"; +Goal "!!k. [| 0 EX j: nat. k = succ(j)"; by (etac rev_mp 1); by (etac nat_induct 1); by (Simp_tac 1); @@ -198,11 +198,11 @@ Addsimps [mult_0_right, mult_succ_right]; -goal Arith.thy "!!n. n:nat ==> 1 #* n = n"; +Goal "!!n. n:nat ==> 1 #* n = n"; by (Asm_simp_tac 1); qed "mult_1"; -goal Arith.thy "!!n. n:nat ==> n #* 1 = n"; +Goal "!!n. n:nat ==> n #* 1 = n"; by (Asm_simp_tac 1); qed "mult_1_right"; @@ -255,7 +255,7 @@ (ALLGOALS (asm_simp_tac (simpset() addsimps prems))) ]); (*Addition is the inverse of subtraction*) -goal Arith.thy "!!m n. [| n le m; m:nat |] ==> n #+ (m#-n) = m"; +Goal "!!m n. [| n le m; m:nat |] ==> n #+ (m#-n) = m"; by (forward_tac [lt_nat_in_nat] 1); by (etac nat_succI 1); by (etac rev_mp 1); @@ -264,7 +264,7 @@ qed "add_diff_inverse"; (*Proof is IDENTICAL to that above*) -goal Arith.thy "!!m n. [| n le m; m:nat |] ==> succ(m) #- n = succ(m#-n)"; +Goal "!!m n. [| n le m; m:nat |] ==> succ(m) #- n = succ(m#-n)"; by (forward_tac [lt_nat_in_nat] 1); by (etac nat_succI 1); by (etac rev_mp 1); @@ -280,19 +280,19 @@ by (ALLGOALS (asm_simp_tac (simpset() addsimps [mnat]))); qed "diff_add_inverse"; -goal Arith.thy +Goal "!!m n. [| m:nat; n:nat |] ==> (m#+n) #- n = m"; by (res_inst_tac [("m1","m")] (add_commute RS ssubst) 1); by (REPEAT (ares_tac [diff_add_inverse] 1)); qed "diff_add_inverse2"; -goal Arith.thy +Goal "!!k. [| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n"; by (nat_ind_tac "k" [] 1); by (ALLGOALS Asm_simp_tac); qed "diff_cancel"; -goal Arith.thy +Goal "!!n. [| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n"; val add_commute_k = read_instantiate [("n","k")] add_commute; by (asm_simp_tac (simpset() addsimps [add_commute_k, diff_cancel]) 1); @@ -306,13 +306,13 @@ (** Difference distributes over multiplication **) -goal Arith.thy +Goal "!!m n. [| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)"; by (res_inst_tac [("m","m"),("n","n")] diff_induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_cancel]))); qed "diff_mult_distrib" ; -goal Arith.thy +Goal "!!m. [| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)"; val mult_commute_k = read_instantiate [("m","k")] mult_commute; by (asm_simp_tac (simpset() addsimps @@ -321,7 +321,7 @@ (*** Remainder ***) -goal Arith.thy "!!m n. [| 0 m #- n < m"; +Goal "!!m n. [| 0 m #- n < m"; by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); by (etac rev_mp 1); by (etac rev_mp 1); @@ -338,16 +338,16 @@ not_lt_iff_le RS iffD2]; (*Type checking depends upon termination!*) -goalw Arith.thy [mod_def] "!!m n. [| 0 m mod n : nat"; +Goalw [mod_def] "!!m n. [| 0 m mod n : nat"; by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); qed "mod_type"; -goal Arith.thy "!!m n. [| 0 m mod n = m"; +Goal "!!m n. [| 0 m mod n = m"; by (rtac (mod_def RS def_transrec RS trans) 1); by (asm_simp_tac div_ss 1); qed "mod_less"; -goal Arith.thy "!!m n. [| 0 m mod n = (m#-n) mod n"; +Goal "!!m n. [| 0 m mod n = (m#-n) mod n"; by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); by (rtac (mod_def RS def_transrec RS trans) 1); by (asm_simp_tac div_ss 1); @@ -358,17 +358,17 @@ (*** Quotient ***) (*Type checking depends upon termination!*) -goalw Arith.thy [div_def] +Goalw [div_def] "!!m n. [| 0 m div n : nat"; by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); qed "div_type"; -goal Arith.thy "!!m n. [| 0 m div n = 0"; +Goal "!!m n. [| 0 m div n = 0"; by (rtac (div_def RS def_transrec RS trans) 1); by (asm_simp_tac div_ss 1); qed "div_less"; -goal Arith.thy +Goal "!!m n. [| 0 m div n = succ((m#-n) div n)"; by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); by (rtac (div_def RS def_transrec RS trans) 1); @@ -378,7 +378,7 @@ Addsimps [div_type, div_less, div_geq]; (*A key result*) -goal Arith.thy +Goal "!!m n. [| 0 (m div n)#*n #+ m mod n = m"; by (etac complete_induct 1); by (excluded_middle_tac "x \ \ succ(m) mod n = if(succ(m mod n) = n, 0, succ(m mod n))"; by (etac complete_induct 1); @@ -410,7 +410,7 @@ by (asm_simp_tac (simpset() addsimps [mod_less, diff_self_eq_0]) 1); qed "mod_succ"; -goal Arith.thy "!!m n. [| 0 m mod n < n"; +Goal "!!m n. [| 0 m mod n < n"; by (etac complete_induct 1); by (excluded_middle_tac "x k mod 2 = b | k mod 2 = if(b=1,0,1)"; by (subgoal_tac "k mod 2: 2" 1); by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); @@ -431,13 +431,13 @@ by (Blast_tac 1); qed "mod2_cases"; -goal Arith.thy "!!m. m:nat ==> succ(succ(m)) mod 2 = m mod 2"; +Goal "!!m. m:nat ==> succ(succ(m)) mod 2 = m mod 2"; by (subgoal_tac "m mod 2: 2" 1); by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2); by (asm_simp_tac (simpset() addsimps [mod_succ] setloop Step_tac) 1); qed "mod2_succ_succ"; -goal Arith.thy "!!m. m:nat ==> (m#+m) mod 2 = 0"; +Goal "!!m. m:nat ==> (m#+m) mod 2 = 0"; by (etac nat_induct 1); by (simp_tac (simpset() addsimps [mod_less]) 1); by (asm_simp_tac (simpset() addsimps [mod2_succ_succ, add_succ_right]) 1); @@ -446,12 +446,12 @@ (**** Additional theorems about "le" ****) -goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le m #+ n"; +Goal "!!m n. [| m:nat; n:nat |] ==> m le m #+ n"; by (etac nat_induct 1); by (ALLGOALS Asm_simp_tac); qed "add_le_self"; -goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le n #+ m"; +Goal "!!m n. [| m:nat; n:nat |] ==> m le n #+ m"; by (stac add_commute 1); by (REPEAT (ares_tac [add_le_self] 1)); qed "add_le_self2"; @@ -459,7 +459,7 @@ (*** Monotonicity of Addition ***) (*strict, in 1st argument; proof is by rule induction on 'less than'*) -goal Arith.thy "!!i j k. [| i i#+k < j#+k"; +Goal "!!i j k. [| i i#+k < j#+k"; by (forward_tac [lt_nat_in_nat] 1); by (assume_tac 1); by (etac succ_lt_induct 1); @@ -467,7 +467,7 @@ qed "add_lt_mono1"; (*strict, in both arguments*) -goal Arith.thy "!!i j k l. [| i i#+k < j#+l"; +Goal "!!i j k l. [| i i#+k < j#+l"; by (rtac (add_lt_mono1 RS lt_trans) 1); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); by (EVERY [stac add_commute 1, @@ -487,14 +487,14 @@ qed "Ord_lt_mono_imp_le_mono"; (*le monotonicity, 1st argument*) -goal Arith.thy +Goal "!!i j k. [| i le j; j:nat; k:nat |] ==> i#+k le j#+k"; by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1); by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1)); qed "add_le_mono1"; (* le monotonicity, BOTH arguments*) -goal Arith.thy +Goal "!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l"; by (rtac (add_le_mono1 RS le_trans) 1); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); @@ -506,14 +506,14 @@ (*** Monotonicity of Multiplication ***) -goal Arith.thy "!!i j k. [| i le j; j:nat; k:nat |] ==> i#*k le j#*k"; +Goal "!!i j k. [| i le j; j:nat; k:nat |] ==> i#*k le j#*k"; by (forward_tac [lt_nat_in_nat] 1); by (nat_ind_tac "k" [] 2); by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono]))); qed "mult_le_mono1"; (* le monotonicity, BOTH arguments*) -goal Arith.thy +Goal "!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l"; by (rtac (mult_le_mono1 RS le_trans) 1); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); @@ -524,7 +524,7 @@ qed "mult_le_mono"; (*strict, in 1st argument; proof is by induction on k>0*) -goal Arith.thy "!!i j k. [| i k#*i < k#*j"; +Goal "!!i j k. [| i k#*i < k#*j"; by (etac zero_lt_natE 1); by (forward_tac [lt_nat_in_nat] 2); by (ALLGOALS Asm_simp_tac); @@ -532,20 +532,20 @@ by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono]))); qed "mult_lt_mono2"; -goal thy "!!k. [| i i#*k < j#*k"; +Goal "!!k. [| i i#*k < j#*k"; by (asm_simp_tac (simpset() addsimps [mult_lt_mono2, mult_commute]) 1); qed "mult_lt_mono1"; -goal Arith.thy "!!k. [| m: nat; n: nat |] ==> 0 < m#*n <-> 0 0 < m#*n <-> 0 m#*n = 1 <-> m=1 & n=1"; +Goal "!!k. [| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1"; by (best_tac (claset() addEs [natE] addss (simpset())) 1); qed "mult_eq_1_iff"; (*Cancellation law for division*) -goal Arith.thy +Goal "!!k. [| 0 (k#*m) div (k#*n) = m div n"; by (eres_inst_tac [("i","m")] complete_induct 1); by (excluded_middle_tac "x \ \ (k#*m) mod (k#*n) = k #* (m mod n)"; by (eres_inst_tac [("i","m")] complete_induct 1); @@ -576,7 +576,7 @@ val mono_lemma = (nat_into_Ord RS Ord_0_lt) RSN (2,mult_lt_mono2); -goal Arith.thy "!!m n. [| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0"; +Goal "!!m n. [| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0"; by (rtac disjCI 1); by (dtac sym 1); by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I])); @@ -587,7 +587,7 @@ (*Thanks to Sten Agerholm*) -goal Arith.thy (* add_le_elim1 *) +Goal (* add_le_elim1 *) "!!m n k. [|m#+n le m#+k; m:nat; n:nat; k:nat|] ==> n le k"; by (etac rev_mp 1); by (eres_inst_tac [("n","n")] nat_induct 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/Bool.ML --- a/src/ZF/Bool.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Bool.ML Mon Jun 22 17:12:27 1998 +0200 @@ -10,21 +10,21 @@ val bool_defs = [bool_def,cond_def]; -goalw Bool.thy [succ_def] "{0} = 1"; +Goalw [succ_def] "{0} = 1"; by (rtac refl 1); qed "singleton_0"; (* Introduction rules *) -goalw Bool.thy bool_defs "1 : bool"; +Goalw bool_defs "1 : bool"; by (rtac (consI1 RS consI2) 1); qed "bool_1I"; -goalw Bool.thy bool_defs "0 : bool"; +Goalw bool_defs "0 : bool"; by (rtac consI1 1); qed "bool_0I"; -goalw Bool.thy bool_defs "1~=0"; +Goalw bool_defs "1~=0"; by (rtac succ_not_0 1); qed "one_not_0"; @@ -40,12 +40,12 @@ (** cond **) (*1 means true*) -goalw Bool.thy bool_defs "cond(1,c,d) = c"; +Goalw bool_defs "cond(1,c,d) = c"; by (rtac (refl RS if_P) 1); qed "cond_1"; (*0 means false*) -goalw Bool.thy bool_defs "cond(0,c,d) = d"; +Goalw bool_defs "cond(0,c,d) = d"; by (rtac (succ_not_0 RS not_sym RS if_not_P) 1); qed "cond_0"; @@ -54,7 +54,7 @@ fun bool_tac i = fast_tac (claset() addSEs [boolE] addss (simpset())) i; -goal Bool.thy +Goal "!!b. [| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)"; by (bool_tac 1); qed "cond_type"; @@ -99,15 +99,15 @@ (*** Laws for 'not' ***) -goal Bool.thy "!!a. a:bool ==> not(not(a)) = a"; +Goal "!!a. a:bool ==> not(not(a)) = a"; by (bool_tac 1); qed "not_not"; -goal Bool.thy "!!a b. a:bool ==> not(a and b) = not(a) or not(b)"; +Goal "!!a b. a:bool ==> not(a and b) = not(a) or not(b)"; by (bool_tac 1); qed "not_and"; -goal Bool.thy "!!a b. a:bool ==> not(a or b) = not(a) and not(b)"; +Goal "!!a b. a:bool ==> not(a or b) = not(a) and not(b)"; by (bool_tac 1); qed "not_or"; @@ -115,21 +115,21 @@ (*** Laws about 'and' ***) -goal Bool.thy "!!a. a: bool ==> a and a = a"; +Goal "!!a. a: bool ==> a and a = a"; by (bool_tac 1); qed "and_absorb"; Addsimps [and_absorb]; -goal Bool.thy "!!a b. [| a: bool; b:bool |] ==> a and b = b and a"; +Goal "!!a b. [| a: bool; b:bool |] ==> a and b = b and a"; by (bool_tac 1); qed "and_commute"; -goal Bool.thy "!!a. a: bool ==> (a and b) and c = a and (b and c)"; +Goal "!!a. a: bool ==> (a and b) and c = a and (b and c)"; by (bool_tac 1); qed "and_assoc"; -goal Bool.thy +Goal "!!a. [| a: bool; b:bool; c:bool |] ==> \ \ (a or b) and c = (a and c) or (b and c)"; by (bool_tac 1); @@ -137,21 +137,21 @@ (** binary orion **) -goal Bool.thy "!!a. a: bool ==> a or a = a"; +Goal "!!a. a: bool ==> a or a = a"; by (bool_tac 1); qed "or_absorb"; Addsimps [or_absorb]; -goal Bool.thy "!!a b. [| a: bool; b:bool |] ==> a or b = b or a"; +Goal "!!a b. [| a: bool; b:bool |] ==> a or b = b or a"; by (bool_tac 1); qed "or_commute"; -goal Bool.thy "!!a. a: bool ==> (a or b) or c = a or (b or c)"; +Goal "!!a. a: bool ==> (a or b) or c = a or (b or c)"; by (bool_tac 1); qed "or_assoc"; -goal Bool.thy +Goal "!!a b c. [| a: bool; b: bool; c: bool |] ==> \ \ (a and b) or c = (a or c) and (b or c)"; by (bool_tac 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/Cardinal.ML --- a/src/ZF/Cardinal.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Cardinal.ML Mon Jun 22 17:12:27 1998 +0200 @@ -14,7 +14,7 @@ (** Lemma: Banach's Decomposition Theorem **) -goal Cardinal.thy "bnd_mono(X, %W. X - g``(Y - f``W))"; +Goal "bnd_mono(X, %W. X - g``(Y - f``W))"; by (rtac bnd_monoI 1); by (REPEAT (ares_tac [Diff_subset, subset_refl, Diff_mono, image_mono] 1)); qed "decomp_bnd_mono"; @@ -54,25 +54,25 @@ (** Equipollence is an equivalence relation **) -goalw Cardinal.thy [eqpoll_def] "!!f A B. f: bij(A,B) ==> A eqpoll B"; +Goalw [eqpoll_def] "!!f A B. f: bij(A,B) ==> A eqpoll B"; by (etac exI 1); qed "bij_imp_eqpoll"; (*A eqpoll A*) bind_thm ("eqpoll_refl", id_bij RS bij_imp_eqpoll); -goalw Cardinal.thy [eqpoll_def] "!!X Y. X eqpoll Y ==> Y eqpoll X"; +Goalw [eqpoll_def] "!!X Y. X eqpoll Y ==> Y eqpoll X"; by (blast_tac (claset() addIs [bij_converse_bij]) 1); qed "eqpoll_sym"; -goalw Cardinal.thy [eqpoll_def] +Goalw [eqpoll_def] "!!X Y. [| X eqpoll Y; Y eqpoll Z |] ==> X eqpoll Z"; by (blast_tac (claset() addIs [comp_bij]) 1); qed "eqpoll_trans"; (** Le-pollence is a partial ordering **) -goalw Cardinal.thy [lepoll_def] "!!X Y. X<=Y ==> X lepoll Y"; +Goalw [lepoll_def] "!!X Y. X<=Y ==> X lepoll Y"; by (rtac exI 1); by (etac id_subset_inj 1); qed "subset_imp_lepoll"; @@ -81,18 +81,18 @@ bind_thm ("le_imp_lepoll", le_imp_subset RS subset_imp_lepoll); -goalw Cardinal.thy [eqpoll_def, bij_def, lepoll_def] +Goalw [eqpoll_def, bij_def, lepoll_def] "!!X Y. X eqpoll Y ==> X lepoll Y"; by (Blast_tac 1); qed "eqpoll_imp_lepoll"; -goalw Cardinal.thy [lepoll_def] +Goalw [lepoll_def] "!!X Y. [| X lepoll Y; Y lepoll Z |] ==> X lepoll Z"; by (blast_tac (claset() addIs [comp_inj]) 1); qed "lepoll_trans"; (*Asymmetry law*) -goalw Cardinal.thy [lepoll_def,eqpoll_def] +Goalw [lepoll_def,eqpoll_def] "!!X Y. [| X lepoll Y; Y lepoll X |] ==> X eqpoll Y"; by (REPEAT (etac exE 1)); by (rtac schroeder_bernstein 1); @@ -105,22 +105,22 @@ by (REPEAT (resolve_tac [major, eqpoll_imp_lepoll, eqpoll_sym] 1)); qed "eqpollE"; -goal Cardinal.thy "X eqpoll Y <-> X lepoll Y & Y lepoll X"; +Goal "X eqpoll Y <-> X lepoll Y & Y lepoll X"; by (blast_tac (claset() addIs [eqpollI] addSEs [eqpollE]) 1); qed "eqpoll_iff"; -goalw Cardinal.thy [lepoll_def, inj_def] "!!A. A lepoll 0 ==> A = 0"; +Goalw [lepoll_def, inj_def] "!!A. A lepoll 0 ==> A = 0"; by (blast_tac (claset() addDs [apply_type]) 1); qed "lepoll_0_is_0"; (*0 lepoll Y*) bind_thm ("empty_lepollI", empty_subsetI RS subset_imp_lepoll); -goal Cardinal.thy "A lepoll 0 <-> A=0"; +Goal "A lepoll 0 <-> A=0"; by (blast_tac (claset() addIs [lepoll_0_is_0, lepoll_refl]) 1); qed "lepoll_0_iff"; -goalw Cardinal.thy [lepoll_def] +Goalw [lepoll_def] "!!A. [| A lepoll B; C lepoll D; B Int D = 0 |] ==> A Un C lepoll B Un D"; by (blast_tac (claset() addIs [inj_disjoint_Un]) 1); qed "Un_lepoll_Un"; @@ -128,11 +128,11 @@ (*A eqpoll 0 ==> A=0*) bind_thm ("eqpoll_0_is_0", eqpoll_imp_lepoll RS lepoll_0_is_0); -goal Cardinal.thy "A eqpoll 0 <-> A=0"; +Goal "A eqpoll 0 <-> A=0"; by (blast_tac (claset() addIs [eqpoll_0_is_0, eqpoll_refl]) 1); qed "eqpoll_0_iff"; -goalw Cardinal.thy [eqpoll_def] +Goalw [eqpoll_def] "!!A. [| A eqpoll B; C eqpoll D; A Int C = 0; B Int D = 0 |] ==> \ \ A Un C eqpoll B Un D"; by (blast_tac (claset() addIs [bij_disjoint_Un]) 1); @@ -141,20 +141,20 @@ (*** lesspoll: contributions by Krzysztof Grabczewski ***) -goalw Cardinal.thy [lesspoll_def] "!!A. A lesspoll B ==> A lepoll B"; +Goalw [lesspoll_def] "!!A. A lesspoll B ==> A lepoll B"; by (Blast_tac 1); qed "lesspoll_imp_lepoll"; -goalw Cardinal.thy [lepoll_def] +Goalw [lepoll_def] "!!A. [| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)"; by (blast_tac (claset() addIs [well_ord_rvimage]) 1); qed "lepoll_well_ord"; -goalw Cardinal.thy [lesspoll_def] "A lepoll B <-> A lesspoll B | A eqpoll B"; +Goalw [lesspoll_def] "A lepoll B <-> A lesspoll B | A eqpoll B"; by (blast_tac (claset() addSIs [eqpollI] addSEs [eqpollE]) 1); qed "lepoll_iff_leqpoll"; -goalw Cardinal.thy [inj_def, surj_def] +Goalw [inj_def, surj_def] "!!f. [| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)"; by (safe_tac (claset_of ZF.thy)); by (swap_res_tac [exI] 1); @@ -168,17 +168,17 @@ (** Variations on transitivity **) -goalw Cardinal.thy [lesspoll_def] +Goalw [lesspoll_def] "!!X. [| X lesspoll Y; Y lesspoll Z |] ==> X lesspoll Z"; by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1); qed "lesspoll_trans"; -goalw Cardinal.thy [lesspoll_def] +Goalw [lesspoll_def] "!!X. [| X lesspoll Y; Y lepoll Z |] ==> X lesspoll Z"; by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1); qed "lesspoll_lepoll_lesspoll"; -goalw Cardinal.thy [lesspoll_def] +Goalw [lesspoll_def] "!!X. [| X lesspoll Y; Z lepoll X |] ==> Z lesspoll Y"; by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1); qed "lepoll_lesspoll_lesspoll"; @@ -195,7 +195,7 @@ by (ALLGOALS (blast_tac (claset() addSIs [premP] addSDs [premNot]))); qed "Least_equality"; -goal Cardinal.thy "!!i. [| P(i); Ord(i) |] ==> P(LEAST x. P(x))"; +Goal "!!i. [| P(i); Ord(i) |] ==> P(LEAST x. P(x))"; by (etac rev_mp 1); by (trans_ind_tac "i" [] 1); by (rtac impI 1); @@ -206,7 +206,7 @@ qed "LeastI"; (*Proof is almost identical to the one above!*) -goal Cardinal.thy "!!i. [| P(i); Ord(i) |] ==> (LEAST x. P(x)) le i"; +Goal "!!i. [| P(i); Ord(i) |] ==> (LEAST x. P(x)) le i"; by (etac rev_mp 1); by (trans_ind_tac "i" [] 1); by (rtac impI 1); @@ -217,7 +217,7 @@ qed "Least_le"; (*LEAST really is the smallest*) -goal Cardinal.thy "!!i. [| P(i); i < (LEAST x. P(x)) |] ==> Q"; +Goal "!!i. [| P(i); i < (LEAST x. P(x)) |] ==> Q"; by (rtac (Least_le RSN (2,lt_trans2) RS lt_irrefl) 1); by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); qed "less_LeastE"; @@ -231,13 +231,13 @@ resolve_tac prems 1 ]); (*If there is no such P then LEAST is vacuously 0*) -goalw Cardinal.thy [Least_def] +Goalw [Least_def] "!!P. [| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0"; by (rtac the_0 1); by (Blast_tac 1); qed "Least_0"; -goal Cardinal.thy "Ord(LEAST x. P(x))"; +Goal "Ord(LEAST x. P(x))"; by (excluded_middle_tac "EX i. Ord(i) & P(i)" 1); by Safe_tac; by (rtac (Least_le RS ltE) 2); @@ -257,13 +257,13 @@ (*Need AC to get X lepoll Y ==> |X| le |Y|; see well_ord_lepoll_imp_Card_le Converse also requires AC, but see well_ord_cardinal_eqE*) -goalw Cardinal.thy [eqpoll_def,cardinal_def] "!!X Y. X eqpoll Y ==> |X| = |Y|"; +Goalw [eqpoll_def,cardinal_def] "!!X Y. X eqpoll Y ==> |X| = |Y|"; by (rtac Least_cong 1); by (blast_tac (claset() addIs [comp_bij, bij_converse_bij]) 1); qed "cardinal_cong"; (*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*) -goalw Cardinal.thy [cardinal_def] +Goalw [cardinal_def] "!!A. well_ord(A,r) ==> |A| eqpoll A"; by (rtac LeastI 1); by (etac Ord_ordertype 2); @@ -273,14 +273,14 @@ (* Ord(A) ==> |A| eqpoll A *) bind_thm ("Ord_cardinal_eqpoll", well_ord_Memrel RS well_ord_cardinal_eqpoll); -goal Cardinal.thy +Goal "!!X Y. [| well_ord(X,r); well_ord(Y,s); |X| = |Y| |] ==> X eqpoll Y"; by (rtac (eqpoll_sym RS eqpoll_trans) 1); by (etac well_ord_cardinal_eqpoll 1); by (asm_simp_tac (simpset() addsimps [well_ord_cardinal_eqpoll]) 1); qed "well_ord_cardinal_eqE"; -goal Cardinal.thy +Goal "!!X Y. [| well_ord(X,r); well_ord(Y,s) |] ==> |X| = |Y| <-> X eqpoll Y"; by (blast_tac (claset() addIs [cardinal_cong, well_ord_cardinal_eqE]) 1); qed "well_ord_cardinal_eqpoll_iff"; @@ -288,11 +288,11 @@ (** Observations from Kunen, page 28 **) -goalw Cardinal.thy [cardinal_def] "!!i. Ord(i) ==> |i| le i"; +Goalw [cardinal_def] "!!i. Ord(i) ==> |i| le i"; by (etac (eqpoll_refl RS Least_le) 1); qed "Ord_cardinal_le"; -goalw Cardinal.thy [Card_def] "!!K. Card(K) ==> |K| = K"; +Goalw [Card_def] "!!K. Card(K) ==> |K| = K"; by (etac sym 1); qed "Card_cardinal_eq"; @@ -303,21 +303,21 @@ by (REPEAT (ares_tac ([refl,eqpoll_refl]@prems) 1)); qed "CardI"; -goalw Cardinal.thy [Card_def, cardinal_def] "!!i. Card(i) ==> Ord(i)"; +Goalw [Card_def, cardinal_def] "!!i. Card(i) ==> Ord(i)"; by (etac ssubst 1); by (rtac Ord_Least 1); qed "Card_is_Ord"; -goal Cardinal.thy "!!K. Card(K) ==> K le |K|"; +Goal "!!K. Card(K) ==> K le |K|"; by (asm_simp_tac (simpset() addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1); qed "Card_cardinal_le"; -goalw Cardinal.thy [cardinal_def] "Ord(|A|)"; +Goalw [cardinal_def] "Ord(|A|)"; by (rtac Ord_Least 1); qed "Ord_cardinal"; (*The cardinals are the initial ordinals*) -goal Cardinal.thy "Card(K) <-> Ord(K) & (ALL j. j ~ j eqpoll K)"; +Goal "Card(K) <-> Ord(K) & (ALL j. j ~ j eqpoll K)"; by (safe_tac (claset() addSIs [CardI, Card_is_Ord])); by (Blast_tac 2); by (rewrite_goals_tac [Card_def, cardinal_def]); @@ -326,12 +326,12 @@ by (ALLGOALS assume_tac); qed "Card_iff_initial"; -goalw Cardinal.thy [lesspoll_def] "!!a. [| Card(a); i i lesspoll a"; +Goalw [lesspoll_def] "!!a. [| Card(a); i i lesspoll a"; by (dresolve_tac [Card_iff_initial RS iffD1] 1); by (blast_tac (claset() addSIs [leI RS le_imp_lepoll]) 1); qed "lt_Card_imp_lesspoll"; -goal Cardinal.thy "Card(0)"; +Goal "Card(0)"; by (rtac (Ord_0 RS CardI) 1); by (blast_tac (claset() addSEs [ltE]) 1); qed "Card_0"; @@ -347,7 +347,7 @@ (*Infinite unions of cardinals? See Devlin, Lemma 6.7, page 98*) -goalw Cardinal.thy [cardinal_def] "Card(|A|)"; +Goalw [cardinal_def] "Card(|A|)"; by (excluded_middle_tac "EX i. Ord(i) & i eqpoll A" 1); by (etac (Least_0 RS ssubst) 1 THEN rtac Card_0 1); by (rtac (Ord_Least RS CardI) 1); @@ -359,7 +359,7 @@ qed "Card_cardinal"; (*Kunen's Lemma 10.5*) -goal Cardinal.thy "!!i j. [| |i| le j; j le i |] ==> |j| = |i|"; +Goal "!!i j. [| |i| le j; j le i |] ==> |j| = |i|"; by (rtac (eqpollI RS cardinal_cong) 1); by (etac le_imp_lepoll 1); by (rtac lepoll_trans 1); @@ -369,7 +369,7 @@ by (REPEAT (eresolve_tac [ltE, Ord_succD] 1)); qed "cardinal_eq_lemma"; -goal Cardinal.thy "!!i j. i le j ==> |i| le |j|"; +Goal "!!i j. i le j ==> |i| le |j|"; by (res_inst_tac [("i","|i|"),("j","|j|")] Ord_linear_le 1); by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI])); by (rtac cardinal_eq_lemma 1); @@ -380,30 +380,30 @@ qed "cardinal_mono"; (*Since we have |succ(nat)| le |nat|, the converse of cardinal_mono fails!*) -goal Cardinal.thy "!!i j. [| |i| < |j|; Ord(i); Ord(j) |] ==> i < j"; +Goal "!!i j. [| |i| < |j|; Ord(i); Ord(j) |] ==> i < j"; by (rtac Ord_linear2 1); by (REPEAT_SOME assume_tac); by (etac (lt_trans2 RS lt_irrefl) 1); by (etac cardinal_mono 1); qed "cardinal_lt_imp_lt"; -goal Cardinal.thy "!!i j. [| |i| < K; Ord(i); Card(K) |] ==> i < K"; +Goal "!!i j. [| |i| < K; Ord(i); Card(K) |] ==> i < K"; by (asm_simp_tac (simpset() addsimps [cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1); qed "Card_lt_imp_lt"; -goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (|i| < K) <-> (i < K)"; +Goal "!!i j. [| Ord(i); Card(K) |] ==> (|i| < K) <-> (i < K)"; by (blast_tac (claset() addIs [Card_lt_imp_lt, Ord_cardinal_le RS lt_trans1]) 1); qed "Card_lt_iff"; -goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (K le |i|) <-> (K le i)"; +Goal "!!i j. [| Ord(i); Card(K) |] ==> (K le |i|) <-> (K le i)"; by (asm_simp_tac (simpset() addsimps [Card_lt_iff, Card_is_Ord, Ord_cardinal, not_lt_iff_le RS iff_sym]) 1); qed "Card_le_iff"; (*Can use AC or finiteness to discharge first premise*) -goal Cardinal.thy +Goal "!!A B. [| well_ord(B,r); A lepoll B |] ==> |A| le |B|"; by (res_inst_tac [("i","|A|"),("j","|B|")] Ord_linear_le 1); by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI])); @@ -421,7 +421,7 @@ qed "well_ord_lepoll_imp_Card_le"; -goal Cardinal.thy "!!A. [| A lepoll i; Ord(i) |] ==> |A| le i"; +Goal "!!A. [| A lepoll i; Ord(i) |] ==> |A| le i"; by (rtac le_trans 1); by (etac (well_ord_Memrel RS well_ord_lepoll_imp_Card_le) 1); by (assume_tac 1); @@ -431,7 +431,7 @@ (*** The finite cardinals ***) -goalw Cardinal.thy [lepoll_def, inj_def] +Goalw [lepoll_def, inj_def] "!!A B. [| cons(u,A) lepoll cons(v,B); u~:A; v~:B |] ==> A lepoll B"; by Safe_tac; by (res_inst_tac [("x", "lam x:A. if(f`x=v, f`u, f`x)")] exI 1); @@ -445,14 +445,14 @@ by (Blast_tac 1); qed "cons_lepoll_consD"; -goal Cardinal.thy +Goal "!!A B. [| cons(u,A) eqpoll cons(v,B); u~:A; v~:B |] ==> A eqpoll B"; by (asm_full_simp_tac (simpset() addsimps [eqpoll_iff]) 1); by (blast_tac (claset() addIs [cons_lepoll_consD]) 1); qed "cons_eqpoll_consD"; (*Lemma suggested by Mike Fourman*) -goalw Cardinal.thy [succ_def] "!!m n. succ(m) lepoll succ(n) ==> m lepoll n"; +Goalw [succ_def] "!!m n. succ(m) lepoll succ(n) ==> m lepoll n"; by (etac cons_lepoll_consD 1); by (REPEAT (rtac mem_not_refl 1)); qed "succ_lepoll_succD"; @@ -470,7 +470,7 @@ bind_thm ("nat_lepoll_imp_le", nat_lepoll_imp_le_lemma RS bspec RS mp); -goal Cardinal.thy +Goal "!!m n. [| m:nat; n: nat |] ==> m eqpoll n <-> m = n"; by (rtac iffI 1); by (asm_simp_tac (simpset() addsimps [eqpoll_refl]) 2); @@ -479,7 +479,7 @@ qed "nat_eqpoll_iff"; (*The object of all this work: every natural number is a (finite) cardinal*) -goalw Cardinal.thy [Card_def,cardinal_def] +Goalw [Card_def,cardinal_def] "!!n. n: nat ==> Card(n)"; by (stac Least_equality 1); by (REPEAT_FIRST (ares_tac [eqpoll_refl, nat_into_Ord, refl])); @@ -488,7 +488,7 @@ qed "nat_into_Card"; (*Part of Kunen's Lemma 10.6*) -goal Cardinal.thy "!!n. [| succ(n) lepoll n; n:nat |] ==> P"; +Goal "!!n. [| succ(n) lepoll n; n:nat |] ==> P"; by (rtac (nat_lepoll_imp_le RS lt_irrefl) 1); by (REPEAT (ares_tac [nat_succI] 1)); qed "succ_lepoll_natE"; @@ -496,7 +496,7 @@ (** lepoll, lesspoll and natural numbers **) -goalw Cardinal.thy [lesspoll_def] +Goalw [lesspoll_def] "!!m. [| A lepoll m; m:nat |] ==> A lesspoll succ(m)"; by (rtac conjI 1); by (blast_tac (claset() addIs [subset_imp_lepoll RSN (2,lepoll_trans)]) 1); @@ -506,18 +506,18 @@ by (etac succ_lepoll_natE 1 THEN assume_tac 1); qed "lepoll_imp_lesspoll_succ"; -goalw Cardinal.thy [lesspoll_def, lepoll_def, eqpoll_def, bij_def] +Goalw [lesspoll_def, lepoll_def, eqpoll_def, bij_def] "!!m. [| A lesspoll succ(m); m:nat |] ==> A lepoll m"; by (Clarify_tac 1); by (blast_tac (claset() addSIs [inj_not_surj_succ]) 1); qed "lesspoll_succ_imp_lepoll"; -goal Cardinal.thy "!!m. m:nat ==> A lesspoll succ(m) <-> A lepoll m"; +Goal "!!m. m:nat ==> A lesspoll succ(m) <-> A lepoll m"; by (blast_tac (claset() addSIs [lepoll_imp_lesspoll_succ, lesspoll_succ_imp_lepoll]) 1); qed "lesspoll_succ_iff"; -goal Cardinal.thy "!!A m. [| A lepoll succ(m); m:nat |] ==> \ +Goal "!!A m. [| A lepoll succ(m); m:nat |] ==> \ \ A lepoll m | A eqpoll succ(m)"; by (rtac disjCI 1); by (rtac lesspoll_succ_imp_lepoll 1); @@ -529,7 +529,7 @@ (*** The first infinite cardinal: Omega, or nat ***) (*This implies Kunen's Lemma 10.6*) -goal Cardinal.thy "!!n. [| n ~ i lepoll n"; +Goal "!!n. [| n ~ i lepoll n"; by (rtac notI 1); by (rtac succ_lepoll_natE 1 THEN assume_tac 2); by (rtac lepoll_trans 1 THEN assume_tac 2); @@ -537,7 +537,7 @@ by (REPEAT (ares_tac [Ord_succ_subsetI RS subset_imp_lepoll] 1)); qed "lt_not_lepoll"; -goal Cardinal.thy "!!i n. [| Ord(i); n:nat |] ==> i eqpoll n <-> i=n"; +Goal "!!i n. [| Ord(i); n:nat |] ==> i eqpoll n <-> i=n"; by (rtac iffI 1); by (asm_simp_tac (simpset() addsimps [eqpoll_refl]) 2); by (rtac Ord_linear_lt 1); @@ -548,7 +548,7 @@ by (etac eqpoll_imp_lepoll 1); qed "Ord_nat_eqpoll_iff"; -goalw Cardinal.thy [Card_def,cardinal_def] "Card(nat)"; +Goalw [Card_def,cardinal_def] "Card(nat)"; by (stac Least_equality 1); by (REPEAT_FIRST (ares_tac [eqpoll_refl, Ord_nat, refl])); by (etac ltE 1); @@ -556,7 +556,7 @@ qed "Card_nat"; (*Allows showing that |i| is a limit cardinal*) -goal Cardinal.thy "!!i. nat le i ==> nat le |i|"; +Goal "!!i. nat le i ==> nat le |i|"; by (rtac (Card_nat RS Card_cardinal_eq RS subst) 1); by (etac cardinal_mono 1); qed "nat_le_cardinal"; @@ -566,7 +566,7 @@ (** Congruence laws for successor, cardinal addition and multiplication **) (*Congruence law for cons under equipollence*) -goalw Cardinal.thy [lepoll_def] +Goalw [lepoll_def] "!!A B. [| A lepoll B; b ~: B |] ==> cons(a,A) lepoll cons(b,B)"; by Safe_tac; by (res_inst_tac [("x", "lam y: cons(a,A).if(y=a, b, f`y)")] exI 1); @@ -578,51 +578,51 @@ setloop etac consE') 1); qed "cons_lepoll_cong"; -goal Cardinal.thy +Goal "!!A B. [| A eqpoll B; a ~: A; b ~: B |] ==> cons(a,A) eqpoll cons(b,B)"; by (asm_full_simp_tac (simpset() addsimps [eqpoll_iff, cons_lepoll_cong]) 1); qed "cons_eqpoll_cong"; -goal Cardinal.thy +Goal "!!A B. [| a ~: A; b ~: B |] ==> \ \ cons(a,A) lepoll cons(b,B) <-> A lepoll B"; by (blast_tac (claset() addIs [cons_lepoll_cong, cons_lepoll_consD]) 1); qed "cons_lepoll_cons_iff"; -goal Cardinal.thy +Goal "!!A B. [| a ~: A; b ~: B |] ==> \ \ cons(a,A) eqpoll cons(b,B) <-> A eqpoll B"; by (blast_tac (claset() addIs [cons_eqpoll_cong, cons_eqpoll_consD]) 1); qed "cons_eqpoll_cons_iff"; -goalw Cardinal.thy [succ_def] "{a} eqpoll 1"; +Goalw [succ_def] "{a} eqpoll 1"; by (blast_tac (claset() addSIs [eqpoll_refl RS cons_eqpoll_cong]) 1); qed "singleton_eqpoll_1"; -goal Cardinal.thy "|{a}| = 1"; +Goal "|{a}| = 1"; by (resolve_tac [singleton_eqpoll_1 RS cardinal_cong RS trans] 1); by (simp_tac (simpset() addsimps [nat_into_Card RS Card_cardinal_eq]) 1); qed "cardinal_singleton"; (*Congruence law for succ under equipollence*) -goalw Cardinal.thy [succ_def] +Goalw [succ_def] "!!A B. A eqpoll B ==> succ(A) eqpoll succ(B)"; by (REPEAT (ares_tac [cons_eqpoll_cong, mem_not_refl] 1)); qed "succ_eqpoll_cong"; (*Congruence law for + under equipollence*) -goalw Cardinal.thy [eqpoll_def] +Goalw [eqpoll_def] "!!A B C D. [| A eqpoll C; B eqpoll D |] ==> A+B eqpoll C+D"; by (blast_tac (claset() addSIs [sum_bij]) 1); qed "sum_eqpoll_cong"; (*Congruence law for * under equipollence*) -goalw Cardinal.thy [eqpoll_def] +Goalw [eqpoll_def] "!!A B C D. [| A eqpoll C; B eqpoll D |] ==> A*B eqpoll C*D"; by (blast_tac (claset() addSIs [prod_bij]) 1); qed "prod_eqpoll_cong"; -goalw Cardinal.thy [eqpoll_def] +Goalw [eqpoll_def] "!!f. [| f: inj(A,B); A Int B = 0 |] ==> A Un (B - range(f)) eqpoll B"; by (rtac exI 1); by (res_inst_tac [("c", "%x. if(x:A, f`x, x)"), @@ -645,7 +645,7 @@ Could easily generalise from succ to cons. ***) (*If A has at most n+1 elements and a:A then A-{a} has at most n.*) -goalw Cardinal.thy [succ_def] +Goalw [succ_def] "!!A a n. [| a:A; A lepoll succ(n) |] ==> A - {a} lepoll n"; by (rtac cons_lepoll_consD 1); by (rtac mem_not_refl 3); @@ -654,7 +654,7 @@ qed "Diff_sing_lepoll"; (*If A has at least n+1 elements then A-{a} has at least n.*) -goalw Cardinal.thy [succ_def] +Goalw [succ_def] "!!A a n. [| succ(n) lepoll A |] ==> n lepoll A - {a}"; by (rtac cons_lepoll_consD 1); by (rtac mem_not_refl 2); @@ -662,19 +662,19 @@ by (blast_tac (claset() addIs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1); qed "lepoll_Diff_sing"; -goal Cardinal.thy "!!A a n. [| a:A; A eqpoll succ(n) |] ==> A - {a} eqpoll n"; +Goal "!!A a n. [| a:A; A eqpoll succ(n) |] ==> A - {a} eqpoll n"; by (blast_tac (claset() addSIs [eqpollI] addSEs [eqpollE] addIs [Diff_sing_lepoll,lepoll_Diff_sing]) 1); qed "Diff_sing_eqpoll"; -goal Cardinal.thy "!!A. [| A lepoll 1; a:A |] ==> A = {a}"; +Goal "!!A. [| A lepoll 1; a:A |] ==> A = {a}"; by (forward_tac [Diff_sing_lepoll] 1); by (assume_tac 1); by (dtac lepoll_0_is_0 1); by (blast_tac (claset() addEs [equalityE]) 1); qed "lepoll_1_is_sing"; -goalw Cardinal.thy [lepoll_def] "A Un B lepoll A+B"; +Goalw [lepoll_def] "A Un B lepoll A+B"; by (res_inst_tac [("x","lam x: A Un B. if (x:A,Inl(x),Inr(x))")] exI 1); by (res_inst_tac [("d","%z. snd(z)")] lam_injective 1); by (split_tac [expand_if] 1); @@ -686,11 +686,11 @@ (*** Finite and infinite sets ***) -goalw Cardinal.thy [Finite_def] "Finite(0)"; +Goalw [Finite_def] "Finite(0)"; by (blast_tac (claset() addSIs [eqpoll_refl, nat_0I]) 1); qed "Finite_0"; -goalw Cardinal.thy [Finite_def] +Goalw [Finite_def] "!!A. [| A lepoll n; n:nat |] ==> Finite(A)"; by (etac rev_mp 1); by (etac nat_induct 1); @@ -698,7 +698,7 @@ by (blast_tac (claset() addSDs [lepoll_succ_disj] addSIs [nat_succI]) 1); qed "lepoll_nat_imp_Finite"; -goalw Cardinal.thy [Finite_def] +Goalw [Finite_def] "!!X. [| Y lepoll X; Finite(X) |] ==> Finite(Y)"; by (blast_tac (claset() addSEs [eqpollE] @@ -710,7 +710,7 @@ bind_thm ("Finite_Diff", Diff_subset RS subset_Finite); -goalw Cardinal.thy [Finite_def] "!!x. Finite(x) ==> Finite(cons(y,x))"; +Goalw [Finite_def] "!!x. Finite(x) ==> Finite(cons(y,x))"; by (excluded_middle_tac "y:x" 1); by (asm_simp_tac (simpset() addsimps [cons_absorb]) 2); by (etac bexE 1); @@ -720,18 +720,18 @@ (simpset() addsimps [succ_def, cons_eqpoll_cong, mem_not_refl]) 1); qed "Finite_cons"; -goalw Cardinal.thy [succ_def] "!!x. Finite(x) ==> Finite(succ(x))"; +Goalw [succ_def] "!!x. Finite(x) ==> Finite(succ(x))"; by (etac Finite_cons 1); qed "Finite_succ"; -goalw Cardinal.thy [Finite_def] +Goalw [Finite_def] "!!i. [| Ord(i); ~ Finite(i) |] ==> nat le i"; by (eresolve_tac [Ord_nat RSN (2,Ord_linear2)] 1); by (assume_tac 2); by (blast_tac (claset() addSIs [eqpoll_refl] addSEs [ltE]) 1); qed "nat_le_infinite_Ord"; -goalw Cardinal.thy [Finite_def, eqpoll_def] +Goalw [Finite_def, eqpoll_def] "!!A. Finite(A) ==> EX r. well_ord(A,r)"; by (blast_tac (claset() addIs [well_ord_rvimage, bij_is_inj, well_ord_Memrel, nat_into_Ord]) 1); @@ -753,14 +753,14 @@ by (Blast.depth_tac (claset()) 4 1); qed "nat_wf_on_converse_Memrel"; -goal Cardinal.thy "!!n. n:nat ==> well_ord(n,converse(Memrel(n)))"; +Goal "!!n. n:nat ==> well_ord(n,converse(Memrel(n)))"; by (forward_tac [transfer thy Ord_nat RS Ord_in_Ord RS well_ord_Memrel] 1); by (rewtac well_ord_def); by (blast_tac (claset() addSIs [tot_ord_converse, nat_wf_on_converse_Memrel]) 1); qed "nat_well_ord_converse_Memrel"; -goal Cardinal.thy +Goal "!!A. [| well_ord(A,r); \ \ well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) \ \ |] ==> well_ord(A,converse(r))"; @@ -772,7 +772,7 @@ ordertype_ord_iso RS ord_iso_rvimage_eq]) 1); qed "well_ord_converse"; -goal Cardinal.thy +Goal "!!A. [| well_ord(A,r); A eqpoll n; n:nat |] ==> ordertype(A,r)=n"; by (rtac (Ord_ordertype RS Ord_nat_eqpoll_iff RS iffD1) 1 THEN REPEAT (assume_tac 1)); @@ -781,7 +781,7 @@ by (blast_tac (claset() addSIs [ordermap_bij RS bij_converse_bij]) 1); qed "ordertype_eq_n"; -goalw Cardinal.thy [Finite_def] +Goalw [Finite_def] "!!A. [| Finite(A); well_ord(A,r) |] ==> well_ord(A,converse(r))"; by (rtac well_ord_converse 1 THEN assume_tac 1); by (blast_tac (claset() addDs [ordertype_eq_n] diff -r 30271d90644f -r 62b6288e6005 src/ZF/CardinalArith.ML --- 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", "%."), ("d", "%.")] 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. ")] 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. ")] 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 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 :A*B. ")] 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", "%. if(x=A, Inl(y), Inr())"), ("d", "case(%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 :K*K. ) : 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 \ \ <, > : 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 pred(K*K, , 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 <, > : csquare_rel(K)"; by (subgoals_tac ["x \ \ <, > : csquare_rel(K) | x=z & y=z"; by (subgoals_tac ["x \ \ ordermap(K*K, csquare_rel(K)) ` < \ \ ordermap(K*K, csquare_rel(K)) ` "; @@ -456,7 +456,7 @@ qed "ordermap_z_lt"; (*Kunen: "each : K*K has no more than z*z predecessors..." (page 29) *) -goalw CardinalArith.thy [cmult_def] +Goalw [cmult_def] "!!K. [| Limit(K); x \ \ | ordermap(K*K, csquare_rel(K)) ` | 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 K |*| L = K"; +Goal "!!K. [| InfCard(K); L le K; 0 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 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)); diff -r 30271d90644f -r 62b6288e6005 src/ZF/Cardinal_AC.ML --- a/src/ZF/Cardinal_AC.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Cardinal_AC.ML Mon Jun 22 17:12:27 1998 +0200 @@ -12,38 +12,38 @@ (*** Strengthened versions of existing theorems about cardinals ***) -goal Cardinal_AC.thy "|A| eqpoll A"; +Goal "|A| eqpoll A"; by (resolve_tac [AC_well_ord RS exE] 1); by (etac well_ord_cardinal_eqpoll 1); qed "cardinal_eqpoll"; val cardinal_idem = cardinal_eqpoll RS cardinal_cong; -goal Cardinal_AC.thy "!!X Y. |X| = |Y| ==> X eqpoll Y"; +Goal "!!X Y. |X| = |Y| ==> X eqpoll Y"; by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (rtac well_ord_cardinal_eqE 1); by (REPEAT_SOME assume_tac); qed "cardinal_eqE"; -goal Cardinal_AC.thy "|X| = |Y| <-> X eqpoll Y"; +Goal "|X| = |Y| <-> X eqpoll Y"; by (blast_tac (claset() addIs [cardinal_cong, cardinal_eqE]) 1); qed "cardinal_eqpoll_iff"; -goal Cardinal_AC.thy +Goal "!!A. [| |A|=|B|; |C|=|D|; A Int C = 0; B Int D = 0 |] ==> \ \ |A Un C| = |B Un D|"; by (asm_full_simp_tac (simpset() addsimps [cardinal_eqpoll_iff, eqpoll_disjoint_Un]) 1); qed "cardinal_disjoint_Un"; -goal Cardinal_AC.thy "!!A B. A lepoll B ==> |A| le |B|"; +Goal "!!A B. A lepoll B ==> |A| le |B|"; by (resolve_tac [AC_well_ord RS exE] 1); by (etac well_ord_lepoll_imp_Card_le 1); by (assume_tac 1); qed "lepoll_imp_Card_le"; -goal Cardinal_AC.thy "(i |+| j) |+| k = i |+| (j |+| k)"; +Goal "(i |+| j) |+| k = i |+| (j |+| k)"; by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); @@ -51,7 +51,7 @@ by (REPEAT_SOME assume_tac); qed "cadd_assoc"; -goal Cardinal_AC.thy "(i |*| j) |*| k = i |*| (j |*| k)"; +Goal "(i |*| j) |*| k = i |*| (j |*| k)"; by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); @@ -59,7 +59,7 @@ by (REPEAT_SOME assume_tac); qed "cmult_assoc"; -goal Cardinal_AC.thy "(i |+| j) |*| k = (i |*| k) |+| (j |*| k)"; +Goal "(i |+| j) |*| k = (i |*| k) |+| (j |*| k)"; by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); @@ -67,7 +67,7 @@ by (REPEAT_SOME assume_tac); qed "cadd_cmult_distrib"; -goal Cardinal_AC.thy "!!A. InfCard(|A|) ==> A*A eqpoll A"; +Goal "!!A. InfCard(|A|) ==> A*A eqpoll A"; by (resolve_tac [AC_well_ord RS exE] 1); by (etac well_ord_InfCard_square_eq 1); by (assume_tac 1); @@ -76,20 +76,20 @@ (*** Other applications of AC ***) -goal Cardinal_AC.thy "!!A B. |A| le |B| ==> A lepoll B"; +Goal "!!A B. |A| le |B| ==> A lepoll B"; by (resolve_tac [cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] 1); by (eresolve_tac [le_imp_subset RS subset_imp_lepoll RS lepoll_trans] 1); by (resolve_tac [cardinal_eqpoll RS eqpoll_imp_lepoll] 1); qed "Card_le_imp_lepoll"; -goal Cardinal_AC.thy "!!A K. Card(K) ==> |A| le K <-> A lepoll K"; +Goal "!!A K. Card(K) ==> |A| le K <-> A lepoll K"; by (eresolve_tac [Card_cardinal_eq RS subst] 1 THEN rtac iffI 1 THEN DEPTH_SOLVE (eresolve_tac [Card_le_imp_lepoll,lepoll_imp_Card_le] 1)); qed "le_Card_iff"; -goalw Cardinal_AC.thy [surj_def] "!!f. f: surj(X,Y) ==> EX g. g: inj(Y,X)"; +Goalw [surj_def] "!!f. f: surj(X,Y) ==> EX g. g: inj(Y,X)"; by (etac CollectE 1); by (res_inst_tac [("A1", "Y"), ("B1", "%y. f-``{y}")] (AC_Pi RS exE) 1); by (fast_tac (claset() addSEs [apply_Pair]) 1); @@ -98,7 +98,7 @@ qed "surj_implies_inj"; (*Kunen's Lemma 10.20*) -goal Cardinal_AC.thy "!!f. f: surj(X,Y) ==> |Y| le |X|"; +Goal "!!f. f: surj(X,Y) ==> |Y| le |X|"; by (rtac lepoll_imp_Card_le 1); by (eresolve_tac [surj_implies_inj RS exE] 1); by (rewtac lepoll_def); @@ -106,7 +106,7 @@ qed "surj_implies_cardinal_le"; (*Kunen's Lemma 10.21*) -goal Cardinal_AC.thy +Goal "!!K. [| InfCard(K); ALL i:K. |X(i)| le K |] ==> |UN i:K. X(i)| le K"; by (asm_full_simp_tac (simpset() addsimps [InfCard_is_Card, le_Card_iff]) 1); by (rtac lepoll_trans 1); @@ -134,7 +134,7 @@ qed "cardinal_UN_le"; (*The same again, using csucc*) -goal Cardinal_AC.thy +Goal "!!K. [| InfCard(K); ALL i:K. |X(i)| < csucc(K) |] ==> \ \ |UN i:K. X(i)| < csucc(K)"; by (asm_full_simp_tac @@ -144,7 +144,7 @@ (*The same again, for a union of ordinals. In use, j(i) is a bit like rank(i), the least ordinal j such that i:Vfrom(A,j). *) -goal Cardinal_AC.thy +Goal "!!K. [| InfCard(K); ALL i:K. j(i) < csucc(K) |] ==> \ \ (UN i:K. j(i)) < csucc(K)"; by (resolve_tac [cardinal_UN_lt_csucc RS Card_lt_imp_lt] 1); @@ -178,7 +178,7 @@ (*Simpler to require |W|=K; we'd have a bijection; but the theorem would be weaker.*) -goal Cardinal_AC.thy +Goal "!!K. [| InfCard(K); |W| le K; ALL w:W. j(w) < csucc(K) |] ==> \ \ (UN w:W. j(w)) < csucc(K)"; by (excluded_middle_tac "W=0" 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/Epsilon.ML --- a/src/ZF/Epsilon.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Epsilon.ML Mon Jun 22 17:12:27 1998 +0200 @@ -10,14 +10,14 @@ (*** Basic closure properties ***) -goalw Epsilon.thy [eclose_def] "A <= eclose(A)"; +Goalw [eclose_def] "A <= eclose(A)"; by (rtac (nat_rec_0 RS equalityD2 RS subset_trans) 1); by (rtac (nat_0I RS UN_upper) 1); qed "arg_subset_eclose"; val arg_into_eclose = arg_subset_eclose RS subsetD; -goalw Epsilon.thy [eclose_def,Transset_def] "Transset(eclose(A))"; +Goalw [eclose_def,Transset_def] "Transset(eclose(A))"; by (rtac (subsetI RS ballI) 1); by (etac UN_E 1); by (rtac (nat_succI RS UN_I) 1); @@ -60,7 +60,7 @@ (** eclose(A) is the least transitive set including A as a subset. **) -goalw Epsilon.thy [Transset_def] +Goalw [Transset_def] "!!X A n. [| Transset(X); A<=X; n: nat |] ==> \ \ nat_rec(n, A, %m r. Union(r)) <= X"; by (etac nat_induct 1); @@ -69,7 +69,7 @@ by (Blast_tac 1); qed "eclose_least_lemma"; -goalw Epsilon.thy [eclose_def] +Goalw [eclose_def] "!!X A. [| Transset(X); A<=X |] ==> eclose(A) <= X"; by (rtac (eclose_least_lemma RS UN_least) 1); by (REPEAT (assume_tac 1)); @@ -89,7 +89,7 @@ by (blast_tac (claset() addIs [step,ecloseD]) 1); qed "eclose_induct_down"; -goal Epsilon.thy "!!X. Transset(X) ==> eclose(X) = X"; +Goal "!!X. Transset(X) ==> eclose(X) = X"; by (etac ([eclose_least, arg_subset_eclose] MRS equalityI) 1); by (rtac subset_refl 1); qed "Transset_eclose_eq_arg"; @@ -98,19 +98,19 @@ (*** Epsilon recursion ***) (*Unused...*) -goal Epsilon.thy "!!A B C. [| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)"; +Goal "!!A B C. [| A: eclose(B); B: eclose(C) |] ==> A: eclose(C)"; by (rtac ([Transset_eclose, eclose_subset] MRS eclose_least RS subsetD) 1); by (REPEAT (assume_tac 1)); qed "mem_eclose_trans"; (*Variant of the previous lemma in a useable form for the sequel*) -goal Epsilon.thy +Goal "!!A B C. [| A: eclose({B}); B: eclose({C}) |] ==> A: eclose({C})"; by (rtac ([Transset_eclose, singleton_subsetI] MRS eclose_least RS subsetD) 1); by (REPEAT (assume_tac 1)); qed "mem_eclose_sing_trans"; -goalw Epsilon.thy [Transset_def] +Goalw [Transset_def] "!!i j. [| Transset(i); j:i |] ==> Memrel(i)-``{j} = j"; by (blast_tac (claset() addSIs [MemrelI] addSEs [MemrelE]) 1); qed "under_Memrel"; @@ -136,7 +136,7 @@ by (rtac (prem RS arg_into_eclose_sing) 1); qed "wfrec_eclose_eq2"; -goalw Epsilon.thy [transrec_def] +Goalw [transrec_def] "transrec(a,H) = H(a, lam x:a. transrec(x,H))"; by (rtac wfrec_ssubst 1); by (simp_tac (simpset() addsimps [wfrec_eclose_eq2, arg_in_eclose_sing, @@ -158,7 +158,7 @@ by (REPEAT (ares_tac (prems @ [lam_type]) 1 ORELSE etac bspec 1)); qed "transrec_type"; -goal Epsilon.thy "!!i. Ord(i) ==> eclose({i}) <= succ(i)"; +Goal "!!i. Ord(i) ==> eclose({i}) <= succ(i)"; by (etac (Ord_is_Transset RS Transset_succ RS eclose_least) 1); by (rtac (succI1 RS singleton_subsetI) 1); qed "eclose_sing_Ord"; @@ -176,12 +176,12 @@ (*** Rank ***) (*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) -goal Epsilon.thy "rank(a) = (UN y:a. succ(rank(y)))"; +Goal "rank(a) = (UN y:a. succ(rank(y)))"; by (stac (rank_def RS def_transrec) 1); by (Simp_tac 1); qed "rank"; -goal Epsilon.thy "Ord(rank(a))"; +Goal "Ord(rank(a))"; by (eps_ind_tac "a" 1); by (stac rank 1); by (rtac (Ord_succ RS Ord_UN) 1); @@ -195,7 +195,7 @@ by (asm_simp_tac (simpset() addsimps [Ord_equality]) 1); qed "rank_of_Ord"; -goal Epsilon.thy "!!a b. a:b ==> rank(a) < rank(b)"; +Goal "!!a b. a:b ==> rank(a) < rank(b)"; by (res_inst_tac [("a1","b")] (rank RS ssubst) 1); by (etac (UN_I RS ltI) 1); by (rtac succI1 1); @@ -209,7 +209,7 @@ by (assume_tac 1); qed "eclose_rank_lt"; -goal Epsilon.thy "!!a b. a<=b ==> rank(a) le rank(b)"; +Goal "!!a b. a<=b ==> rank(a) le rank(b)"; by (rtac subset_imp_le 1); by (stac rank 1); by (stac rank 1); @@ -217,7 +217,7 @@ by (REPEAT (resolve_tac [subset_refl, Ord_rank] 1)); qed "rank_mono"; -goal Epsilon.thy "rank(Pow(a)) = succ(rank(a))"; +Goal "rank(Pow(a)) = succ(rank(a))"; by (rtac (rank RS trans) 1); by (rtac le_anti_sym 1); by (DO_GOAL [rtac (Ord_rank RS Ord_succ RS UN_least_le), @@ -226,12 +226,12 @@ REPEAT o rtac (Ord_rank RS Ord_succ)] 1); qed "rank_Pow"; -goal Epsilon.thy "rank(0) = 0"; +Goal "rank(0) = 0"; by (rtac (rank RS trans) 1); by (Blast_tac 1); qed "rank_0"; -goal Epsilon.thy "rank(succ(x)) = succ(rank(x))"; +Goal "rank(succ(x)) = succ(rank(x))"; by (rtac (rank RS trans) 1); by (rtac ([UN_least, succI1 RS UN_upper] MRS equalityI) 1); by (etac succE 1); @@ -239,7 +239,7 @@ by (etac (rank_lt RS leI RS succ_leI RS le_imp_subset) 1); qed "rank_succ"; -goal Epsilon.thy "rank(Union(A)) = (UN x:A. rank(x))"; +Goal "rank(Union(A)) = (UN x:A. rank(x))"; by (rtac equalityI 1); by (rtac (rank_mono RS le_imp_subset RS UN_least) 2); by (etac Union_upper 2); @@ -251,7 +251,7 @@ by (etac (rank_lt RS succ_leI RS le_imp_subset) 1); qed "rank_Union"; -goal Epsilon.thy "rank(eclose(a)) = rank(a)"; +Goal "rank(eclose(a)) = rank(a)"; by (rtac le_anti_sym 1); by (rtac (arg_subset_eclose RS rank_mono) 2); by (res_inst_tac [("a1","eclose(a)")] (rank RS ssubst) 1); @@ -259,24 +259,24 @@ by (etac (eclose_rank_lt RS succ_leI) 1); qed "rank_eclose"; -goalw Epsilon.thy [Pair_def] "rank(a) < rank()"; +Goalw [Pair_def] "rank(a) < rank()"; by (rtac (consI1 RS rank_lt RS lt_trans) 1); by (rtac (consI1 RS consI2 RS rank_lt) 1); qed "rank_pair1"; -goalw Epsilon.thy [Pair_def] "rank(b) < rank()"; +Goalw [Pair_def] "rank(b) < rank()"; by (rtac (consI1 RS consI2 RS rank_lt RS lt_trans) 1); by (rtac (consI1 RS consI2 RS rank_lt) 1); qed "rank_pair2"; (*** Corollaries of leastness ***) -goal Epsilon.thy "!!A B. A:B ==> eclose(A)<=eclose(B)"; +Goal "!!A B. A:B ==> eclose(A)<=eclose(B)"; by (rtac (Transset_eclose RS eclose_least) 1); by (etac (arg_into_eclose RS eclose_subset) 1); qed "mem_eclose_subset"; -goal Epsilon.thy "!!A B. A<=B ==> eclose(A) <= eclose(B)"; +Goal "!!A B. A<=B ==> eclose(A) <= eclose(B)"; by (rtac (Transset_eclose RS eclose_least) 1); by (etac subset_trans 1); by (rtac arg_subset_eclose 1); @@ -284,7 +284,7 @@ (** Idempotence of eclose **) -goal Epsilon.thy "eclose(eclose(A)) = eclose(A)"; +Goal "eclose(eclose(A)) = eclose(A)"; by (rtac equalityI 1); by (rtac ([Transset_eclose, subset_refl] MRS eclose_least) 1); by (rtac arg_subset_eclose 1); @@ -293,23 +293,23 @@ (*Transfinite recursion for definitions based on the three cases of ordinals. *) -goal thy "transrec2(0,a,b) = a"; +Goal "transrec2(0,a,b) = a"; by (rtac (transrec2_def RS def_transrec RS trans) 1); by (Simp_tac 1); qed "transrec2_0"; -goal thy "(THE j. i=j) = i"; +Goal "(THE j. i=j) = i"; by (blast_tac (claset() addSIs [the_equality]) 1); qed "THE_eq"; -goal thy "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))"; +Goal "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))"; by (rtac (transrec2_def RS def_transrec RS trans) 1); by (simp_tac (simpset() addsimps [succ_not_0, THE_eq, if_P] setloop split_tac [expand_if]) 1); by (Blast_tac 1); qed "transrec2_succ"; -goal thy "!!i. Limit(i) ==> transrec2(i,a,b) = (UN j transrec2(i,a,b) = (UN j converse(r) O r = r **) -goalw EquivClass.thy [trans_def,sym_def] +Goalw [trans_def,sym_def] "!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r"; by (Blast_tac 1); qed "sym_trans_comp_subset"; -goalw EquivClass.thy [refl_def] +Goalw [refl_def] "!!A r. [| refl(A,r); r <= A*A |] ==> r <= converse(r) O r"; by (Blast_tac 1); qed "refl_comp_subset"; -goalw EquivClass.thy [equiv_def] +Goalw [equiv_def] "!!A r. 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 EquivClass.thy [equiv_def,refl_def,sym_def,trans_def] +Goalw [equiv_def,refl_def,sym_def,trans_def] "!!A r. [| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)"; by (etac equalityE 1); by (subgoal_tac "ALL x y. : r --> : r" 1); @@ -40,25 +40,25 @@ (** Equivalence classes **) (*Lemma for the next result*) -goalw EquivClass.thy [trans_def,sym_def] +Goalw [trans_def,sym_def] "!!A r. [| sym(r); trans(r); : r |] ==> r``{a} <= r``{b}"; by (Blast_tac 1); qed "equiv_class_subset"; -goalw EquivClass.thy [equiv_def] +Goalw [equiv_def] "!!A r. [| equiv(A,r); : 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 EquivClass.thy [equiv_def,refl_def] +Goalw [equiv_def,refl_def] "!!A r. [| equiv(A,r); a: A |] ==> a: r``{a}"; by (Blast_tac 1); qed "equiv_class_self"; (*Lemma for the next result*) -goalw EquivClass.thy [equiv_def,refl_def] +Goalw [equiv_def,refl_def] "!!A r. [| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> : r"; by (Blast_tac 1); qed "subset_equiv_class"; @@ -69,22 +69,22 @@ qed "eq_equiv_class"; (*thus r``{a} = r``{b} as well*) -goalw EquivClass.thy [equiv_def,trans_def,sym_def] +Goalw [equiv_def,trans_def,sym_def] "!!A r. [| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> : r"; by (Blast_tac 1); qed "equiv_class_nondisjoint"; -goalw EquivClass.thy [equiv_def] "!!A r. equiv(A,r) ==> r <= A*A"; +Goalw [equiv_def] "!!A r. equiv(A,r) ==> r <= A*A"; by (safe_tac subset_cs); qed "equiv_type"; -goal EquivClass.thy +Goal "!!A r. equiv(A,r) ==> : 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 EquivClass.thy +Goal "!!A r. [| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> : r"; by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq] addDs [equiv_type]) 1); @@ -107,12 +107,12 @@ by (assume_tac 1); qed "quotientE"; -goalw EquivClass.thy [equiv_def,refl_def,quotient_def] +Goalw [equiv_def,refl_def,quotient_def] "!!A r. equiv(A,r) ==> Union(A/r) = A"; by (Blast_tac 1); qed "Union_quotient"; -goalw EquivClass.thy [quotient_def] +Goalw [quotient_def] "!!A r. [| 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); @@ -167,7 +167,7 @@ (**** Defining binary operations upon equivalence classes ****) -goalw EquivClass.thy [congruent_def,congruent2_def,equiv_def,refl_def] +Goalw [congruent_def,congruent2_def,equiv_def,refl_def] "!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))"; by (Blast_tac 1); qed "congruent2_implies_congruent"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Finite.ML --- a/src/ZF/Finite.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Finite.ML Mon Jun 22 17:12:27 1998 +0200 @@ -14,7 +14,7 @@ (*** Finite powerset operator ***) -goalw Finite.thy Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)"; +Goalw Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)"; by (rtac lfp_mono 1); by (REPEAT (rtac Fin.bnd_mono 1)); by (REPEAT (ares_tac (Pow_mono::basic_monos) 1)); @@ -41,7 +41,7 @@ Addsimps Fin.intrs; (*The union of two finite sets is finite.*) -goal Finite.thy +Goal "!!b c. [| b: Fin(A); c: Fin(A) |] ==> b Un c : Fin(A)"; by (etac Fin_induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [Un_cons]))); @@ -56,7 +56,7 @@ qed "Fin_UnionI"; (*Every subset of a finite set is finite.*) -goal Finite.thy "!!b A. b: Fin(A) ==> ALL z. z<=b --> z: Fin(A)"; +Goal "!!b A. b: Fin(A) ==> ALL z. z<=b --> z: Fin(A)"; by (etac Fin_induct 1); by (simp_tac (simpset() addsimps [subset_empty_iff]) 1); by (asm_simp_tac (simpset() addsimps subset_cons_iff::distrib_simps) 1); @@ -65,7 +65,7 @@ by (Asm_simp_tac 1); qed "Fin_subset_lemma"; -goal Finite.thy "!!c b A. [| c<=b; b: Fin(A) |] ==> c: Fin(A)"; +Goal "!!c b A. [| c<=b; b: Fin(A) |] ==> c: Fin(A)"; by (REPEAT (ares_tac [Fin_subset_lemma RS spec RS mp] 1)); qed "Fin_subset"; @@ -101,24 +101,24 @@ (*** Finite function space ***) -goalw Finite.thy FiniteFun.defs +Goalw FiniteFun.defs "!!A B C D. [| A<=C; B<=D |] ==> A -||> B <= C -||> D"; by (rtac lfp_mono 1); by (REPEAT (rtac FiniteFun.bnd_mono 1)); by (REPEAT (ares_tac (Fin_mono::Sigma_mono::basic_monos) 1)); qed "FiniteFun_mono"; -goal Finite.thy "!!A B. A<=B ==> A -||> A <= B -||> B"; +Goal "!!A B. A<=B ==> A -||> A <= B -||> B"; by (REPEAT (ares_tac [FiniteFun_mono] 1)); qed "FiniteFun_mono1"; -goal Finite.thy "!!h. h: A -||>B ==> h: domain(h) -> B"; +Goal "!!h. h: A -||>B ==> h: domain(h) -> B"; by (etac FiniteFun.induct 1); by (simp_tac (simpset() addsimps [empty_fun, domain_0]) 1); by (asm_simp_tac (simpset() addsimps [fun_extend3, domain_cons]) 1); qed "FiniteFun_is_fun"; -goal Finite.thy "!!h. h: A -||>B ==> domain(h) : Fin(A)"; +Goal "!!h. h: A -||>B ==> domain(h) : Fin(A)"; by (etac FiniteFun.induct 1); by (simp_tac (simpset() addsimps [domain_0]) 1); by (asm_simp_tac (simpset() addsimps [domain_cons]) 1); @@ -127,7 +127,7 @@ bind_thm ("FiniteFun_apply_type", FiniteFun_is_fun RS apply_type); (*Every subset of a finite function is a finite function.*) -goal Finite.thy "!!b A. b: A-||>B ==> ALL z. z<=b --> z: A-||>B"; +Goal "!!b A. b: A-||>B ==> ALL z. z<=b --> z: A-||>B"; by (etac FiniteFun.induct 1); by (simp_tac (simpset() addsimps subset_empty_iff::FiniteFun.intrs) 1); by (asm_simp_tac (simpset() addsimps subset_cons_iff::distrib_simps) 1); @@ -137,7 +137,7 @@ by (fast_tac (claset() addSIs FiniteFun.intrs) 1); qed "FiniteFun_subset_lemma"; -goal Finite.thy "!!c b A. [| c<=b; b: A-||>B |] ==> c: A-||>B"; +Goal "!!c b A. [| c<=b; b: A-||>B |] ==> c: A-||>B"; by (REPEAT (ares_tac [FiniteFun_subset_lemma RS spec RS mp] 1)); qed "FiniteFun_subset"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Fixedpt.ML --- a/src/ZF/Fixedpt.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Fixedpt.ML Mon Jun 22 17:12:27 1998 +0200 @@ -38,14 +38,14 @@ by (rtac subset_refl 1); qed "bnd_mono_subset"; -goal Fixedpt.thy "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ +Goal "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ \ h(A) Un h(B) <= h(A Un B)"; by (REPEAT (ares_tac [Un_upper1, Un_upper2, Un_least] 1 ORELSE etac bnd_monoD2 1)); qed "bnd_mono_Un"; (*Useful??*) -goal Fixedpt.thy "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ +Goal "!!A B. [| bnd_mono(D,h); A <= D; B <= D |] ==> \ \ h(A Int B) <= h(A) Int h(B)"; by (REPEAT (ares_tac [Int_lower1, Int_lower2, Int_greatest] 1 ORELSE etac bnd_monoD2 1)); @@ -54,14 +54,14 @@ (**** Proof of Knaster-Tarski Theorem for the lfp ****) (*lfp is contained in each pre-fixedpoint*) -goalw Fixedpt.thy [lfp_def] +Goalw [lfp_def] "!!A. [| h(A) <= A; A<=D |] ==> lfp(D,h) <= A"; by (Blast_tac 1); (*or by (rtac (PowI RS CollectI RS Inter_lower) 1); *) qed "lfp_lowerbound"; (*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*) -goalw Fixedpt.thy [lfp_def,Inter_def] "lfp(D,h) <= D"; +Goalw [lfp_def,Inter_def] "lfp(D,h) <= D"; by (Blast_tac 1); qed "lfp_subset"; @@ -189,7 +189,7 @@ by (REPEAT (resolve_tac prems 1)); qed "gfp_upperbound"; -goalw Fixedpt.thy [gfp_def] "gfp(D,h) <= D"; +Goalw [gfp_def] "gfp(D,h) <= D"; by (Blast_tac 1); qed "gfp_subset"; @@ -244,7 +244,7 @@ (*** Coinduction rules for greatest fixed points ***) (*weak version*) -goal Fixedpt.thy "!!X h. [| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)"; +Goal "!!X h. [| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)"; by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1)); qed "weak_coinduct"; @@ -258,7 +258,7 @@ qed "coinduct_lemma"; (*strong version*) -goal Fixedpt.thy +Goal "!!X D. [| bnd_mono(D,h); a: X; X <= h(X Un gfp(D,h)); X <= D |] ==> \ \ a : gfp(D,h)"; by (rtac weak_coinduct 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/InfDatatype.ML --- a/src/ZF/InfDatatype.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/InfDatatype.ML Mon Jun 22 17:12:27 1998 +0200 @@ -11,7 +11,7 @@ transfer InfDatatype.thy Limit_VfromE |> standard; -goal InfDatatype.thy +Goal "!!K. [| f: W -> Vfrom(A,csucc(K)); |W| le K; InfCard(K) \ \ |] ==> EX j. f: W -> Vfrom(A,j) & j < csucc(K)"; by (res_inst_tac [("x", "UN w:W. LEAST i. f`w : Vfrom(A,i)")] exI 1); @@ -29,14 +29,14 @@ by (assume_tac 1); qed "fun_Vcsucc_lemma"; -goal InfDatatype.thy +Goal "!!K. [| W <= Vfrom(A,csucc(K)); |W| le K; InfCard(K) \ \ |] ==> EX j. W <= Vfrom(A,j) & j < csucc(K)"; by (asm_full_simp_tac (simpset() addsimps [subset_iff_id, fun_Vcsucc_lemma]) 1); qed "subset_Vcsucc"; (*Version for arbitrary index sets*) -goal InfDatatype.thy +Goal "!!K. [| |W| le K; InfCard(K); W <= Vfrom(A,csucc(K)) |] ==> \ \ W -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"; by (safe_tac (claset() addSDs [fun_Vcsucc_lemma, subset_Vcsucc])); @@ -50,7 +50,7 @@ Limit_has_succ, Un_least_lt] 1)); qed "fun_Vcsucc"; -goal InfDatatype.thy +Goal "!!K. [| f: W -> Vfrom(A, csucc(K)); |W| le K; InfCard(K); \ \ W <= Vfrom(A,csucc(K)) \ \ |] ==> f: Vfrom(A,csucc(K))"; @@ -62,7 +62,7 @@ (** Version where K itself is the index set **) -goal InfDatatype.thy +Goal "!!K. InfCard(K) ==> K -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"; by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); by (REPEAT (ares_tac [fun_Vcsucc, Ord_cardinal_le, @@ -70,7 +70,7 @@ lt_csucc RS leI RS le_imp_subset RS subset_trans] 1)); qed "Card_fun_Vcsucc"; -goal InfDatatype.thy +Goal "!!K. [| f: K -> Vfrom(A, csucc(K)); InfCard(K) \ \ |] ==> f: Vfrom(A,csucc(K))"; by (REPEAT (ares_tac [Card_fun_Vcsucc RS subsetD] 1)); diff -r 30271d90644f -r 62b6288e6005 src/ZF/List.ML --- a/src/ZF/List.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/List.ML Mon Jun 22 17:12:27 1998 +0200 @@ -26,7 +26,7 @@ rename_last_tac a ["1"] (i+2), ares_tac prems i]; -goal List.thy "list(A) = {0} + (A * list(A))"; +Goal "list(A) = {0} + (A * list(A))"; let open list; val rew = rewrite_rule con_defs in by (blast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1) end; @@ -34,14 +34,14 @@ (** Lemmas to justify using "list" in other recursive type definitions **) -goalw List.thy list.defs "!!A B. A<=B ==> list(A) <= list(B)"; +Goalw list.defs "!!A B. A<=B ==> list(A) <= list(B)"; by (rtac lfp_mono 1); by (REPEAT (rtac list.bnd_mono 1)); by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); qed "list_mono"; (*There is a similar proof by list induction.*) -goalw List.thy (list.defs@list.con_defs) "list(univ(A)) <= univ(A)"; +Goalw (list.defs@list.con_defs) "list(univ(A)) <= univ(A)"; by (rtac lfp_lowerbound 1); by (rtac (A_subset_univ RS univ_mono) 2); by (blast_tac (claset() addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, @@ -51,7 +51,7 @@ (*These two theorems justify datatypes involving list(nat), list(A), ...*) bind_thm ("list_subset_univ", ([list_mono, list_univ] MRS subset_trans)); -goal List.thy "!!l A B. [| l: list(A); A <= univ(B) |] ==> l: univ(B)"; +Goal "!!l A B. [| l: list(A); A <= univ(B) |] ==> l: univ(B)"; by (REPEAT (ares_tac [list_subset_univ RS subsetD] 1)); qed "list_into_univ"; @@ -67,11 +67,11 @@ (** For recursion **) -goalw List.thy list.con_defs "rank(a) < rank(Cons(a,l))"; +Goalw list.con_defs "rank(a) < rank(Cons(a,l))"; by (simp_tac rank_ss 1); qed "rank_Cons1"; -goalw List.thy list.con_defs "rank(l) < rank(Cons(a,l))"; +Goalw list.con_defs "rank(l) < rank(Cons(a,l))"; by (simp_tac rank_ss 1); qed "rank_Cons2"; @@ -80,37 +80,37 @@ (** hd and tl **) -goalw List.thy [hd_def] "hd(Cons(a,l)) = a"; +Goalw [hd_def] "hd(Cons(a,l)) = a"; by (resolve_tac list.case_eqns 1); qed "hd_Cons"; -goalw List.thy [tl_def] "tl(Nil) = Nil"; +Goalw [tl_def] "tl(Nil) = Nil"; by (resolve_tac list.case_eqns 1); qed "tl_Nil"; -goalw List.thy [tl_def] "tl(Cons(a,l)) = l"; +Goalw [tl_def] "tl(Cons(a,l)) = l"; by (resolve_tac list.case_eqns 1); qed "tl_Cons"; Addsimps [hd_Cons, tl_Nil, tl_Cons]; -goal List.thy "!!l. l: list(A) ==> tl(l) : list(A)"; +Goal "!!l. l: list(A) ==> tl(l) : list(A)"; by (etac list.elim 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps list.intrs))); qed "tl_type"; (** drop **) -goalw List.thy [drop_def] "drop(0, l) = l"; +Goalw [drop_def] "drop(0, l) = l"; by (rtac rec_0 1); qed "drop_0"; -goalw List.thy [drop_def] "!!i. i:nat ==> drop(i, Nil) = Nil"; +Goalw [drop_def] "!!i. i:nat ==> drop(i, Nil) = Nil"; by (etac nat_induct 1); by (ALLGOALS Asm_simp_tac); qed "drop_Nil"; -goalw List.thy [drop_def] +Goalw [drop_def] "!!i. i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)"; by (rtac sym 1); by (etac nat_induct 1); @@ -120,7 +120,7 @@ Addsimps [drop_0, drop_Nil, drop_succ_Cons]; -goalw List.thy [drop_def] +Goalw [drop_def] "!!i l. [| i:nat; l: list(A) |] ==> drop(i,l) : list(A)"; by (etac nat_induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [tl_type]))); @@ -128,12 +128,12 @@ (** list_rec -- by Vset recursion **) -goal List.thy "list_rec(Nil,c,h) = c"; +Goal "list_rec(Nil,c,h) = c"; by (rtac (list_rec_def RS def_Vrec RS trans) 1); by (simp_tac (simpset() addsimps list.case_eqns) 1); qed "list_rec_Nil"; -goal List.thy "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))"; +Goal "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))"; by (rtac (list_rec_def RS def_Vrec RS trans) 1); by (simp_tac (rank_ss addsimps (rank_Cons2::list.case_eqns)) 1); qed "list_rec_Cons"; @@ -188,7 +188,7 @@ val [length_Nil,length_Cons] = list_recs length_def; Addsimps [length_Nil,length_Cons]; -goalw List.thy [length_def] +Goalw [length_def] "!!l. l: list(A) ==> length(l) : nat"; by (REPEAT (ares_tac [list_rec_type, nat_0I, nat_succI] 1)); qed "length_type"; @@ -198,7 +198,7 @@ val [app_Nil,app_Cons] = list_recs app_def; Addsimps [app_Nil, app_Cons]; -goalw List.thy [app_def] +Goalw [app_def] "!!xs ys. [| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)"; by (REPEAT (ares_tac [list_rec_type, list.Cons_I] 1)); qed "app_type"; @@ -208,7 +208,7 @@ val [rev_Nil,rev_Cons] = list_recs rev_def; Addsimps [rev_Nil,rev_Cons] ; -goalw List.thy [rev_def] +Goalw [rev_def] "!!xs. xs: list(A) ==> rev(xs) : list(A)"; by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1)); qed "rev_type"; @@ -219,7 +219,7 @@ val [flat_Nil,flat_Cons] = list_recs flat_def; Addsimps [flat_Nil,flat_Cons]; -goalw List.thy [flat_def] +Goalw [flat_def] "!!ls. ls: list(list(A)) ==> flat(ls) : list(A)"; by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1)); qed "flat_type"; @@ -230,13 +230,13 @@ val [set_of_list_Nil,set_of_list_Cons] = list_recs set_of_list_def; Addsimps [set_of_list_Nil,set_of_list_Cons]; -goalw List.thy [set_of_list_def] +Goalw [set_of_list_def] "!!l. l: list(A) ==> set_of_list(l) : Pow(A)"; by (etac list_rec_type 1); by (ALLGOALS (Blast_tac)); qed "set_of_list_type"; -goal List.thy +Goal "!!l. xs: list(A) ==> \ \ set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)"; by (etac list.induct 1); @@ -249,7 +249,7 @@ val [list_add_Nil,list_add_Cons] = list_recs list_add_def; Addsimps [list_add_Nil,list_add_Cons]; -goalw List.thy [list_add_def] +Goalw [list_add_def] "!!xs. xs: list(nat) ==> list_add(xs) : nat"; by (REPEAT (ares_tac [list_rec_type, nat_0I, add_type] 1)); qed "list_add_type"; @@ -340,7 +340,7 @@ (** Length and drop **) (*Lemma for the inductive step of drop_length*) -goal List.thy +Goal "!!xs. xs: list(A) ==> \ \ ALL x. EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)"; by (etac list.induct 1); @@ -349,7 +349,7 @@ qed "drop_length_Cons_lemma"; bind_thm ("drop_length_Cons", (drop_length_Cons_lemma RS spec)); -goal List.thy +Goal "!!l. l: list(A) ==> ALL i: length(l). EX z zs. drop(i,l) = Cons(z,zs)"; by (etac list.induct 1); by (ALLGOALS Asm_simp_tac); @@ -394,7 +394,7 @@ instantiate the variable ?A in the rules' typing conditions; note that rev_type does not instantiate ?A. Only the premises do. *) -goal List.thy +Goal "!!xs. [| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)"; by (etac list.induct 1); by (ALLGOALS (asm_simp_tac (simpset() addsimps [app_right_Nil,app_assoc]))); diff -r 30271d90644f -r 62b6288e6005 src/ZF/Nat.ML --- a/src/ZF/Nat.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Nat.ML Mon Jun 22 17:12:27 1998 +0200 @@ -8,7 +8,7 @@ open Nat; -goal Nat.thy "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"; +Goal "bnd_mono(Inf, %X. {0} Un {succ(i). i:X})"; by (rtac bnd_monoI 1); by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2)); by (cut_facts_tac [infinity] 1); @@ -20,7 +20,7 @@ (** Type checking of 0 and successor **) -goal Nat.thy "0 : nat"; +Goal "0 : nat"; by (stac nat_unfold 1); by (rtac (singletonI RS UnI1) 1); qed "nat_0I"; @@ -31,18 +31,18 @@ by (resolve_tac prems 1); qed "nat_succI"; -goal Nat.thy "1 : nat"; +Goal "1 : nat"; by (rtac (nat_0I RS nat_succI) 1); qed "nat_1I"; -goal Nat.thy "2 : nat"; +Goal "2 : nat"; by (rtac (nat_1I RS nat_succI) 1); qed "nat_2I"; Addsimps [nat_0I, nat_1I, nat_2I, nat_succI]; AddSIs [nat_0I, nat_1I, nat_2I]; -goal Nat.thy "bool <= nat"; +Goal "bool <= nat"; by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 ORELSE eresolve_tac [boolE,ssubst] 1)); qed "bool_subset_nat"; @@ -83,7 +83,7 @@ (* i: nat ==> i le i; same thing as i nat le i"; +Goal "!!i. Limit(i) ==> nat le i"; by (rtac subset_imp_le 1); by (rtac subsetI 1); by (etac nat_induct 1); @@ -115,7 +115,7 @@ (* [| succ(i): k; k: nat |] ==> i: k *) val succ_in_naturalD = [succI1, asm_rl, nat_into_Ord] MRS Ord_trans; -goal Nat.thy "!!m n. [| m m: nat"; +Goal "!!m n. [| m m: nat"; by (etac ltE 1); by (etac (Ord_nat RSN (3,Ord_trans)) 1); by (assume_tac 1); @@ -164,7 +164,7 @@ (** Induction principle analogous to trancl_induct **) -goal Nat.thy +Goal "!!m. m: nat ==> P(m,succ(m)) --> (ALL x: nat. P(m,x) --> P(m,succ(x))) --> \ \ (ALL n:nat. m P(m,n))"; by (etac nat_induct 1); @@ -185,11 +185,11 @@ (** nat_case **) -goalw Nat.thy [nat_case_def] "nat_case(a,b,0) = a"; +Goalw [nat_case_def] "nat_case(a,b,0) = a"; by (blast_tac (claset() addIs [the_equality]) 1); qed "nat_case_0"; -goalw Nat.thy [nat_case_def] "nat_case(a,b,succ(m)) = b(m)"; +Goalw [nat_case_def] "nat_case(a,b,succ(m)) = b(m)"; by (blast_tac (claset() addIs [the_equality]) 1); qed "nat_case_succ"; @@ -208,7 +208,7 @@ val nat_rec_trans = wf_Memrel RS (nat_rec_def RS def_wfrec RS trans); -goal Nat.thy "nat_rec(0,a,b) = a"; +Goal "nat_rec(0,a,b) = a"; by (rtac nat_rec_trans 1); by (rtac nat_case_0 1); qed "nat_rec_0"; @@ -221,12 +221,12 @@ (** The union of two natural numbers is a natural number -- their maximum **) -goal Nat.thy "!!i j. [| i: nat; j: nat |] ==> i Un j: nat"; +Goal "!!i j. [| i: nat; j: nat |] ==> i Un j: nat"; by (rtac (Un_least_lt RS ltD) 1); by (REPEAT (ares_tac [ltI, Ord_nat] 1)); qed "Un_nat_type"; -goal Nat.thy "!!i j. [| i: nat; j: nat |] ==> i Int j: nat"; +Goal "!!i j. [| i: nat; j: nat |] ==> i Int j: nat"; by (rtac (Int_greatest_lt RS ltD) 1); by (REPEAT (ares_tac [ltI, Ord_nat] 1)); qed "Int_nat_type"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Order.ML --- a/src/ZF/Order.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Order.ML Mon Jun 22 17:12:27 1998 +0200 @@ -14,7 +14,7 @@ (** Basic properties of the definitions **) (*needed?*) -goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, asym_def] +Goalw [part_ord_def, irrefl_def, trans_on_def, asym_def] "!!r. part_ord(A,r) ==> asym(r Int A*A)"; by (Blast_tac 1); qed "part_ord_Imp_asym"; @@ -35,24 +35,24 @@ (** General properties of well_ord **) -goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def, +Goalw [irrefl_def, part_ord_def, tot_ord_def, trans_on_def, well_ord_def] "!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)"; by (asm_simp_tac (simpset() addsimps [wf_on_not_refl]) 1); by (fast_tac (claset() addEs [linearE, wf_on_asym, wf_on_chain3]) 1); qed "well_ordI"; -goalw Order.thy [well_ord_def] +Goalw [well_ord_def] "!!r. well_ord(A,r) ==> wf[A](r)"; by Safe_tac; qed "well_ord_is_wf"; -goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def] +Goalw [well_ord_def, tot_ord_def, part_ord_def] "!!r. well_ord(A,r) ==> trans[A](r)"; by Safe_tac; qed "well_ord_is_trans_on"; -goalw Order.thy [well_ord_def, tot_ord_def] +Goalw [well_ord_def, tot_ord_def] "!!r. well_ord(A,r) ==> linear(A,r)"; by (Blast_tac 1); qed "well_ord_is_linear"; @@ -60,7 +60,7 @@ (** Derived rules for pred(A,x,r) **) -goalw Order.thy [pred_def] "y : pred(A,x,r) <-> :r & y:A"; +Goalw [pred_def] "y : pred(A,x,r) <-> :r & y:A"; by (Blast_tac 1); qed "pred_iff"; @@ -72,20 +72,20 @@ by (REPEAT (ares_tac [minor] 1)); qed "predE"; -goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}"; +Goalw [pred_def] "pred(A,x,r) <= r -`` {x}"; by (Blast_tac 1); qed "pred_subset_under"; -goalw Order.thy [pred_def] "pred(A,x,r) <= A"; +Goalw [pred_def] "pred(A,x,r) <= A"; by (Blast_tac 1); qed "pred_subset"; -goalw Order.thy [pred_def] +Goalw [pred_def] "pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)"; by (Blast_tac 1); qed "pred_pred_eq"; -goalw Order.thy [trans_on_def, pred_def] +Goalw [trans_on_def, pred_def] "!!r. [| trans[A](r); :r; x:A; y:A \ \ |] ==> pred(pred(A,x,r), y, r) = pred(A,y,r)"; by (Blast_tac 1); @@ -96,22 +96,22 @@ [including initial segments of the form pred(A,x,r) **) (*Note: a relation s such that s<=r need not be a partial ordering*) -goalw Order.thy [part_ord_def, irrefl_def, trans_on_def] +Goalw [part_ord_def, irrefl_def, trans_on_def] "!!A B r. [| part_ord(A,r); B<=A |] ==> part_ord(B,r)"; by (Blast_tac 1); qed "part_ord_subset"; -goalw Order.thy [linear_def] +Goalw [linear_def] "!!A B r. [| linear(A,r); B<=A |] ==> linear(B,r)"; by (Blast_tac 1); qed "linear_subset"; -goalw Order.thy [tot_ord_def] +Goalw [tot_ord_def] "!!A B r. [| tot_ord(A,r); B<=A |] ==> tot_ord(B,r)"; by (fast_tac (claset() addSEs [part_ord_subset, linear_subset]) 1); qed "tot_ord_subset"; -goalw Order.thy [well_ord_def] +Goalw [well_ord_def] "!!A B r. [| well_ord(A,r); B<=A |] ==> well_ord(B,r)"; by (fast_tac (claset() addSEs [tot_ord_subset, wf_on_subset_A]) 1); qed "well_ord_subset"; @@ -119,74 +119,74 @@ (** Relations restricted to a smaller domain, by Krzysztof Grabczewski **) -goalw Order.thy [irrefl_def] "irrefl(A,r Int A*A) <-> irrefl(A,r)"; +Goalw [irrefl_def] "irrefl(A,r Int A*A) <-> irrefl(A,r)"; by (Blast_tac 1); qed "irrefl_Int_iff"; -goalw Order.thy [trans_on_def] "trans[A](r Int A*A) <-> trans[A](r)"; +Goalw [trans_on_def] "trans[A](r Int A*A) <-> trans[A](r)"; by (Blast_tac 1); qed "trans_on_Int_iff"; -goalw Order.thy [part_ord_def] "part_ord(A,r Int A*A) <-> part_ord(A,r)"; +Goalw [part_ord_def] "part_ord(A,r Int A*A) <-> part_ord(A,r)"; by (simp_tac (simpset() addsimps [irrefl_Int_iff, trans_on_Int_iff]) 1); qed "part_ord_Int_iff"; -goalw Order.thy [linear_def] "linear(A,r Int A*A) <-> linear(A,r)"; +Goalw [linear_def] "linear(A,r Int A*A) <-> linear(A,r)"; by (Blast_tac 1); qed "linear_Int_iff"; -goalw Order.thy [tot_ord_def] "tot_ord(A,r Int A*A) <-> tot_ord(A,r)"; +Goalw [tot_ord_def] "tot_ord(A,r Int A*A) <-> tot_ord(A,r)"; by (simp_tac (simpset() addsimps [part_ord_Int_iff, linear_Int_iff]) 1); qed "tot_ord_Int_iff"; -goalw Order.thy [wf_on_def, wf_def] "wf[A](r Int A*A) <-> wf[A](r)"; +Goalw [wf_on_def, wf_def] "wf[A](r Int A*A) <-> wf[A](r)"; by (Fast_tac 1); (*10 times faster than Blast_tac!*) qed "wf_on_Int_iff"; -goalw Order.thy [well_ord_def] "well_ord(A,r Int A*A) <-> well_ord(A,r)"; +Goalw [well_ord_def] "well_ord(A,r Int A*A) <-> well_ord(A,r)"; by (simp_tac (simpset() addsimps [tot_ord_Int_iff, wf_on_Int_iff]) 1); qed "well_ord_Int_iff"; (** Relations over the Empty Set **) -goalw Order.thy [irrefl_def] "irrefl(0,r)"; +Goalw [irrefl_def] "irrefl(0,r)"; by (Blast_tac 1); qed "irrefl_0"; -goalw Order.thy [trans_on_def] "trans[0](r)"; +Goalw [trans_on_def] "trans[0](r)"; by (Blast_tac 1); qed "trans_on_0"; -goalw Order.thy [part_ord_def] "part_ord(0,r)"; +Goalw [part_ord_def] "part_ord(0,r)"; by (simp_tac (simpset() addsimps [irrefl_0, trans_on_0]) 1); qed "part_ord_0"; -goalw Order.thy [linear_def] "linear(0,r)"; +Goalw [linear_def] "linear(0,r)"; by (Blast_tac 1); qed "linear_0"; -goalw Order.thy [tot_ord_def] "tot_ord(0,r)"; +Goalw [tot_ord_def] "tot_ord(0,r)"; by (simp_tac (simpset() addsimps [part_ord_0, linear_0]) 1); qed "tot_ord_0"; -goalw Order.thy [wf_on_def, wf_def] "wf[0](r)"; +Goalw [wf_on_def, wf_def] "wf[0](r)"; by (Blast_tac 1); qed "wf_on_0"; -goalw Order.thy [well_ord_def] "well_ord(0,r)"; +Goalw [well_ord_def] "well_ord(0,r)"; by (simp_tac (simpset() addsimps [tot_ord_0, wf_on_0]) 1); qed "well_ord_0"; (** Order-preserving (monotone) maps **) -goalw Order.thy [mono_map_def] +Goalw [mono_map_def] "!!f. f: mono_map(A,r,B,s) ==> f: A->B"; by (etac CollectD1 1); qed "mono_map_is_fun"; -goalw Order.thy [mono_map_def, inj_def] +Goalw [mono_map_def, inj_def] "!!f. [| linear(A,r); wf[B](s); f: mono_map(A,r,B,s) |] ==> f: inj(A,B)"; by (Clarify_tac 1); by (linear_case_tac 1); @@ -207,24 +207,24 @@ by (blast_tac (claset() addSIs prems) 1); qed "ord_isoI"; -goalw Order.thy [ord_iso_def, mono_map_def] +Goalw [ord_iso_def, mono_map_def] "!!f. f: ord_iso(A,r,B,s) ==> f: mono_map(A,r,B,s)"; by (blast_tac (claset() addSDs [bij_is_fun]) 1); qed "ord_iso_is_mono_map"; -goalw Order.thy [ord_iso_def] +Goalw [ord_iso_def] "!!f. f: ord_iso(A,r,B,s) ==> f: bij(A,B)"; by (etac CollectD1 1); qed "ord_iso_is_bij"; (*Needed? But ord_iso_converse is!*) -goalw Order.thy [ord_iso_def] +Goalw [ord_iso_def] "!!f. [| f: ord_iso(A,r,B,s); : r; x:A; y:A |] ==> \ \ : s"; by (Blast_tac 1); qed "ord_iso_apply"; -goalw Order.thy [ord_iso_def] +Goalw [ord_iso_def] "!!f. [| f: ord_iso(A,r,B,s); : s; x:B; y:B |] ==> \ \ : r"; by (etac CollectE 1); @@ -244,26 +244,26 @@ (** Symmetry and Transitivity Rules **) (*Reflexivity of similarity*) -goal Order.thy "id(A): ord_iso(A,r,A,r)"; +Goal "id(A): ord_iso(A,r,A,r)"; by (resolve_tac [id_bij RS ord_isoI] 1); by (Asm_simp_tac 1); qed "ord_iso_refl"; (*Symmetry of similarity*) -goalw Order.thy [ord_iso_def] +Goalw [ord_iso_def] "!!f. f: ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)"; by (fast_tac (claset() addss bij_inverse_ss) 1); qed "ord_iso_sym"; (*Transitivity of similarity*) -goalw Order.thy [mono_map_def] +Goalw [mono_map_def] "!!f. [| g: mono_map(A,r,B,s); f: mono_map(B,s,C,t) |] ==> \ \ (f O g): mono_map(A,r,C,t)"; by (fast_tac (claset() addss bij_inverse_ss) 1); qed "mono_map_trans"; (*Transitivity of similarity: the order-isomorphism relation*) -goalw Order.thy [ord_iso_def] +Goalw [ord_iso_def] "!!f. [| g: ord_iso(A,r,B,s); f: ord_iso(B,s,C,t) |] ==> \ \ (f O g): ord_iso(A,r,C,t)"; by (fast_tac (claset() addss bij_inverse_ss) 1); @@ -271,7 +271,7 @@ (** Two monotone maps can make an order-isomorphism **) -goalw Order.thy [ord_iso_def, mono_map_def] +Goalw [ord_iso_def, mono_map_def] "!!f g. [| f: mono_map(A,r,B,s); g: mono_map(B,s,A,r); \ \ f O g = id(B); g O f = id(A) |] ==> f: ord_iso(A,r,B,s)"; by Safe_tac; @@ -282,7 +282,7 @@ by (asm_full_simp_tac (simpset() addsimps [comp_eq_id_iff RS iffD1]) 1); qed "mono_ord_isoI"; -goal Order.thy +Goal "!!B. [| well_ord(A,r); well_ord(B,s); \ \ f: mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) \ \ |] ==> f: ord_iso(A,r,B,s)"; @@ -296,13 +296,13 @@ (** Order-isomorphisms preserve the ordering's properties **) -goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, ord_iso_def] +Goalw [part_ord_def, irrefl_def, trans_on_def, ord_iso_def] "!!A B r. [| part_ord(B,s); f: ord_iso(A,r,B,s) |] ==> part_ord(A,r)"; by (Asm_simp_tac 1); by (fast_tac (claset() addIs [bij_is_fun RS apply_type]) 1); qed "part_ord_ord_iso"; -goalw Order.thy [linear_def, ord_iso_def] +Goalw [linear_def, ord_iso_def] "!!A B r. [| linear(B,s); f: ord_iso(A,r,B,s) |] ==> linear(A,r)"; by (Asm_simp_tac 1); by Safe_tac; @@ -312,7 +312,7 @@ by (asm_full_simp_tac (simpset() addsimps [left_inverse_bij]) 1); qed "linear_ord_iso"; -goalw Order.thy [wf_on_def, wf_def, ord_iso_def] +Goalw [wf_on_def, wf_def, ord_iso_def] "!!A B r. [| wf[B](s); f: ord_iso(A,r,B,s) |] ==> wf[A](r)"; (*reversed &-congruence rule handles context of membership in A*) by (asm_full_simp_tac (simpset() addcongs [conj_cong2]) 1); @@ -323,7 +323,7 @@ (claset() addSDs [equalityD1] addIs [bij_is_fun RS apply_type]))); qed "wf_on_ord_iso"; -goalw Order.thy [well_ord_def, tot_ord_def] +Goalw [well_ord_def, tot_ord_def] "!!A B r. [| well_ord(B,s); f: ord_iso(A,r,B,s) |] ==> well_ord(A,r)"; by (fast_tac (claset() addSEs [part_ord_ord_iso, linear_ord_iso, wf_on_ord_iso]) 1); @@ -334,7 +334,7 @@ (*Inductive argument for Kunen's Lemma 6.1, etc. Simple proof from Halmos, page 72*) -goalw Order.thy [well_ord_def, ord_iso_def] +Goalw [well_ord_def, ord_iso_def] "!!r. [| well_ord(A,r); f: ord_iso(A,r, A',r); A'<= A; y: A |] \ \ ==> ~ : r"; by (REPEAT (eresolve_tac [conjE, CollectE] 1)); @@ -346,7 +346,7 @@ (*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment of a well-ordering*) -goal Order.thy +Goal "!!r. [| well_ord(A,r); f : ord_iso(A, r, pred(A,x,r), r); x:A |] ==> P"; by (metacut_tac well_ord_iso_subset_lemma 1); by (REPEAT_FIRST (ares_tac [pred_subset])); @@ -358,7 +358,7 @@ qed "well_ord_iso_predE"; (*Simple consequence of Lemma 6.1*) -goal Order.thy +Goal "!!r. [| well_ord(A,r); f : ord_iso(pred(A,a,r), r, pred(A,c,r), r); \ \ a:A; c:A |] ==> a=c"; by (forward_tac [well_ord_is_trans_on] 1); @@ -373,7 +373,7 @@ qed "well_ord_iso_pred_eq"; (*Does not assume r is a wellordering!*) -goalw Order.thy [ord_iso_def, pred_def] +Goalw [ord_iso_def, pred_def] "!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ \ f `` pred(A,a,r) = pred(B, f`a, s)"; by (etac CollectE 1); @@ -388,7 +388,7 @@ (*But in use, A and B may themselves be initial segments. Then use trans_pred_pred_eq to simplify the pred(pred...) terms. See just below.*) -goal Order.thy +Goal "!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ \ restrict(f, pred(A,a,r)) : ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)"; by (asm_simp_tac (simpset() addsimps [ord_iso_image_pred RS sym]) 1); @@ -400,7 +400,7 @@ qed "ord_iso_restrict_pred"; (*Tricky; a lot of forward proof!*) -goal Order.thy +Goal "!!r. [| well_ord(A,r); well_ord(B,s); : 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); \ @@ -423,7 +423,7 @@ addIs [ord_iso_is_bij, bij_is_fun, apply_funtype]; (*See Halmos, page 72*) -goal Order.thy +Goal "!!r. [| well_ord(A,r); \ \ f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s); y: A |] \ \ ==> ~ : s"; @@ -437,7 +437,7 @@ qed "well_ord_iso_unique_lemma"; (*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*) -goal Order.thy +Goal "!!r. [| well_ord(A,r); \ \ f: ord_iso(A,r, B,s); g: ord_iso(A,r, B,s) |] ==> f = g"; by (rtac fun_extension 1); @@ -453,30 +453,30 @@ (** Towards Kunen's Theorem 6.3: linearity of the similarity relation **) -goalw Order.thy [ord_iso_map_def] "ord_iso_map(A,r,B,s) <= A*B"; +Goalw [ord_iso_map_def] "ord_iso_map(A,r,B,s) <= A*B"; by (Blast_tac 1); qed "ord_iso_map_subset"; -goalw Order.thy [ord_iso_map_def] "domain(ord_iso_map(A,r,B,s)) <= A"; +Goalw [ord_iso_map_def] "domain(ord_iso_map(A,r,B,s)) <= A"; by (Blast_tac 1); qed "domain_ord_iso_map"; -goalw Order.thy [ord_iso_map_def] "range(ord_iso_map(A,r,B,s)) <= B"; +Goalw [ord_iso_map_def] "range(ord_iso_map(A,r,B,s)) <= B"; by (Blast_tac 1); qed "range_ord_iso_map"; -goalw Order.thy [ord_iso_map_def] +Goalw [ord_iso_map_def] "converse(ord_iso_map(A,r,B,s)) = ord_iso_map(B,s,A,r)"; by (blast_tac (claset() addIs [ord_iso_sym]) 1); qed "converse_ord_iso_map"; -goalw Order.thy [ord_iso_map_def, function_def] +Goalw [ord_iso_map_def, function_def] "!!B. well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))"; by (blast_tac (claset() addIs [well_ord_iso_pred_eq, ord_iso_sym, ord_iso_trans]) 1); qed "function_ord_iso_map"; -goal Order.thy +Goal "!!B. well_ord(B,s) ==> ord_iso_map(A,r,B,s) \ \ : domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))"; by (asm_simp_tac @@ -484,7 +484,7 @@ ord_iso_map_subset RS domain_times_range]) 1); qed "ord_iso_map_fun"; -goalw Order.thy [mono_map_def] +Goalw [mono_map_def] "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ \ : mono_map(domain(ord_iso_map(A,r,B,s)), r, \ \ range(ord_iso_map(A,r,B,s)), s)"; @@ -500,7 +500,7 @@ by (rtac well_ord_iso_preserving 1 THEN REPEAT_FIRST assume_tac); qed "ord_iso_map_mono_map"; -goalw Order.thy [mono_map_def] +Goalw [mono_map_def] "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ \ : ord_iso(domain(ord_iso_map(A,r,B,s)), r, \ \ range(ord_iso_map(A,r,B,s)), s)"; @@ -514,7 +514,7 @@ qed "ord_iso_map_ord_iso"; (*One way of saying that domain(ord_iso_map(A,r,B,s)) is downwards-closed*) -goalw Order.thy [ord_iso_map_def] +Goalw [ord_iso_map_def] "!!B. [| well_ord(A,r); well_ord(B,s); \ \ a: A; a ~: domain(ord_iso_map(A,r,B,s)) \ \ |] ==> domain(ord_iso_map(A,r,B,s)) <= pred(A, a, r)"; @@ -535,7 +535,7 @@ qed "domain_ord_iso_map_subset"; (*For the 4-way case analysis in the main result*) -goal Order.thy +Goal "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ \ domain(ord_iso_map(A,r,B,s)) = A | \ \ (EX x:A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))"; @@ -556,7 +556,7 @@ qed "domain_ord_iso_map_cases"; (*As above, by duality*) -goal Order.thy +Goal "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ \ range(ord_iso_map(A,r,B,s)) = B | \ \ (EX y:B. range(ord_iso_map(A,r,B,s))= pred(B,y,s))"; @@ -566,7 +566,7 @@ qed "range_ord_iso_map_cases"; (*Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets*) -goal Order.thy +Goal "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ \ ord_iso_map(A,r,B,s) : ord_iso(A, r, B, s) | \ \ (EX x:A. ord_iso_map(A,r,B,s) : ord_iso(pred(A,x,r), r, B, s)) | \ @@ -589,27 +589,27 @@ (*** Properties of converse(r), by Krzysztof Grabczewski ***) -goalw Order.thy [irrefl_def] +Goalw [irrefl_def] "!!A. irrefl(A,r) ==> irrefl(A,converse(r))"; by (Blast_tac 1); qed "irrefl_converse"; -goalw Order.thy [trans_on_def] +Goalw [trans_on_def] "!!A. trans[A](r) ==> trans[A](converse(r))"; by (Blast_tac 1); qed "trans_on_converse"; -goalw Order.thy [part_ord_def] +Goalw [part_ord_def] "!!A. part_ord(A,r) ==> part_ord(A,converse(r))"; by (blast_tac (claset() addSIs [irrefl_converse, trans_on_converse]) 1); qed "part_ord_converse"; -goalw Order.thy [linear_def] +Goalw [linear_def] "!!A. linear(A,r) ==> linear(A,converse(r))"; by (Blast_tac 1); qed "linear_converse"; -goalw Order.thy [tot_ord_def] +Goalw [tot_ord_def] "!!A. tot_ord(A,r) ==> tot_ord(A,converse(r))"; by (blast_tac (claset() addSIs [part_ord_converse, linear_converse]) 1); qed "tot_ord_converse"; @@ -618,11 +618,11 @@ (** By Krzysztof Grabczewski. Lemmas involving the first element of a well ordered set **) -goalw thy [first_def] "!!b. first(b,B,r) ==> b:B"; +Goalw [first_def] "!!b. first(b,B,r) ==> b:B"; by (Blast_tac 1); qed "first_is_elem"; -goalw thy [well_ord_def, wf_on_def, wf_def, first_def] +Goalw [well_ord_def, wf_on_def, wf_def, first_def] "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))"; by (REPEAT (eresolve_tac [conjE,allE,disjE] 1)); by (contr_tac 1); @@ -633,7 +633,7 @@ by (Blast.depth_tac (claset()) 7 1); qed "well_ord_imp_ex1_first"; -goal thy "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B"; +Goal "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B"; by (dtac well_ord_imp_ex1_first 1 THEN REPEAT (assume_tac 1)); by (rtac first_is_elem 1); by (etac theI 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/OrderArith.ML --- a/src/ZF/OrderArith.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/OrderArith.ML Mon Jun 22 17:12:27 1998 +0200 @@ -13,22 +13,22 @@ (** Rewrite rules. Can be used to obtain introduction rules **) -goalw OrderArith.thy [radd_def] +Goalw [radd_def] " : radd(A,r,B,s) <-> a:A & b:B"; by (Blast_tac 1); qed "radd_Inl_Inr_iff"; -goalw OrderArith.thy [radd_def] +Goalw [radd_def] " : radd(A,r,B,s) <-> a':A & a:A & :r"; by (Blast_tac 1); qed "radd_Inl_iff"; -goalw OrderArith.thy [radd_def] +Goalw [radd_def] " : radd(A,r,B,s) <-> b':B & b:B & :s"; by (Blast_tac 1); qed "radd_Inr_iff"; -goalw OrderArith.thy [radd_def] +Goalw [radd_def] " : radd(A,r,B,s) <-> False"; by (Blast_tac 1); qed "radd_Inr_Inl_iff"; @@ -54,7 +54,7 @@ (** Type checking **) -goalw OrderArith.thy [radd_def] "radd(A,r,B,s) <= (A+B) * (A+B)"; +Goalw [radd_def] "radd(A,r,B,s) <= (A+B) * (A+B)"; by (rtac Collect_subset 1); qed "radd_type"; @@ -65,7 +65,7 @@ Addsimps [radd_Inl_iff, radd_Inr_iff, radd_Inl_Inr_iff, radd_Inr_Inl_iff]; -goalw OrderArith.thy [linear_def] +Goalw [linear_def] "!!r s. [| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))"; by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac sumE)); by (ALLGOALS Asm_simp_tac); @@ -74,7 +74,7 @@ (** Well-foundedness **) -goal OrderArith.thy +Goal "!!r s. [| wf[A](r); wf[B](s) |] ==> wf[A+B](radd(A,r,B,s))"; by (rtac wf_onI2 1); by (subgoal_tac "ALL x:A. Inl(x): Ba" 1); @@ -90,14 +90,14 @@ by (best_tac (claset() addSEs [raddE, bspec RS mp]) 1); qed "wf_on_radd"; -goal OrderArith.thy +Goal "!!r s. [| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))"; by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1); by (rtac (field_radd RSN (2, wf_on_subset_A)) 1); by (REPEAT (ares_tac [wf_on_radd] 1)); qed "wf_radd"; -goal OrderArith.thy +Goal "!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \ \ well_ord(A+B, radd(A,r,B,s))"; by (rtac well_ordI 1); @@ -108,7 +108,7 @@ (** An ord_iso congruence law **) -goal OrderArith.thy +Goal "!!f g. [| f: bij(A,C); g: bij(B,D) |] ==> \ \ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) : bij(A+B, C+D)"; by (res_inst_tac @@ -118,7 +118,7 @@ by (ALLGOALS (asm_simp_tac bij_inverse_ss)); qed "sum_bij"; -goalw OrderArith.thy [ord_iso_def] +Goalw [ord_iso_def] "!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \ \ (lam 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'))"; @@ -134,7 +134,7 @@ (*Could we prove an ord_iso result? Perhaps ord_iso(A+B, radd(A,r,B,s), A Un B, r Un s) *) -goal OrderArith.thy +Goal "!!A B. A Int B = 0 ==> \ \ (lam z:A+B. case(%x. x, %y. y, z)) : bij(A+B, A Un B)"; by (res_inst_tac [("d", "%z. if(z:A, Inl(z), Inr(z))")] @@ -148,7 +148,7 @@ (** Associativity **) -goal OrderArith.thy +Goal "(lam z:(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z)) \ \ : bij((A+B)+C, A+(B+C))"; by (res_inst_tac [("d", "case(%x. Inl(Inl(x)), case(%x. Inl(Inr(x)), Inr))")] @@ -156,7 +156,7 @@ by (ALLGOALS (asm_simp_tac (simpset() setloop etac sumE))); qed "sum_assoc_bij"; -goal OrderArith.thy +Goal "(lam 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)))"; @@ -170,7 +170,7 @@ (** Rewrite rule. Can be used to obtain introduction rules **) -goalw OrderArith.thy [rmult_def] +Goalw [rmult_def] "!!r s. <, > : rmult(A,r,B,s) <-> \ \ (: r & a':A & a:A & b': B & b: B) | \ \ (: s & a'=a & a:A & b': B & b: B)"; @@ -190,7 +190,7 @@ (** Type checking **) -goalw OrderArith.thy [rmult_def] "rmult(A,r,B,s) <= (A*B) * (A*B)"; +Goalw [rmult_def] "rmult(A,r,B,s) <= (A*B) * (A*B)"; by (rtac Collect_subset 1); qed "rmult_type"; @@ -211,7 +211,7 @@ (** Well-foundedness **) -goal OrderArith.thy +Goal "!!r s. [| wf[A](r); wf[B](s) |] ==> wf[A*B](rmult(A,r,B,s))"; by (rtac wf_onI2 1); by (etac SigmaE 1); @@ -225,14 +225,14 @@ qed "wf_on_rmult"; -goal OrderArith.thy +Goal "!!r s. [| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))"; by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1); by (rtac (field_rmult RSN (2, wf_on_subset_A)) 1); by (REPEAT (ares_tac [wf_on_rmult] 1)); qed "wf_rmult"; -goal OrderArith.thy +Goal "!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \ \ well_ord(A*B, rmult(A,r,B,s))"; by (rtac well_ordI 1); @@ -244,7 +244,7 @@ (** An ord_iso congruence law **) -goal OrderArith.thy +Goal "!!f g. [| f: bij(A,C); g: bij(B,D) |] ==> \ \ (lam :A*B. ) : bij(A*B, C*D)"; by (res_inst_tac [("d", "%. ")] @@ -253,7 +253,7 @@ by (ALLGOALS (asm_simp_tac bij_inverse_ss)); qed "prod_bij"; -goalw OrderArith.thy [ord_iso_def] +Goalw [ord_iso_def] "!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \ \ (lam :A*B. ) \ \ : ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))"; @@ -264,14 +264,14 @@ by (blast_tac (claset() addIs [bij_is_inj RS inj_apply_equality]) 1); qed "prod_ord_iso_cong"; -goal OrderArith.thy "(lam z:A. ) : bij(A, {x}*A)"; +Goal "(lam z:A. ) : bij(A, {x}*A)"; by (res_inst_tac [("d", "snd")] lam_bijective 1); by Safe_tac; by (ALLGOALS Asm_simp_tac); qed "singleton_prod_bij"; (*Used??*) -goal OrderArith.thy +Goal "!!x xr. well_ord({x},xr) ==> \ \ (lam z:A. ) : ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))"; by (resolve_tac [singleton_prod_bij RS ord_isoI] 1); @@ -281,7 +281,7 @@ (*Here we build a complicated function term, then simplify it using case_cong, id_conv, comp_lam, case_case.*) -goal OrderArith.thy +Goal "!!a. a~:C ==> \ \ (lam x:C*B + D. case(%x. x, %y., x)) \ \ : bij(C*B + D, C*B Un {a}*D)"; @@ -296,7 +296,7 @@ by (asm_simp_tac (simpset() addsimps [case_case]) 1); qed "prod_sum_singleton_bij"; -goal OrderArith.thy +Goal "!!A. [| a:A; well_ord(A,r) |] ==> \ \ (lam x:pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y., x)) \ \ : ord_iso(pred(A,a,r)*B + pred(B,b,s), \ @@ -313,7 +313,7 @@ (** Distributive law **) -goal OrderArith.thy +Goal "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), x)) \ \ : bij((A+B)*C, (A*C)+(B*C))"; by (res_inst_tac @@ -322,7 +322,7 @@ by (ALLGOALS Asm_simp_tac); qed "sum_prod_distrib_bij"; -goal OrderArith.thy +Goal "(lam :(A+B)*C. case(%y. Inl(), %y. Inr(), 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)))"; @@ -333,13 +333,13 @@ (** Associativity **) -goal OrderArith.thy +Goal "(lam <, z>:(A*B)*C. >) : bij((A*B)*C, A*(B*C))"; by (res_inst_tac [("d", "%>. <, z>")] lam_bijective 1); by (ALLGOALS (asm_simp_tac (simpset() setloop etac SigmaE))); qed "prod_assoc_bij"; -goal OrderArith.thy +Goal "(lam <, z>:(A*B)*C. >) \ \ : 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)))"; @@ -353,20 +353,20 @@ (** Rewrite rule **) -goalw OrderArith.thy [rvimage_def] +Goalw [rvimage_def] " : rvimage(A,f,r) <-> : r & a:A & b:A"; by (Blast_tac 1); qed "rvimage_iff"; (** Type checking **) -goalw OrderArith.thy [rvimage_def] "rvimage(A,f,r) <= A*A"; +Goalw [rvimage_def] "rvimage(A,f,r) <= A*A"; by (rtac Collect_subset 1); qed "rvimage_type"; bind_thm ("field_rvimage", (rvimage_type RS field_rel_subset)); -goalw OrderArith.thy [rvimage_def] +Goalw [rvimage_def] "rvimage(A,f, converse(r)) = converse(rvimage(A,f,r))"; by (Blast_tac 1); qed "rvimage_converse"; @@ -374,17 +374,17 @@ (** Partial Ordering Properties **) -goalw OrderArith.thy [irrefl_def, rvimage_def] +Goalw [irrefl_def, rvimage_def] "!!A B. [| f: inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))"; by (blast_tac (claset() addIs [inj_is_fun RS apply_type]) 1); qed "irrefl_rvimage"; -goalw OrderArith.thy [trans_on_def, rvimage_def] +Goalw [trans_on_def, rvimage_def] "!!A B. [| f: inj(A,B); trans[B](r) |] ==> trans[A](rvimage(A,f,r))"; by (blast_tac (claset() addIs [inj_is_fun RS apply_type]) 1); qed "trans_on_rvimage"; -goalw OrderArith.thy [part_ord_def] +Goalw [part_ord_def] "!!A B. [| f: inj(A,B); part_ord(B,r) |] ==> part_ord(A, rvimage(A,f,r))"; by (blast_tac (claset() addSIs [irrefl_rvimage, trans_on_rvimage]) 1); qed "part_ord_rvimage"; @@ -401,7 +401,7 @@ by (REPEAT_SOME (blast_tac (claset() addIs [apply_funtype]))); qed "linear_rvimage"; -goalw OrderArith.thy [tot_ord_def] +Goalw [tot_ord_def] "!!A B. [| f: inj(A,B); tot_ord(B,r) |] ==> tot_ord(A, rvimage(A,f,r))"; by (blast_tac (claset() addSIs [part_ord_rvimage, linear_rvimage]) 1); qed "tot_ord_rvimage"; @@ -409,7 +409,7 @@ (** Well-foundedness **) -goal OrderArith.thy +Goal "!!r. [| f: A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))"; by (rtac wf_onI2 1); by (subgoal_tac "ALL z:A. f`z=f`y --> z: Ba" 1); @@ -421,7 +421,7 @@ qed "wf_on_rvimage"; (*Note that we need only wf[A](...) and linear(A,...) to get the result!*) -goal OrderArith.thy +Goal "!!r. [| f: inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))"; by (rtac well_ordI 1); by (rewrite_goals_tac [well_ord_def, tot_ord_def]); @@ -429,12 +429,12 @@ by (blast_tac (claset() addSIs [linear_rvimage]) 1); qed "well_ord_rvimage"; -goalw OrderArith.thy [ord_iso_def] +Goalw [ord_iso_def] "!!A B. f: bij(A,B) ==> f: ord_iso(A, rvimage(A,f,s), B, s)"; by (asm_full_simp_tac (simpset() addsimps [rvimage_iff]) 1); qed "ord_iso_rvimage"; -goalw OrderArith.thy [ord_iso_def, rvimage_def] +Goalw [ord_iso_def, rvimage_def] "!!A B. f: ord_iso(A,r, B,s) ==> rvimage(A,f,s) = r Int A*A"; by (Blast_tac 1); qed "ord_iso_rvimage_eq"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/OrderType.ML --- a/src/ZF/OrderType.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/OrderType.ML Mon Jun 22 17:12:27 1998 +0200 @@ -29,18 +29,18 @@ (*Kunen's Theorem 7.3 (i), page 16; see also Ordinal/Ord_in_Ord The smaller ordinal is an initial segment of the larger *) -goalw OrderType.thy [pred_def, lt_def] +Goalw [pred_def, lt_def] "!!i j. j pred(i, j, Memrel(i)) = j"; by (asm_simp_tac (simpset() addsimps [Memrel_iff]) 1); by (blast_tac (claset() addIs [Ord_trans]) 1); qed "lt_pred_Memrel"; -goalw OrderType.thy [pred_def,Memrel_def] +Goalw [pred_def,Memrel_def] "!!A x. x:A ==> pred(A, x, Memrel(A)) = A Int x"; by (Blast_tac 1); qed "pred_Memrel"; -goal OrderType.thy +Goal "!!i. [| j R"; by (forward_tac [lt_pred_Memrel] 1); by (etac ltE 1); @@ -53,7 +53,7 @@ qed "Ord_iso_implies_eq_lemma"; (*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*) -goal OrderType.thy +Goal "!!i. [| Ord(i); Ord(j); f: ord_iso(i,Memrel(i),j,Memrel(j)) \ \ |] ==> i=j"; by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1); @@ -63,7 +63,7 @@ (**** Ordermap and ordertype ****) -goalw OrderType.thy [ordermap_def,ordertype_def] +Goalw [ordermap_def,ordertype_def] "ordermap(A,r) : A -> ordertype(A,r)"; by (rtac lam_type 1); by (rtac (lamI RS imageI) 1); @@ -73,7 +73,7 @@ (*** Unfolding of ordermap ***) (*Useful for cardinality reasoning; see CardinalArith.ML*) -goalw OrderType.thy [ordermap_def, pred_def] +Goalw [ordermap_def, pred_def] "!!r. [| wf[A](r); x:A |] ==> \ \ ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)"; by (Asm_simp_tac 1); @@ -84,7 +84,7 @@ qed "ordermap_eq_image"; (*Useful for rewriting PROVIDED pred is not unfolded until later!*) -goal OrderType.thy +Goal "!!r. [| wf[A](r); x:A |] ==> \ \ ordermap(A,r) ` x = {ordermap(A,r)`y . y : pred(A,x,r)}"; by (asm_simp_tac (simpset() addsimps [ordermap_eq_image, pred_subset, @@ -101,7 +101,7 @@ assume_tac (i+1), assume_tac i]; -goalw OrderType.thy [well_ord_def, tot_ord_def, part_ord_def] +Goalw [well_ord_def, tot_ord_def, part_ord_def] "!!r. [| well_ord(A,r); x:A |] ==> Ord(ordermap(A,r) ` x)"; by Safe_tac; by (wf_on_ind_tac "x" [] 1); @@ -114,7 +114,7 @@ by (fast_tac (claset() addSEs [trans_onD]) 1); qed "Ord_ordermap"; -goalw OrderType.thy [ordertype_def] +Goalw [ordertype_def] "!!r. well_ord(A,r) ==> Ord(ordertype(A,r))"; by (stac ([ordermap_type, subset_refl] MRS image_fun) 1); by (rtac (Ord_is_Transset RSN (2,OrdI)) 1); @@ -127,7 +127,7 @@ (*** ordermap preserves the orderings in both directions ***) -goal OrderType.thy +Goal "!!r. [| : r; wf[A](r); w: A; x: A |] ==> \ \ ordermap(A,r)`w : ordermap(A,r)`x"; by (eres_inst_tac [("x1", "x")] (ordermap_unfold RS ssubst) 1); @@ -136,7 +136,7 @@ qed "ordermap_mono"; (*linearity of r is crucial here*) -goalw OrderType.thy [well_ord_def, tot_ord_def] +Goalw [well_ord_def, tot_ord_def] "!!r. [| ordermap(A,r)`w : ordermap(A,r)`x; well_ord(A,r); \ \ w: A; x: A |] ==> : r"; by Safe_tac; @@ -152,7 +152,7 @@ rewrite_rule [symmetric ordertype_def] (ordermap_type RS surj_image)); -goalw OrderType.thy [well_ord_def, tot_ord_def, bij_def, inj_def] +Goalw [well_ord_def, tot_ord_def, bij_def, inj_def] "!!r. well_ord(A,r) ==> ordermap(A,r) : bij(A, ordertype(A,r))"; by (fast_tac (claset() addSIs [ordermap_type, ordermap_surj] addEs [linearE] @@ -162,7 +162,7 @@ (*** Isomorphisms involving ordertype ***) -goalw OrderType.thy [ord_iso_def] +Goalw [ord_iso_def] "!!r. well_ord(A,r) ==> \ \ ordermap(A,r) : ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))"; by (safe_tac (claset() addSEs [well_ord_is_wf] @@ -171,7 +171,7 @@ by (blast_tac (claset() addSDs [converse_ordermap_mono]) 1); qed "ordertype_ord_iso"; -goal OrderType.thy +Goal "!!f. [| f: ord_iso(A,r,B,s); well_ord(B,s) |] ==> \ \ ordertype(A,r) = ordertype(B,s)"; by (forward_tac [well_ord_ord_iso] 1 THEN assume_tac 1); @@ -181,7 +181,7 @@ addSEs [ordertype_ord_iso]) 0 1); qed "ordertype_eq"; -goal OrderType.thy +Goal "!!A B. [| ordertype(A,r) = ordertype(B,s); \ \ well_ord(A,r); well_ord(B,s) \ \ |] ==> EX f. f: ord_iso(A,r,B,s)"; @@ -195,7 +195,7 @@ (*** Basic equalities for ordertype ***) (*Ordertype of Memrel*) -goal OrderType.thy "!!i. j le i ==> ordertype(j,Memrel(i)) = j"; +Goal "!!i. j le i ==> ordertype(j,Memrel(i)) = j"; by (resolve_tac [Ord_iso_implies_eq RS sym] 1); by (etac ltE 1); by (REPEAT (ares_tac [le_well_ord_Memrel, Ord_ordertype] 1)); @@ -209,7 +209,7 @@ (*"Ord(i) ==> ordertype(i, Memrel(i)) = i"*) bind_thm ("ordertype_Memrel", le_refl RS le_ordertype_Memrel); -goal OrderType.thy "ordertype(0,r) = 0"; +Goal "ordertype(0,r) = 0"; by (resolve_tac [id_bij RS ord_isoI RS ordertype_eq RS trans] 1); by (etac emptyE 1); by (rtac well_ord_0 1); @@ -225,7 +225,7 @@ (*** A fundamental unfolding law for ordertype. ***) (*Ordermap returns the same result if applied to an initial segment*) -goal OrderType.thy +Goal "!!r. [| well_ord(A,r); y:A; z: pred(A,y,r) |] ==> \ \ ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z"; by (forward_tac [[well_ord_is_wf, pred_subset] MRS wf_on_subset_A] 1); @@ -241,14 +241,14 @@ by (fast_tac (claset() addSEs [trans_onD]) 1); qed "ordermap_pred_eq_ordermap"; -goalw OrderType.thy [ordertype_def] +Goalw [ordertype_def] "ordertype(A,r) = {ordermap(A,r)`y . y : A}"; by (rtac ([ordermap_type, subset_refl] MRS image_fun) 1); qed "ordertype_unfold"; (** Theorems by Krzysztof Grabczewski; proofs simplified by lcp **) -goal OrderType.thy +Goal "!!r. [| well_ord(A,r); x:A |] ==> \ \ ordertype(pred(A,x,r),r) <= ordertype(A,r)"; by (asm_simp_tac (simpset() addsimps [ordertype_unfold, @@ -257,7 +257,7 @@ addEs [predE]) 1); qed "ordertype_pred_subset"; -goal OrderType.thy +Goal "!!r. [| well_ord(A,r); x:A |] ==> \ \ ordertype(pred(A,x,r),r) < ordertype(A,r)"; by (resolve_tac [ordertype_pred_subset RS subset_imp_le RS leE] 1); @@ -269,7 +269,7 @@ (*May rewrite with this -- provided no rules are supplied for proving that well_ord(pred(A,x,r), r) *) -goal OrderType.thy +Goal "!!A r. well_ord(A,r) ==> \ \ ordertype(A,r) = {ordertype(pred(A,x,r),r). x:A}"; by (rtac equalityI 1); @@ -288,7 +288,7 @@ (**** Alternative definition of ordinal ****) (*proof by Krzysztof Grabczewski*) -goalw OrderType.thy [Ord_alt_def] "!!i. Ord(i) ==> Ord_alt(i)"; +Goalw [Ord_alt_def] "!!i. Ord(i) ==> Ord_alt(i)"; by (rtac conjI 1); by (etac well_ord_Memrel 1); by (rewrite_goals_tac [Ord_def, Transset_def, pred_def, Memrel_def]); @@ -296,7 +296,7 @@ qed "Ord_is_Ord_alt"; (*proof by lcp*) -goalw OrderType.thy [Ord_alt_def, Ord_def, Transset_def, well_ord_def, +Goalw [Ord_alt_def, Ord_def, Transset_def, well_ord_def, tot_ord_def, part_ord_def, trans_on_def] "!!i. Ord_alt(i) ==> Ord(i)"; by (asm_full_simp_tac (simpset() addsimps [Memrel_iff, pred_Memrel]) 1); @@ -310,26 +310,26 @@ (** Addition with 0 **) -goal OrderType.thy "(lam z:A+0. case(%x. x, %y. y, z)) : bij(A+0, A)"; +Goal "(lam z:A+0. case(%x. x, %y. y, z)) : bij(A+0, A)"; by (res_inst_tac [("d", "Inl")] lam_bijective 1); by Safe_tac; by (ALLGOALS Asm_simp_tac); qed "bij_sum_0"; -goal OrderType.thy +Goal "!!A r. well_ord(A,r) ==> ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)"; by (resolve_tac [bij_sum_0 RS ord_isoI RS ordertype_eq] 1); by (assume_tac 2); by (fast_tac (claset() addss (simpset() addsimps [radd_Inl_iff, Memrel_iff])) 1); qed "ordertype_sum_0_eq"; -goal OrderType.thy "(lam z:0+A. case(%x. x, %y. y, z)) : bij(0+A, A)"; +Goal "(lam z:0+A. case(%x. x, %y. y, z)) : bij(0+A, A)"; by (res_inst_tac [("d", "Inr")] lam_bijective 1); by Safe_tac; by (ALLGOALS Asm_simp_tac); qed "bij_0_sum"; -goal OrderType.thy +Goal "!!A r. well_ord(A,r) ==> ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)"; by (resolve_tac [bij_0_sum RS ord_isoI RS ordertype_eq] 1); by (assume_tac 2); @@ -339,7 +339,7 @@ (** Initial segments of radd. Statements by Grabczewski **) (*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *) -goalw OrderType.thy [pred_def] +Goalw [pred_def] "!!A B. a:A ==> \ \ (lam x:pred(A,a,r). Inl(x)) \ \ : bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))"; @@ -350,7 +350,7 @@ (simpset() addsimps [radd_Inl_iff, radd_Inr_Inl_iff]))); qed "pred_Inl_bij"; -goal OrderType.thy +Goal "!!A B. [| a:A; well_ord(A,r) |] ==> \ \ ordertype(pred(A+B, Inl(a), radd(A,r,B,s)), radd(A,r,B,s)) = \ \ ordertype(pred(A,a,r), r)"; @@ -359,7 +359,7 @@ by (asm_full_simp_tac (simpset() addsimps [radd_Inl_iff, pred_def]) 1); qed "ordertype_pred_Inl_eq"; -goalw OrderType.thy [pred_def, id_def] +Goalw [pred_def, id_def] "!!A B. b:B ==> \ \ id(A+pred(B,b,s)) \ \ : bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))"; @@ -368,7 +368,7 @@ by (ALLGOALS (Asm_full_simp_tac)); qed "pred_Inr_bij"; -goal OrderType.thy +Goal "!!A B. [| b:B; well_ord(A,r); well_ord(B,s) |] ==> \ \ ordertype(pred(A+B, Inr(b), radd(A,r,B,s)), radd(A,r,B,s)) = \ \ ordertype(A+pred(B,b,s), radd(A,r,pred(B,b,s),s))"; @@ -379,19 +379,19 @@ (*** Basic laws for ordinal addition ***) -goalw OrderType.thy [oadd_def] +Goalw [oadd_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i++j)"; by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 1)); qed "Ord_oadd"; (** Ordinal addition with zero **) -goalw OrderType.thy [oadd_def] "!!i. Ord(i) ==> i++0 = i"; +Goalw [oadd_def] "!!i. Ord(i) ==> i++0 = i"; by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_sum_0_eq, ordertype_Memrel, well_ord_Memrel]) 1); qed "oadd_0"; -goalw OrderType.thy [oadd_def] "!!i. Ord(i) ==> 0++i = i"; +Goalw [oadd_def] "!!i. Ord(i) ==> 0++i = i"; by (asm_simp_tac (simpset() addsimps [Memrel_0, ordertype_0_sum_eq, ordertype_Memrel, well_ord_Memrel]) 1); qed "oadd_0_left"; @@ -401,7 +401,7 @@ (*** Further properties of ordinal addition. Statements by Grabczewski, proofs by lcp. ***) -goalw OrderType.thy [oadd_def] "!!i j k. [| k k < i++j"; +Goalw [oadd_def] "!!i j k. [| k k < i++j"; by (rtac ltE 1 THEN assume_tac 1); by (rtac ltI 1); by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 2)); @@ -414,21 +414,21 @@ qed "lt_oadd1"; (*Thus also we obtain the rule i++j = k ==> i le k *) -goal OrderType.thy "!!i j. [| Ord(i); Ord(j) |] ==> i le i++j"; +Goal "!!i j. [| Ord(i); Ord(j) |] ==> i le i++j"; by (rtac all_lt_imp_le 1); by (REPEAT (ares_tac [Ord_oadd, lt_oadd1] 1)); qed "oadd_le_self"; (** A couple of strange but necessary results! **) -goal OrderType.thy +Goal "!!A B. A<=B ==> id(A) : ord_iso(A, Memrel(A), A, Memrel(B))"; by (resolve_tac [id_bij RS ord_isoI] 1); by (asm_simp_tac (simpset() addsimps [id_conv, Memrel_iff]) 1); by (Blast_tac 1); qed "id_ord_iso_Memrel"; -goal OrderType.thy +Goal "!!k. [| well_ord(A,r); k \ \ ordertype(A+k, radd(A, r, k, Memrel(j))) = \ \ ordertype(A+k, radd(A, r, k, Memrel(k)))"; @@ -438,7 +438,7 @@ by (REPEAT_FIRST (ares_tac [well_ord_radd, well_ord_Memrel])); qed "ordertype_sum_Memrel"; -goalw OrderType.thy [oadd_def] "!!i j k. [| k i++k < i++j"; +Goalw [oadd_def] "!!i j k. [| k i++k < i++j"; by (rtac ltE 1 THEN assume_tac 1); by (resolve_tac [ordertype_pred_unfold RS equalityD2 RS subsetD RS ltI] 1); by (REPEAT_FIRST (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel])); @@ -450,7 +450,7 @@ ordertype_sum_Memrel]) 1); qed "oadd_lt_mono2"; -goal OrderType.thy +Goal "!!i j k. [| i++j < i++k; Ord(i); Ord(j); Ord(k) |] ==> j i++j < i++k <-> j j=k"; by (rtac Ord_linear_lt 1); by (REPEAT_SOME assume_tac); @@ -472,7 +472,7 @@ addss (simpset() addsimps [lt_not_refl])))); qed "oadd_inject"; -goalw OrderType.thy [oadd_def] +Goalw [oadd_def] "!!i j k. [| k < i++j; Ord(i); Ord(j) |] ==> k (i++j)++k = i++(j++k)"; by (resolve_tac [ordertype_eq RS trans] 1); by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS @@ -502,7 +502,7 @@ by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1)); qed "oadd_assoc"; -goal OrderType.thy +Goal "!!i j. [| Ord(i); Ord(j) |] ==> i++j = i Un (UN k:j. {i++k})"; by (rtac (subsetI RS equalityI) 1); by (eresolve_tac [ltI RS lt_oadd_disj RS disjE] 1); @@ -513,12 +513,12 @@ by (blast_tac (claset() addSEs [ltE]) 1); qed "oadd_unfold"; -goal OrderType.thy "!!i. Ord(i) ==> i++1 = succ(i)"; +Goal "!!i. Ord(i) ==> i++1 = succ(i)"; by (asm_simp_tac (simpset() addsimps [oadd_unfold, Ord_1, oadd_0]) 1); by (Blast_tac 1); qed "oadd_1"; -goal OrderType.thy +Goal "!!i. [| Ord(i); Ord(j) |] ==> i++succ(j) = succ(i++j)"; (*ZF_ss prevents looping*) by (asm_simp_tac (ZF_ss addsimps [Ord_oadd, oadd_1 RS sym]) 1); @@ -536,7 +536,7 @@ addSEs [ltE, ltI RS lt_oadd_disj RS disjE]) 1); qed "oadd_UN"; -goal OrderType.thy +Goal "!!i j. [| Ord(i); Limit(j) |] ==> i++j = (UN k:j. i++k)"; by (forward_tac [Limit_has_0 RS ltD] 1); by (asm_simp_tac (simpset() addsimps [Limit_is_Ord RS Ord_in_Ord, @@ -546,7 +546,7 @@ (** Order/monotonicity properties of ordinal addition **) -goal OrderType.thy "!!i j. [| Ord(i); Ord(j) |] ==> i le j++i"; +Goal "!!i j. [| Ord(i); Ord(j) |] ==> i le j++i"; by (eres_inst_tac [("i","i")] trans_induct3 1); by (asm_simp_tac (simpset() addsimps [Ord_0_le]) 1); by (asm_simp_tac (simpset() addsimps [oadd_succ, succ_leI]) 1); @@ -558,7 +558,7 @@ le_refl, Limit_is_Ord]) 1); qed "oadd_le_self2"; -goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> k++i le j++i"; +Goal "!!i j k. [| k le j; Ord(i) |] ==> k++i le j++i"; by (forward_tac [lt_Ord] 1); by (forward_tac [le_Ord2] 1); by (etac trans_induct3 1); @@ -569,17 +569,17 @@ by (Blast_tac 1); qed "oadd_le_mono1"; -goal OrderType.thy "!!i j. [| i' le i; j' i'++j' < i++j"; +Goal "!!i j. [| i' le i; j' i'++j' < i++j"; by (rtac lt_trans1 1); by (REPEAT (eresolve_tac [asm_rl, oadd_le_mono1, oadd_lt_mono2, ltE, Ord_succD] 1)); qed "oadd_lt_mono"; -goal OrderType.thy "!!i j. [| i' le i; j' le j |] ==> i'++j' le i++j"; +Goal "!!i j. [| i' le i; j' le j |] ==> i'++j' le i++j"; by (asm_simp_tac (simpset() addsimps [oadd_succ RS sym, le_Ord2, oadd_lt_mono]) 1); qed "oadd_le_mono"; -goal OrderType.thy +Goal "!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> i++j le i++k <-> j le k"; by (asm_simp_tac (simpset() addsimps [oadd_lt_iff2, oadd_succ RS sym, Ord_succ]) 1); @@ -590,7 +590,7 @@ Probably simpler to define the difference recursively! **) -goal OrderType.thy +Goal "!!A B. A<=B ==> (lam y:B. if(y:A, Inl(y), Inr(y))) : bij(B, A+(B-A))"; by (res_inst_tac [("d", "case(%x. x, %y. y)")] lam_bijective 1); by (blast_tac (claset() addSIs [if_type]) 1); @@ -599,7 +599,7 @@ by (ALLGOALS (asm_simp_tac (simpset() setloop split_tac [expand_if]))); qed "bij_sum_Diff"; -goal OrderType.thy +Goal "!!i j. i le j ==> \ \ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = \ \ ordertype(j, Memrel(j))"; @@ -615,7 +615,7 @@ by (blast_tac (claset() addIs [lt_trans2, lt_trans]) 1); qed "ordertype_sum_Diff"; -goalw OrderType.thy [oadd_def, odiff_def] +Goalw [oadd_def, odiff_def] "!!i j. i le j ==> \ \ i ++ (j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))"; by (safe_tac (claset() addSDs [le_subset_iff RS iffD1])); @@ -626,12 +626,12 @@ Diff_subset] 1)); qed "oadd_ordertype_Diff"; -goal OrderType.thy "!!i j. i le j ==> i ++ (j--i) = j"; +Goal "!!i j. i le j ==> i ++ (j--i) = j"; by (asm_simp_tac (simpset() addsimps [oadd_ordertype_Diff, ordertype_sum_Diff, ordertype_Memrel, lt_Ord2 RS Ord_succD]) 1); qed "oadd_odiff_inverse"; -goalw OrderType.thy [odiff_def] +Goalw [odiff_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i--j)"; by (REPEAT (ares_tac [Ord_ordertype, well_ord_Memrel RS well_ord_subset, Diff_subset] 1)); @@ -639,7 +639,7 @@ (*By oadd_inject, the difference between i and j is unique. Note that we get i++j = k ==> j = k--i. *) -goal OrderType.thy +Goal "!!i j. [| Ord(i); Ord(j) |] ==> (i++j) -- i = j"; by (rtac oadd_inject 1); by (REPEAT (ares_tac [Ord_ordertype, Ord_oadd, Ord_odiff] 2)); @@ -659,14 +659,14 @@ (**** Ordinal Multiplication ****) -goalw OrderType.thy [omult_def] +Goalw [omult_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i**j)"; by (REPEAT (ares_tac [Ord_ordertype, well_ord_rmult, well_ord_Memrel] 1)); qed "Ord_omult"; (*** A useful unfolding law ***) -goalw OrderType.thy [pred_def] +Goalw [pred_def] "!!A B. [| a:A; b:B |] ==> \ \ pred(A*B, , rmult(A,r,B,s)) = \ \ pred(A,a,r)*B Un ({a} * pred(B,b,s))"; @@ -676,7 +676,7 @@ by (ALLGOALS (Blast_tac)); qed "pred_Pair_eq"; -goal OrderType.thy +Goal "!!A B. [| 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,a,r)*B + pred(B,b,s), \ @@ -688,7 +688,7 @@ by (blast_tac (claset() addSEs [predE]) 1); qed "ordertype_pred_Pair_eq"; -goalw OrderType.thy [oadd_def, omult_def] +Goalw [oadd_def, omult_def] "!!i j. [| i' \ \ ordertype(pred(i*j, , rmult(i,Memrel(i),j,Memrel(j))), \ \ rmult(i,Memrel(i),j,Memrel(j))) = \ @@ -707,7 +707,7 @@ by (ALLGOALS (blast_tac (claset() addIs [Ord_trans]))); qed "ordertype_pred_Pair_lemma"; -goalw OrderType.thy [omult_def] +Goalw [omult_def] "!!i j. [| Ord(i); Ord(j); k \ \ EX j' i'. k = j**i' ++ j' & j' j**i' ++ j' < j**i"; by (rtac ltI 1); by (asm_simp_tac @@ -734,7 +734,7 @@ symmetric omult_def]) 1); qed "omult_oadd_lt"; -goal OrderType.thy +Goal "!!i j. [| Ord(i); Ord(j) |] ==> j**i = (UN j':j. UN i':i. {j**i' ++ j'})"; by (rtac (subsetI RS equalityI) 1); by (resolve_tac [lt_omult RS exE] 1); @@ -748,11 +748,11 @@ (** Ordinal multiplication by zero **) -goalw OrderType.thy [omult_def] "i**0 = 0"; +Goalw [omult_def] "i**0 = 0"; by (Asm_simp_tac 1); qed "omult_0"; -goalw OrderType.thy [omult_def] "0**i = 0"; +Goalw [omult_def] "0**i = 0"; by (Asm_simp_tac 1); qed "omult_0_left"; @@ -760,7 +760,7 @@ (** Ordinal multiplication by 1 **) -goalw OrderType.thy [omult_def] "!!i. Ord(i) ==> i**1 = i"; +Goalw [omult_def] "!!i. Ord(i) ==> i**1 = i"; by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1); by (res_inst_tac [("c", "snd"), ("d", "%z.<0,z>")] lam_bijective 1); by (REPEAT_FIRST (eresolve_tac [snd_type, SigmaE, succE, emptyE, @@ -768,7 +768,7 @@ by (ALLGOALS (asm_simp_tac (simpset() addsimps [rmult_iff, Memrel_iff]))); qed "omult_1"; -goalw OrderType.thy [omult_def] "!!i. Ord(i) ==> 1**i = i"; +Goalw [omult_def] "!!i. Ord(i) ==> 1**i = i"; by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1); by (res_inst_tac [("c", "fst"), ("d", "%z.")] lam_bijective 1); by (REPEAT_FIRST (eresolve_tac [fst_type, SigmaE, succE, emptyE, @@ -780,7 +780,7 @@ (** Distributive law for ordinal multiplication and addition **) -goalw OrderType.thy [omult_def, oadd_def] +Goalw [omult_def, oadd_def] "!!i. [| Ord(i); Ord(j); Ord(k) |] ==> i**(j++k) = (i**j)++(i**k)"; by (resolve_tac [ordertype_eq RS trans] 1); by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS @@ -794,7 +794,7 @@ Ord_ordertype] 1)); qed "oadd_omult_distrib"; -goal OrderType.thy "!!i. [| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i"; +Goal "!!i. [| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i"; (*ZF_ss prevents looping*) by (asm_simp_tac (ZF_ss addsimps [oadd_1 RS sym]) 1); by (asm_simp_tac @@ -803,7 +803,7 @@ (** Associative law **) -goalw OrderType.thy [omult_def] +Goalw [omult_def] "!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> (i**j)**k = i**(j**k)"; by (resolve_tac [ordertype_eq RS trans] 1); by (rtac ([ord_iso_refl, ordertype_ord_iso RS ord_iso_sym] MRS @@ -826,7 +826,7 @@ by (Blast_tac 1); qed "omult_UN"; -goal OrderType.thy +Goal "!!i j. [| Ord(i); Limit(j) |] ==> i**j = (UN k:j. i**k)"; by (asm_simp_tac (simpset() addsimps [Limit_is_Ord RS Ord_in_Ord, omult_UN RS sym, @@ -837,19 +837,19 @@ (*** Ordering/monotonicity properties of ordinal multiplication ***) (*As a special case we have "[| 0 0 < i**j" *) -goal OrderType.thy "!!i j. [| k k < i**j"; +Goal "!!i j. [| k k < i**j"; by (safe_tac (claset() addSEs [ltE] addSIs [ltI, Ord_omult])); by (asm_simp_tac (simpset() addsimps [omult_unfold]) 1); by (REPEAT_FIRST (ares_tac [bexI])); by (Asm_simp_tac 1); qed "lt_omult1"; -goal OrderType.thy "!!i j. [| Ord(i); 0 i le i**j"; +Goal "!!i j. [| Ord(i); 0 i le i**j"; by (rtac all_lt_imp_le 1); by (REPEAT (ares_tac [Ord_omult, lt_omult1, lt_Ord2] 1)); qed "omult_le_self"; -goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> k**i le j**i"; +Goal "!!i j k. [| k le j; Ord(i) |] ==> k**i le j**i"; by (forward_tac [lt_Ord] 1); by (forward_tac [le_Ord2] 1); by (etac trans_induct3 1); @@ -860,7 +860,7 @@ by (Blast_tac 1); qed "omult_le_mono1"; -goal OrderType.thy "!!i j k. [| k i**k < i**j"; +Goal "!!i j k. [| k i**k < i**j"; by (rtac ltI 1); by (asm_simp_tac (simpset() addsimps [omult_unfold, lt_Ord2]) 1); by (safe_tac (claset() addSEs [ltE] addSIs [Ord_omult])); @@ -868,27 +868,27 @@ by (asm_simp_tac (simpset() addsimps [Ord_omult]) 1); qed "omult_lt_mono2"; -goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> i**k le i**j"; +Goal "!!i j k. [| k le j; Ord(i) |] ==> i**k le i**j"; by (rtac subset_imp_le 1); by (safe_tac (claset() addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult])); by (asm_full_simp_tac (simpset() addsimps [omult_unfold]) 1); by (deepen_tac (claset() addEs [Ord_trans]) 0 1); qed "omult_le_mono2"; -goal OrderType.thy "!!i j. [| i' le i; j' le j |] ==> i'**j' le i**j"; +Goal "!!i j. [| i' le i; j' le j |] ==> i'**j' le i**j"; by (rtac le_trans 1); by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_le_mono2, ltE, Ord_succD] 1)); qed "omult_le_mono"; -goal OrderType.thy +Goal "!!i j. [| i' le i; j' i'**j' < i**j"; by (rtac lt_trans1 1); by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_lt_mono2, ltE, Ord_succD] 1)); qed "omult_lt_mono"; -goal OrderType.thy "!!i j. [| Ord(i); 0 i le j**i"; +Goal "!!i j. [| Ord(i); 0 i le j**i"; by (forward_tac [lt_Ord2] 1); by (eres_inst_tac [("i","i")] trans_induct3 1); by (asm_simp_tac (simpset() addsimps [omult_0, Ord_0 RS le_refl]) 1); @@ -908,7 +908,7 @@ (** Further properties of ordinal multiplication **) -goal OrderType.thy "!!i j. [| i**j = i**k; 0 j=k"; +Goal "!!i j. [| i**j = i**k; 0 j=k"; by (rtac Ord_linear_lt 1); by (REPEAT_SOME assume_tac); by (ALLGOALS diff -r 30271d90644f -r 62b6288e6005 src/ZF/Ordinal.ML --- a/src/ZF/Ordinal.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Ordinal.ML Mon Jun 22 17:12:27 1998 +0200 @@ -12,17 +12,17 @@ (** Two neat characterisations of Transset **) -goalw Ordinal.thy [Transset_def] "Transset(A) <-> A<=Pow(A)"; +Goalw [Transset_def] "Transset(A) <-> A<=Pow(A)"; by (Blast_tac 1); qed "Transset_iff_Pow"; -goalw Ordinal.thy [Transset_def] "Transset(A) <-> Union(succ(A)) = A"; +Goalw [Transset_def] "Transset(A) <-> Union(succ(A)) = A"; by (blast_tac (claset() addSEs [equalityE]) 1); qed "Transset_iff_Union_succ"; (** Consequences of downwards closure **) -goalw Ordinal.thy [Transset_def] +Goalw [Transset_def] "!!C a b. [| Transset(C); {a,b}: C |] ==> a:C & b: C"; by (Blast_tac 1); qed "Transset_doubleton_D"; @@ -47,29 +47,29 @@ (** Closure properties **) -goalw Ordinal.thy [Transset_def] "Transset(0)"; +Goalw [Transset_def] "Transset(0)"; by (Blast_tac 1); qed "Transset_0"; -goalw Ordinal.thy [Transset_def] +Goalw [Transset_def] "!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Un j)"; by (Blast_tac 1); qed "Transset_Un"; -goalw Ordinal.thy [Transset_def] +Goalw [Transset_def] "!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Int j)"; by (Blast_tac 1); qed "Transset_Int"; -goalw Ordinal.thy [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))"; +Goalw [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))"; by (Blast_tac 1); qed "Transset_succ"; -goalw Ordinal.thy [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))"; +Goalw [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))"; by (Blast_tac 1); qed "Transset_Pow"; -goalw Ordinal.thy [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))"; +Goalw [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))"; by (Blast_tac 1); qed "Transset_Union"; @@ -103,7 +103,7 @@ (*** Lemmas for ordinals ***) -goalw Ordinal.thy [Ord_def,Transset_def] "!!i j.[| Ord(i); j:i |] ==> Ord(j)"; +Goalw [Ord_def,Transset_def] "!!i j.[| Ord(i); j:i |] ==> Ord(j)"; by (Blast_tac 1); qed "Ord_in_Ord"; @@ -112,31 +112,31 @@ AddSDs [Ord_succD]; -goal Ordinal.thy "!!i j. [| Ord(i); Transset(j); j<=i |] ==> Ord(j)"; +Goal "!!i j. [| Ord(i); Transset(j); j<=i |] ==> Ord(j)"; by (REPEAT (ares_tac [OrdI] 1 ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1)); qed "Ord_subset_Ord"; -goalw Ordinal.thy [Ord_def,Transset_def] "!!i j. [| j:i; Ord(i) |] ==> j<=i"; +Goalw [Ord_def,Transset_def] "!!i j. [| j:i; Ord(i) |] ==> j<=i"; by (Blast_tac 1); qed "OrdmemD"; -goal Ordinal.thy "!!i j k. [| i:j; j:k; Ord(k) |] ==> i:k"; +Goal "!!i j k. [| i:j; j:k; Ord(k) |] ==> i:k"; by (REPEAT (ares_tac [OrdmemD RS subsetD] 1)); qed "Ord_trans"; -goal Ordinal.thy "!!i j. [| i:j; Ord(j) |] ==> succ(i) <= j"; +Goal "!!i j. [| i:j; Ord(j) |] ==> succ(i) <= j"; by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1)); qed "Ord_succ_subsetI"; (*** The construction of ordinals: 0, succ, Union ***) -goal Ordinal.thy "Ord(0)"; +Goal "Ord(0)"; by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1)); qed "Ord_0"; -goal Ordinal.thy "!!i. Ord(i) ==> Ord(succ(i))"; +Goal "!!i. Ord(i) ==> Ord(succ(i))"; by (REPEAT (ares_tac [OrdI,Transset_succ] 1 ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset, Ord_contains_Transset] 1)); @@ -144,18 +144,18 @@ bind_thm ("Ord_1", Ord_0 RS Ord_succ); -goal Ordinal.thy "Ord(succ(i)) <-> Ord(i)"; +Goal "Ord(succ(i)) <-> Ord(i)"; by (blast_tac (claset() addIs [Ord_succ] addDs [Ord_succD]) 1); qed "Ord_succ_iff"; Addsimps [Ord_0, Ord_succ_iff]; AddSIs [Ord_0, Ord_succ]; -goalw Ordinal.thy [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Un j)"; +Goalw [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Un j)"; by (blast_tac (claset() addSIs [Transset_Un]) 1); qed "Ord_Un"; -goalw Ordinal.thy [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Int j)"; +Goalw [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Int j)"; by (blast_tac (claset() addSIs [Transset_Int]) 1); qed "Ord_Int"; @@ -176,7 +176,7 @@ qed "Ord_INT"; (*There is no set of all ordinals, for then it would contain itself*) -goal Ordinal.thy "~ (ALL i. i:X <-> Ord(i))"; +Goal "~ (ALL i. i:X <-> Ord(i))"; by (rtac notI 1); by (forw_inst_tac [("x", "X")] spec 1); by (safe_tac (claset() addSEs [mem_irrefl])); @@ -190,7 +190,7 @@ (*** < is 'less than' for ordinals ***) -goalw Ordinal.thy [lt_def] "!!i j. [| i:j; Ord(j) |] ==> i i i:j"; +Goal "!!i j. i i:j"; by (etac ltE 1); by (assume_tac 1); qed "ltD"; -goalw Ordinal.thy [lt_def] "~ i<0"; +Goalw [lt_def] "~ i<0"; by (Blast_tac 1); qed "not_lt0"; Addsimps [not_lt0]; -goal Ordinal.thy "!!i j. j Ord(j)"; +Goal "!!i j. j Ord(j)"; by (etac ltE 1 THEN assume_tac 1); qed "lt_Ord"; -goal Ordinal.thy "!!i j. j Ord(i)"; +Goal "!!i j. j Ord(i)"; by (etac ltE 1 THEN assume_tac 1); qed "lt_Ord2"; @@ -225,11 +225,11 @@ (* i<0 ==> R *) bind_thm ("lt0E", not_lt0 RS notE); -goal Ordinal.thy "!!i j k. [| i i i P"; +Goalw [lt_def] "!!i j. [| i P"; by (REPEAT (eresolve_tac [asm_rl, conjE, mem_asym] 1)); qed "lt_asym"; @@ -243,16 +243,16 @@ (** le is less than or equals; recall i le j abbrevs i i i i < succ(j)*) -goal Ordinal.thy "!!i j. i i le j"; +Goal "!!i j. i i le j"; by (asm_simp_tac (simpset() addsimps [le_iff]) 1); qed "leI"; -goal Ordinal.thy "!!i. [| i=j; Ord(j) |] ==> i le j"; +Goal "!!i. [| i=j; Ord(j) |] ==> i le j"; by (asm_simp_tac (simpset() addsimps [le_iff]) 1); qed "le_eqI"; @@ -269,12 +269,12 @@ by (DEPTH_SOLVE (ares_tac prems 1 ORELSE etac conjE 1)); qed "leE"; -goal Ordinal.thy "!!i j. [| i le j; j le i |] ==> i=j"; +Goal "!!i j. [| i le j; j le i |] ==> i=j"; by (asm_full_simp_tac (simpset() addsimps [le_iff]) 1); by (blast_tac (claset() addEs [lt_asym]) 1); qed "le_anti_sym"; -goal Ordinal.thy "i le 0 <-> i=0"; +Goal "i le 0 <-> i=0"; by (blast_tac (claset() addSIs [Ord_0 RS le_refl] addSEs [leE]) 1); qed "le0_iff"; @@ -290,11 +290,11 @@ (*** Natural Deduction rules for Memrel ***) -goalw Ordinal.thy [Memrel_def] " : Memrel(A) <-> a:b & a:A & b:A"; +Goalw [Memrel_def] " : Memrel(A) <-> a:b & a:A & b:A"; by (Blast_tac 1); qed "Memrel_iff"; -goal Ordinal.thy "!!A. [| a: b; a: A; b: A |] ==> : Memrel(A)"; +Goal "!!A. [| a: b; a: A; b: A |] ==> : Memrel(A)"; by (REPEAT (ares_tac [conjI, Memrel_iff RS iffD2] 1)); qed "MemrelI"; @@ -311,19 +311,19 @@ AddSIs [MemrelI]; AddSEs [MemrelE]; -goalw Ordinal.thy [Memrel_def] "Memrel(A) <= A*A"; +Goalw [Memrel_def] "Memrel(A) <= A*A"; by (Blast_tac 1); qed "Memrel_type"; -goalw Ordinal.thy [Memrel_def] "!!A B. A<=B ==> Memrel(A) <= Memrel(B)"; +Goalw [Memrel_def] "!!A B. A<=B ==> Memrel(A) <= Memrel(B)"; by (Blast_tac 1); qed "Memrel_mono"; -goalw Ordinal.thy [Memrel_def] "Memrel(0) = 0"; +Goalw [Memrel_def] "Memrel(0) = 0"; by (Blast_tac 1); qed "Memrel_0"; -goalw Ordinal.thy [Memrel_def] "Memrel(1) = 0"; +Goalw [Memrel_def] "Memrel(1) = 0"; by (Blast_tac 1); qed "Memrel_1"; @@ -331,7 +331,7 @@ (*The membership relation (as a set) is well-founded. Proof idea: show A<=B by applying the foundation axiom to A-B *) -goalw Ordinal.thy [wf_def] "wf(Memrel(A))"; +Goalw [wf_def] "wf(Memrel(A))"; by (EVERY1 [rtac (foundation RS disjE RS allI), etac disjI1, etac bexE, @@ -342,13 +342,13 @@ qed "wf_Memrel"; (*Transset(i) does not suffice, though ALL j:i.Transset(j) does*) -goalw Ordinal.thy [Ord_def, Transset_def, trans_def] +Goalw [Ord_def, Transset_def, trans_def] "!!i. Ord(i) ==> trans(Memrel(i))"; by (Blast_tac 1); qed "trans_Memrel"; (*If Transset(A) then Memrel(A) internalizes the membership relation below A*) -goalw Ordinal.thy [Transset_def] +Goalw [Transset_def] "!!A. Transset(A) ==> : Memrel(A) <-> a:b & b:A"; by (Blast_tac 1); qed "Transset_Memrel_iff"; @@ -425,11 +425,11 @@ by (DEPTH_SOLVE (ares_tac ([leI,le_eqI] @ prems) 1)); qed "Ord_linear_le"; -goal Ordinal.thy "!!i j. j le i ==> ~ i ~ i j le i"; +Goal "!!i j. [| ~ i j le i"; by (res_inst_tac [("i","i"),("j","j")] Ord_linear2 1); by (REPEAT (SOMEGOAL assume_tac)); by (blast_tac le_cs 1); @@ -437,25 +437,25 @@ (** Some rewrite rules for <, le **) -goalw Ordinal.thy [lt_def] "!!i j. Ord(j) ==> i:j <-> i i:j <-> i ~ i j le i"; +Goal "!!i j. [| Ord(i); Ord(j) |] ==> ~ i j le i"; by (REPEAT (ares_tac [iffI, le_imp_not_lt, not_lt_imp_le] 1)); qed "not_lt_iff_le"; -goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> ~ i le j <-> j ~ i le j <-> j 0 le i"; +Goal "!!i. Ord(i) ==> 0 le i"; by (etac (not_lt_iff_le RS iffD1) 1); by (REPEAT (resolve_tac [Ord_0, not_lt0] 1)); qed "Ord_0_le"; -goal Ordinal.thy "!!i. [| Ord(i); i~=0 |] ==> 0 0 j le i"; +Goal "!!i j. [| j<=i; Ord(i); Ord(j) |] ==> j le i"; by (rtac (not_lt_iff_le RS iffD1) 1); by (assume_tac 1); by (assume_tac 1); by (blast_tac (claset() addEs [ltE, mem_irrefl]) 1); qed "subset_imp_le"; -goal Ordinal.thy "!!i j. i le j ==> i<=j"; +Goal "!!i j. i le j ==> i<=j"; by (etac leE 1); by (Blast_tac 2); by (blast_tac (subset_cs addIs [OrdmemD] addEs [ltE]) 1); qed "le_imp_subset"; -goal Ordinal.thy "j le i <-> j<=i & Ord(i) & Ord(j)"; +Goal "j le i <-> j<=i & Ord(i) & Ord(j)"; by (blast_tac (claset() addDs [Ord_succD, subset_imp_le, le_imp_subset] addEs [ltE]) 1); qed "le_subset_iff"; -goal Ordinal.thy "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"; +Goal "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"; by (simp_tac (simpset() addsimps [le_iff]) 1); by (blast_tac (claset() addIs [Ord_succ] addDs [Ord_succD]) 1); qed "le_succ_iff"; @@ -497,55 +497,55 @@ (** Transitive laws **) -goal Ordinal.thy "!!i j. [| i le j; j i i i i i le k"; +Goal "!!i j. [| i le j; j le k |] ==> i le k"; by (REPEAT (ares_tac [lt_trans1] 1)); qed "le_trans"; -goal Ordinal.thy "!!i j. i succ(i) le j"; +Goal "!!i j. i succ(i) le j"; by (rtac (not_lt_iff_le RS iffD1) 1); by (blast_tac le_cs 3); by (ALLGOALS (blast_tac (claset() addEs [ltE]))); qed "succ_leI"; (*Identical to succ(i) < succ(j) ==> i i i i i i le j"; +Goal "!!i j. succ(i) le succ(j) ==> i le j"; by (blast_tac (claset() addSDs [succ_leE]) 1); qed "succ_le_imp_le"; (** Union and Intersection **) -goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> i le i Un j"; +Goal "!!i j. [| Ord(i); Ord(j) |] ==> i le i Un j"; by (rtac (Un_upper1 RS subset_imp_le) 1); by (REPEAT (ares_tac [Ord_Un] 1)); qed "Un_upper1_le"; -goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> j le i Un j"; +Goal "!!i j. [| Ord(i); Ord(j) |] ==> j le i Un j"; by (rtac (Un_upper2 RS subset_imp_le) 1); by (REPEAT (ares_tac [Ord_Un] 1)); qed "Un_upper2_le"; (*Replacing k by succ(k') yields the similar rule for le!*) -goal Ordinal.thy "!!i j k. [| i i Un j < k"; +Goal "!!i j k. [| i i Un j < k"; by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1); by (stac Un_commute 4); by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Un_iff]) 4); @@ -553,7 +553,7 @@ by (REPEAT (eresolve_tac [asm_rl, ltE] 1)); qed "Un_least_lt"; -goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> i Un j < k <-> i i Un j < k <-> i i Int j < k"; +Goal "!!i j k. [| i i Int j < k"; by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1); by (stac Int_commute 4); by (asm_full_simp_tac (simpset() addsimps [le_subset_iff, subset_Int_iff]) 4); @@ -625,35 +625,35 @@ ORELSE dtac Ord_succD 1)); qed "le_implies_UN_le_UN"; -goal Ordinal.thy "!!i. Ord(i) ==> (UN y:i. succ(y)) = i"; +Goal "!!i. Ord(i) ==> (UN y:i. succ(y)) = i"; by (blast_tac (claset() addIs [Ord_trans]) 1); qed "Ord_equality"; (*Holds for all transitive sets, not just ordinals*) -goal Ordinal.thy "!!i. Ord(i) ==> Union(i) <= i"; +Goal "!!i. Ord(i) ==> Union(i) <= i"; by (blast_tac (claset() addIs [Ord_trans]) 1); qed "Ord_Union_subset"; (*** Limit ordinals -- general properties ***) -goalw Ordinal.thy [Limit_def] "!!i. Limit(i) ==> Union(i) = i"; +Goalw [Limit_def] "!!i. Limit(i) ==> Union(i) = i"; by (fast_tac (claset() addSIs [ltI] addSEs [ltE] addEs [Ord_trans]) 1); qed "Limit_Union_eq"; -goalw Ordinal.thy [Limit_def] "!!i. Limit(i) ==> Ord(i)"; +Goalw [Limit_def] "!!i. Limit(i) ==> Ord(i)"; by (etac conjunct1 1); qed "Limit_is_Ord"; -goalw Ordinal.thy [Limit_def] "!!i. Limit(i) ==> 0 < i"; +Goalw [Limit_def] "!!i. Limit(i) ==> 0 < i"; by (etac (conjunct2 RS conjunct1) 1); qed "Limit_has_0"; -goalw Ordinal.thy [Limit_def] "!!i. [| Limit(i); j succ(j) < i"; +Goalw [Limit_def] "!!i. [| Limit(i); j succ(j) < i"; by (Blast_tac 1); qed "Limit_has_succ"; -goalw Ordinal.thy [Limit_def] +Goalw [Limit_def] "!!i. [| 0 Limit(i)"; by (safe_tac subset_cs); by (rtac (not_le_iff_lt RS iffD1) 2); @@ -661,20 +661,20 @@ by (REPEAT (eresolve_tac [asm_rl, ltE, Ord_succ] 1)); qed "non_succ_LimitI"; -goal Ordinal.thy "!!i. Limit(succ(i)) ==> P"; +Goal "!!i. Limit(succ(i)) ==> P"; by (rtac lt_irrefl 1); by (rtac Limit_has_succ 1); by (assume_tac 1); by (etac (Limit_is_Ord RS Ord_succD RS le_refl) 1); qed "succ_LimitE"; -goal Ordinal.thy "!!i. [| Limit(i); i le succ(j) |] ==> i le j"; +Goal "!!i. [| Limit(i); i le succ(j) |] ==> i le j"; by (safe_tac (claset() addSEs [succ_LimitE, leE])); qed "Limit_le_succD"; (** Traditional 3-way case analysis on ordinals **) -goal Ordinal.thy "!!i. Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)"; +Goal "!!i. Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)"; by (blast_tac (claset() addSIs [non_succ_LimitI, Ord_0_lt] addSDs [Ord_succD]) 1); qed "Ord_cases_disj"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Perm.ML --- a/src/ZF/Perm.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Perm.ML Mon Jun 22 17:12:27 1998 +0200 @@ -13,15 +13,15 @@ (** Surjective function space **) -goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B"; +Goalw [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B"; by (etac CollectD1 1); qed "surj_is_fun"; -goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))"; +Goalw [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))"; by (blast_tac (claset() addIs [apply_equality, range_of_fun, domain_type]) 1); qed "fun_is_surj"; -goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B"; +Goalw [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B"; by (best_tac (claset() addIs [apply_Pair] addEs [range_type]) 1); qed "surj_range"; @@ -43,7 +43,7 @@ qed "lam_surjective"; (*Cantor's theorem revisited*) -goalw Perm.thy [surj_def] "f ~: surj(A,Pow(A))"; +Goalw [surj_def] "f ~: surj(A,Pow(A))"; by Safe_tac; by (cut_facts_tac [cantor] 1); by (fast_tac subset_cs 1); @@ -52,24 +52,24 @@ (** Injective function space **) -goalw Perm.thy [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B"; +Goalw [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B"; by (etac CollectD1 1); qed "inj_is_fun"; (*Good for dealing with sets of pairs, but a bit ugly in use [used in AC]*) -goalw Perm.thy [inj_def] +Goalw [inj_def] "!!f A B. [| :f; :f; f: inj(A,B) |] ==> a=c"; by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1)); by (Blast_tac 1); qed "inj_equality"; -goalw thy [inj_def] "!!A B f. [| f:inj(A,B); a:A; b:A; f`a=f`b |] ==> a=b"; +Goalw [inj_def] "!!A B f. [| f:inj(A,B); a:A; b:A; f`a=f`b |] ==> a=b"; by (Blast_tac 1); val inj_apply_equality = result(); (** A function with a left inverse is an injection **) -goal Perm.thy "!!f. [| f: A->B; ALL x:A. d(f`x)=x |] ==> f: inj(A,B)"; +Goal "!!f. [| f: A->B; ALL x:A. d(f`x)=x |] ==> f: inj(A,B)"; by (asm_simp_tac (simpset() addsimps [inj_def]) 1); by (deepen_tac (claset() addEs [subst_context RS box_equals]) 0 1); bind_thm ("f_imp_injective", ballI RSN (2,result())); @@ -84,11 +84,11 @@ (** Bijections **) -goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)"; +Goalw [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)"; by (etac IntD1 1); qed "bij_is_inj"; -goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)"; +Goalw [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)"; by (etac IntD2 1); qed "bij_is_surj"; @@ -118,12 +118,12 @@ by (REPEAT (ares_tac prems 1)); qed "idE"; -goalw Perm.thy [id_def] "id(A) : A->A"; +Goalw [id_def] "id(A) : A->A"; by (rtac lam_type 1); by (assume_tac 1); qed "id_type"; -goalw Perm.thy [id_def] "!!A x. x:A ==> id(A)`x = x"; +Goalw [id_def] "!!A x. x:A ==> id(A)`x = x"; by (Asm_simp_tac 1); qed "id_conv"; @@ -133,7 +133,7 @@ by (rtac (prem RS lam_mono) 1); qed "id_mono"; -goalw Perm.thy [inj_def,id_def] "!!A B. A<=B ==> id(A): inj(A,B)"; +Goalw [inj_def,id_def] "!!A B. A<=B ==> id(A): inj(A,B)"; by (REPEAT (ares_tac [CollectI,lam_type] 1)); by (etac subsetD 1 THEN assume_tac 1); by (Simp_tac 1); @@ -141,15 +141,15 @@ val id_inj = subset_refl RS id_subset_inj; -goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)"; +Goalw [id_def,surj_def] "id(A): surj(A,A)"; by (blast_tac (claset() addIs [lam_type, beta]) 1); qed "id_surj"; -goalw Perm.thy [bij_def] "id(A): bij(A,A)"; +Goalw [bij_def] "id(A): bij(A,A)"; by (blast_tac (claset() addIs [id_inj, id_surj]) 1); qed "id_bij"; -goalw Perm.thy [id_def] "A <= B <-> id(A) : A->B"; +Goalw [id_def] "A <= B <-> id(A) : A->B"; by (fast_tac (claset() addSIs [lam_type] addDs [apply_type] addss (simpset())) 1); qed "subset_iff_id"; @@ -157,7 +157,7 @@ (*** Converse of a function ***) -goalw Perm.thy [inj_def] "!!f. f: inj(A,B) ==> converse(f) : range(f)->A"; +Goalw [inj_def] "!!f. f: inj(A,B) ==> converse(f) : range(f)->A"; by (asm_simp_tac (simpset() addsimps [Pi_iff, function_def]) 1); by (etac CollectE 1); by (asm_simp_tac (simpset() addsimps [apply_iff]) 1); @@ -167,12 +167,12 @@ (** Equations for converse(f) **) (*The premises are equivalent to saying that f is injective...*) -goal Perm.thy +Goal "!!f. [| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a"; by (blast_tac (claset() addIs [apply_Pair, apply_equality, converseI]) 1); qed "left_inverse_lemma"; -goal Perm.thy +Goal "!!f. [| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a"; by (blast_tac (claset() addIs [left_inverse_lemma, inj_converse_fun, inj_is_fun]) 1); @@ -188,7 +188,7 @@ (*Should the premises be f:surj(A,B), b:B for symmetry with left_inverse? No: they would not imply that converse(f) was a function! *) -goal Perm.thy "!!f. [| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b"; +Goal "!!f. [| f: inj(A,B); b: range(f) |] ==> f`(converse(f)`b) = b"; by (rtac right_inverse_lemma 1); by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1)); qed "right_inverse"; @@ -196,26 +196,26 @@ (*Cannot add [left_inverse, right_inverse] to default simpset: there are too many ways of expressing sufficient conditions.*) -goal Perm.thy "!!f. [| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b"; +Goal "!!f. [| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b"; by (fast_tac (claset() addss (simpset() addsimps [bij_def, right_inverse, surj_range])) 1); qed "right_inverse_bij"; (** Converses of injections, surjections, bijections **) -goal Perm.thy "!!f A B. f: inj(A,B) ==> converse(f): inj(range(f), A)"; +Goal "!!f A B. f: inj(A,B) ==> converse(f): inj(range(f), A)"; by (rtac f_imp_injective 1); by (etac inj_converse_fun 1); by (rtac right_inverse 1); by (REPEAT (assume_tac 1)); qed "inj_converse_inj"; -goal Perm.thy "!!f A B. f: inj(A,B) ==> converse(f): surj(range(f), A)"; +Goal "!!f A B. f: inj(A,B) ==> converse(f): surj(range(f), A)"; by (blast_tac (claset() addIs [f_imp_surjective, inj_converse_fun, left_inverse, inj_is_fun, range_of_fun RS apply_type]) 1); qed "inj_converse_surj"; -goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)"; +Goalw [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)"; by (fast_tac (claset() addEs [surj_range RS subst, inj_converse_inj, inj_converse_surj]) 1); qed "bij_converse_bij"; @@ -226,7 +226,7 @@ (*The inductive definition package could derive these theorems for (r O s)*) -goalw Perm.thy [comp_def] "!!r s. [| :s; :r |] ==> : r O s"; +Goalw [comp_def] "!!r s. [| :s; :r |] ==> : r O s"; by (Blast_tac 1); qed "compI"; @@ -250,68 +250,68 @@ (** Domain and Range -- see Suppes, section 3.1 **) (*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*) -goal Perm.thy "range(r O s) <= range(r)"; +Goal "range(r O s) <= range(r)"; by (Blast_tac 1); qed "range_comp"; -goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)"; +Goal "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)"; by (rtac (range_comp RS equalityI) 1); by (Blast_tac 1); qed "range_comp_eq"; -goal Perm.thy "domain(r O s) <= domain(s)"; +Goal "domain(r O s) <= domain(s)"; by (Blast_tac 1); qed "domain_comp"; -goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)"; +Goal "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)"; by (rtac (domain_comp RS equalityI) 1); by (Blast_tac 1); qed "domain_comp_eq"; -goal Perm.thy "(r O s)``A = r``(s``A)"; +Goal "(r O s)``A = r``(s``A)"; by (Blast_tac 1); qed "image_comp"; (** Other results **) -goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; +Goal "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; by (Blast_tac 1); qed "comp_mono"; (*composition preserves relations*) -goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"; +Goal "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C"; by (Blast_tac 1); qed "comp_rel"; (*associative law for composition*) -goal Perm.thy "(r O s) O t = r O (s O t)"; +Goal "(r O s) O t = r O (s O t)"; by (Blast_tac 1); qed "comp_assoc"; (*left identity of composition; provable inclusions are id(A) O r <= r and [| r<=A*B; B<=C |] ==> r <= id(C) O r *) -goal Perm.thy "!!r A B. r<=A*B ==> id(B) O r = r"; +Goal "!!r A B. r<=A*B ==> id(B) O r = r"; by (Blast_tac 1); qed "left_comp_id"; (*right identity of composition; provable inclusions are r O id(A) <= r and [| r<=A*B; A<=C |] ==> r <= r O id(C) *) -goal Perm.thy "!!r A B. r<=A*B ==> r O id(A) = r"; +Goal "!!r A B. r<=A*B ==> r O id(A) = r"; by (Blast_tac 1); qed "right_comp_id"; (** Composition preserves functions, injections, and surjections **) -goalw Perm.thy [function_def] +Goalw [function_def] "!!f g. [| function(g); function(f) |] ==> function(f O g)"; by (Blast_tac 1); qed "comp_function"; -goal Perm.thy "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C"; +Goal "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C"; by (asm_full_simp_tac (simpset() addsimps [Pi_def, comp_function, Pow_iff, comp_rel] setloop etac conjE) 1); @@ -319,7 +319,7 @@ by (Blast_tac 1); qed "comp_fun"; -goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; +Goal "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; by (REPEAT (ares_tac [comp_fun,apply_equality,compI, apply_Pair,apply_type] 1)); qed "comp_fun_apply"; @@ -338,7 +338,7 @@ setSolver type_auto_tac [lam_type, lam_funtype, prem]) 1); qed "comp_lam"; -goal Perm.thy "!!f g. [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)"; +Goal "!!f g. [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)"; by (res_inst_tac [("d", "%y. converse(g) ` (converse(f) ` y)")] f_imp_injective 1); by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1)); @@ -346,12 +346,12 @@ setSolver type_auto_tac [inj_is_fun, apply_type]) 1); qed "comp_inj"; -goalw Perm.thy [surj_def] +Goalw [surj_def] "!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)"; by (blast_tac (claset() addSIs [comp_fun,comp_fun_apply]) 1); qed "comp_surj"; -goalw Perm.thy [bij_def] +Goalw [bij_def] "!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)"; by (blast_tac (claset() addIs [comp_inj,comp_surj]) 1); qed "comp_bij"; @@ -361,14 +361,14 @@ D Pastre. Automatic theorem proving in set theory. Artificial Intelligence, 10:1--27, 1978. **) -goalw Perm.thy [inj_def] +Goalw [inj_def] "!!f g. [| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)"; by Safe_tac; by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1)); by (asm_simp_tac (simpset() ) 1); qed "comp_mem_injD1"; -goalw Perm.thy [inj_def,surj_def] +Goalw [inj_def,surj_def] "!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)"; by Safe_tac; by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1); @@ -380,17 +380,17 @@ by (asm_simp_tac (simpset() ) 1); qed "comp_mem_injD2"; -goalw Perm.thy [surj_def] +Goalw [surj_def] "!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)"; by (blast_tac (claset() addSIs [comp_fun_apply RS sym, apply_funtype]) 1); qed "comp_mem_surjD1"; -goal Perm.thy +Goal "!!f g. [| (f O g)`a = c; g: A->B; f: B->C; a:A |] ==> f`(g`a) = c"; by (REPEAT (ares_tac [comp_fun_apply RS sym RS trans] 1)); qed "comp_fun_applyD"; -goalw Perm.thy [inj_def,surj_def] +Goalw [inj_def,surj_def] "!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)"; by Safe_tac; by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1); @@ -403,7 +403,7 @@ (*left inverse of composition; one inclusion is f: A->B ==> id(A) <= converse(f) O f *) -goalw Perm.thy [inj_def] "!!f. f: inj(A,B) ==> converse(f) O f = id(A)"; +Goalw [inj_def] "!!f. f: inj(A,B) ==> converse(f) O f = id(A)"; by (fast_tac (claset() addIs [apply_Pair] addEs [domain_type] addss (simpset() addsimps [apply_iff])) 1); @@ -423,7 +423,7 @@ (** Proving that a function is a bijection **) -goalw Perm.thy [id_def] +Goalw [id_def] "!!f A B. [| f: A->B; g: B->A |] ==> \ \ f O g = id(B) <-> (ALL y:B. f`(g`y)=y)"; by Safe_tac; @@ -434,7 +434,7 @@ by Auto_tac; qed "comp_eq_id_iff"; -goalw Perm.thy [bij_def] +Goalw [bij_def] "!!f A B. [| f: A->B; g: B->A; f O g = id(B); g O f = id(A) \ \ |] ==> f : bij(A,B)"; by (asm_full_simp_tac (simpset() addsimps [comp_eq_id_iff]) 1); @@ -442,11 +442,11 @@ ORELSE eresolve_tac [bspec, apply_type] 1)); qed "fg_imp_bijective"; -goal Perm.thy "!!f A. [| f: A->A; f O f = id(A) |] ==> f : bij(A,A)"; +Goal "!!f A. [| f: A->A; f O f = id(A) |] ==> f : bij(A,A)"; by (REPEAT (ares_tac [fg_imp_bijective] 1)); qed "nilpotent_imp_bijective"; -goal Perm.thy "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; +Goal "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; by (asm_simp_tac (simpset() addsimps [fg_imp_bijective, comp_eq_id_iff, left_inverse_lemma, right_inverse_lemma]) 1); qed "invertible_imp_bijective"; @@ -454,7 +454,7 @@ (** Unions of functions -- cf similar theorems on func.ML **) (*Theorem by KG, proof by LCP*) -goal Perm.thy +Goal "!!f g. [| f: inj(A,B); g: inj(C,D); B Int D = 0 |] ==> \ \ (lam a: A Un C. if(a:A, f`a, g`a)) : inj(A Un C, B Un D)"; by (res_inst_tac [("d","%z. if(z:B, converse(f)`z, converse(g)`z)")] @@ -465,7 +465,7 @@ by (blast_tac (claset() addIs [inj_is_fun RS apply_type] addDs [equals0D]) 1); qed "inj_disjoint_Un"; -goalw Perm.thy [surj_def] +Goalw [surj_def] "!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \ \ (f Un g) : surj(A Un C, B Un D)"; by (blast_tac (claset() addIs [fun_disjoint_apply1, fun_disjoint_apply2, @@ -474,7 +474,7 @@ (*A simple, high-level proof; the version for injections follows from it, using f:inj(A,B) <-> f:bij(A,range(f)) *) -goal Perm.thy +Goal "!!f g. [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==> \ \ (f Un g) : bij(A Un C, B Un D)"; by (rtac invertible_imp_bijective 1); @@ -491,7 +491,7 @@ by (fast_tac (claset() addIs rls) 1); qed "surj_image"; -goal Perm.thy "!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A"; +Goal "!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A"; by (rtac equalityI 1); by (SELECT_GOAL (rewtac restrict_def) 2); by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2 @@ -499,7 +499,7 @@ by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1)); qed "restrict_image"; -goalw Perm.thy [inj_def] +Goalw [inj_def] "!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)"; by (safe_tac (claset() addSEs [restrict_type2])); by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD, @@ -513,7 +513,7 @@ by (REPEAT (resolve_tac prems 1)); qed "restrict_surj"; -goalw Perm.thy [inj_def,bij_def] +Goalw [inj_def,bij_def] "!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)"; by (blast_tac (claset() addSIs [restrict, restrict_surj] addIs [box_equals, surj_is_fun]) 1); @@ -522,7 +522,7 @@ (*** Lemmas for Ramsey's Theorem ***) -goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)"; +Goalw [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)"; by (blast_tac (claset() addIs [fun_weaken_type]) 1); qed "inj_weaken_type"; @@ -536,7 +536,7 @@ addDs [apply_equality]) 1); qed "inj_succ_restrict"; -goalw Perm.thy [inj_def] +Goalw [inj_def] "!!f. [| f: inj(A,B); a~:A; b~:B |] ==> \ \ cons(,f) : inj(cons(a,A), cons(b,B))"; by (fast_tac (claset() addIs [apply_type] diff -r 30271d90644f -r 62b6288e6005 src/ZF/QPair.ML --- a/src/ZF/QPair.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/QPair.ML Mon Jun 22 17:12:27 1998 +0200 @@ -116,7 +116,7 @@ "!!p. p:QSigma(A,B) ==> qsnd(p) : B(qfst(p))" (fn _=> [ Auto_tac ]); -goal thy "!!a A B. a: QSigma(A,B) ==> = a"; +Goal "!!a A B. a: QSigma(A,B) ==> = a"; by Auto_tac; qed "QPair_qfst_qsnd_eq"; @@ -139,7 +139,7 @@ [ (rtac (major RS QSigmaE) 1), (asm_simp_tac (simpset() addsimps prems) 1) ]); -goalw thy [qsplit_def] +Goalw [qsplit_def] "!!u. u: A<*>B ==> \ \ R(qsplit(c,u)) <-> (ALL x:A. ALL y:B. u = --> R(c(x,y)))"; by Auto_tac; @@ -148,7 +148,7 @@ (*** qsplit for predicates: result type o ***) -goalw thy [qsplit_def] "!!R a b. R(a,b) ==> qsplit(R, )"; +Goalw [qsplit_def] "!!R a b. R(a,b) ==> qsplit(R, )"; by (Asm_simp_tac 1); qed "qsplitI"; @@ -162,7 +162,7 @@ by (Asm_full_simp_tac 1); qed "qsplitE"; -goalw thy [qsplit_def] "!!R a b. qsplit(R,) ==> R(a,b)"; +Goalw [qsplit_def] "!!R a b. qsplit(R,) ==> R(a,b)"; by (Asm_full_simp_tac 1); qed "qsplitD"; @@ -211,11 +211,11 @@ (** Introduction rules for the injections **) -goalw thy qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B"; +Goalw qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B"; by (Blast_tac 1); qed "QInlI"; -goalw thy qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B"; +Goalw qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B"; by (Blast_tac 1); qed "QInrI"; @@ -235,23 +235,23 @@ (** Injection and freeness equivalences, for rewriting **) -goalw thy qsum_defs "QInl(a)=QInl(b) <-> a=b"; +Goalw qsum_defs "QInl(a)=QInl(b) <-> a=b"; by (Simp_tac 1); qed "QInl_iff"; -goalw thy qsum_defs "QInr(a)=QInr(b) <-> a=b"; +Goalw qsum_defs "QInr(a)=QInr(b) <-> a=b"; by (Simp_tac 1); qed "QInr_iff"; -goalw thy qsum_defs "QInl(a)=QInr(b) <-> False"; +Goalw qsum_defs "QInl(a)=QInr(b) <-> False"; by (Simp_tac 1); qed "QInl_QInr_iff"; -goalw thy qsum_defs "QInr(b)=QInl(a) <-> False"; +Goalw qsum_defs "QInr(b)=QInl(a) <-> False"; by (Simp_tac 1); qed "QInr_QInl_iff"; -goalw thy qsum_defs "0<+>0 = 0"; +Goalw qsum_defs "0<+>0 = 0"; by (Simp_tac 1); qed "qsum_empty"; @@ -266,37 +266,37 @@ AddSDs [QInl_inject, QInr_inject]; Addsimps [QInl_iff, QInr_iff, QInl_QInr_iff, QInr_QInl_iff, qsum_empty]; -goal thy "!!A B. QInl(a): A<+>B ==> a: A"; +Goal "!!A B. QInl(a): A<+>B ==> a: A"; by (Blast_tac 1); qed "QInlD"; -goal thy "!!A B. QInr(b): A<+>B ==> b: B"; +Goal "!!A B. QInr(b): A<+>B ==> b: B"; by (Blast_tac 1); qed "QInrD"; (** <+> is itself injective... who cares?? **) -goal thy +Goal "u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))"; by (Blast_tac 1); qed "qsum_iff"; -goal thy "A <+> B <= C <+> D <-> A<=C & B<=D"; +Goal "A <+> B <= C <+> D <-> A<=C & B<=D"; by (Blast_tac 1); qed "qsum_subset_iff"; -goal thy "A <+> B = C <+> D <-> A=C & B=D"; +Goal "A <+> B = C <+> D <-> A=C & B=D"; by (simp_tac (simpset() addsimps [extension,qsum_subset_iff]) 1); by (Blast_tac 1); qed "qsum_equal_iff"; (*** Eliminator -- qcase ***) -goalw thy qsum_defs "qcase(c, d, QInl(a)) = c(a)"; +Goalw qsum_defs "qcase(c, d, QInl(a)) = c(a)"; by (Simp_tac 1); qed "qcase_QInl"; -goalw thy qsum_defs "qcase(c, d, QInr(b)) = d(b)"; +Goalw qsum_defs "qcase(c, d, QInr(b)) = d(b)"; by (Simp_tac 1); qed "qcase_QInr"; @@ -314,18 +314,18 @@ (** Rules for the Part primitive **) -goal thy "Part(A <+> B,QInl) = {QInl(x). x: A}"; +Goal "Part(A <+> B,QInl) = {QInl(x). x: A}"; by (Blast_tac 1); qed "Part_QInl"; -goal thy "Part(A <+> B,QInr) = {QInr(y). y: B}"; +Goal "Part(A <+> B,QInr) = {QInr(y). y: B}"; by (Blast_tac 1); qed "Part_QInr"; -goal thy "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y: Part(B,h)}"; +Goal "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y: Part(B,h)}"; by (Blast_tac 1); qed "Part_QInr2"; -goal thy "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C"; +Goal "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C"; by (Blast_tac 1); qed "Part_qsum_equality"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/QUniv.ML --- a/src/ZF/QUniv.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/QUniv.ML Mon Jun 22 17:12:27 1998 +0200 @@ -25,36 +25,36 @@ (** Introduction and elimination rules avoid tiresome folding/unfolding **) -goalw QUniv.thy [quniv_def] +Goalw [quniv_def] "!!X A. X <= univ(eclose(A)) ==> X : quniv(A)"; by (etac PowI 1); qed "qunivI"; -goalw QUniv.thy [quniv_def] +Goalw [quniv_def] "!!X A. X : quniv(A) ==> X <= univ(eclose(A))"; by (etac PowD 1); qed "qunivD"; -goalw QUniv.thy [quniv_def] "!!A B. A<=B ==> quniv(A) <= quniv(B)"; +Goalw [quniv_def] "!!A B. A<=B ==> quniv(A) <= quniv(B)"; by (etac (eclose_mono RS univ_mono RS Pow_mono) 1); qed "quniv_mono"; (*** Closure properties ***) -goalw QUniv.thy [quniv_def] "univ(eclose(A)) <= quniv(A)"; +Goalw [quniv_def] "univ(eclose(A)) <= quniv(A)"; by (rtac (Transset_iff_Pow RS iffD1) 1); by (rtac (Transset_eclose RS Transset_univ) 1); qed "univ_eclose_subset_quniv"; (*Key property for proving A_subset_quniv; requires eclose in def of quniv*) -goal QUniv.thy "univ(A) <= quniv(A)"; +Goal "univ(A) <= quniv(A)"; by (rtac (arg_subset_eclose RS univ_mono RS subset_trans) 1); by (rtac univ_eclose_subset_quniv 1); qed "univ_subset_quniv"; bind_thm ("univ_into_quniv", univ_subset_quniv RS subsetD); -goalw QUniv.thy [quniv_def] "Pow(univ(A)) <= quniv(A)"; +Goalw [quniv_def] "Pow(univ(A)) <= quniv(A)"; by (rtac (arg_subset_eclose RS univ_mono RS Pow_mono) 1); qed "Pow_univ_subset_quniv"; @@ -73,14 +73,14 @@ (*** univ(A) closure for Quine-inspired pairs and injections ***) (*Quine ordered pairs*) -goalw QUniv.thy [QPair_def] +Goalw [QPair_def] "!!A a. [| a <= univ(A); b <= univ(A) |] ==> <= univ(A)"; by (REPEAT (ares_tac [sum_subset_univ] 1)); qed "QPair_subset_univ"; (** Quine disjoint sum **) -goalw QUniv.thy [QInl_def] "!!A a. a <= univ(A) ==> QInl(a) <= univ(A)"; +Goalw [QInl_def] "!!A a. a <= univ(A) ==> QInl(a) <= univ(A)"; by (etac (empty_subsetI RS QPair_subset_univ) 1); qed "QInl_subset_univ"; @@ -91,20 +91,20 @@ val naturals_subset_univ = [naturals_subset_nat, nat_subset_univ] MRS subset_trans; -goalw QUniv.thy [QInr_def] "!!A a. a <= univ(A) ==> QInr(a) <= univ(A)"; +Goalw [QInr_def] "!!A a. a <= univ(A) ==> QInr(a) <= univ(A)"; by (etac (nat_1I RS naturals_subset_univ RS QPair_subset_univ) 1); qed "QInr_subset_univ"; (*** Closure for Quine-inspired products and sums ***) (*Quine ordered pairs*) -goalw QUniv.thy [quniv_def,QPair_def] +Goalw [quniv_def,QPair_def] "!!A a. [| a: quniv(A); b: quniv(A) |] ==> : quniv(A)"; by (REPEAT (dtac PowD 1)); by (REPEAT (ares_tac [PowI, sum_subset_univ] 1)); qed "QPair_in_quniv"; -goal QUniv.thy "quniv(A) <*> quniv(A) <= quniv(A)"; +Goal "quniv(A) <*> quniv(A) <= quniv(A)"; by (REPEAT (ares_tac [subsetI, QPair_in_quniv] 1 ORELSE eresolve_tac [QSigmaE, ssubst] 1)); qed "QSigma_quniv"; @@ -113,7 +113,7 @@ [QSigma_mono, QSigma_quniv] MRS subset_trans); (*The opposite inclusion*) -goalw QUniv.thy [quniv_def,QPair_def] +Goalw [quniv_def,QPair_def] "!!A a b. : quniv(A) ==> a: quniv(A) & b: quniv(A)"; by (etac ([Transset_eclose RS Transset_univ, PowD] MRS Transset_includes_summands RS conjE) 1); @@ -122,7 +122,7 @@ bind_thm ("quniv_QPair_E", quniv_QPair_D RS conjE); -goal QUniv.thy " : quniv(A) <-> a: quniv(A) & b: quniv(A)"; +Goal " : quniv(A) <-> a: quniv(A) & b: quniv(A)"; by (REPEAT (ares_tac [iffI, QPair_in_quniv, quniv_QPair_D] 1 ORELSE etac conjE 1)); qed "quniv_QPair_iff"; @@ -130,15 +130,15 @@ (** Quine disjoint sum **) -goalw QUniv.thy [QInl_def] "!!A a. a: quniv(A) ==> QInl(a) : quniv(A)"; +Goalw [QInl_def] "!!A a. a: quniv(A) ==> QInl(a) : quniv(A)"; by (REPEAT (ares_tac [zero_in_quniv,QPair_in_quniv] 1)); qed "QInl_in_quniv"; -goalw QUniv.thy [QInr_def] "!!A b. b: quniv(A) ==> QInr(b) : quniv(A)"; +Goalw [QInr_def] "!!A b. b: quniv(A) ==> QInr(b) : quniv(A)"; by (REPEAT (ares_tac [one_in_quniv, QPair_in_quniv] 1)); qed "QInr_in_quniv"; -goal QUniv.thy "quniv(C) <+> quniv(C) <= quniv(C)"; +Goal "quniv(C) <+> quniv(C) <= quniv(C)"; by (REPEAT (ares_tac [subsetI, QInl_in_quniv, QInr_in_quniv] 1 ORELSE eresolve_tac [qsumE, ssubst] 1)); qed "qsum_quniv"; @@ -198,7 +198,7 @@ (*** Intersecting with Vfrom... ***) -goalw QUniv.thy [QPair_def,sum_def] +Goalw [QPair_def,sum_def] "!!X. Transset(X) ==> \ \ Int Vfrom(X, succ(i)) <= "; by (stac Int_Un_distrib 1); @@ -211,7 +211,7 @@ (*Rule for level i -- preserving the level, not decreasing it*) -goalw QUniv.thy [QPair_def] +Goalw [QPair_def] "!!X. Transset(X) ==> \ \ Int Vfrom(X,i) <= "; by (etac (Transset_Vfrom RS Transset_sum_Int_subset) 1); @@ -221,7 +221,7 @@ bind_thm ("QPair_Int_Vset_subset_trans", [Transset_0 RS QPair_Int_Vfrom_subset, QPair_mono] MRS subset_trans); -goal QUniv.thy +Goal "!!i. [| Ord(i) \ \ |] ==> Int Vset(i) <= (UN j:i. )"; by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac); diff -r 30271d90644f -r 62b6288e6005 src/ZF/Rel.ML --- a/src/ZF/Rel.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Rel.ML Mon Jun 22 17:12:27 1998 +0200 @@ -29,7 +29,7 @@ by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); qed "symI"; -goalw Rel.thy [sym_def] "!!r. [| sym(r); : r |] ==> : r"; +Goalw [sym_def] "!!r. [| sym(r); : r |] ==> : r"; by (Blast_tac 1); qed "symE"; @@ -48,12 +48,12 @@ (* transitivity *) -goalw Rel.thy [trans_def] +Goalw [trans_def] "!!r. [| trans(r); :r; :r |] ==> :r"; by (Blast_tac 1); qed "transD"; -goalw Rel.thy [trans_on_def] +Goalw [trans_on_def] "!!r. [| trans[A](r); :r; :r; a:A; b:A; c:A |] ==> :r"; by (Blast_tac 1); qed "trans_onD"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Sum.ML --- a/src/ZF/Sum.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Sum.ML Mon Jun 22 17:12:27 1998 +0200 @@ -10,12 +10,12 @@ (*** Rules for the Part primitive ***) -goalw Sum.thy [Part_def] +Goalw [Part_def] "a : Part(A,h) <-> a:A & (EX y. a=h(y))"; by (rtac separation 1); qed "Part_iff"; -goalw Sum.thy [Part_def] +Goalw [Part_def] "!!a b A h. [| a : A; a=h(b) |] ==> a : Part(A,h)"; by (Blast_tac 1); qed "Part_eqI"; @@ -33,7 +33,7 @@ AddIs [Part_eqI]; AddSEs [PartE]; -goalw Sum.thy [Part_def] "Part(A,h) <= A"; +Goalw [Part_def] "Part(A,h) <= A"; by (rtac Collect_subset 1); qed "Part_subset"; @@ -42,17 +42,17 @@ val sum_defs = [sum_def,Inl_def,Inr_def,case_def]; -goalw Sum.thy (bool_def::sum_defs) "Sigma(bool,C) = C(0) + C(1)"; +Goalw (bool_def::sum_defs) "Sigma(bool,C) = C(0) + C(1)"; by (Blast_tac 1); qed "Sigma_bool"; (** Introduction rules for the injections **) -goalw Sum.thy sum_defs "!!a A B. a : A ==> Inl(a) : A+B"; +Goalw sum_defs "!!a A B. a : A ==> Inl(a) : A+B"; by (Blast_tac 1); qed "InlI"; -goalw Sum.thy sum_defs "!!b A B. b : B ==> Inr(b) : A+B"; +Goalw sum_defs "!!b A B. b : B ==> Inr(b) : A+B"; by (Blast_tac 1); qed "InrI"; @@ -70,23 +70,23 @@ (** Injection and freeness equivalences, for rewriting **) -goalw Sum.thy sum_defs "Inl(a)=Inl(b) <-> a=b"; +Goalw sum_defs "Inl(a)=Inl(b) <-> a=b"; by (Simp_tac 1); qed "Inl_iff"; -goalw Sum.thy sum_defs "Inr(a)=Inr(b) <-> a=b"; +Goalw sum_defs "Inr(a)=Inr(b) <-> a=b"; by (Simp_tac 1); qed "Inr_iff"; -goalw Sum.thy sum_defs "Inl(a)=Inr(b) <-> False"; +Goalw sum_defs "Inl(a)=Inr(b) <-> False"; by (Simp_tac 1); qed "Inl_Inr_iff"; -goalw Sum.thy sum_defs "Inr(b)=Inl(a) <-> False"; +Goalw sum_defs "Inr(b)=Inl(a) <-> False"; by (Simp_tac 1); qed "Inr_Inl_iff"; -goalw Sum.thy sum_defs "0+0 = 0"; +Goalw sum_defs "0+0 = 0"; by (Simp_tac 1); qed "sum_empty"; @@ -103,39 +103,39 @@ Addsimps [InlI, InrI, Inl_iff, Inr_iff, Inl_Inr_iff, Inr_Inl_iff, sum_empty]; -goal Sum.thy "!!A B. Inl(a): A+B ==> a: A"; +Goal "!!A B. Inl(a): A+B ==> a: A"; by (Blast_tac 1); qed "InlD"; -goal Sum.thy "!!A B. Inr(b): A+B ==> b: B"; +Goal "!!A B. Inr(b): A+B ==> b: B"; by (Blast_tac 1); qed "InrD"; -goal Sum.thy "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"; +Goal "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"; by (Blast_tac 1); qed "sum_iff"; -goal Sum.thy "A+B <= C+D <-> A<=C & B<=D"; +Goal "A+B <= C+D <-> A<=C & B<=D"; by (Blast_tac 1); qed "sum_subset_iff"; -goal Sum.thy "A+B = C+D <-> A=C & B=D"; +Goal "A+B = C+D <-> A=C & B=D"; by (simp_tac (simpset() addsimps [extension,sum_subset_iff]) 1); by (Blast_tac 1); qed "sum_equal_iff"; -goalw Sum.thy [sum_def] "A+A = 2*A"; +Goalw [sum_def] "A+A = 2*A"; by (Blast_tac 1); qed "sum_eq_2_times"; (*** Eliminator -- case ***) -goalw Sum.thy sum_defs "case(c, d, Inl(a)) = c(a)"; +Goalw sum_defs "case(c, d, Inl(a)) = c(a)"; by (simp_tac (simpset() addsimps [cond_0]) 1); qed "case_Inl"; -goalw Sum.thy sum_defs "case(c, d, Inr(b)) = d(b)"; +Goalw sum_defs "case(c, d, Inr(b)) = d(b)"; by (simp_tac (simpset() addsimps [cond_1]) 1); qed "case_Inr"; @@ -151,7 +151,7 @@ by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); qed "case_type"; -goal Sum.thy +Goal "!!u. u: A+B ==> \ \ R(case(c,d,u)) <-> \ \ ((ALL x:A. u = Inl(x) --> R(c(x))) & \ @@ -168,7 +168,7 @@ by (ALLGOALS (asm_simp_tac (simpset() addsimps prems))); qed "case_cong"; -goal Sum.thy +Goal "!!z. 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)"; @@ -178,36 +178,36 @@ (*** More rules for Part(A,h) ***) -goal Sum.thy "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)"; +Goal "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)"; by (Blast_tac 1); qed "Part_mono"; -goal Sum.thy "Part(Collect(A,P), h) = Collect(Part(A,h), P)"; +Goal "Part(Collect(A,P), h) = Collect(Part(A,h), P)"; by (Blast_tac 1); qed "Part_Collect"; bind_thm ("Part_CollectE", Part_Collect RS equalityD1 RS subsetD RS CollectE); -goal Sum.thy "Part(A+B,Inl) = {Inl(x). x: A}"; +Goal "Part(A+B,Inl) = {Inl(x). x: A}"; by (Blast_tac 1); qed "Part_Inl"; -goal Sum.thy "Part(A+B,Inr) = {Inr(y). y: B}"; +Goal "Part(A+B,Inr) = {Inr(y). y: B}"; by (Blast_tac 1); qed "Part_Inr"; -goalw Sum.thy [Part_def] "!!a. a : Part(A,h) ==> a : A"; +Goalw [Part_def] "!!a. a : Part(A,h) ==> a : A"; by (etac CollectD1 1); qed "PartD1"; -goal Sum.thy "Part(A,%x. x) = A"; +Goal "Part(A,%x. x) = A"; by (Blast_tac 1); qed "Part_id"; -goal Sum.thy "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}"; +Goal "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}"; by (Blast_tac 1); qed "Part_Inr2"; -goal Sum.thy "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"; +Goal "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"; by (Blast_tac 1); qed "Part_sum_equality"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Trancl.ML --- a/src/ZF/Trancl.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Trancl.ML Mon Jun 22 17:12:27 1998 +0200 @@ -8,7 +8,7 @@ open Trancl; -goal Trancl.thy "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"; +Goal "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))"; by (rtac bnd_monoI 1); by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2)); by (Blast_tac 1); @@ -54,7 +54,7 @@ by (blast_tac (claset() addIs [r_into_rtrancl]) 1); qed "r_subset_rtrancl"; -goal Trancl.thy "field(r^*) = field(r)"; +Goal "field(r^*) = field(r)"; by (blast_tac (claset() addIs [r_into_rtrancl] addSDs [rtrancl_type RS subsetD]) 1); qed "rtrancl_field"; @@ -89,7 +89,7 @@ qed "rtrancl_induct"; (*transitivity of transitive closure!! -- by induction.*) -goalw Trancl.thy [trans_def] "trans(r^*)"; +Goalw [trans_def] "trans(r^*)"; by (REPEAT (resolve_tac [allI,impI] 1)); by (eres_inst_tac [("b","z")] rtrancl_induct 1); by (DEPTH_SOLVE (eresolve_tac [asm_rl, rtrancl_into_rtrancl] 1)); @@ -110,29 +110,29 @@ (**** The relation trancl ****) (*Transitivity of r^+ is proved by transitivity of r^* *) -goalw Trancl.thy [trans_def,trancl_def] "trans(r^+)"; +Goalw [trans_def,trancl_def] "trans(r^+)"; by (blast_tac (claset() addIs [rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)]) 1); qed "trans_trancl"; (** Conversions between trancl and rtrancl **) -goalw Trancl.thy [trancl_def] "!!r. : r^+ ==> : r^*"; +Goalw [trancl_def] "!!r. : r^+ ==> : r^*"; by (blast_tac (claset() addIs [rtrancl_into_rtrancl]) 1); qed "trancl_into_rtrancl"; (*r^+ contains all pairs in r *) -goalw Trancl.thy [trancl_def] "!!r. : r ==> : r^+"; +Goalw [trancl_def] "!!r. : r ==> : r^+"; by (blast_tac (claset() addSIs [rtrancl_refl]) 1); qed "r_into_trancl"; (*The premise ensures that r consists entirely of pairs*) -goal Trancl.thy "!!r. r <= Sigma(A,B) ==> r <= r^+"; +Goal "!!r. r <= Sigma(A,B) ==> r <= r^+"; by (blast_tac (claset() addIs [r_into_trancl]) 1); qed "r_subset_trancl"; (*intro rule by definition: from r^* and r *) -goalw Trancl.thy [trancl_def] +Goalw [trancl_def] "!!r. [| : r^*; : r |] ==> : r^+"; by (Blast_tac 1); qed "rtrancl_into_trancl1"; @@ -174,7 +174,7 @@ by (ALLGOALS (blast_tac (claset() addIs [rtrancl_into_trancl1]))); qed "tranclE"; -goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)"; +Goalw [trancl_def] "r^+ <= field(r)*field(r)"; by (blast_tac (claset() addEs [rtrancl_type RS subsetD RS SigmaE2]) 1); qed "trancl_type"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Univ.ML --- a/src/ZF/Univ.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Univ.ML Mon Jun 22 17:12:27 1998 +0200 @@ -9,14 +9,14 @@ open Univ; (*NOT SUITABLE FOR REWRITING -- RECURSIVE!*) -goal Univ.thy "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))"; +Goal "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))"; by (stac (Vfrom_def RS def_transrec) 1); by (Simp_tac 1); qed "Vfrom"; (** Monotonicity **) -goal Univ.thy "!!A B. A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"; +Goal "!!A B. A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"; by (eps_ind_tac "i" 1); by (rtac (impI RS allI) 1); by (stac Vfrom 1); @@ -33,14 +33,14 @@ (** A fundamental equality: Vfrom does not require ordinals! **) -goal Univ.thy "Vfrom(A,x) <= Vfrom(A,rank(x))"; +Goal "Vfrom(A,x) <= Vfrom(A,rank(x))"; by (eps_ind_tac "x" 1); by (stac Vfrom 1); by (stac Vfrom 1); by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1); qed "Vfrom_rank_subset1"; -goal Univ.thy "Vfrom(A,rank(x)) <= Vfrom(A,x)"; +Goal "Vfrom(A,rank(x)) <= Vfrom(A,x)"; by (eps_ind_tac "x" 1); by (stac Vfrom 1); by (stac Vfrom 1); @@ -57,7 +57,7 @@ by (assume_tac 1); qed "Vfrom_rank_subset2"; -goal Univ.thy "Vfrom(A,rank(x)) = Vfrom(A,x)"; +Goal "Vfrom(A,rank(x)) = Vfrom(A,x)"; by (rtac equalityI 1); by (rtac Vfrom_rank_subset2 1); by (rtac Vfrom_rank_subset1 1); @@ -66,43 +66,43 @@ (*** Basic closure properties ***) -goal Univ.thy "!!x y. y:x ==> 0 : Vfrom(A,x)"; +Goal "!!x y. y:x ==> 0 : Vfrom(A,x)"; by (stac Vfrom 1); by (Blast_tac 1); qed "zero_in_Vfrom"; -goal Univ.thy "i <= Vfrom(A,i)"; +Goal "i <= Vfrom(A,i)"; by (eps_ind_tac "i" 1); by (stac Vfrom 1); by (Blast_tac 1); qed "i_subset_Vfrom"; -goal Univ.thy "A <= Vfrom(A,i)"; +Goal "A <= Vfrom(A,i)"; by (stac Vfrom 1); by (rtac Un_upper1 1); qed "A_subset_Vfrom"; bind_thm ("A_into_Vfrom", A_subset_Vfrom RS subsetD); -goal Univ.thy "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))"; +Goal "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))"; by (stac Vfrom 1); by (Blast_tac 1); qed "subset_mem_Vfrom"; (** Finite sets and ordered pairs **) -goal Univ.thy "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))"; +Goal "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))"; by (rtac subset_mem_Vfrom 1); by Safe_tac; qed "singleton_in_Vfrom"; -goal Univ.thy +Goal "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))"; by (rtac subset_mem_Vfrom 1); by Safe_tac; qed "doubleton_in_Vfrom"; -goalw Univ.thy [Pair_def] +Goalw [Pair_def] "!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> \ \ : Vfrom(A,succ(succ(i)))"; by (REPEAT (ares_tac [doubleton_in_Vfrom] 1)); @@ -117,12 +117,12 @@ (*** 0, successor and limit equations fof Vfrom ***) -goal Univ.thy "Vfrom(A,0) = A"; +Goal "Vfrom(A,0) = A"; by (stac Vfrom 1); by (Blast_tac 1); qed "Vfrom_0"; -goal Univ.thy "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; +Goal "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; by (rtac (Vfrom RS trans) 1); by (rtac (succI1 RS RepFunI RS Union_upper RSN (2, equalityI RS subst_context)) 1); @@ -132,7 +132,7 @@ by (etac Ord_succ 1); qed "Vfrom_succ_lemma"; -goal Univ.thy "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; +Goal "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; by (res_inst_tac [("x1", "succ(i)")] (Vfrom_rank_eq RS subst) 1); by (res_inst_tac [("x1", "i")] (Vfrom_rank_eq RS subst) 1); by (stac rank_succ 1); @@ -171,7 +171,7 @@ by (rtac refl 1); qed "Limit_Vfrom_eq"; -goal Univ.thy "!!a. [| a: Vfrom(A,j); Limit(i); j a : Vfrom(A,i)"; +Goal "!!a. [| a: Vfrom(A,j); Limit(i); j a : Vfrom(A,i)"; by (rtac (Limit_Vfrom_eq RS equalityD2 RS subsetD) 1); by (REPEAT (ares_tac [ltD RS UN_I] 1)); qed "Limit_VfromI"; @@ -221,7 +221,7 @@ by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1)); qed "Pair_in_VLimit"; -goal Univ.thy "!!i. Limit(i) ==> Vfrom(A,i)*Vfrom(A,i) <= Vfrom(A,i)"; +Goal "!!i. Limit(i) ==> Vfrom(A,i)*Vfrom(A,i) <= Vfrom(A,i)"; by (REPEAT (ares_tac [subsetI,Pair_in_VLimit] 1 ORELSE eresolve_tac [SigmaE, ssubst] 1)); qed "product_VLimit"; @@ -232,7 +232,7 @@ bind_thm ("nat_subset_VLimit", [nat_le_Limit RS le_imp_subset, i_subset_Vfrom] MRS subset_trans); -goal Univ.thy "!!i. [| n: nat; Limit(i) |] ==> n : Vfrom(A,i)"; +Goal "!!i. [| n: nat; Limit(i) |] ==> n : Vfrom(A,i)"; by (REPEAT (ares_tac [nat_subset_VLimit RS subsetD] 1)); qed "nat_into_VLimit"; @@ -240,21 +240,21 @@ bind_thm ("zero_in_VLimit", Limit_has_0 RS ltD RS zero_in_Vfrom); -goal Univ.thy "!!i. Limit(i) ==> 1 : Vfrom(A,i)"; +Goal "!!i. Limit(i) ==> 1 : Vfrom(A,i)"; by (REPEAT (ares_tac [nat_into_VLimit, nat_0I, nat_succI] 1)); qed "one_in_VLimit"; -goalw Univ.thy [Inl_def] +Goalw [Inl_def] "!!A a. [| a: Vfrom(A,i); Limit(i) |] ==> Inl(a) : Vfrom(A,i)"; by (REPEAT (ares_tac [zero_in_VLimit, Pair_in_VLimit] 1)); qed "Inl_in_VLimit"; -goalw Univ.thy [Inr_def] +Goalw [Inr_def] "!!A b. [| b: Vfrom(A,i); Limit(i) |] ==> Inr(b) : Vfrom(A,i)"; by (REPEAT (ares_tac [one_in_VLimit, Pair_in_VLimit] 1)); qed "Inr_in_VLimit"; -goal Univ.thy "!!i. Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)"; +Goal "!!i. Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)"; by (blast_tac (claset() addSIs [Inl_in_VLimit, Inr_in_VLimit]) 1); qed "sum_VLimit"; @@ -264,14 +264,14 @@ (*** Properties assuming Transset(A) ***) -goal Univ.thy "!!i A. Transset(A) ==> Transset(Vfrom(A,i))"; +Goal "!!i A. Transset(A) ==> Transset(Vfrom(A,i))"; by (eps_ind_tac "i" 1); by (stac Vfrom 1); by (blast_tac (claset() addSIs [Transset_Union_family, Transset_Un, Transset_Pow]) 1); qed "Transset_Vfrom"; -goal Univ.thy "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"; +Goal "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"; by (rtac (Vfrom_succ RS trans) 1); by (rtac (Un_upper2 RSN (2,equalityI)) 1); by (rtac (subset_refl RSN (2,Un_least)) 1); @@ -284,7 +284,7 @@ by (Blast_tac 1); qed "Transset_Pair_subset"; -goal Univ.thy +Goal "!!a b.[| <= Vfrom(A,i); Transset(A); Limit(i) |] ==> \ \ : Vfrom(A,i)"; by (etac (Transset_Pair_subset RS conjE) 1); @@ -317,7 +317,7 @@ (** products **) -goal Univ.thy +Goal "!!A. [| a: Vfrom(A,j); b: Vfrom(A,j); Transset(A) |] ==> \ \ a*b : Vfrom(A, succ(succ(succ(j))))"; by (dtac Transset_Vfrom 1); @@ -336,7 +336,7 @@ (** Disjoint sums, aka Quine ordered pairs **) -goalw Univ.thy [sum_def] +Goalw [sum_def] "!!A. [| a: Vfrom(A,j); b: Vfrom(A,j); Transset(A); 1:j |] ==> \ \ a+b : Vfrom(A, succ(succ(succ(j))))"; by (dtac Transset_Vfrom 1); @@ -356,7 +356,7 @@ (** function space! **) -goalw Univ.thy [Pi_def] +Goalw [Pi_def] "!!A. [| a: Vfrom(A,j); b: Vfrom(A,j); Transset(A) |] ==> \ \ a->b : Vfrom(A, succ(succ(succ(succ(j)))))"; by (dtac Transset_Vfrom 1); @@ -379,7 +379,7 @@ limiti RS Limit_has_succ] 1)); qed "fun_in_VLimit"; -goalw Univ.thy [Pi_def] +Goalw [Pi_def] "!!A. [| a: Vfrom(A,j); Transset(A) |] ==> \ \ Pow(a) : Vfrom(A, succ(succ(j)))"; by (dtac Transset_Vfrom 1); @@ -389,7 +389,7 @@ by (Blast_tac 1); qed "Pow_in_Vfrom"; -goal Univ.thy +Goal "!!a. [| a: Vfrom(A,i); Limit(i); Transset(A) |] ==> Pow(a) : Vfrom(A,i)"; (*Blast_tac: PROOF FAILED*) by (fast_tac (claset() addEs [Limit_VfromE] @@ -399,7 +399,7 @@ (*** The set Vset(i) ***) -goal Univ.thy "Vset(i) = (UN j:i. Pow(Vset(j)))"; +Goal "Vset(i) = (UN j:i. Pow(Vset(j)))"; by (stac Vfrom 1); by (Blast_tac 1); qed "Vset"; @@ -427,23 +427,23 @@ by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1); val lemma = result(); -goal Univ.thy "!!x i. rank(x) x : Vset(i)"; +Goal "!!x i. rank(x) x : Vset(i)"; by (etac ltE 1); by (etac (lemma RS spec RS mp) 1); by (assume_tac 1); qed "VsetI"; -goal Univ.thy "!!i. Ord(i) ==> b : Vset(i) <-> rank(b) < i"; +Goal "!!i. Ord(i) ==> b : Vset(i) <-> rank(b) < i"; by (rtac iffI 1); by (REPEAT (eresolve_tac [asm_rl, VsetD, VsetI] 1)); qed "Vset_Ord_rank_iff"; -goal Univ.thy "b : Vset(a) <-> rank(b) < rank(a)"; +Goal "b : Vset(a) <-> rank(b) < rank(a)"; by (rtac (Vfrom_rank_eq RS subst) 1); by (rtac (Ord_rank RS Vset_Ord_rank_iff) 1); qed "Vset_rank_iff"; -goal Univ.thy "!!i. Ord(i) ==> rank(Vset(i)) = i"; +Goal "!!i. Ord(i) ==> rank(Vset(i)) = i"; by (stac rank 1); by (rtac equalityI 1); by Safe_tac; @@ -457,7 +457,7 @@ (** Lemmas for reasoning about sets in terms of their elements' ranks **) -goal Univ.thy "a <= Vset(rank(a))"; +Goal "a <= Vset(rank(a))"; by (rtac subsetI 1); by (etac (rank_lt RS VsetI) 1); qed "arg_subset_Vset_rank"; @@ -471,11 +471,11 @@ (** Set up an environment for simplification **) -goalw Univ.thy [Inl_def] "rank(a) < rank(Inl(a))"; +Goalw [Inl_def] "rank(a) < rank(Inl(a))"; by (rtac rank_pair2 1); qed "rank_Inl"; -goalw Univ.thy [Inr_def] "rank(a) < rank(Inr(a))"; +Goalw [Inr_def] "rank(a) < rank(Inr(a))"; by (rtac rank_pair2 1); qed "rank_Inr"; @@ -487,7 +487,7 @@ (** Recursion over Vset levels! **) (*NOT SUITABLE FOR REWRITING: recursive!*) -goalw Univ.thy [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"; +Goalw [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"; by (stac transrec 1); by (simp_tac (simpset() addsimps [Ord_rank, Ord_succ, VsetD RS ltD RS beta, VsetI RS beta, le_refl]) 1); @@ -504,22 +504,22 @@ (*** univ(A) ***) -goalw Univ.thy [univ_def] "!!A B. A<=B ==> univ(A) <= univ(B)"; +Goalw [univ_def] "!!A B. A<=B ==> univ(A) <= univ(B)"; by (etac Vfrom_mono 1); by (rtac subset_refl 1); qed "univ_mono"; -goalw Univ.thy [univ_def] "!!A. Transset(A) ==> Transset(univ(A))"; +Goalw [univ_def] "!!A. Transset(A) ==> Transset(univ(A))"; by (etac Transset_Vfrom 1); qed "Transset_univ"; (** univ(A) as a limit **) -goalw Univ.thy [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))"; +Goalw [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))"; by (rtac (Limit_nat RS Limit_Vfrom_eq) 1); qed "univ_eq_UN"; -goal Univ.thy "!!c. c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))"; +Goal "!!c. c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))"; by (rtac (subset_UN_iff_eq RS iffD1) 1); by (etac (univ_eq_UN RS subst) 1); qed "subset_univ_eq_Int"; @@ -546,11 +546,11 @@ (** Closure properties **) -goalw Univ.thy [univ_def] "0 : univ(A)"; +Goalw [univ_def] "0 : univ(A)"; by (rtac (nat_0I RS zero_in_Vfrom) 1); qed "zero_in_univ"; -goalw Univ.thy [univ_def] "A <= univ(A)"; +Goalw [univ_def] "A <= univ(A)"; by (rtac A_subset_Vfrom 1); qed "A_subset_univ"; @@ -558,28 +558,28 @@ (** Closure under unordered and ordered pairs **) -goalw Univ.thy [univ_def] "!!A a. a: univ(A) ==> {a} : univ(A)"; +Goalw [univ_def] "!!A a. a: univ(A) ==> {a} : univ(A)"; by (REPEAT (ares_tac [singleton_in_VLimit, Limit_nat] 1)); qed "singleton_in_univ"; -goalw Univ.thy [univ_def] +Goalw [univ_def] "!!A a. [| a: univ(A); b: univ(A) |] ==> {a,b} : univ(A)"; by (REPEAT (ares_tac [doubleton_in_VLimit, Limit_nat] 1)); qed "doubleton_in_univ"; -goalw Univ.thy [univ_def] +Goalw [univ_def] "!!A a. [| a: univ(A); b: univ(A) |] ==> : univ(A)"; by (REPEAT (ares_tac [Pair_in_VLimit, Limit_nat] 1)); qed "Pair_in_univ"; -goalw Univ.thy [univ_def] "univ(A)*univ(A) <= univ(A)"; +Goalw [univ_def] "univ(A)*univ(A) <= univ(A)"; by (rtac (Limit_nat RS product_VLimit) 1); qed "product_univ"; (** The natural numbers **) -goalw Univ.thy [univ_def] "nat <= univ(A)"; +Goalw [univ_def] "nat <= univ(A)"; by (rtac i_subset_Vfrom 1); qed "nat_subset_univ"; @@ -588,16 +588,16 @@ (** instances for 1 and 2 **) -goalw Univ.thy [univ_def] "1 : univ(A)"; +Goalw [univ_def] "1 : univ(A)"; by (rtac (Limit_nat RS one_in_VLimit) 1); qed "one_in_univ"; (*unused!*) -goal Univ.thy "2 : univ(A)"; +Goal "2 : univ(A)"; by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1)); qed "two_in_univ"; -goalw Univ.thy [bool_def] "bool <= univ(A)"; +Goalw [bool_def] "bool <= univ(A)"; by (blast_tac (claset() addSIs [zero_in_univ,one_in_univ]) 1); qed "bool_subset_univ"; @@ -606,15 +606,15 @@ (** Closure under disjoint union **) -goalw Univ.thy [univ_def] "!!A a. a: univ(A) ==> Inl(a) : univ(A)"; +Goalw [univ_def] "!!A a. a: univ(A) ==> Inl(a) : univ(A)"; by (etac (Limit_nat RSN (2,Inl_in_VLimit)) 1); qed "Inl_in_univ"; -goalw Univ.thy [univ_def] "!!A b. b: univ(A) ==> Inr(b) : univ(A)"; +Goalw [univ_def] "!!A b. b: univ(A) ==> Inr(b) : univ(A)"; by (etac (Limit_nat RSN (2,Inr_in_VLimit)) 1); qed "Inr_in_univ"; -goalw Univ.thy [univ_def] "univ(C)+univ(C) <= univ(C)"; +Goalw [univ_def] "univ(C)+univ(C) <= univ(C)"; by (rtac (Limit_nat RS sum_VLimit) 1); qed "sum_univ"; @@ -630,7 +630,7 @@ (** Closure under finite powerset **) -goal Univ.thy +Goal "!!i. [| b: Fin(Vfrom(A,i)); Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j Fin(Vfrom(A,i)) <= Vfrom(A,i)"; +Goal "!!i. Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)"; by (rtac subsetI 1); by (dtac Fin_Vfrom_lemma 1); by Safe_tac; @@ -650,13 +650,13 @@ bind_thm ("Fin_subset_VLimit", [Fin_mono, Fin_VLimit] MRS subset_trans); -goalw Univ.thy [univ_def] "Fin(univ(A)) <= univ(A)"; +Goalw [univ_def] "Fin(univ(A)) <= univ(A)"; by (rtac (Limit_nat RS Fin_VLimit) 1); val Fin_univ = result(); (** Closure under finite powers (functions from a fixed natural number) **) -goal Univ.thy +Goal "!!i. [| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) <= Vfrom(A,i)"; by (eresolve_tac [nat_fun_subset_Fin RS subset_trans] 1); by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit, @@ -665,7 +665,7 @@ bind_thm ("nat_fun_subset_VLimit", [Pi_mono, nat_fun_VLimit] MRS subset_trans); -goalw Univ.thy [univ_def] "!!i. n: nat ==> n -> univ(A) <= univ(A)"; +Goalw [univ_def] "!!i. n: nat ==> n -> univ(A) <= univ(A)"; by (etac (Limit_nat RSN (2,nat_fun_VLimit)) 1); val nat_fun_univ = result(); @@ -673,29 +673,29 @@ (** Closure under finite function space **) (*General but seldom-used version; normally the domain is fixed*) -goal Univ.thy +Goal "!!i. Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) <= Vfrom(A,i)"; by (resolve_tac [FiniteFun.dom_subset RS subset_trans] 1); by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit, subset_refl] 1)); val FiniteFun_VLimit1 = result(); -goalw Univ.thy [univ_def] "univ(A) -||> univ(A) <= univ(A)"; +Goalw [univ_def] "univ(A) -||> univ(A) <= univ(A)"; by (rtac (Limit_nat RS FiniteFun_VLimit1) 1); val FiniteFun_univ1 = result(); (*Version for a fixed domain*) -goal Univ.thy +Goal "!!i. [| W <= Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) <= Vfrom(A,i)"; by (eresolve_tac [subset_refl RSN (2, FiniteFun_mono) RS subset_trans] 1); by (etac FiniteFun_VLimit1 1); val FiniteFun_VLimit = result(); -goalw Univ.thy [univ_def] +Goalw [univ_def] "!!W. W <= univ(A) ==> W -||> univ(A) <= univ(A)"; by (etac (Limit_nat RSN (2, FiniteFun_VLimit)) 1); val FiniteFun_univ = result(); -goal Univ.thy +Goal "!!W. [| f: W -||> univ(A); W <= univ(A) |] ==> f : univ(A)"; by (eresolve_tac [FiniteFun_univ RS subsetD] 1); by (assume_tac 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/WF.ML --- a/src/ZF/WF.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/WF.ML Mon Jun 22 17:12:27 1998 +0200 @@ -24,25 +24,25 @@ (** Equivalences between wf and wf_on **) -goalw WF.thy [wf_def, wf_on_def] "!!A r. wf(r) ==> wf[A](r)"; +Goalw [wf_def, wf_on_def] "!!A r. wf(r) ==> wf[A](r)"; by (Clarify_tac 1); (*essential for Blast_tac's efficiency*) by (Blast_tac 1); qed "wf_imp_wf_on"; -goalw WF.thy [wf_def, wf_on_def] "!!r. wf[field(r)](r) ==> wf(r)"; +Goalw [wf_def, wf_on_def] "!!r. wf[field(r)](r) ==> wf(r)"; by (Fast_tac 1); qed "wf_on_field_imp_wf"; -goal WF.thy "wf(r) <-> wf[field(r)](r)"; +Goal "wf(r) <-> wf[field(r)](r)"; by (blast_tac (claset() addIs [wf_imp_wf_on, wf_on_field_imp_wf]) 1); qed "wf_iff_wf_on_field"; -goalw WF.thy [wf_on_def, wf_def] "!!A B r. [| wf[A](r); B<=A |] ==> wf[B](r)"; +Goalw [wf_on_def, wf_def] "!!A B r. [| wf[A](r); B<=A |] ==> wf[B](r)"; by (Clarify_tac 1); (*essential for Blast_tac's efficiency*) by (Blast_tac 1); qed "wf_on_subset_A"; -goalw WF.thy [wf_on_def, wf_def] "!!A r s. [| wf[A](r); s<=r |] ==> wf[A](s)"; +Goalw [wf_on_def, wf_def] "!!A r s. [| wf[A](r); s<=r |] ==> wf[A](s)"; by (Clarify_tac 1); (*essential for Blast_tac's efficiency*) by (Blast_tac 1); qed "wf_on_subset_r"; @@ -136,24 +136,24 @@ (*** Properties of well-founded relations ***) -goal WF.thy "!!r. wf(r) ==> ~: r"; +Goal "!!r. wf(r) ==> ~: r"; by (wf_ind_tac "a" [] 1); by (Blast_tac 1); qed "wf_not_refl"; -goal WF.thy "!!r. [| wf(r); :r; :r |] ==> P"; +Goal "!!r. [| wf(r); :r; :r |] ==> P"; by (subgoal_tac "ALL x. :r --> :r --> P" 1); by (wf_ind_tac "a" [] 2); by (Blast_tac 2); by (Blast_tac 1); qed "wf_asym"; -goal WF.thy "!!r. [| wf[A](r); a: A |] ==> ~: r"; +Goal "!!r. [| wf[A](r); a: A |] ==> ~: r"; by (wf_on_ind_tac "a" [] 1); by (Blast_tac 1); qed "wf_on_not_refl"; -goal WF.thy "!!r. [| wf[A](r); :r; :r; a:A; b:A |] ==> P"; +Goal "!!r. [| wf[A](r); :r; :r; a:A; b:A |] ==> P"; by (subgoal_tac "ALL y:A. :r --> :r --> P" 1); by (wf_on_ind_tac "a" [] 2); by (Blast_tac 2); @@ -162,7 +162,7 @@ (*Needed to prove well_ordI. Could also reason that wf[A](r) means wf(r Int A*A); thus wf( (r Int A*A)^+ ) and use wf_not_refl *) -goal WF.thy +Goal "!!r. [| wf[A](r); :r; :r; :r; a:A; b:A; c:A |] ==> P"; by (subgoal_tac "ALL y:A. ALL z:A. :r --> :r --> :r --> P" 1); @@ -185,7 +185,7 @@ by (blast_tac (claset() addEs [tranclE]) 1); qed "wf_on_trancl"; -goal WF.thy "!!r. wf(r) ==> wf(r^+)"; +Goal "!!r. wf(r) ==> wf(r^+)"; by (asm_full_simp_tac (simpset() addsimps [wf_iff_wf_on_field]) 1); by (rtac (trancl_type RS field_rel_subset RSN (2, wf_on_subset_A)) 1); by (etac wf_on_trancl 1); @@ -293,7 +293,7 @@ qed "the_recfun_cut"; (*NOT SUITABLE FOR REWRITING: it is recursive!*) -goalw WF.thy [wftrec_def] +Goalw [wftrec_def] "!!r. [| wf(r); trans(r) |] ==> \ \ wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))"; by (stac (rewrite_rule [is_recfun_def] unfold_the_recfun) 1); @@ -333,7 +333,7 @@ qed "wfrec_type"; -goalw WF.thy [wf_on_def, wfrec_on_def] +Goalw [wf_on_def, wfrec_on_def] "!!A r. [| wf[A](r); a: A |] ==> \ \ wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))"; by (etac (wfrec RS trans) 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/ZF.ML --- a/src/ZF/ZF.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/ZF.ML Mon Jun 22 17:12:27 1998 +0200 @@ -9,7 +9,7 @@ open ZF; (*Useful examples: singletonI RS subst_elem, subst_elem RSN (2,IntI) *) -goal ZF.thy "!!a b A. [| b:A; a=b |] ==> a:A"; +Goal "!!a b A. [| b:A; a=b |] ==> a:A"; by (etac ssubst 1); by (assume_tac 1); val subst_elem = result(); @@ -271,7 +271,7 @@ "b : {f(x). x:A} <-> (EX x:A. b=f(x))" (fn _ => [Blast_tac 1]); -goal ZF.thy "{x. x:A} = A"; +Goal "{x. x:A} = A"; by (Blast_tac 1); qed "triv_RepFun"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/Zorn.ML --- a/src/ZF/Zorn.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/Zorn.ML Mon Jun 22 17:12:27 1998 +0200 @@ -25,12 +25,12 @@ (*** Section 2. The Transfinite Construction ***) -goalw Zorn.thy [increasing_def] +Goalw [increasing_def] "!!f A. f: increasing(A) ==> f: Pow(A)->Pow(A)"; by (etac CollectD1 1); qed "increasingD1"; -goalw Zorn.thy [increasing_def] +Goalw [increasing_def] "!!f A. [| f: increasing(A); x<=A |] ==> x <= f`x"; by (eresolve_tac [CollectD2 RS spec RS mp] 1); by (assume_tac 1); @@ -106,7 +106,7 @@ qed "TFin_linear_lemma2"; (*a more convenient form for Lemma 2*) -goal Zorn.thy +Goal "!!m n. [| n<=m; m: TFin(S,next); n: TFin(S,next); next: increasing(S) \ \ |] ==> n=m | next`n<=m"; by (rtac (TFin_linear_lemma2 RS bspec RS mp) 1); @@ -114,7 +114,7 @@ qed "TFin_subsetD"; (*Consequences from section 3.3 -- Property 3.2, the ordering is total*) -goal Zorn.thy +Goal "!!m n. [| m: TFin(S,next); n: TFin(S,next); next: increasing(S) \ \ |] ==> n<=m | m<=n"; by (rtac (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1); @@ -136,7 +136,7 @@ qed "equal_next_upper"; (*Property 3.3 of section 3.3*) -goal Zorn.thy +Goal "!!m. [| m: TFin(S,next); next: increasing(S) \ \ |] ==> m = next`m <-> m = Union(TFin(S,next))"; by (rtac iffI 1); @@ -155,19 +155,19 @@ (** Defining the "next" operation for Hausdorff's Theorem **) -goalw Zorn.thy [chain_def] "chain(A) <= Pow(A)"; +Goalw [chain_def] "chain(A) <= Pow(A)"; by (rtac Collect_subset 1); qed "chain_subset_Pow"; -goalw Zorn.thy [super_def] "super(A,c) <= chain(A)"; +Goalw [super_def] "super(A,c) <= chain(A)"; by (rtac Collect_subset 1); qed "super_subset_chain"; -goalw Zorn.thy [maxchain_def] "maxchain(A) <= chain(A)"; +Goalw [maxchain_def] "maxchain(A) <= chain(A)"; by (rtac Collect_subset 1); qed "maxchain_subset_chain"; -goal Zorn.thy +Goal "!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \ \ X : chain(S); X ~: maxchain(S) \ \ |] ==> ch ` super(S,X) : super(S,X)"; @@ -176,7 +176,7 @@ by (Blast_tac 1); qed "choice_super"; -goal Zorn.thy +Goal "!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \ \ X : chain(S); X ~: maxchain(S) \ \ |] ==> ch ` super(S,X) ~= X"; @@ -188,7 +188,7 @@ qed "choice_not_equals"; (*This justifies Definition 4.4*) -goal Zorn.thy +Goal "!!S. ch: (PROD X: Pow(chain(S))-{0}. X) ==> \ \ EX next: increasing(S). ALL X: Pow(S). \ \ next`X = if(X: chain(S)-maxchain(S), ch`super(S,X), X)"; @@ -214,7 +214,7 @@ qed "Hausdorff_next_exists"; (*Lemma 4*) -goal Zorn.thy +Goal "!!S. [| c: TFin(S,next); \ \ ch: (PROD X: Pow(chain(S))-{0}. X); \ \ next: increasing(S); \ @@ -233,7 +233,7 @@ by (ALLGOALS Fast_tac); qed "TFin_chain_lemma4"; -goal Zorn.thy "EX c. c : maxchain(S)"; +Goal "EX c. c : maxchain(S)"; by (rtac (AC_Pi_Pow RS exE) 1); by (rtac (Hausdorff_next_exists RS bexE) 1); by (assume_tac 1); @@ -260,12 +260,12 @@ then S contains a maximal element ***) (*Used in the proof of Zorn's Lemma*) -goalw Zorn.thy [chain_def] +Goalw [chain_def] "!!c. [| c: chain(A); z: A; ALL x:c. x<=z |] ==> cons(z,c) : chain(A)"; by (Blast_tac 1); qed "chain_extend"; -goal Zorn.thy +Goal "!!S. ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z"; by (resolve_tac [Hausdorff RS exE] 1); by (asm_full_simp_tac (simpset() addsimps [maxchain_def]) 1); @@ -302,7 +302,7 @@ qed "TFin_well_lemma5"; (*Well-ordering of TFin(S,next)*) -goal Zorn.thy "!!Z. [| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z"; +Goal "!!Z. [| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z"; by (rtac classical 1); by (subgoal_tac "Z = {Union(TFin(S,next))}" 1); by (asm_simp_tac (simpset() addsimps [Inter_singleton]) 1); @@ -313,7 +313,7 @@ qed "well_ord_TFin_lemma"; (*This theorem just packages the previous result*) -goal Zorn.thy +Goal "!!S. next: increasing(S) ==> \ \ well_ord(TFin(S,next), Subset_rel(TFin(S,next)))"; by (rtac well_ordI 1); @@ -344,7 +344,7 @@ qed "choice_Diff"; (*This justifies Definition 6.1*) -goal Zorn.thy +Goal "!!S. ch: (PROD X: Pow(S)-{0}. X) ==> \ \ EX next: increasing(S). ALL X: Pow(S). \ \ next`X = if(X=S, S, cons(ch`(S-X), X))"; @@ -368,7 +368,7 @@ (*The construction of the injection*) -goal Zorn.thy +Goal "!!S. [| ch: (PROD X: Pow(S)-{0}. X); \ \ next: increasing(S); \ \ ALL X: Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) \ @@ -404,7 +404,7 @@ qed "choice_imp_injection"; (*The wellordering theorem*) -goal Zorn.thy "EX r. well_ord(S,r)"; +Goal "EX r. well_ord(S,r)"; by (rtac (AC_Pi_Pow RS exE) 1); by (rtac (Zermelo_next_exists RS bexE) 1); by (assume_tac 1); diff -r 30271d90644f -r 62b6288e6005 src/ZF/equalities.ML --- a/src/ZF/equalities.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/equalities.ML Mon Jun 22 17:12:27 1998 +0200 @@ -490,12 +490,12 @@ by (Blast_tac 1); qed "vimage_Int_subset"; -goalw thy [function_def] +Goalw [function_def] "!!f. function(f) ==> f-``(A Int B) = (f-``A) Int (f-``B)"; by (Blast_tac 1); qed "function_vimage_Int"; -goalw thy [function_def] +Goalw [function_def] "!!f. function(f) ==> f-``(A-B) = (f-``A) - (f-``B)"; by (Blast_tac 1); qed "function_vimage_Diff"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/func.ML --- a/src/ZF/func.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/func.ML Mon Jun 22 17:12:27 1998 +0200 @@ -253,7 +253,7 @@ addcongs [RepFun_cong]) 1); qed "image_fun"; -goal thy "!!f. [| f: Pi(A,B); x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)"; +Goal "!!f. [| f: Pi(A,B); x: A |] ==> f `` cons(x,y) = cons(f`x, f``y)"; by (blast_tac (claset() addDs [apply_equality, apply_Pair]) 1); qed "Pi_image_cons"; diff -r 30271d90644f -r 62b6288e6005 src/ZF/mono.ML --- a/src/ZF/mono.ML Mon Jun 22 15:53:24 1998 +0200 +++ b/src/ZF/mono.ML Mon Jun 22 17:12:27 1998 +0200 @@ -12,19 +12,19 @@ (*Not easy to express monotonicity in P, since any "bigger" predicate would have to be single-valued*) -goal thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)"; +Goal "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)"; by (blast_tac (claset() addSEs [ReplaceE]) 1); qed "Replace_mono"; -goal thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}"; +Goal "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}"; by (Blast_tac 1); qed "RepFun_mono"; -goal thy "!!A B. A<=B ==> Pow(A) <= Pow(B)"; +Goal "!!A B. A<=B ==> Pow(A) <= Pow(B)"; by (Blast_tac 1); qed "Pow_mono"; -goal thy "!!A B. A<=B ==> Union(A) <= Union(B)"; +Goal "!!A B. A<=B ==> Union(A) <= Union(B)"; by (Blast_tac 1); qed "Union_mono"; @@ -35,95 +35,95 @@ qed "UN_mono"; (*Intersection is ANTI-monotonic. There are TWO premises! *) -goal thy "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)"; +Goal "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)"; by (Blast_tac 1); qed "Inter_anti_mono"; -goal thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)"; +Goal "!!C D. C<=D ==> cons(a,C) <= cons(a,D)"; by (Blast_tac 1); qed "cons_mono"; -goal thy "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D"; +Goal "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D"; by (Blast_tac 1); qed "Un_mono"; -goal thy "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D"; +Goal "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D"; by (Blast_tac 1); qed "Int_mono"; -goal thy "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D"; +Goal "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D"; by (Blast_tac 1); qed "Diff_mono"; (** Standard products, sums and function spaces **) -goal thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ +Goal "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ \ Sigma(A,B) <= Sigma(C,D)"; by (Blast_tac 1); qed "Sigma_mono_lemma"; val Sigma_mono = ballI RSN (2,Sigma_mono_lemma); -goalw thy sum_defs "!!A B C D. [| A<=C; B<=D |] ==> A+B <= C+D"; +Goalw sum_defs "!!A B C D. [| A<=C; B<=D |] ==> A+B <= C+D"; by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1)); qed "sum_mono"; (*Note that B->A and C->A are typically disjoint!*) -goal thy "!!A B C. B<=C ==> A->B <= A->C"; +Goal "!!A B C. B<=C ==> A->B <= A->C"; by (blast_tac (claset() addIs [lam_type] addEs [Pi_lamE]) 1); qed "Pi_mono"; -goalw thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)"; +Goalw [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)"; by (etac RepFun_mono 1); qed "lam_mono"; (** Quine-inspired ordered pairs, products, injections and sums **) -goalw thy [QPair_def] "!!a b c d. [| a<=c; b<=d |] ==> <= "; +Goalw [QPair_def] "!!a b c d. [| a<=c; b<=d |] ==> <= "; by (REPEAT (ares_tac [sum_mono] 1)); qed "QPair_mono"; -goal thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ +Goal "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \ \ QSigma(A,B) <= QSigma(C,D)"; by (Blast_tac 1); qed "QSigma_mono_lemma"; val QSigma_mono = ballI RSN (2,QSigma_mono_lemma); -goalw thy [QInl_def] "!!a b. a<=b ==> QInl(a) <= QInl(b)"; +Goalw [QInl_def] "!!a b. a<=b ==> QInl(a) <= QInl(b)"; by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1)); qed "QInl_mono"; -goalw thy [QInr_def] "!!a b. a<=b ==> QInr(a) <= QInr(b)"; +Goalw [QInr_def] "!!a b. a<=b ==> QInr(a) <= QInr(b)"; by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1)); qed "QInr_mono"; -goal thy "!!A B C D. [| A<=C; B<=D |] ==> A <+> B <= C <+> D"; +Goal "!!A B C D. [| A<=C; B<=D |] ==> A <+> B <= C <+> D"; by (Blast_tac 1); qed "qsum_mono"; (** Converse, domain, range, field **) -goal thy "!!r s. r<=s ==> converse(r) <= converse(s)"; +Goal "!!r s. r<=s ==> converse(r) <= converse(s)"; by (Blast_tac 1); qed "converse_mono"; -goal thy "!!r s. r<=s ==> domain(r)<=domain(s)"; +Goal "!!r s. r<=s ==> domain(r)<=domain(s)"; by (Blast_tac 1); qed "domain_mono"; bind_thm ("domain_rel_subset", [domain_mono, domain_subset] MRS subset_trans); -goal thy "!!r s. r<=s ==> range(r)<=range(s)"; +Goal "!!r s. r<=s ==> range(r)<=range(s)"; by (Blast_tac 1); qed "range_mono"; bind_thm ("range_rel_subset", [range_mono, range_subset] MRS subset_trans); -goal thy "!!r s. r<=s ==> field(r)<=field(s)"; +Goal "!!r s. r<=s ==> field(r)<=field(s)"; by (Blast_tac 1); qed "field_mono"; -goal thy "!!r A. r <= A*A ==> field(r) <= A"; +Goal "!!r A. r <= A*A ==> field(r) <= A"; by (etac (field_mono RS subset_trans) 1); by (Blast_tac 1); qed "field_rel_subset"; @@ -141,11 +141,11 @@ by (blast_tac (claset() addIs [prem1, prem2 RS subsetD]) 1); qed "vimage_pair_mono"; -goal thy "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B"; +Goal "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B"; by (Blast_tac 1); qed "image_mono"; -goal thy "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B"; +Goal "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B"; by (Blast_tac 1); qed "vimage_mono"; @@ -156,7 +156,7 @@ (** Monotonicity of implications -- some could go to FOL **) -goal thy "!!A B x. A<=B ==> x:A --> x:B"; +Goal "!!A B x. A<=B ==> x:A --> x:B"; by (Blast_tac 1); qed "in_mono";