--- a/src/ZF/AC.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/AC.ML Wed Dec 07 13:12:04 1994 +0100
@@ -15,20 +15,20 @@
by (asm_simp_tac (ZF_ss addsimps [Pi_empty1]) 2 THEN fast_tac ZF_cs 2);
(*The non-trivial case*)
by (fast_tac (eq_cs addSIs [AC, nonempty]) 1);
-val AC_Pi = result();
+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)";
by (resolve_tac [AC_Pi] 1);
by (eresolve_tac [bspec] 1);
by (assume_tac 1);
-val AC_ball_Pi = result();
+qed "AC_ball_Pi";
goal AC.thy "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 (fast_tac eq_cs 1);
-val AC_Pi_Pow = result();
+qed "AC_Pi_Pow";
val [nonempty] = goal AC.thy
"[| !!x. x:A ==> (EX y. y:x) \
@@ -36,17 +36,17 @@
by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
by (etac nonempty 1);
by (fast_tac (ZF_cs addDs [apply_type] addIs [Pi_type]) 1);
-val AC_func = result();
+qed "AC_func";
goal ZF.thy "!!x A. [| 0 ~: A; x: A |] ==> EX y. y:x";
by (subgoal_tac "x ~= 0" 1);
by (ALLGOALS (fast_tac eq_cs));
-val non_empty_family = result();
+qed "non_empty_family";
goal AC.thy "!!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));
-val AC_func0 = result();
+qed "AC_func0";
goal AC.thy "EX f: (Pow(C)-{0}) -> C. ALL x:(Pow(C)-{0}). f`x : x";
by (resolve_tac [AC_func0 RS bexE] 1);
@@ -54,5 +54,5 @@
by (assume_tac 2);
by (eresolve_tac [fun_weaken_type] 2);
by (ALLGOALS (fast_tac ZF_cs));
-val AC_func_Pow = result();
+qed "AC_func_Pow";
--- a/src/ZF/Arith.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Arith.ML Wed Dec 07 13:12:04 1994 +0100
@@ -25,12 +25,12 @@
goal Arith.thy "rec(0,a,b) = a";
by (rtac rec_trans 1);
by (rtac nat_case_0 1);
-val rec_0 = result();
+qed "rec_0";
goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))";
by (rtac rec_trans 1);
by (simp_tac (ZF_ss addsimps [nat_case_succ, nat_succI]) 1);
-val rec_succ = result();
+qed "rec_succ";
val major::prems = goal Arith.thy
"[| n: nat; \
@@ -40,7 +40,7 @@
by (rtac (major RS nat_induct) 1);
by (ALLGOALS
(asm_simp_tac (ZF_ss addsimps (prems@[rec_0,rec_succ]))));
-val rec_type = result();
+qed "rec_type";
val nat_le_refl = nat_into_Ord RS le_refl;
@@ -54,44 +54,44 @@
(** Addition **)
-val add_type = prove_goalw Arith.thy [add_def]
+qed_goalw "add_type" Arith.thy [add_def]
"[| m:nat; n:nat |] ==> m #+ n : nat"
(fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]);
-val add_0 = prove_goalw Arith.thy [add_def]
+qed_goalw "add_0" Arith.thy [add_def]
"0 #+ n = n"
(fn _ => [ (rtac rec_0 1) ]);
-val add_succ = prove_goalw Arith.thy [add_def]
+qed_goalw "add_succ" Arith.thy [add_def]
"succ(m) #+ n = succ(m #+ n)"
(fn _=> [ (rtac rec_succ 1) ]);
(** Multiplication **)
-val mult_type = prove_goalw Arith.thy [mult_def]
+qed_goalw "mult_type" Arith.thy [mult_def]
"[| m:nat; n:nat |] ==> m #* n : nat"
(fn prems=>
[ (typechk_tac (prems@[add_type]@nat_typechecks@ZF_typechecks)) ]);
-val mult_0 = prove_goalw Arith.thy [mult_def]
+qed_goalw "mult_0" Arith.thy [mult_def]
"0 #* n = 0"
(fn _ => [ (rtac rec_0 1) ]);
-val mult_succ = prove_goalw Arith.thy [mult_def]
+qed_goalw "mult_succ" Arith.thy [mult_def]
"succ(m) #* n = n #+ (m #* n)"
(fn _ => [ (rtac rec_succ 1) ]);
(** Difference **)
-val diff_type = prove_goalw Arith.thy [diff_def]
+qed_goalw "diff_type" Arith.thy [diff_def]
"[| m:nat; n:nat |] ==> m #- n : nat"
(fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]);
-val diff_0 = prove_goalw Arith.thy [diff_def]
+qed_goalw "diff_0" Arith.thy [diff_def]
"m #- 0 = m"
(fn _ => [ (rtac rec_0 1) ]);
-val diff_0_eq_0 = prove_goalw Arith.thy [diff_def]
+qed_goalw "diff_0_eq_0" Arith.thy [diff_def]
"n:nat ==> 0 #- n = 0"
(fn [prem]=>
[ (rtac (prem RS nat_induct) 1),
@@ -99,7 +99,7 @@
(*Must simplify BEFORE the induction!! (Else we get a critical pair)
succ(m) #- succ(n) rewrites to pred(succ(m) #- n) *)
-val diff_succ_succ = prove_goalw Arith.thy [diff_def]
+qed_goalw "diff_succ_succ" Arith.thy [diff_def]
"[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n"
(fn prems=>
[ (asm_simp_tac (nat_ss addsimps prems) 1),
@@ -114,7 +114,7 @@
(asm_simp_tac
(nat_ss addsimps (prems @ [le_iff, diff_0, diff_0_eq_0,
diff_succ_succ, nat_into_Ord]))));
-val diff_le_self = result();
+qed "diff_le_self";
(*** Simplification over add, mult, diff ***)
@@ -128,7 +128,7 @@
(*** Addition ***)
(*Associative law for addition*)
-val add_assoc = prove_goal Arith.thy
+qed_goal "add_assoc" Arith.thy
"m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
@@ -136,20 +136,20 @@
(*The following two lemmas are used for add_commute and sometimes
elsewhere, since they are safe for rewriting.*)
-val add_0_right = prove_goal Arith.thy
+qed_goal "add_0_right" Arith.thy
"m:nat ==> m #+ 0 = m"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
(ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]);
-val add_succ_right = prove_goal Arith.thy
+qed_goal "add_succ_right" Arith.thy
"m:nat ==> m #+ succ(n) = succ(m #+ n)"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
(ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]);
(*Commutative law for addition*)
-val add_commute = prove_goal Arith.thy
+qed_goal "add_commute" Arith.thy
"!!m n. [| m:nat; n:nat |] ==> m #+ n = n #+ m"
(fn _ =>
[ (nat_ind_tac "n" [] 1),
@@ -157,7 +157,7 @@
(asm_simp_tac (arith_ss addsimps [add_0_right, add_succ_right]))) ]);
(*for a/c rewriting*)
-val add_left_commute = prove_goal Arith.thy
+qed_goal "add_left_commute" Arith.thy
"!!m n k. [| m:nat; n:nat |] ==> m#+(n#+k)=n#+(m#+k)"
(fn _ => [asm_simp_tac (ZF_ss addsimps [add_assoc RS sym, add_commute]) 1]);
@@ -171,26 +171,26 @@
by (nat_ind_tac "k" [knat] 1);
by (ALLGOALS (simp_tac arith_ss));
by (fast_tac ZF_cs 1);
-val add_left_cancel = result();
+qed "add_left_cancel";
(*** Multiplication ***)
(*right annihilation in product*)
-val mult_0_right = prove_goal Arith.thy
+qed_goal "mult_0_right" Arith.thy
"!!m. m:nat ==> m #* 0 = 0"
(fn _=>
[ (nat_ind_tac "m" [] 1),
(ALLGOALS (asm_simp_tac arith_ss)) ]);
(*right successor law for multiplication*)
-val mult_succ_right = prove_goal Arith.thy
+qed_goal "mult_succ_right" Arith.thy
"!!m n. [| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)"
(fn _ =>
[ (nat_ind_tac "m" [] 1),
(ALLGOALS (asm_simp_tac (arith_ss addsimps add_ac))) ]);
(*Commutative law for multiplication*)
-val mult_commute = prove_goal Arith.thy
+qed_goal "mult_commute" Arith.thy
"[| m:nat; n:nat |] ==> m #* n = n #* m"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
@@ -198,14 +198,14 @@
(arith_ss addsimps (prems@[mult_0_right, mult_succ_right])))) ]);
(*addition distributes over multiplication*)
-val add_mult_distrib = prove_goal Arith.thy
+qed_goal "add_mult_distrib" Arith.thy
"!!m n. [| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)"
(fn _=>
[ (etac nat_induct 1),
(ALLGOALS (asm_simp_tac (arith_ss addsimps [add_assoc RS sym]))) ]);
(*Distributive law on the left; requires an extra typing premise*)
-val add_mult_distrib_left = prove_goal Arith.thy
+qed_goal "add_mult_distrib_left" Arith.thy
"!!m. [| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)"
(fn prems=>
[ (nat_ind_tac "m" [] 1),
@@ -213,14 +213,14 @@
(asm_simp_tac (arith_ss addsimps ([mult_succ_right] @ add_ac)) 1) ]);
(*Associative law for multiplication*)
-val mult_assoc = prove_goal Arith.thy
+qed_goal "mult_assoc" Arith.thy
"!!m n k. [| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)"
(fn _=>
[ (etac nat_induct 1),
(ALLGOALS (asm_simp_tac (arith_ss addsimps [add_mult_distrib]))) ]);
(*for a/c rewriting*)
-val mult_left_commute = prove_goal Arith.thy
+qed_goal "mult_left_commute" Arith.thy
"!!m n k. [| m:nat; n:nat; k:nat |] ==> m #* (n #* k) = n #* (m #* k)"
(fn _ => [rtac (mult_commute RS trans) 1,
rtac (mult_assoc RS trans) 3,
@@ -232,7 +232,7 @@
(*** Difference ***)
-val diff_self_eq_0 = prove_goal Arith.thy
+qed_goal "diff_self_eq_0" Arith.thy
"m:nat ==> m #- m = 0"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
@@ -245,26 +245,26 @@
by (etac rev_mp 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS (asm_simp_tac arith_ss));
-val add_diff_inverse = result();
+qed "add_diff_inverse";
(*Subtraction is the inverse of addition. *)
val [mnat,nnat] = goal Arith.thy
"[| m:nat; n:nat |] ==> (n#+m) #- n = m";
by (rtac (nnat RS nat_induct) 1);
by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mnat])));
-val diff_add_inverse = result();
+qed "diff_add_inverse";
goal Arith.thy
"!!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));
-val diff_add_inverse2 = result();
+qed "diff_add_inverse2";
val [mnat,nnat] = goal Arith.thy
"[| m:nat; n:nat |] ==> n #- (n#+m) = 0";
by (rtac (nnat RS nat_induct) 1);
by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mnat])));
-val diff_add_0 = result();
+qed "diff_add_0";
(*** Remainder ***)
@@ -275,7 +275,7 @@
by (etac rev_mp 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS (asm_simp_tac (nat_ss addsimps [diff_le_self,diff_succ_succ])));
-val div_termination = result();
+qed "div_termination";
val div_rls = (*for mod and div*)
nat_typechecks @
@@ -288,18 +288,18 @@
(*Type checking depends upon termination!*)
goalw Arith.thy [mod_def] "!!m n. [| 0<n; m:nat; n:nat |] ==> m mod n : nat";
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1));
-val mod_type = result();
+qed "mod_type";
goal Arith.thy "!!m n. [| 0<n; m<n |] ==> m mod n = m";
by (rtac (mod_def RS def_transrec RS trans) 1);
by (asm_simp_tac div_ss 1);
-val mod_less = result();
+qed "mod_less";
goal Arith.thy "!!m n. [| 0<n; n le m; m:nat |] ==> 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);
-val mod_geq = result();
+qed "mod_geq";
(*** Quotient ***)
@@ -307,19 +307,19 @@
goalw Arith.thy [div_def]
"!!m n. [| 0<n; m:nat; n:nat |] ==> m div n : nat";
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1));
-val div_type = result();
+qed "div_type";
goal Arith.thy "!!m n. [| 0<n; m<n |] ==> m div n = 0";
by (rtac (div_def RS def_transrec RS trans) 1);
by (asm_simp_tac div_ss 1);
-val div_less = result();
+qed "div_less";
goal Arith.thy
"!!m n. [| 0<n; n le m; m:nat |] ==> 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);
by (asm_simp_tac div_ss 1);
-val div_geq = result();
+qed "div_geq";
(*Main Result.*)
goal Arith.thy
@@ -333,7 +333,7 @@
(arith_ss addsimps [not_lt_iff_le, nat_into_Ord,
mod_geq, div_geq, add_assoc,
div_termination RS ltD, add_diff_inverse]) 1);
-val mod_div_equality = result();
+qed "mod_div_equality";
(**** Additional theorems about "le" ****)
@@ -341,12 +341,12 @@
goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le m #+ n";
by (etac nat_induct 1);
by (ALLGOALS (asm_simp_tac arith_ss));
-val add_le_self = result();
+qed "add_le_self";
goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le n #+ m";
by (rtac (add_commute RS ssubst) 1);
by (REPEAT (ares_tac [add_le_self] 1));
-val add_le_self2 = result();
+qed "add_le_self2";
(** Monotonicity of addition **)
@@ -356,7 +356,7 @@
by (assume_tac 1);
by (etac succ_lt_induct 1);
by (ALLGOALS (asm_simp_tac (arith_ss addsimps [leI])));
-val add_lt_mono1 = result();
+qed "add_lt_mono1";
(*strict, in both arguments*)
goal Arith.thy "!!i j k l. [| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l";
@@ -366,7 +366,7 @@
rtac (add_commute RS ssubst) 3,
rtac add_lt_mono1 5]);
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1));
-val add_lt_mono = result();
+qed "add_lt_mono";
(*A [clumsy] way of lifting < monotonicity to le monotonicity *)
val lt_mono::ford::prems = goal Ordinal.thy
@@ -376,14 +376,14 @@
\ |] ==> f(i) le f(j)";
by (cut_facts_tac prems 1);
by (fast_tac (lt_cs addSIs [lt_mono,ford] addSEs [leE]) 1);
-val Ord_lt_mono_imp_le_mono = result();
+qed "Ord_lt_mono_imp_le_mono";
(*le monotonicity, 1st argument*)
goal Arith.thy
"!!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));
-val add_le_mono1 = result();
+qed "add_le_mono1";
(* le monotonicity, BOTH arguments*)
goal Arith.thy
@@ -394,4 +394,4 @@
rtac (add_commute RS ssubst) 3,
rtac add_le_mono1 5]);
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
-val add_le_mono = result();
+qed "add_le_mono";
--- a/src/ZF/Bool.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Bool.ML Wed Dec 07 13:12:04 1994 +0100
@@ -14,15 +14,15 @@
goalw Bool.thy bool_defs "1 : bool";
by (rtac (consI1 RS consI2) 1);
-val bool_1I = result();
+qed "bool_1I";
goalw Bool.thy bool_defs "0 : bool";
by (rtac consI1 1);
-val bool_0I = result();
+qed "bool_0I";
goalw Bool.thy bool_defs "1~=0";
by (rtac succ_not_0 1);
-val one_not_0 = result();
+qed "one_not_0";
(** 1=0 ==> R **)
val one_neq_0 = one_not_0 RS notE;
@@ -31,36 +31,36 @@
"[| c: bool; c=1 ==> P; c=0 ==> P |] ==> P";
by (rtac (major RS consE) 1);
by (REPEAT (eresolve_tac (singletonE::prems) 1));
-val boolE = result();
+qed "boolE";
(** cond **)
(*1 means true*)
goalw Bool.thy bool_defs "cond(1,c,d) = c";
by (rtac (refl RS if_P) 1);
-val cond_1 = result();
+qed "cond_1";
(*0 means false*)
goalw Bool.thy bool_defs "cond(0,c,d) = d";
by (rtac (succ_not_0 RS not_sym RS if_not_P) 1);
-val cond_0 = result();
+qed "cond_0";
val major::prems = goal Bool.thy
"[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)";
by (rtac (major RS boolE) 1);
by (asm_simp_tac (ZF_ss addsimps (cond_1::prems)) 1);
by (asm_simp_tac (ZF_ss addsimps (cond_0::prems)) 1);
-val cond_type = result();
+qed "cond_type";
val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c";
by (rewtac rew);
by (rtac cond_1 1);
-val def_cond_1 = result();
+qed "def_cond_1";
val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d";
by (rewtac rew);
by (rtac cond_0 1);
-val def_cond_0 = result();
+qed "def_cond_0";
fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)];
@@ -72,19 +72,19 @@
val [xor_1,xor_0] = conds xor_def;
-val not_type = prove_goalw Bool.thy [not_def]
+qed_goalw "not_type" Bool.thy [not_def]
"a:bool ==> not(a) : bool"
(fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]);
-val and_type = prove_goalw Bool.thy [and_def]
+qed_goalw "and_type" Bool.thy [and_def]
"[| a:bool; b:bool |] ==> a and b : bool"
(fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]);
-val or_type = prove_goalw Bool.thy [or_def]
+qed_goalw "or_type" Bool.thy [or_def]
"[| a:bool; b:bool |] ==> a or b : bool"
(fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]);
-val xor_type = prove_goalw Bool.thy [xor_def]
+qed_goalw "xor_type" Bool.thy [xor_def]
"[| a:bool; b:bool |] ==> a xor b : bool"
(fn prems=> [ (typechk_tac(prems@[bool_1I, bool_0I, cond_type, not_type])) ]);
@@ -102,58 +102,58 @@
goal Bool.thy "!!a. a:bool ==> not(not(a)) = a";
by (bool0_tac 1);
-val not_not = result();
+qed "not_not";
goal Bool.thy "!!a b. a:bool ==> not(a and b) = not(a) or not(b)";
by (bool0_tac 1);
-val not_and = result();
+qed "not_and";
goal Bool.thy "!!a b. a:bool ==> not(a or b) = not(a) and not(b)";
by (bool0_tac 1);
-val not_or = result();
+qed "not_or";
(*** Laws about 'and' ***)
goal Bool.thy "!!a. a: bool ==> a and a = a";
by (bool0_tac 1);
-val and_absorb = result();
+qed "and_absorb";
goal Bool.thy "!!a b. [| a: bool; b:bool |] ==> a and b = b and a";
by (etac boolE 1);
by (bool0_tac 1);
by (bool0_tac 1);
-val and_commute = result();
+qed "and_commute";
goal Bool.thy
"!!a. a: bool ==> (a and b) and c = a and (b and c)";
by (bool0_tac 1);
-val and_assoc = result();
+qed "and_assoc";
goal Bool.thy
"!!a. [| a: bool; b:bool; c:bool |] ==> \
\ (a or b) and c = (a and c) or (b and c)";
by (REPEAT (bool0_tac 1));
-val and_or_distrib = result();
+qed "and_or_distrib";
(** binary orion **)
goal Bool.thy "!!a. a: bool ==> a or a = a";
by (bool0_tac 1);
-val or_absorb = result();
+qed "or_absorb";
goal Bool.thy "!!a b. [| a: bool; b:bool |] ==> a or b = b or a";
by (etac boolE 1);
by (bool0_tac 1);
by (bool0_tac 1);
-val or_commute = result();
+qed "or_commute";
goal Bool.thy "!!a. a: bool ==> (a or b) or c = a or (b or c)";
by (bool0_tac 1);
-val or_assoc = result();
+qed "or_assoc";
goal Bool.thy
"!!a b c. [| a: bool; b: bool; c: bool |] ==> \
\ (a and b) or c = (a or c) and (b or c)";
by (REPEAT (bool0_tac 1));
-val or_and_distrib = result();
+qed "or_and_distrib";
--- a/src/ZF/Cardinal.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Cardinal.ML Wed Dec 07 13:12:04 1994 +0100
@@ -17,7 +17,7 @@
goal Cardinal.thy "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));
-val decomp_bnd_mono = result();
+qed "decomp_bnd_mono";
val [gfun] = goal Cardinal.thy
"g: Y->X ==> \
@@ -27,7 +27,7 @@
(decomp_bnd_mono RS lfp_Tarski RS ssubst) 1);
by (simp_tac (ZF_ss addsimps [subset_refl, double_complement,
gfun RS fun_is_rel RS image_subset]) 1);
-val Banach_last_equation = result();
+qed "Banach_last_equation";
val prems = goal Cardinal.thy
"[| f: X->Y; g: Y->X |] ==> \
@@ -39,7 +39,7 @@
(resolve_tac [refl, exI, conjI, Diff_disjoint, Diff_partition])));
by (rtac Banach_last_equation 3);
by (REPEAT (resolve_tac (prems@[fun_is_rel, image_subset, lfp_subset]) 1));
-val decomposition = result();
+qed "decomposition";
val prems = goal Cardinal.thy
"[| f: inj(X,Y); g: inj(Y,X) |] ==> EX h. h: bij(X,Y)";
@@ -49,7 +49,7 @@
addIs [bij_converse_bij]) 1);
(* The instantiation of exI to "restrict(f,XA) Un converse(restrict(g,YB))"
is forced by the context!! *)
-val schroeder_bernstein = result();
+qed "schroeder_bernstein";
(** Equipollence is an equivalence relation **)
@@ -57,35 +57,35 @@
goalw Cardinal.thy [eqpoll_def] "X eqpoll X";
by (rtac exI 1);
by (rtac id_bij 1);
-val eqpoll_refl = result();
+qed "eqpoll_refl";
goalw Cardinal.thy [eqpoll_def] "!!X Y. X eqpoll Y ==> Y eqpoll X";
by (fast_tac (ZF_cs addIs [bij_converse_bij]) 1);
-val eqpoll_sym = result();
+qed "eqpoll_sym";
goalw Cardinal.thy [eqpoll_def]
"!!X Y. [| X eqpoll Y; Y eqpoll Z |] ==> X eqpoll Z";
by (fast_tac (ZF_cs addIs [comp_bij]) 1);
-val eqpoll_trans = result();
+qed "eqpoll_trans";
(** Le-pollence is a partial ordering **)
goalw Cardinal.thy [lepoll_def] "!!X Y. X<=Y ==> X lepoll Y";
by (rtac exI 1);
by (etac id_subset_inj 1);
-val subset_imp_lepoll = result();
+qed "subset_imp_lepoll";
val lepoll_refl = subset_refl RS subset_imp_lepoll;
goalw Cardinal.thy [eqpoll_def, bij_def, lepoll_def]
"!!X Y. X eqpoll Y ==> X lepoll Y";
by (fast_tac ZF_cs 1);
-val eqpoll_imp_lepoll = result();
+qed "eqpoll_imp_lepoll";
goalw Cardinal.thy [lepoll_def]
"!!X Y. [| X lepoll Y; Y lepoll Z |] ==> X lepoll Z";
by (fast_tac (ZF_cs addIs [comp_inj]) 1);
-val lepoll_trans = result();
+qed "lepoll_trans";
(*Asymmetry law*)
goalw Cardinal.thy [lepoll_def,eqpoll_def]
@@ -93,17 +93,17 @@
by (REPEAT (etac exE 1));
by (rtac schroeder_bernstein 1);
by (REPEAT (assume_tac 1));
-val eqpollI = result();
+qed "eqpollI";
val [major,minor] = goal Cardinal.thy
"[| X eqpoll Y; [| X lepoll Y; Y lepoll X |] ==> P |] ==> P";
by (rtac minor 1);
by (REPEAT (resolve_tac [major, eqpoll_imp_lepoll, eqpoll_sym] 1));
-val eqpollE = result();
+qed "eqpollE";
goal Cardinal.thy "X eqpoll Y <-> X lepoll Y & Y lepoll X";
by (fast_tac (ZF_cs addIs [eqpollI] addSEs [eqpollE]) 1);
-val eqpoll_iff = result();
+qed "eqpoll_iff";
(** LEAST -- the least number operator [from HOL/Univ.ML] **)
@@ -115,7 +115,7 @@
by (REPEAT (etac conjE 1));
by (etac (premOrd RS Ord_linear_lt) 1);
by (ALLGOALS (fast_tac (ZF_cs addSIs [premP] addSDs [premNot])));
-val Least_equality = result();
+qed "Least_equality";
goal Cardinal.thy "!!i. [| P(i); Ord(i) |] ==> P(LEAST x.P(x))";
by (etac rev_mp 1);
@@ -125,7 +125,7 @@
by (EVERY1 [rtac (Least_equality RS ssubst), assume_tac, assume_tac]);
by (assume_tac 2);
by (fast_tac (ZF_cs addSEs [ltE]) 1);
-val LeastI = result();
+qed "LeastI";
(*Proof is almost identical to the one above!*)
goal Cardinal.thy "!!i. [| P(i); Ord(i) |] ==> (LEAST x.P(x)) le i";
@@ -136,20 +136,20 @@
by (EVERY1 [rtac (Least_equality RS ssubst), assume_tac, assume_tac]);
by (etac le_refl 2);
by (fast_tac (ZF_cs addEs [ltE, lt_trans1] addIs [leI, ltI]) 1);
-val Least_le = result();
+qed "Least_le";
(*LEAST really is the smallest*)
goal Cardinal.thy "!!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));
-val less_LeastE = result();
+qed "less_LeastE";
(*If there is no such P then LEAST is vacuously 0*)
goalw Cardinal.thy [Least_def]
"!!P. [| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x.P(x)) = 0";
by (rtac the_0 1);
by (fast_tac ZF_cs 1);
-val Least_0 = result();
+qed "Least_0";
goal Cardinal.thy "Ord(LEAST x.P(x))";
by (excluded_middle_tac "EX i. Ord(i) & P(i)" 1);
@@ -158,7 +158,7 @@
by (REPEAT_SOME assume_tac);
by (etac (Least_0 RS ssubst) 1);
by (rtac Ord_0 1);
-val Ord_Least = result();
+qed "Ord_Least";
(** Basic properties of cardinals **)
@@ -167,13 +167,13 @@
val prems = goal Cardinal.thy
"[| !!y. P(y) <-> Q(y) |] ==> (LEAST x.P(x)) = (LEAST x.Q(x))";
by (simp_tac (FOL_ss addsimps prems) 1);
-val Least_cong = result();
+qed "Least_cong";
(*Need AC to prove X lepoll Y ==> |X| le |Y| ; see well_ord_lepoll_imp_le *)
goalw Cardinal.thy [eqpoll_def,cardinal_def] "!!X Y. X eqpoll Y ==> |X| = |Y|";
by (rtac Least_cong 1);
by (fast_tac (ZF_cs addEs [comp_bij,bij_converse_bij]) 1);
-val cardinal_cong = result();
+qed "cardinal_cong";
(*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*)
goalw Cardinal.thy [eqpoll_def, cardinal_def]
@@ -182,7 +182,7 @@
by (etac Ord_ordertype 2);
by (rtac exI 1);
by (etac (ordermap_bij RS bij_converse_bij) 1);
-val well_ord_cardinal_eqpoll = result();
+qed "well_ord_cardinal_eqpoll";
val Ord_cardinal_eqpoll = well_ord_Memrel RS well_ord_cardinal_eqpoll
|> standard;
@@ -192,38 +192,38 @@
by (rtac (eqpoll_sym RS eqpoll_trans) 1);
by (etac well_ord_cardinal_eqpoll 1);
by (asm_simp_tac (ZF_ss addsimps [well_ord_cardinal_eqpoll]) 1);
-val well_ord_cardinal_eqE = result();
+qed "well_ord_cardinal_eqE";
(** Observations from Kunen, page 28 **)
goalw Cardinal.thy [cardinal_def] "!!i. Ord(i) ==> |i| le i";
by (etac (eqpoll_refl RS Least_le) 1);
-val Ord_cardinal_le = result();
+qed "Ord_cardinal_le";
goalw Cardinal.thy [Card_def] "!!K. Card(K) ==> |K| = K";
by (etac sym 1);
-val Card_cardinal_eq = result();
+qed "Card_cardinal_eq";
val prems = goalw Cardinal.thy [Card_def,cardinal_def]
"[| Ord(i); !!j. j<i ==> ~(j eqpoll i) |] ==> Card(i)";
by (rtac (Least_equality RS ssubst) 1);
by (REPEAT (ares_tac ([refl,eqpoll_refl]@prems) 1));
-val CardI = result();
+qed "CardI";
goalw Cardinal.thy [Card_def, cardinal_def] "!!i. Card(i) ==> Ord(i)";
by (etac ssubst 1);
by (rtac Ord_Least 1);
-val Card_is_Ord = result();
+qed "Card_is_Ord";
goalw Cardinal.thy [cardinal_def] "Ord(|A|)";
by (rtac Ord_Least 1);
-val Ord_cardinal = result();
+qed "Ord_cardinal";
goal Cardinal.thy "Card(0)";
by (rtac (Ord_0 RS CardI) 1);
by (fast_tac (ZF_cs addSEs [ltE]) 1);
-val Card_0 = result();
+qed "Card_0";
val [premK,premL] = goal Cardinal.thy
"[| Card(K); Card(L) |] ==> Card(K Un L)";
@@ -232,7 +232,7 @@
(ZF_ss addsimps [premL, le_imp_subset, subset_Un_iff RS iffD1]) 1);
by (asm_simp_tac
(ZF_ss addsimps [premK, le_imp_subset, subset_Un_iff2 RS iffD1]) 1);
-val Card_Un = result();
+qed "Card_Un";
(*Infinite unions of cardinals? See Devlin, Lemma 6.7, page 98*)
@@ -245,7 +245,7 @@
by (assume_tac 2);
by (etac eqpoll_trans 1);
by (REPEAT (ares_tac [LeastI] 1));
-val Card_cardinal = result();
+qed "Card_cardinal";
(*Kunen's Lemma 10.5*)
goal Cardinal.thy "!!i j. [| |i| le j; j le i |] ==> |j| = |i|";
@@ -256,7 +256,7 @@
by (rtac (eqpoll_sym RS eqpoll_imp_lepoll) 1);
by (rtac Ord_cardinal_eqpoll 1);
by (REPEAT (eresolve_tac [ltE, Ord_succD] 1));
-val cardinal_eq_lemma = result();
+qed "cardinal_eq_lemma";
goal Cardinal.thy "!!i j. i le j ==> |i| le |j|";
by (res_inst_tac [("i","|i|"),("j","|j|")] Ord_linear_le 1);
@@ -266,7 +266,7 @@
by (etac le_trans 1);
by (etac ltE 1);
by (etac Ord_cardinal_le 1);
-val cardinal_mono = result();
+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";
@@ -274,22 +274,22 @@
by (REPEAT_SOME assume_tac);
by (etac (lt_trans2 RS lt_irrefl) 1);
by (etac cardinal_mono 1);
-val cardinal_lt_imp_lt = result();
+qed "cardinal_lt_imp_lt";
goal Cardinal.thy "!!i j. [| |i| < K; Ord(i); Card(K) |] ==> i < K";
by (asm_simp_tac (ZF_ss addsimps
[cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1);
-val Card_lt_imp_lt = result();
+qed "Card_lt_imp_lt";
goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (|i| < K) <-> (i < K)";
by (fast_tac (ZF_cs addEs [Card_lt_imp_lt, Ord_cardinal_le RS lt_trans1]) 1);
-val Card_lt_iff = result();
+qed "Card_lt_iff";
goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (K le |i|) <-> (K le i)";
by (asm_simp_tac (ZF_ss addsimps
[Card_lt_iff, Card_is_Ord, Ord_cardinal,
not_lt_iff_le RS iff_sym]) 1);
-val Card_le_iff = result();
+qed "Card_le_iff";
(** The swap operator [NOT USED] **)
@@ -297,18 +297,18 @@
goalw Cardinal.thy [swap_def]
"!!A. [| x:A; y:A |] ==> swap(A,x,y) : A->A";
by (REPEAT (ares_tac [lam_type,if_type] 1));
-val swap_type = result();
+qed "swap_type";
goalw Cardinal.thy [swap_def]
"!!A. [| x:A; y:A; z:A |] ==> swap(A,x,y)`(swap(A,x,y)`z) = z";
by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1);
-val swap_swap_identity = result();
+qed "swap_swap_identity";
goal Cardinal.thy "!!A. [| x:A; y:A |] ==> swap(A,x,y) : bij(A,A)";
by (rtac nilpotent_imp_bijective 1);
by (REPEAT (ares_tac [swap_type, comp_eq_id_iff RS iffD2,
ballI, swap_swap_identity] 1));
-val swap_bij = result();
+qed "swap_bij";
(*** The finite cardinals ***)
@@ -326,7 +326,7 @@
(*Adding prem earlier would cause the simplifier to loop*)
by (cut_facts_tac [prem] 1);
by (fast_tac (ZF_cs addSEs [mem_irrefl]) 1);
-val inj_succ_succD = result();
+qed "inj_succ_succD";
val [prem] = goalw Cardinal.thy [lepoll_def]
"m:nat ==> ALL n: nat. m lepoll n --> m le n";
@@ -336,7 +336,7 @@
by (eres_inst_tac [("n","n")] natE 1);
by (asm_simp_tac (ZF_ss addsimps [inj_def, succI1 RS Pi_empty2]) 1);
by (fast_tac (ZF_cs addSIs [succ_leI] addSDs [inj_succ_succD]) 1);
-val nat_lepoll_imp_le_lemma = result();
+qed "nat_lepoll_imp_le_lemma";
val nat_lepoll_imp_le = nat_lepoll_imp_le_lemma RS bspec RS mp |> standard;
goal Cardinal.thy
@@ -345,7 +345,7 @@
by (asm_simp_tac (ZF_ss addsimps [eqpoll_refl]) 2);
by (fast_tac (ZF_cs addIs [nat_lepoll_imp_le, le_anti_sym]
addSEs [eqpollE]) 1);
-val nat_eqpoll_iff = result();
+qed "nat_eqpoll_iff";
goalw Cardinal.thy [Card_def,cardinal_def]
"!!n. n: nat ==> Card(n)";
@@ -353,13 +353,13 @@
by (REPEAT_FIRST (ares_tac [eqpoll_refl, nat_into_Ord, refl]));
by (asm_simp_tac (ZF_ss addsimps [lt_nat_in_nat RS nat_eqpoll_iff]) 1);
by (fast_tac (ZF_cs addSEs [lt_irrefl]) 1);
-val nat_into_Card = result();
+qed "nat_into_Card";
(*Part of Kunen's Lemma 10.6*)
goal Cardinal.thy "!!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));
-val succ_lepoll_natE = result();
+qed "succ_lepoll_natE";
(*** The first infinite cardinal: Omega, or nat ***)
@@ -371,7 +371,7 @@
by (rtac lepoll_trans 1 THEN assume_tac 2);
by (etac ltE 1);
by (REPEAT (ares_tac [Ord_succ_subsetI RS subset_imp_lepoll] 1));
-val lt_not_lepoll = result();
+qed "lt_not_lepoll";
goal Cardinal.thy "!!i n. [| Ord(i); n:nat |] ==> i eqpoll n <-> i=n";
by (rtac iffI 1);
@@ -382,20 +382,20 @@
REPEAT (assume_tac 1));
by (rtac (lt_not_lepoll RS notE) 1 THEN (REPEAT (assume_tac 1)));
by (etac eqpoll_imp_lepoll 1);
-val Ord_nat_eqpoll_iff = result();
+qed "Ord_nat_eqpoll_iff";
goalw Cardinal.thy [Card_def,cardinal_def] "Card(nat)";
by (rtac (Least_equality RS ssubst) 1);
by (REPEAT_FIRST (ares_tac [eqpoll_refl, Ord_nat, refl]));
by (etac ltE 1);
by (asm_simp_tac (ZF_ss addsimps [eqpoll_iff, lt_not_lepoll, ltI]) 1);
-val Card_nat = result();
+qed "Card_nat";
(*Allows showing that |i| is a limit cardinal*)
goal Cardinal.thy "!!i. nat le i ==> nat le |i|";
by (rtac (Card_nat RS Card_cardinal_eq RS subst) 1);
by (etac cardinal_mono 1);
-val nat_le_cardinal = result();
+qed "nat_le_cardinal";
(*** Towards Cardinal Arithmetic ***)
@@ -421,18 +421,18 @@
setloop etac consE') 1);
by (asm_simp_tac (ZF_ss addsimps [inj_is_fun RS apply_type, left_inverse]
setloop etac consE') 1);
-val cons_lepoll_cong = result();
+qed "cons_lepoll_cong";
goal Cardinal.thy
"!!A B. [| A eqpoll B; a ~: A; b ~: B |] ==> cons(a,A) eqpoll cons(b,B)";
by (asm_full_simp_tac (ZF_ss addsimps [eqpoll_iff, cons_lepoll_cong]) 1);
-val cons_eqpoll_cong = result();
+qed "cons_eqpoll_cong";
(*Congruence law for succ under equipollence*)
goalw Cardinal.thy [succ_def]
"!!A B. A eqpoll B ==> succ(A) eqpoll succ(B)";
by (REPEAT (ares_tac [cons_eqpoll_cong, mem_not_refl] 1));
-val succ_eqpoll_cong = result();
+qed "succ_eqpoll_cong";
(*Congruence law for + under equipollence*)
goalw Cardinal.thy [eqpoll_def]
@@ -444,7 +444,7 @@
lam_bijective 1);
by (safe_tac (ZF_cs addSEs [sumE]));
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
-val sum_eqpoll_cong = result();
+qed "sum_eqpoll_cong";
(*Congruence law for * under equipollence*)
goalw Cardinal.thy [eqpoll_def]
@@ -456,7 +456,7 @@
lam_bijective 1);
by (safe_tac ZF_cs);
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
-val prod_eqpoll_cong = result();
+qed "prod_eqpoll_cong";
goalw Cardinal.thy [eqpoll_def]
"!!f. [| f: inj(A,B); A Int B = 0 |] ==> A Un (B - range(f)) eqpoll B";
@@ -474,4 +474,4 @@
(ZF_ss addsimps [inj_converse_fun RS apply_funtype, right_inverse]
setloop split_tac [expand_if]) 1);
by (fast_tac (ZF_cs addEs [equals0D]) 1);
-val inj_disjoint_eqpoll = result();
+qed "inj_disjoint_eqpoll";
--- a/src/ZF/CardinalArith.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/CardinalArith.ML Wed Dec 07 13:12:04 1994 +0100
@@ -26,7 +26,7 @@
by (rtac well_ord_cardinal_eqpoll 1);
by (etac well_ord_rvimage 1);
by (assume_tac 1);
-val well_ord_lepoll_imp_le = result();
+qed "well_ord_lepoll_imp_le";
(*Each element of Fin(A) is equivalent to a natural number*)
goal CardinalArith.thy
@@ -36,7 +36,7 @@
by (fast_tac (ZF_cs addSIs [cons_eqpoll_cong,
rewrite_rule [succ_def] nat_succI]
addSEs [mem_irrefl]) 1);
-val Fin_eqpoll = result();
+qed "Fin_eqpoll";
(*** Cardinal addition ***)
@@ -48,11 +48,11 @@
lam_bijective 1);
by (safe_tac (ZF_cs addSEs [sumE]));
by (ALLGOALS (asm_simp_tac case_ss));
-val sum_commute_eqpoll = result();
+qed "sum_commute_eqpoll";
goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i";
by (rtac (sum_commute_eqpoll RS cardinal_cong) 1);
-val cadd_commute = result();
+qed "cadd_commute";
(** Cardinal addition is associative **)
@@ -62,7 +62,7 @@
("d", "case(%x.Inl(Inl(x)), case(%x.Inl(Inr(x)), Inr))")]
lam_bijective 1);
by (ALLGOALS (asm_simp_tac (case_ss setloop etac sumE)));
-val sum_assoc_eqpoll = result();
+qed "sum_assoc_eqpoll";
(*Unconditional version requires AC*)
goalw CardinalArith.thy [cadd_def]
@@ -75,7 +75,7 @@
br ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong RS
eqpoll_sym) 2;
by (REPEAT (ares_tac [well_ord_radd] 1));
-val well_ord_cadd_assoc = result();
+qed "well_ord_cadd_assoc";
(** 0 is the identity for addition **)
@@ -84,12 +84,12 @@
by (res_inst_tac [("c", "case(%x.x, %y.y)"), ("d", "Inr")]
lam_bijective 1);
by (ALLGOALS (asm_simp_tac (case_ss setloop eresolve_tac [sumE,emptyE])));
-val sum_0_eqpoll = result();
+qed "sum_0_eqpoll";
goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
by (asm_simp_tac (ZF_ss addsimps [sum_0_eqpoll RS cardinal_cong,
Card_cardinal_eq]) 1);
-val cadd_0 = result();
+qed "cadd_0";
(** Addition of finite cardinals is "ordinary" addition **)
@@ -101,7 +101,7 @@
by (ALLGOALS
(asm_simp_tac (case_ss addsimps [succI2, mem_imp_not_eq]
setloop eresolve_tac [sumE,succE])));
-val sum_succ_eqpoll = result();
+qed "sum_succ_eqpoll";
(*Pulling the succ(...) outside the |...| requires m, n: nat *)
(*Unconditional version requires AC*)
@@ -111,7 +111,7 @@
by (rtac (succ_eqpoll_cong RS cardinal_cong) 1);
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1);
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1));
-val cadd_succ_lemma = result();
+qed "cadd_succ_lemma";
val [mnat,nnat] = goal CardinalArith.thy
"[| m: nat; n: nat |] ==> m |+| n = m#+n";
@@ -120,7 +120,7 @@
by (asm_simp_tac (arith_ss addsimps [nat_into_Card RS cadd_0]) 1);
by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cadd_succ_lemma,
nat_into_Card RS Card_cardinal_eq]) 1);
-val nat_cadd_eq_add = result();
+qed "nat_cadd_eq_add";
(*** Cardinal multiplication ***)
@@ -134,11 +134,11 @@
lam_bijective 1);
by (safe_tac ZF_cs);
by (ALLGOALS (asm_simp_tac ZF_ss));
-val prod_commute_eqpoll = result();
+qed "prod_commute_eqpoll";
goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i";
by (rtac (prod_commute_eqpoll RS cardinal_cong) 1);
-val cmult_commute = result();
+qed "cmult_commute";
(** Cardinal multiplication is associative **)
@@ -149,7 +149,7 @@
lam_bijective 1);
by (safe_tac ZF_cs);
by (ALLGOALS (asm_simp_tac ZF_ss));
-val prod_assoc_eqpoll = result();
+qed "prod_assoc_eqpoll";
(*Unconditional version requires AC*)
goalw CardinalArith.thy [cmult_def]
@@ -162,7 +162,7 @@
br ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS prod_eqpoll_cong RS
eqpoll_sym) 2;
by (REPEAT (ares_tac [well_ord_rmult] 1));
-val well_ord_cmult_assoc = result();
+qed "well_ord_cmult_assoc";
(** Cardinal multiplication distributes over addition **)
@@ -174,12 +174,12 @@
lam_bijective 1);
by (safe_tac (ZF_cs addSEs [sumE]));
by (ALLGOALS (asm_simp_tac case_ss));
-val sum_prod_distrib_eqpoll = result();
+qed "sum_prod_distrib_eqpoll";
goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A";
by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1);
by (simp_tac (ZF_ss addsimps [lam_type]) 1);
-val prod_square_lepoll = result();
+qed "prod_square_lepoll";
goalw CardinalArith.thy [cmult_def] "!!K. Card(K) ==> K le K |*| K";
by (rtac le_trans 1);
@@ -187,7 +187,7 @@
by (rtac prod_square_lepoll 3);
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2));
by (asm_simp_tac (ZF_ss addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
-val cmult_square_le = result();
+qed "cmult_square_le";
(** Multiplication by 0 yields 0 **)
@@ -195,12 +195,12 @@
by (rtac exI 1);
by (rtac lam_bijective 1);
by (safe_tac ZF_cs);
-val prod_0_eqpoll = result();
+qed "prod_0_eqpoll";
goalw CardinalArith.thy [cmult_def] "0 |*| i = 0";
by (asm_simp_tac (ZF_ss addsimps [prod_0_eqpoll RS cardinal_cong,
Card_0 RS Card_cardinal_eq]) 1);
-val cmult_0 = result();
+qed "cmult_0";
(** 1 is the identity for multiplication **)
@@ -209,12 +209,12 @@
by (res_inst_tac [("c", "snd"), ("d", "%z.<x,z>")] lam_bijective 1);
by (safe_tac ZF_cs);
by (ALLGOALS (asm_simp_tac ZF_ss));
-val prod_singleton_eqpoll = result();
+qed "prod_singleton_eqpoll";
goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
by (asm_simp_tac (ZF_ss addsimps [prod_singleton_eqpoll RS cardinal_cong,
Card_cardinal_eq]) 1);
-val cmult_1 = result();
+qed "cmult_1";
(** Multiplication of finite cardinals is "ordinary" multiplication **)
@@ -226,7 +226,7 @@
by (safe_tac (ZF_cs addSEs [sumE]));
by (ALLGOALS
(asm_simp_tac (case_ss addsimps [succI2, if_type, mem_imp_not_eq])));
-val prod_succ_eqpoll = result();
+qed "prod_succ_eqpoll";
(*Unconditional version requires AC*)
@@ -236,7 +236,7 @@
by (rtac (cardinal_cong RS sym) 1);
by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong) 1);
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
-val cmult_succ_lemma = result();
+qed "cmult_succ_lemma";
val [mnat,nnat] = goal CardinalArith.thy
"[| m: nat; n: nat |] ==> m |*| n = m#*n";
@@ -245,7 +245,7 @@
by (asm_simp_tac (arith_ss addsimps [cmult_0]) 1);
by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cmult_succ_lemma,
nat_cadd_eq_add]) 1);
-val nat_cmult_eq_mult = result();
+qed "nat_cmult_eq_mult";
(*** Infinite Cardinals are Limit Ordinals ***)
@@ -272,31 +272,31 @@
left_inverse, right_inverse, nat_0I, nat_succI,
nat_case_0, nat_case_succ]
setloop split_tac [expand_if]) 1);
-val nat_cons_lepoll = result();
+qed "nat_cons_lepoll";
goal CardinalArith.thy "!!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);
-val nat_cons_eqpoll = result();
+qed "nat_cons_eqpoll";
(*Specialized version required below*)
goalw CardinalArith.thy [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A";
by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1);
-val nat_succ_eqpoll = result();
+qed "nat_succ_eqpoll";
goalw CardinalArith.thy [InfCard_def] "InfCard(nat)";
by (fast_tac (ZF_cs addIs [Card_nat, le_refl, Card_is_Ord]) 1);
-val InfCard_nat = result();
+qed "InfCard_nat";
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
by (etac conjunct1 1);
-val InfCard_is_Card = result();
+qed "InfCard_is_Card";
goalw CardinalArith.thy [InfCard_def]
"!!K L. [| InfCard(K); Card(L) |] ==> InfCard(K Un L)";
by (asm_simp_tac (ZF_ss addsimps [Card_Un, Un_upper1_le RSN (2,le_trans),
Card_is_Ord]) 1);
-val InfCard_Un = result();
+qed "InfCard_Un";
(*Kunen's Lemma 10.11*)
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
@@ -312,7 +312,7 @@
(*Tricky combination of substitutions; backtracking needed*)
by (etac ssubst 1 THEN etac ssubst 1 THEN rtac Ord_cardinal_le 1);
by (assume_tac 1);
-val InfCard_is_Limit = result();
+qed "InfCard_is_Limit";
@@ -325,7 +325,7 @@
by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, well_ord_is_wf]) 1);
by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1);
by (rtac pred_subset 1);
-val ordermap_eqpoll_pred = result();
+qed "ordermap_eqpoll_pred";
(** Establishing the well-ordering **)
@@ -336,13 +336,13 @@
by (fast_tac (ZF_cs addIs [lam_type, Un_least_lt RS ltD, ltI]
addSEs [split_type]) 1);
by (asm_full_simp_tac ZF_ss 1);
-val csquare_lam_inj = result();
+qed "csquare_lam_inj";
goalw CardinalArith.thy [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));
-val well_ord_csquare = result();
+qed "well_ord_csquare";
(** Characterising initial segments of the well-ordering **)
@@ -355,7 +355,7 @@
by (safe_tac (ZF_cs addSEs [mem_irrefl]
addSIs [Un_upper1_le, Un_upper2_le]));
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [lt_def, succI2, Ord_succ])));
-val csquareD_lemma = result();
+qed "csquareD_lemma";
val csquareD = csquareD_lemma RS mp |> standard;
goalw CardinalArith.thy [pred_def]
@@ -365,7 +365,7 @@
by (rewtac lt_def);
by (assume_tac 4);
by (ALLGOALS (fast_tac ZF_cs));
-val pred_csquare_subset = result();
+qed "pred_csquare_subset";
goalw CardinalArith.thy [csquare_rel_def]
"!!K. [| x<z; y<z; z<K |] ==> \
@@ -375,7 +375,7 @@
by (REPEAT (etac ltE 1));
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
Un_absorb, Un_least_mem_iff, ltD]) 1);
-val csquare_ltI = result();
+qed "csquare_ltI";
(*Part of the traditional proof. UNUSED since it's harder to prove & apply *)
goalw CardinalArith.thy [csquare_rel_def]
@@ -390,7 +390,7 @@
by (ALLGOALS
(asm_simp_tac (ZF_ss addsimps [subset_Un_iff RS iff_sym,
subset_Un_iff2 RS iff_sym, OrdmemD])));
-val csquare_or_eqI = result();
+qed "csquare_or_eqI";
(** The cardinality of initial segments **)
@@ -407,7 +407,7 @@
by (ALLGOALS
(fast_tac (ZF_cs addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap]
addSEs [ltE])));
-val ordermap_z_lepoll = result();
+qed "ordermap_z_lepoll";
(*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *)
goalw CardinalArith.thy [cmult_def]
@@ -428,7 +428,7 @@
by (REPEAT_FIRST (etac ltE));
by (rtac (prod_eqpoll_cong RS eqpoll_sym RS eqpoll_imp_lepoll) 1);
by (REPEAT_FIRST (etac (Ord_succ RS Ord_cardinal_eqpoll)));
-val ordermap_csquare_le = result();
+qed "ordermap_csquare_le";
(*Kunen: "... so the order type <= K" *)
goal CardinalArith.thy
@@ -461,7 +461,7 @@
le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt,
Ord_Un, ltI, nat_le_cardinal,
Ord_cardinal_le RS lt_trans1 RS ltD]) 1);
-val ordertype_csquare_le = result();
+qed "ordertype_csquare_le";
(*This lemma can easily be generalized to premise well_ord(A*A,r) *)
goalw CardinalArith.thy [cmult_def]
@@ -470,7 +470,7 @@
by (rewtac eqpoll_def);
by (rtac exI 1);
by (etac (well_ord_csquare RS ordermap_bij) 1);
-val csquare_eq_ordertype = result();
+qed "csquare_eq_ordertype";
(*Main result: Kunen's Theorem 10.12*)
goal CardinalArith.thy "!!K. InfCard(K) ==> K |*| K = K";
@@ -486,7 +486,7 @@
by (asm_simp_tac
(ZF_ss addsimps [csquare_eq_ordertype, Ord_cardinal_le,
well_ord_csquare RS Ord_ordertype]) 1);
-val InfCard_csquare_eq = result();
+qed "InfCard_csquare_eq";
goal CardinalArith.thy
@@ -496,7 +496,7 @@
by (resolve_tac [well_ord_cardinal_eqE] 1);
by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1));
by (asm_simp_tac (ZF_ss addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1);
-val well_ord_InfCard_square_eq = result();
+qed "well_ord_InfCard_square_eq";
(*** For every cardinal number there exists a greater one
@@ -511,14 +511,14 @@
by (resolve_tac [UN_I] 1);
by (resolve_tac [ReplaceI] 2);
by (ALLGOALS (fast_tac (ZF_cs addSEs [well_ord_subset])));
-val Ord_jump_cardinal = result();
+qed "Ord_jump_cardinal";
(*Allows selective unfolding. Less work than deriving intro/elim rules*)
goalw CardinalArith.thy [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 <=*)
-val jump_cardinal_iff = result();
+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)";
@@ -528,7 +528,7 @@
by (resolve_tac [subset_refl] 2);
by (asm_simp_tac (ZF_ss addsimps [Memrel_def, subset_iff]) 1);
by (asm_simp_tac (ZF_ss addsimps [ordertype_Memrel]) 1);
-val K_lt_jump_cardinal = result();
+qed "K_lt_jump_cardinal";
(*The proof by contradiction: the bijection f yields a wellordering of X
whose ordertype is jump_cardinal(K). *)
@@ -547,7 +547,7 @@
by (asm_simp_tac
(ZF_ss addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage),
ordertype_Memrel, Ord_jump_cardinal]) 1);
-val Card_jump_cardinal_lemma = result();
+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))";
@@ -555,7 +555,7 @@
by (rewrite_goals_tac [eqpoll_def]);
by (safe_tac (ZF_cs addSDs [ltD, jump_cardinal_iff RS iffD1]));
by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1));
-val Card_jump_cardinal = result();
+qed "Card_jump_cardinal";
(*** Basic properties of successor cardinals ***)
@@ -564,7 +564,7 @@
by (rtac LeastI 1);
by (REPEAT (ares_tac [conjI, Card_jump_cardinal, K_lt_jump_cardinal,
Ord_jump_cardinal] 1));
-val csucc_basic = result();
+qed "csucc_basic";
val Card_csucc = csucc_basic RS conjunct1 |> standard;
@@ -573,13 +573,13 @@
goal CardinalArith.thy "!!K. Ord(K) ==> 0 < csucc(K)";
by (resolve_tac [[Ord_0_le, lt_csucc] MRS lt_trans1] 1);
by (REPEAT (assume_tac 1));
-val Ord_0_lt_csucc = result();
+qed "Ord_0_lt_csucc";
goalw CardinalArith.thy [csucc_def]
"!!K L. [| Card(L); K<L |] ==> csucc(K) le L";
by (rtac Least_le 1);
by (REPEAT (ares_tac [conjI, Card_is_Ord] 1));
-val csucc_le = result();
+qed "csucc_le";
goal CardinalArith.thy
"!!K. [| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K";
@@ -593,18 +593,18 @@
by (resolve_tac [Ord_cardinal_le RS lt_trans1] 1);
by (REPEAT (ares_tac [Ord_cardinal] 1
ORELSE eresolve_tac [ltE, Card_is_Ord] 1));
-val lt_csucc_iff = result();
+qed "lt_csucc_iff";
goal CardinalArith.thy
"!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K";
by (asm_simp_tac
(ZF_ss addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
-val Card_lt_csucc_iff = result();
+qed "Card_lt_csucc_iff";
goalw CardinalArith.thy [InfCard_def]
"!!K. InfCard(K) ==> InfCard(csucc(K))";
by (asm_simp_tac (ZF_ss addsimps [Card_csucc, Card_is_Ord,
lt_csucc RS leI RSN (2,le_trans)]) 1);
-val InfCard_csucc = result();
+qed "InfCard_csucc";
val Limit_csucc = InfCard_csucc RS InfCard_is_Limit |> standard;
--- a/src/ZF/Cardinal_AC.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Cardinal_AC.ML Wed Dec 07 13:12:04 1994 +0100
@@ -15,7 +15,7 @@
goal Cardinal_AC.thy "|A| eqpoll A";
by (resolve_tac [AC_well_ord RS exE] 1);
by (eresolve_tac [well_ord_cardinal_eqpoll] 1);
-val cardinal_eqpoll = result();
+qed "cardinal_eqpoll";
val cardinal_idem = cardinal_eqpoll RS cardinal_cong;
@@ -24,13 +24,13 @@
by (resolve_tac [AC_well_ord RS exE] 1);
by (resolve_tac [well_ord_cardinal_eqE] 1);
by (REPEAT_SOME assume_tac);
-val cardinal_eqE = result();
+qed "cardinal_eqE";
goal Cardinal_AC.thy "!!A B. A lepoll B ==> |A| le |B|";
by (resolve_tac [AC_well_ord RS exE] 1);
by (eresolve_tac [well_ord_lepoll_imp_le] 1);
by (assume_tac 1);
-val lepoll_imp_le = result();
+qed "lepoll_imp_le";
goal Cardinal_AC.thy "(i |+| j) |+| k = i |+| (j |+| k)";
by (resolve_tac [AC_well_ord RS exE] 1);
@@ -38,7 +38,7 @@
by (resolve_tac [AC_well_ord RS exE] 1);
by (resolve_tac [well_ord_cadd_assoc] 1);
by (REPEAT_SOME assume_tac);
-val cadd_assoc = result();
+qed "cadd_assoc";
goal Cardinal_AC.thy "(i |*| j) |*| k = i |*| (j |*| k)";
by (resolve_tac [AC_well_ord RS exE] 1);
@@ -46,13 +46,13 @@
by (resolve_tac [AC_well_ord RS exE] 1);
by (resolve_tac [well_ord_cmult_assoc] 1);
by (REPEAT_SOME assume_tac);
-val cmult_assoc = result();
+qed "cmult_assoc";
goal Cardinal_AC.thy "!!A. InfCard(|A|) ==> A*A eqpoll A";
by (resolve_tac [AC_well_ord RS exE] 1);
by (eresolve_tac [well_ord_InfCard_square_eq] 1);
by (assume_tac 1);
-val InfCard_square_eq = result();
+qed "InfCard_square_eq";
(*** Other applications of AC ***)
@@ -62,13 +62,13 @@
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);
-val le_imp_lepoll = result();
+qed "le_imp_lepoll";
goal Cardinal_AC.thy "!!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 [le_imp_lepoll,lepoll_imp_le] 1));
-val le_Card_iff = result();
+qed "le_Card_iff";
goalw Cardinal_AC.thy [surj_def] "!!f. f: surj(X,Y) ==> EX g. g: inj(Y,X)";
by (etac CollectE 1);
@@ -79,7 +79,7 @@
by (resolve_tac [Pi_type] 1 THEN assume_tac 1);
by (fast_tac (ZF_cs addDs [apply_type] addDs [Pi_memberD]) 1);
by (fast_tac (ZF_cs addDs [apply_type] addEs [apply_equality]) 1);
-val surj_implies_inj = result();
+qed "surj_implies_inj";
(*Kunen's Lemma 10.20*)
goal Cardinal_AC.thy "!!f. f: surj(X,Y) ==> |Y| le |X|";
@@ -87,7 +87,7 @@
by (eresolve_tac [surj_implies_inj RS exE] 1);
by (rewtac lepoll_def);
by (eresolve_tac [exI] 1);
-val surj_implies_cardinal_le = result();
+qed "surj_implies_cardinal_le";
(*Kunen's Lemma 10.21*)
goal Cardinal_AC.thy
@@ -115,7 +115,7 @@
by (dresolve_tac [apply_type] 1);
by (eresolve_tac [conjunct2] 1);
by (asm_simp_tac (ZF_ss addsimps [left_inverse]) 1);
-val cardinal_UN_le = result();
+qed "cardinal_UN_le";
(*The same again, using csucc*)
goal Cardinal_AC.thy
@@ -124,7 +124,7 @@
by (asm_full_simp_tac
(ZF_ss addsimps [Card_lt_csucc_iff, cardinal_UN_le,
InfCard_is_Card, Card_cardinal]) 1);
-val cardinal_UN_lt_csucc = result();
+qed "cardinal_UN_lt_csucc";
(*The same again, for a union of ordinals*)
goal Cardinal_AC.thy
@@ -135,14 +135,14 @@
by (fast_tac (ZF_cs addIs [Ord_cardinal_le RS lt_trans1] addEs [ltE]) 1);
by (fast_tac (ZF_cs addSIs [Ord_UN] addEs [ltE]) 1);
by (eresolve_tac [InfCard_is_Card RS Card_is_Ord RS Card_csucc] 1);
-val cardinal_UN_Ord_lt_csucc = result();
+qed "cardinal_UN_Ord_lt_csucc";
(*Saves checking Ord(j) below*)
goal Ordinal.thy "!!i j. [| i <= j; j<k; Ord(i) |] ==> i<k";
by (resolve_tac [subset_imp_le RS lt_trans1] 1);
by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
-val lt_subset_trans = result();
+qed "lt_subset_trans";
(*The same yet again, but the index set need not be a cardinal.
Surprisingly complicated proof!*)
@@ -168,5 +168,5 @@
by (fast_tac (ZF_cs addSIs [Ord_UN] addEs [ltE]) 2);
by (asm_simp_tac (ZF_ss addsimps [inj_converse_fun RS apply_type]
setloop split_tac [expand_if]) 1);
-val le_UN_Ord_lt_csucc = result();
+qed "le_UN_Ord_lt_csucc";
--- a/src/ZF/Epsilon.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Epsilon.ML Wed Dec 07 13:12:04 1994 +0100
@@ -13,7 +13,7 @@
goalw Epsilon.thy [eclose_def] "A <= eclose(A)";
by (rtac (nat_rec_0 RS equalityD2 RS subset_trans) 1);
by (rtac (nat_0I RS UN_upper) 1);
-val arg_subset_eclose = result();
+qed "arg_subset_eclose";
val arg_into_eclose = arg_subset_eclose RS subsetD;
@@ -25,7 +25,7 @@
by (etac (nat_rec_succ RS ssubst) 1);
by (etac UnionI 1);
by (assume_tac 1);
-val Transset_eclose = result();
+qed "Transset_eclose";
(* x : eclose(A) ==> x <= eclose(A) *)
val eclose_subset =
@@ -48,7 +48,7 @@
"[| !!x. ALL y:x. P(y) ==> P(x) |] ==> P(a)";
by (rtac (arg_in_eclose_sing RS eclose_induct) 1);
by (eresolve_tac prems 1);
-val eps_induct = result();
+qed "eps_induct";
(*Perform epsilon-induction on i. *)
fun eps_ind_tac a =
@@ -67,13 +67,13 @@
by (asm_simp_tac (ZF_ss addsimps [nat_rec_0]) 1);
by (asm_simp_tac (ZF_ss addsimps [nat_rec_succ]) 1);
by (fast_tac ZF_cs 1);
-val eclose_least_lemma = result();
+qed "eclose_least_lemma";
goalw Epsilon.thy [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));
-val eclose_least = result();
+qed "eclose_least";
(*COMPLETELY DIFFERENT induction principle from eclose_induct!!*)
val [major,base,step] = goal Epsilon.thy
@@ -87,12 +87,12 @@
by (etac base 2);
by (rewtac Transset_def);
by (fast_tac (ZF_cs addEs [step,ecloseD]) 1);
-val eclose_induct_down = result();
+qed "eclose_induct_down";
goal Epsilon.thy "!!X. Transset(X) ==> eclose(X) = X";
by (etac ([eclose_least, arg_subset_eclose] MRS equalityI) 1);
by (rtac subset_refl 1);
-val Transset_eclose_eq_arg = result();
+qed "Transset_eclose_eq_arg";
(*** Epsilon recursion ***)
@@ -101,19 +101,19 @@
goal Epsilon.thy "!!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));
-val mem_eclose_trans = result();
+qed "mem_eclose_trans";
(*Variant of the previous lemma in a useable form for the sequel*)
goal Epsilon.thy
"!!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));
-val mem_eclose_sing_trans = result();
+qed "mem_eclose_sing_trans";
goalw Epsilon.thy [Transset_def]
"!!i j. [| Transset(i); j:i |] ==> Memrel(i)-``{j} = j";
by (fast_tac (eq_cs addSIs [MemrelI] addSEs [MemrelE]) 1);
-val under_Memrel = result();
+qed "under_Memrel";
(* j : eclose(A) ==> Memrel(eclose(A)) -`` j = j *)
val under_Memrel_eclose = Transset_eclose RS under_Memrel;
@@ -128,27 +128,27 @@
by (rtac wfrec_ssubst 1);
by (asm_simp_tac (ZF_ss addsimps [under_Memrel_eclose,
jmemi RSN (2,mem_eclose_sing_trans)]) 1);
-val wfrec_eclose_eq = result();
+qed "wfrec_eclose_eq";
val [prem] = goal Epsilon.thy
"k: i ==> wfrec(Memrel(eclose({i})),k,H) = wfrec(Memrel(eclose({k})),k,H)";
by (rtac (arg_in_eclose_sing RS wfrec_eclose_eq) 1);
by (rtac (prem RS arg_into_eclose_sing) 1);
-val wfrec_eclose_eq2 = result();
+qed "wfrec_eclose_eq2";
goalw Epsilon.thy [transrec_def]
"transrec(a,H) = H(a, lam x:a. transrec(x,H))";
by (rtac wfrec_ssubst 1);
by (simp_tac (ZF_ss addsimps [wfrec_eclose_eq2, arg_in_eclose_sing,
under_Memrel_eclose]) 1);
-val transrec = result();
+qed "transrec";
(*Avoids explosions in proofs; resolve it with a meta-level definition.*)
val rew::prems = goal Epsilon.thy
"[| !!x. f(x)==transrec(x,H) |] ==> f(a) = H(a, lam x:a. f(x))";
by (rewtac rew);
by (REPEAT (resolve_tac (prems@[transrec]) 1));
-val def_transrec = result();
+qed "def_transrec";
val prems = goal Epsilon.thy
"[| !!x u. [| x:eclose({a}); u: Pi(x,B) |] ==> H(x,u) : B(x) \
@@ -156,12 +156,12 @@
by (res_inst_tac [("i", "a")] (arg_in_eclose_sing RS eclose_induct) 1);
by (rtac (transrec RS ssubst) 1);
by (REPEAT (ares_tac (prems @ [lam_type]) 1 ORELSE etac bspec 1));
-val transrec_type = result();
+qed "transrec_type";
goal Epsilon.thy "!!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);
-val eclose_sing_Ord = result();
+qed "eclose_sing_Ord";
val prems = goal Epsilon.thy
"[| j: i; Ord(i); \
@@ -171,7 +171,7 @@
by (resolve_tac prems 1);
by (rtac (Ord_in_Ord RS eclose_sing_Ord RS subsetD RS succE) 1);
by (DEPTH_SOLVE (ares_tac prems 1 ORELSE eresolve_tac [ssubst,Ord_trans] 1));
-val Ord_transrec_type = result();
+qed "Ord_transrec_type";
(*** Rank ***)
@@ -179,7 +179,7 @@
goal Epsilon.thy "rank(a) = (UN y:a. succ(rank(y)))";
by (rtac (rank_def RS def_transrec RS ssubst) 1);
by (simp_tac ZF_ss 1);
-val rank = result();
+qed "rank";
goal Epsilon.thy "Ord(rank(a))";
by (eps_ind_tac "a" 1);
@@ -187,27 +187,27 @@
by (rtac (Ord_succ RS Ord_UN) 1);
by (etac bspec 1);
by (assume_tac 1);
-val Ord_rank = result();
+qed "Ord_rank";
val [major] = goal Epsilon.thy "Ord(i) ==> rank(i) = i";
by (rtac (major RS trans_induct) 1);
by (rtac (rank RS ssubst) 1);
by (asm_simp_tac (ZF_ss addsimps [Ord_equality]) 1);
-val rank_of_Ord = result();
+qed "rank_of_Ord";
goal Epsilon.thy "!!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);
by (REPEAT (ares_tac [Ord_UN, Ord_succ, Ord_rank] 1));
-val rank_lt = result();
+qed "rank_lt";
val [major] = goal Epsilon.thy "a: eclose(b) ==> rank(a) < rank(b)";
by (rtac (major RS eclose_induct_down) 1);
by (etac rank_lt 1);
by (etac (rank_lt RS lt_trans) 1);
by (assume_tac 1);
-val eclose_rank_lt = result();
+qed "eclose_rank_lt";
goal Epsilon.thy "!!a b. a<=b ==> rank(a) le rank(b)";
by (rtac subset_imp_le 1);
@@ -215,7 +215,7 @@
by (rtac (rank RS ssubst) 1);
by (etac UN_mono 1);
by (REPEAT (resolve_tac [subset_refl, Ord_rank] 1));
-val rank_mono = result();
+qed "rank_mono";
goal Epsilon.thy "rank(Pow(a)) = succ(rank(a))";
by (rtac (rank RS trans) 1);
@@ -224,12 +224,12 @@
etac (PowD RS rank_mono RS succ_leI)] 1);
by (DO_GOAL [rtac ([Pow_top, le_refl] MRS UN_upper_le),
REPEAT o rtac (Ord_rank RS Ord_succ)] 1);
-val rank_Pow = result();
+qed "rank_Pow";
goal Epsilon.thy "rank(0) = 0";
by (rtac (rank RS trans) 1);
by (fast_tac (ZF_cs addSIs [equalityI]) 1);
-val rank_0 = result();
+qed "rank_0";
goal Epsilon.thy "rank(succ(x)) = succ(rank(x))";
by (rtac (rank RS trans) 1);
@@ -237,7 +237,7 @@
by (etac succE 1);
by (fast_tac ZF_cs 1);
by (etac (rank_lt RS leI RS succ_leI RS le_imp_subset) 1);
-val rank_succ = result();
+qed "rank_succ";
goal Epsilon.thy "rank(Union(A)) = (UN x:A. rank(x))";
by (rtac equalityI 1);
@@ -249,7 +249,7 @@
by (rtac subset_trans 1);
by (etac (RepFunI RS Union_upper) 2);
by (etac (rank_lt RS succ_leI RS le_imp_subset) 1);
-val rank_Union = result();
+qed "rank_Union";
goal Epsilon.thy "rank(eclose(a)) = rank(a)";
by (rtac le_anti_sym 1);
@@ -257,17 +257,17 @@
by (res_inst_tac [("a1","eclose(a)")] (rank RS ssubst) 1);
by (rtac (Ord_rank RS UN_least_le) 1);
by (etac (eclose_rank_lt RS succ_leI) 1);
-val rank_eclose = result();
+qed "rank_eclose";
goalw Epsilon.thy [Pair_def] "rank(a) < rank(<a,b>)";
by (rtac (consI1 RS rank_lt RS lt_trans) 1);
by (rtac (consI1 RS consI2 RS rank_lt) 1);
-val rank_pair1 = result();
+qed "rank_pair1";
goalw Epsilon.thy [Pair_def] "rank(b) < rank(<a,b>)";
by (rtac (consI1 RS consI2 RS rank_lt RS lt_trans) 1);
by (rtac (consI1 RS consI2 RS rank_lt) 1);
-val rank_pair2 = result();
+qed "rank_pair2";
goalw (merge_theories(Epsilon.thy,Sum.thy)) [Inl_def] "rank(a) < rank(Inl(a))";
by (rtac rank_pair2 1);
@@ -282,13 +282,13 @@
goal Epsilon.thy "!!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);
-val mem_eclose_subset = result();
+qed "mem_eclose_subset";
goal Epsilon.thy "!!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);
-val eclose_mono = result();
+qed "eclose_mono";
(** Idempotence of eclose **)
@@ -296,4 +296,4 @@
by (rtac equalityI 1);
by (rtac ([Transset_eclose, subset_refl] MRS eclose_least) 1);
by (rtac arg_subset_eclose 1);
-val eclose_idem = result();
+qed "eclose_idem";
--- a/src/ZF/EquivClass.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/EquivClass.ML Wed Dec 07 13:12:04 1994 +0100
@@ -17,18 +17,18 @@
goalw EquivClass.thy [trans_def,sym_def]
"!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r";
by (fast_tac (ZF_cs addSEs [converseD,compE]) 1);
-val sym_trans_comp_subset = result();
+qed "sym_trans_comp_subset";
goalw EquivClass.thy [refl_def]
"!!A r. [| refl(A,r); r <= A*A |] ==> r <= converse(r) O r";
by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 1);
-val refl_comp_subset = result();
+qed "refl_comp_subset";
goalw EquivClass.thy [equiv_def]
"!!A r. equiv(A,r) ==> converse(r) O r = r";
by (fast_tac (subset_cs addSIs [equalityI, sym_trans_comp_subset,
refl_comp_subset]) 1);
-val equiv_comp_eq = result();
+qed "equiv_comp_eq";
(*second half*)
goalw EquivClass.thy [equiv_def,refl_def,sym_def,trans_def]
@@ -40,7 +40,7 @@
by (ALLGOALS (fast_tac
(ZF_cs addSIs [converseI] addIs [compI] addSEs [compE])));
by flexflex_tac;
-val comp_equivI = result();
+qed "comp_equivI";
(** Equivalence classes **)
@@ -48,52 +48,52 @@
goalw EquivClass.thy [trans_def,sym_def]
"!!A r. [| sym(r); trans(r); <a,b>: r |] ==> r``{a} <= r``{b}";
by (fast_tac ZF_cs 1);
-val equiv_class_subset = result();
+qed "equiv_class_subset";
goalw EquivClass.thy [equiv_def]
"!!A r. [| equiv(A,r); <a,b>: r |] ==> r``{a} = r``{b}";
by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset]));
by (rewrite_goals_tac [sym_def]);
by (fast_tac ZF_cs 1);
-val equiv_class_eq = result();
+qed "equiv_class_eq";
goalw EquivClass.thy [equiv_def,refl_def]
"!!A r. [| equiv(A,r); a: A |] ==> a: r``{a}";
by (fast_tac ZF_cs 1);
-val equiv_class_self = result();
+qed "equiv_class_self";
(*Lemma for the next result*)
goalw EquivClass.thy [equiv_def,refl_def]
"!!A r. [| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> <a,b>: r";
by (fast_tac ZF_cs 1);
-val subset_equiv_class = result();
+qed "subset_equiv_class";
val prems = goal EquivClass.thy
"[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> <a,b>: r";
by (REPEAT (resolve_tac (prems @ [equalityD2, subset_equiv_class]) 1));
-val eq_equiv_class = result();
+qed "eq_equiv_class";
(*thus r``{a} = r``{b} as well*)
goalw EquivClass.thy [equiv_def,trans_def,sym_def]
"!!A r. [| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> <a,b>: r";
by (fast_tac ZF_cs 1);
-val equiv_class_nondisjoint = result();
+qed "equiv_class_nondisjoint";
goalw EquivClass.thy [equiv_def] "!!A r. equiv(A,r) ==> r <= A*A";
by (safe_tac ZF_cs);
-val equiv_type = result();
+qed "equiv_type";
goal EquivClass.thy
"!!A r. equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A";
by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq]
addDs [equiv_type]) 1);
-val equiv_class_eq_iff = result();
+qed "equiv_class_eq_iff";
goal EquivClass.thy
"!!A r. [| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> <x,y>: r";
by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq]
addDs [equiv_type]) 1);
-val eq_equiv_class_iff = result();
+qed "eq_equiv_class_iff";
(*** Quotients ***)
@@ -102,7 +102,7 @@
val prems = goalw EquivClass.thy [quotient_def] "x:A ==> r``{x}: A/r";
by (rtac RepFunI 1);
by (resolve_tac prems 1);
-val quotientI = result();
+qed "quotientI";
val major::prems = goalw EquivClass.thy [quotient_def]
"[| X: A/r; !!x. [| X = r``{x}; x:A |] ==> P |] \
@@ -110,12 +110,12 @@
by (rtac (major RS RepFunE) 1);
by (eresolve_tac prems 1);
by (assume_tac 1);
-val quotientE = result();
+qed "quotientE";
goalw EquivClass.thy [equiv_def,refl_def,quotient_def]
"!!A r. equiv(A,r) ==> Union(A/r) = A";
by (fast_tac eq_cs 1);
-val Union_quotient = result();
+qed "Union_quotient";
goalw EquivClass.thy [quotient_def]
"!!A r. [| equiv(A,r); X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)";
@@ -123,7 +123,7 @@
by (assume_tac 1);
by (rewrite_goals_tac [equiv_def,trans_def,sym_def]);
by (fast_tac ZF_cs 1);
-val quotient_disj = result();
+qed "quotient_disj";
(**** Defining unary operations upon equivalence classes ****)
@@ -140,7 +140,7 @@
by (assume_tac 2);
by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]);
by (fast_tac ZF_cs 1);
-val UN_equiv_class = result();
+qed "UN_equiv_class";
(*Resolve th against the "local" premises*)
val localize = RSLIST [equivA,bcong];
@@ -154,7 +154,7 @@
by (safe_tac ZF_cs);
by (rtac (localize UN_equiv_class RS ssubst) 1);
by (REPEAT (ares_tac prems 1));
-val UN_equiv_class_type = result();
+qed "UN_equiv_class_type";
(*Sufficient conditions for injectiveness. Could weaken premises!
major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B
@@ -170,7 +170,7 @@
by (REPEAT (ares_tac prems 1));
by (etac box_equals 1);
by (REPEAT (ares_tac [localize UN_equiv_class] 1));
-val UN_equiv_class_inject = result();
+qed "UN_equiv_class_inject";
(**** Defining binary operations upon equivalence classes ****)
@@ -179,7 +179,7 @@
goalw EquivClass.thy [congruent_def,congruent2_def,equiv_def,refl_def]
"!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))";
by (fast_tac ZF_cs 1);
-val congruent2_implies_congruent = result();
+qed "congruent2_implies_congruent";
val equivA::prems = goalw EquivClass.thy [congruent_def]
"[| equiv(A,r); congruent2(r,b); a: A |] ==> \
@@ -192,7 +192,7 @@
congruent2_implies_congruent]) 1);
by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
by (fast_tac ZF_cs 1);
-val congruent2_implies_congruent_UN = result();
+qed "congruent2_implies_congruent_UN";
val prems as equivA::_ = goal EquivClass.thy
"[| equiv(A,r); congruent2(r,b); a1: A; a2: A |] \
@@ -201,7 +201,7 @@
by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class,
congruent2_implies_congruent,
congruent2_implies_congruent_UN]) 1);
-val UN_equiv_class2 = result();
+qed "UN_equiv_class2";
(*type checking*)
val prems = goalw EquivClass.thy [quotient_def]
@@ -214,7 +214,7 @@
by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
congruent2_implies_congruent_UN,
congruent2_implies_congruent, quotientI]) 1));
-val UN_equiv_class_type2 = result();
+qed "UN_equiv_class_type2";
(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
@@ -229,7 +229,7 @@
by (rtac trans 1);
by (REPEAT (ares_tac prems 1
ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
-val congruent2I = result();
+qed "congruent2I";
val [equivA,commute,congt] = goal EquivClass.thy
"[| equiv(A,r); \
@@ -242,7 +242,7 @@
by (rtac sym 5);
by (REPEAT (ares_tac [congt] 1
ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1));
-val congruent2_commuteI = result();
+qed "congruent2_commuteI";
(*Obsolete?*)
val [equivA,ZinA,congt,commute] = goalw EquivClass.thy [quotient_def]
@@ -259,4 +259,4 @@
by (asm_simp_tac (ZF_ss addsimps [commute,
[equivA, congt] MRS UN_equiv_class]) 1);
by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1));
-val congruent_commuteI = result();
+qed "congruent_commuteI";
--- a/src/ZF/Finite.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Finite.ML Wed Dec 07 13:12:04 1994 +0100
@@ -18,7 +18,7 @@
by (rtac lfp_mono 1);
by (REPEAT (rtac Fin.bnd_mono 1));
by (REPEAT (ares_tac (Pow_mono::basic_monos) 1));
-val Fin_mono = result();
+qed "Fin_mono";
(* A : Fin(B) ==> A <= B *)
val FinD = Fin.dom_subset RS subsetD RS PowD;
@@ -35,7 +35,7 @@
by (excluded_middle_tac "a:b" 2);
by (etac (cons_absorb RS ssubst) 3 THEN assume_tac 3); (*backtracking!*)
by (REPEAT (ares_tac prems 1));
-val Fin_induct = result();
+qed "Fin_induct";
(** Simplification for Fin **)
val Fin_ss = arith_ss addsimps Fin.intrs;
@@ -45,13 +45,13 @@
"[| b: Fin(A); c: Fin(A) |] ==> b Un c : Fin(A)";
by (rtac (major RS Fin_induct) 1);
by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Un_0, Un_cons]))));
-val Fin_UnI = result();
+qed "Fin_UnI";
(*The union of a set of finite sets is finite.*)
val [major] = goal Finite.thy "C : Fin(Fin(A)) ==> Union(C) : Fin(A)";
by (rtac (major RS Fin_induct) 1);
by (ALLGOALS (asm_simp_tac (Fin_ss addsimps [Union_0, Union_cons, Fin_UnI])));
-val Fin_UnionI = result();
+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)";
@@ -61,11 +61,11 @@
by (safe_tac ZF_cs);
by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 1);
by (asm_simp_tac Fin_ss 1);
-val Fin_subset_lemma = result();
+qed "Fin_subset_lemma";
goal Finite.thy "!!c b A. [| c<=b; b: Fin(A) |] ==> c: Fin(A)";
by (REPEAT (ares_tac [Fin_subset_lemma RS spec RS mp] 1));
-val Fin_subset = result();
+qed "Fin_subset";
val major::prems = goal Finite.thy
"[| c: Fin(A); b: Fin(A); \
@@ -76,7 +76,7 @@
by (rtac (Diff_cons RS ssubst) 2);
by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Diff_0, cons_subset_iff,
Diff_subset RS Fin_subset]))));
-val Fin_0_induct_lemma = result();
+qed "Fin_0_induct_lemma";
val prems = goal Finite.thy
"[| b: Fin(A); \
@@ -86,7 +86,7 @@
by (rtac (Diff_cancel RS subst) 1);
by (rtac (Fin_0_induct_lemma RS mp) 1);
by (REPEAT (ares_tac (subset_refl::prems) 1));
-val Fin_0_induct = result();
+qed "Fin_0_induct";
(*Functions from a finite ordinal*)
val prems = goal Finite.thy "n: nat ==> n->A <= Fin(nat*A)";
@@ -94,7 +94,7 @@
by (simp_tac (ZF_ss addsimps [Pi_empty1, Fin.emptyI, subset_iff, cons_iff]) 1);
by (asm_simp_tac (ZF_ss addsimps [succ_def, mem_not_refl RS cons_fun_eq]) 1);
by (fast_tac (ZF_cs addSIs [Fin.consI]) 1);
-val nat_fun_subset_Fin = result();
+qed "nat_fun_subset_Fin";
(*** Finite function space ***)
@@ -104,23 +104,23 @@
by (rtac lfp_mono 1);
by (REPEAT (rtac FiniteFun.bnd_mono 1));
by (REPEAT (ares_tac (Fin_mono::Sigma_mono::basic_monos) 1));
-val FiniteFun_mono = result();
+qed "FiniteFun_mono";
goal Finite.thy "!!A B. A<=B ==> A -||> A <= B -||> B";
by (REPEAT (ares_tac [FiniteFun_mono] 1));
-val FiniteFun_mono1 = result();
+qed "FiniteFun_mono1";
goal Finite.thy "!!h. h: A -||>B ==> h: domain(h) -> B";
by (etac FiniteFun.induct 1);
by (simp_tac (ZF_ss addsimps [empty_fun, domain_0]) 1);
by (asm_simp_tac (ZF_ss addsimps [fun_extend3, domain_cons]) 1);
-val FiniteFun_is_fun = result();
+qed "FiniteFun_is_fun";
goal Finite.thy "!!h. h: A -||>B ==> domain(h) : Fin(A)";
by (etac FiniteFun.induct 1);
by (simp_tac (Fin_ss addsimps [domain_0]) 1);
by (asm_simp_tac (Fin_ss addsimps [domain_cons]) 1);
-val FiniteFun_domain_Fin = result();
+qed "FiniteFun_domain_Fin";
val FiniteFun_apply_type = FiniteFun_is_fun RS apply_type |> standard;
@@ -133,9 +133,9 @@
by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 1);
by (dtac (spec RS mp) 1 THEN assume_tac 1);
by (fast_tac (ZF_cs addSIs FiniteFun.intrs) 1);
-val FiniteFun_subset_lemma = result();
+qed "FiniteFun_subset_lemma";
goal Finite.thy "!!c b A. [| c<=b; b: A-||>B |] ==> c: A-||>B";
by (REPEAT (ares_tac [FiniteFun_subset_lemma RS spec RS mp] 1));
-val FiniteFun_subset = result();
+qed "FiniteFun_subset";
--- a/src/ZF/Fixedpt.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Fixedpt.ML Wed Dec 07 13:12:04 1994 +0100
@@ -18,17 +18,17 @@
\ |] ==> bnd_mono(D,h)";
by (REPEAT (ares_tac (prems@[conjI,allI,impI]) 1
ORELSE etac subset_trans 1));
-val bnd_monoI = result();
+qed "bnd_monoI";
val [major] = goalw Fixedpt.thy [bnd_mono_def] "bnd_mono(D,h) ==> h(D) <= D";
by (rtac (major RS conjunct1) 1);
-val bnd_monoD1 = result();
+qed "bnd_monoD1";
val major::prems = goalw Fixedpt.thy [bnd_mono_def]
"[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) <= h(X)";
by (rtac (major RS conjunct2 RS spec RS spec RS mp RS mp) 1);
by (REPEAT (resolve_tac prems 1));
-val bnd_monoD2 = result();
+qed "bnd_monoD2";
val [major,minor] = goal Fixedpt.thy
"[| bnd_mono(D,h); X<=D |] ==> h(X) <= D";
@@ -36,20 +36,20 @@
by (rtac (major RS bnd_monoD1) 3);
by (rtac minor 1);
by (rtac subset_refl 1);
-val bnd_mono_subset = result();
+qed "bnd_mono_subset";
goal Fixedpt.thy "!!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));
-val bnd_mono_Un = result();
+qed "bnd_mono_Un";
(*Useful??*)
goal Fixedpt.thy "!!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));
-val bnd_mono_Int = result();
+qed "bnd_mono_Int";
(**** Proof of Knaster-Tarski Theorem for the lfp ****)
@@ -58,39 +58,39 @@
"!!A. [| h(A) <= A; A<=D |] ==> lfp(D,h) <= A";
by (fast_tac ZF_cs 1);
(*or by (rtac (PowI RS CollectI RS Inter_lower) 1); *)
-val lfp_lowerbound = result();
+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";
by (fast_tac ZF_cs 1);
-val lfp_subset = result();
+qed "lfp_subset";
(*Used in datatype package*)
val [rew] = goal Fixedpt.thy "A==lfp(D,h) ==> A <= D";
by (rewtac rew);
by (rtac lfp_subset 1);
-val def_lfp_subset = result();
+qed "def_lfp_subset";
val prems = goalw Fixedpt.thy [lfp_def]
"[| h(D) <= D; !!X. [| h(X) <= X; X<=D |] ==> A<=X |] ==> \
\ A <= lfp(D,h)";
by (rtac (Pow_top RS CollectI RS Inter_greatest) 1);
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [CollectE,PowD] 1));
-val lfp_greatest = result();
+qed "lfp_greatest";
val hmono::prems = goal Fixedpt.thy
"[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) <= A";
by (rtac (hmono RS bnd_monoD2 RS subset_trans) 1);
by (rtac lfp_lowerbound 1);
by (REPEAT (resolve_tac prems 1));
-val lfp_lemma1 = result();
+qed "lfp_lemma1";
val [hmono] = goal Fixedpt.thy
"bnd_mono(D,h) ==> h(lfp(D,h)) <= lfp(D,h)";
by (rtac (bnd_monoD1 RS lfp_greatest) 1);
by (rtac lfp_lemma1 2);
by (REPEAT (ares_tac [hmono] 1));
-val lfp_lemma2 = result();
+qed "lfp_lemma2";
val [hmono] = goal Fixedpt.thy
"bnd_mono(D,h) ==> lfp(D,h) <= h(lfp(D,h))";
@@ -99,19 +99,19 @@
by (rtac (hmono RS lfp_lemma2) 1);
by (rtac (hmono RS bnd_mono_subset) 2);
by (REPEAT (rtac lfp_subset 1));
-val lfp_lemma3 = result();
+qed "lfp_lemma3";
val prems = goal Fixedpt.thy
"bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))";
by (REPEAT (resolve_tac (prems@[equalityI,lfp_lemma2,lfp_lemma3]) 1));
-val lfp_Tarski = result();
+qed "lfp_Tarski";
(*Definition form, to control unfolding*)
val [rew,mono] = goal Fixedpt.thy
"[| A==lfp(D,h); bnd_mono(D,h) |] ==> A = h(A)";
by (rewtac rew);
by (rtac (mono RS lfp_Tarski) 1);
-val def_lfp_Tarski = result();
+qed "def_lfp_Tarski";
(*** General induction rule for least fixedpoints ***)
@@ -124,7 +124,7 @@
by (rtac (hmono RS lfp_lemma2 RS subsetD) 1);
by (rtac (hmono RS bnd_monoD2 RS subsetD) 1);
by (REPEAT (ares_tac [Collect_subset, lfp_subset] 1));
-val Collect_is_pre_fixedpt = result();
+qed "Collect_is_pre_fixedpt";
(*This rule yields an induction hypothesis in which the components of a
data structure may be assumed to be elements of lfp(D,h)*)
@@ -135,7 +135,7 @@
by (rtac (Collect_is_pre_fixedpt RS lfp_lowerbound RS subsetD RS CollectD2) 1);
by (rtac (lfp_subset RS (Collect_subset RS subset_trans)) 3);
by (REPEAT (ares_tac prems 1));
-val induct = result();
+qed "induct";
(*Definition form, to control unfolding*)
val rew::prems = goal Fixedpt.thy
@@ -144,7 +144,7 @@
\ |] ==> P(a)";
by (rtac induct 1);
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1));
-val def_induct = result();
+qed "def_induct";
(*This version is useful when "A" is not a subset of D;
second premise could simply be h(D Int A) <= D or !!X. X<=D ==> h(X)<=D *)
@@ -153,7 +153,7 @@
by (rtac (lfp_lowerbound RS subset_trans) 1);
by (rtac (hmono RS bnd_mono_subset RS Int_greatest) 1);
by (REPEAT (resolve_tac [hsub,Int_lower1,Int_lower2] 1));
-val lfp_Int_lowerbound = result();
+qed "lfp_Int_lowerbound";
(*Monotonicity of lfp, where h precedes i under a domain-like partial order
monotonicity of h is not strictly necessary; h must be bounded by D*)
@@ -166,7 +166,7 @@
by (rtac (Int_lower1 RS subhi RS subset_trans) 1);
by (rtac (imono RS bnd_monoD2 RS subset_trans) 1);
by (REPEAT (ares_tac [Int_lower2] 1));
-val lfp_mono = result();
+qed "lfp_mono";
(*This (unused) version illustrates that monotonicity is not really needed,
but both lfp's must be over the SAME set D; Inter is anti-monotonic!*)
@@ -177,7 +177,7 @@
by (rtac lfp_lowerbound 1);
by (etac (subhi RS subset_trans) 1);
by (REPEAT (assume_tac 1));
-val lfp_mono2 = result();
+qed "lfp_mono2";
(**** Proof of Knaster-Tarski Theorem for the gfp ****)
@@ -187,23 +187,23 @@
"[| A <= h(A); A<=D |] ==> A <= gfp(D,h)";
by (rtac (PowI RS CollectI RS Union_upper) 1);
by (REPEAT (resolve_tac prems 1));
-val gfp_upperbound = result();
+qed "gfp_upperbound";
goalw Fixedpt.thy [gfp_def] "gfp(D,h) <= D";
by (fast_tac ZF_cs 1);
-val gfp_subset = result();
+qed "gfp_subset";
(*Used in datatype package*)
val [rew] = goal Fixedpt.thy "A==gfp(D,h) ==> A <= D";
by (rewtac rew);
by (rtac gfp_subset 1);
-val def_gfp_subset = result();
+qed "def_gfp_subset";
val hmono::prems = goalw Fixedpt.thy [gfp_def]
"[| bnd_mono(D,h); !!X. [| X <= h(X); X<=D |] ==> X<=A |] ==> \
\ gfp(D,h) <= A";
by (fast_tac (subset_cs addIs ((hmono RS bnd_monoD1)::prems)) 1);
-val gfp_least = result();
+qed "gfp_least";
val hmono::prems = goal Fixedpt.thy
"[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A <= h(gfp(D,h))";
@@ -211,14 +211,14 @@
by (rtac gfp_subset 3);
by (rtac gfp_upperbound 2);
by (REPEAT (resolve_tac prems 1));
-val gfp_lemma1 = result();
+qed "gfp_lemma1";
val [hmono] = goal Fixedpt.thy
"bnd_mono(D,h) ==> gfp(D,h) <= h(gfp(D,h))";
by (rtac gfp_least 1);
by (rtac gfp_lemma1 2);
by (REPEAT (ares_tac [hmono] 1));
-val gfp_lemma2 = result();
+qed "gfp_lemma2";
val [hmono] = goal Fixedpt.thy
"bnd_mono(D,h) ==> h(gfp(D,h)) <= gfp(D,h)";
@@ -226,19 +226,19 @@
by (rtac (hmono RS bnd_monoD2) 1);
by (rtac (hmono RS gfp_lemma2) 1);
by (REPEAT (rtac ([hmono, gfp_subset] MRS bnd_mono_subset) 1));
-val gfp_lemma3 = result();
+qed "gfp_lemma3";
val prems = goal Fixedpt.thy
"bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))";
by (REPEAT (resolve_tac (prems@[equalityI,gfp_lemma2,gfp_lemma3]) 1));
-val gfp_Tarski = result();
+qed "gfp_Tarski";
(*Definition form, to control unfolding*)
val [rew,mono] = goal Fixedpt.thy
"[| A==gfp(D,h); bnd_mono(D,h) |] ==> A = h(A)";
by (rewtac rew);
by (rtac (mono RS gfp_Tarski) 1);
-val def_gfp_Tarski = result();
+qed "def_gfp_Tarski";
(*** Coinduction rules for greatest fixed points ***)
@@ -246,7 +246,7 @@
(*weak version*)
goal Fixedpt.thy "!!X h. [| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)";
by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1));
-val weak_coinduct = result();
+qed "weak_coinduct";
val [subs_h,subs_D,mono] = goal Fixedpt.thy
"[| X <= h(X Un gfp(D,h)); X <= D; bnd_mono(D,h) |] ==> \
@@ -255,7 +255,7 @@
by (rtac (mono RS gfp_lemma2 RS subset_trans) 1);
by (rtac (Un_upper2 RS subset_trans) 1);
by (rtac ([mono, subs_D, gfp_subset] MRS bnd_mono_Un) 1);
-val coinduct_lemma = result();
+qed "coinduct_lemma";
(*strong version*)
goal Fixedpt.thy
@@ -264,7 +264,7 @@
by (rtac weak_coinduct 1);
by (etac coinduct_lemma 2);
by (REPEAT (ares_tac [gfp_subset, UnI1, Un_least] 1));
-val coinduct = result();
+qed "coinduct";
(*Definition form, to control unfolding*)
val rew::prems = goal Fixedpt.thy
@@ -273,7 +273,7 @@
by (rewtac rew);
by (rtac coinduct 1);
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1));
-val def_coinduct = result();
+qed "def_coinduct";
(*Lemma used immediately below!*)
val [subsA,XimpP] = goal ZF.thy
@@ -281,7 +281,7 @@
by (rtac (subsA RS subsetD RS CollectI RS subsetI) 1);
by (assume_tac 1);
by (etac XimpP 1);
-val subset_Collect = result();
+qed "subset_Collect";
(*The version used in the induction/coinduction package*)
val prems = goal Fixedpt.thy
@@ -290,7 +290,7 @@
\ a : A";
by (rtac def_coinduct 1);
by (REPEAT (ares_tac (subset_Collect::prems) 1));
-val def_Collect_coinduct = result();
+qed "def_Collect_coinduct";
(*Monotonicity of gfp!*)
val [hmono,subde,subhi] = goal Fixedpt.thy
@@ -300,5 +300,5 @@
by (rtac (hmono RS gfp_lemma2 RS subset_trans) 1);
by (rtac (gfp_subset RS subhi) 1);
by (rtac ([gfp_subset, subde] MRS subset_trans) 1);
-val gfp_mono = result();
+qed "gfp_mono";
--- a/src/ZF/IMP/Equiv.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/IMP/Equiv.ML Wed Dec 07 13:12:04 1994 +0100
@@ -12,7 +12,7 @@
by (rewrite_goals_tac A_rewrite_rules); (* rewr. Den. *)
by (TRYALL (fast_tac (ZF_cs addSIs (evala.intrs@prems)
addSEs aexp_elim_cases)));
-val aexp_iff = result();
+qed "aexp_iff";
val aexp1 = prove_goal Equiv.thy
@@ -41,7 +41,7 @@
by (rewrite_goals_tac B_rewrite_rules); (* rewr. Den. *)
by (TRYALL (fast_tac (ZF_cs addSIs (evalb.intrs@prems@[aexp2])
addSDs [aexp1] addSEs bexp_elim_cases)));
-val bexp_iff = result();
+qed "bexp_iff";
val bexp1 = prove_goal Equiv.thy
"[| <b,sigma> -b-> w; b: bexp; sigma: loc -> nat |]\
--- a/src/ZF/InfDatatype.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/InfDatatype.ML Wed Dec 07 13:12:04 1994 +0100
@@ -113,13 +113,13 @@
by (fast_tac (ZF_cs addEs [LeastI, ltE]) 2);
by (eresolve_tac [[subset_refl, UN_upper] MRS Vfrom_mono RS subsetD] 1);
by (assume_tac 1);
-val fun_Vcsucc_lemma = result();
+qed "fun_Vcsucc_lemma";
goal InfDatatype.thy
"!!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 (ZF_ss addsimps [subset_iff_id, fun_Vcsucc_lemma]) 1);
-val subset_Vcsucc = result();
+qed "subset_Vcsucc";
(*Version for arbitrary index sets*)
goal InfDatatype.thy
@@ -134,14 +134,14 @@
by (fast_tac (ZF_cs addIs [Pair_in_Vfrom, Vfrom_UnI1, Vfrom_UnI2]) 2);
by (REPEAT (ares_tac [ltD, InfCard_csucc, InfCard_is_Limit,
Limit_has_succ, Un_least_lt] 1));
-val fun_Vcsucc = result();
+qed "fun_Vcsucc";
goal InfDatatype.thy
"!!K. [| f: W -> Vfrom(A, csucc(K)); |W| le K; InfCard(K); \
\ W <= Vfrom(A,csucc(K)) \
\ |] ==> f: Vfrom(A,csucc(K))";
by (REPEAT (ares_tac [fun_Vcsucc RS subsetD] 1));
-val fun_in_Vcsucc = result();
+qed "fun_in_Vcsucc";
(*Remove <= from the rule above*)
val fun_in_Vcsucc' = subsetI RSN (4, fun_in_Vcsucc);
@@ -154,13 +154,13 @@
by (REPEAT (ares_tac [fun_Vcsucc, Ord_cardinal_le,
i_subset_Vfrom,
lt_csucc RS leI RS le_imp_subset RS subset_trans] 1));
-val Card_fun_Vcsucc = result();
+qed "Card_fun_Vcsucc";
goal InfDatatype.thy
"!!K. [| f: K -> Vfrom(A, csucc(K)); InfCard(K) \
\ |] ==> f: Vfrom(A,csucc(K))";
by (REPEAT (ares_tac [Card_fun_Vcsucc RS subsetD] 1));
-val Card_fun_in_Vcsucc = result();
+qed "Card_fun_in_Vcsucc";
val Pair_in_Vcsucc = Limit_csucc RSN (3, Pair_in_VLimit) |> standard;
val Inl_in_Vcsucc = Limit_csucc RSN (2, Inl_in_VLimit) |> standard;
--- a/src/ZF/List.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/List.ML Wed Dec 07 13:12:04 1994 +0100
@@ -28,7 +28,7 @@
by (fast_tac (sum_cs addSIs (equalityI :: map rew intrs)
addEs [rew elim]) 1)
end;
-val list_unfold = result();
+qed "list_unfold";
(** Lemmas to justify using "list" in other recursive type definitions **)
@@ -36,7 +36,7 @@
by (rtac lfp_mono 1);
by (REPEAT (rtac list.bnd_mono 1));
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));
-val list_mono = result();
+qed "list_mono";
(*There is a similar proof by list induction.*)
goalw List.thy (list.defs@list.con_defs) "list(univ(A)) <= univ(A)";
@@ -44,14 +44,14 @@
by (rtac (A_subset_univ RS univ_mono) 2);
by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
Pair_in_univ]) 1);
-val list_univ = result();
+qed "list_univ";
(*These two theorems are unused -- useful??*)
val list_subset_univ = standard ([list_mono, list_univ] MRS subset_trans);
goal List.thy "!!l A B. [| l: list(A); A <= univ(B) |] ==> l: univ(B)";
by (REPEAT (ares_tac [list_subset_univ RS subsetD] 1));
-val list_into_univ = result();
+qed "list_into_univ";
val major::prems = goal List.thy
"[| l: list(A); \
@@ -60,18 +60,18 @@
\ |] ==> list_case(c,h,l) : C(l)";
by (rtac (major RS list.induct) 1);
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps (list.case_eqns @ prems))));
-val list_case_type = result();
+qed "list_case_type";
(** For recursion **)
goalw List.thy list.con_defs "rank(a) < rank(Cons(a,l))";
by (simp_tac rank_ss 1);
-val rank_Cons1 = result();
+qed "rank_Cons1";
goalw List.thy list.con_defs "rank(l) < rank(Cons(a,l))";
by (simp_tac rank_ss 1);
-val rank_Cons2 = result();
+qed "rank_Cons2";
(*** List functions ***)
@@ -80,55 +80,55 @@
goalw List.thy [hd_def] "hd(Cons(a,l)) = a";
by (resolve_tac list.case_eqns 1);
-val hd_Cons = result();
+qed "hd_Cons";
goalw List.thy [tl_def] "tl(Nil) = Nil";
by (resolve_tac list.case_eqns 1);
-val tl_Nil = result();
+qed "tl_Nil";
goalw List.thy [tl_def] "tl(Cons(a,l)) = l";
by (resolve_tac list.case_eqns 1);
-val tl_Cons = result();
+qed "tl_Cons";
goal List.thy "!!l. l: list(A) ==> tl(l) : list(A)";
by (etac list.elim 1);
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps (list.intrs @ [tl_Nil,tl_Cons]))));
-val tl_type = result();
+qed "tl_type";
(** drop **)
goalw List.thy [drop_def] "drop(0, l) = l";
by (rtac rec_0 1);
-val drop_0 = result();
+qed "drop_0";
goalw List.thy [drop_def] "!!i. i:nat ==> drop(i, Nil) = Nil";
by (etac nat_induct 1);
by (ALLGOALS (asm_simp_tac (nat_ss addsimps [tl_Nil])));
-val drop_Nil = result();
+qed "drop_Nil";
goalw List.thy [drop_def]
"!!i. i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)";
by (etac nat_induct 1);
by (ALLGOALS (asm_simp_tac (nat_ss addsimps [tl_Cons])));
-val drop_succ_Cons = result();
+qed "drop_succ_Cons";
goalw List.thy [drop_def]
"!!i l. [| i:nat; l: list(A) |] ==> drop(i,l) : list(A)";
by (etac nat_induct 1);
by (ALLGOALS (asm_simp_tac (nat_ss addsimps [tl_type])));
-val drop_type = result();
+qed "drop_type";
(** list_rec -- by Vset recursion **)
goal List.thy "list_rec(Nil,c,h) = c";
by (rtac (list_rec_def RS def_Vrec RS trans) 1);
by (simp_tac (ZF_ss addsimps list.case_eqns) 1);
-val list_rec_Nil = result();
+qed "list_rec_Nil";
goal List.thy "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);
-val list_rec_Cons = result();
+qed "list_rec_Cons";
(*Type checking -- proved by induction, as usual*)
val prems = goal List.thy
@@ -139,7 +139,7 @@
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac
(ZF_ss addsimps (prems@[list_rec_Nil,list_rec_Cons]))));
-val list_rec_type = result();
+qed "list_rec_type";
(** Versions for use with definitions **)
@@ -147,13 +147,13 @@
"[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Nil) = c";
by (rewtac rew);
by (rtac list_rec_Nil 1);
-val def_list_rec_Nil = result();
+qed "def_list_rec_Nil";
val [rew] = goal List.thy
"[| !!l. j(l)==list_rec(l, c, h) |] ==> j(Cons(a,l)) = h(a,l,j(l))";
by (rewtac rew);
by (rtac list_rec_Cons 1);
-val def_list_rec_Cons = result();
+qed "def_list_rec_Cons";
fun list_recs def = map standard
([def] RL [def_list_rec_Nil, def_list_rec_Cons]);
@@ -165,12 +165,12 @@
val prems = goalw List.thy [map_def]
"[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)";
by (REPEAT (ares_tac (prems @ list.intrs @ [list_rec_type]) 1));
-val map_type = result();
+qed "map_type";
val [major] = goal List.thy "l: list(A) ==> map(h,l) : list({h(u). u:A})";
by (rtac (major RS map_type) 1);
by (etac RepFunI 1);
-val map_type2 = result();
+qed "map_type2";
(** length **)
@@ -179,7 +179,7 @@
goalw List.thy [length_def]
"!!l. l: list(A) ==> length(l) : nat";
by (REPEAT (ares_tac [list_rec_type, nat_0I, nat_succI] 1));
-val length_type = result();
+qed "length_type";
(** app **)
@@ -188,7 +188,7 @@
goalw List.thy [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));
-val app_type = result();
+qed "app_type";
(** rev **)
@@ -197,7 +197,7 @@
goalw List.thy [rev_def]
"!!xs. xs: list(A) ==> rev(xs) : list(A)";
by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1));
-val rev_type = result();
+qed "rev_type";
(** flat **)
@@ -207,7 +207,7 @@
goalw List.thy [flat_def]
"!!ls. ls: list(list(A)) ==> flat(ls) : list(A)";
by (REPEAT (ares_tac (list.intrs @ [list_rec_type, app_type]) 1));
-val flat_type = result();
+qed "flat_type";
(** list_add **)
@@ -217,7 +217,7 @@
goalw List.thy [list_add_def]
"!!xs. xs: list(nat) ==> list_add(xs) : nat";
by (REPEAT (ares_tac [list_rec_type, nat_0I, add_type] 1));
-val list_add_type = result();
+qed "list_add_type";
(** List simplification **)
@@ -241,25 +241,25 @@
"l: list(A) ==> map(%u.u, l) = l";
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val map_ident = result();
+qed "map_ident";
val prems = goal List.thy
"l: list(A) ==> map(h, map(j,l)) = map(%u.h(j(u)), l)";
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val map_compose = result();
+qed "map_compose";
val prems = goal List.thy
"xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)";
by (list_ind_tac "xs" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val map_app_distrib = result();
+qed "map_app_distrib";
val prems = goal List.thy
"ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))";
by (list_ind_tac "ls" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib])));
-val map_flat = result();
+qed "map_flat";
val prems = goal List.thy
"l: list(A) ==> \
@@ -267,7 +267,7 @@
\ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))";
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val list_rec_map = result();
+qed "list_rec_map";
(** theorems about list(Collect(A,P)) -- used in ex/term.ML **)
@@ -278,7 +278,7 @@
"l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)";
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val map_list_Collect = result();
+qed "map_list_Collect";
(*** theorems about length ***)
@@ -286,13 +286,13 @@
"xs: list(A) ==> length(map(h,xs)) = length(xs)";
by (list_ind_tac "xs" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val length_map = result();
+qed "length_map";
val prems = goal List.thy
"xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)";
by (list_ind_tac "xs" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val length_app = result();
+qed "length_app";
(* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m
Used for rewriting below*)
@@ -302,13 +302,13 @@
"xs: list(A) ==> length(rev(xs)) = length(xs)";
by (list_ind_tac "xs" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_app, add_commute_succ])));
-val length_rev = result();
+qed "length_rev";
val prems = goal List.thy
"ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))";
by (list_ind_tac "ls" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_app])));
-val length_flat = result();
+qed "length_flat";
(** Length and drop **)
@@ -319,7 +319,7 @@
by (etac list.induct 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [drop_0,drop_succ_Cons])));
by (fast_tac ZF_cs 1);
-val drop_length_Cons_lemma = result();
+qed "drop_length_Cons_lemma";
val drop_length_Cons = standard (drop_length_Cons_lemma RS spec);
goal List.thy
@@ -338,7 +338,7 @@
by (dtac bspec 1);
by (fast_tac ZF_cs 2);
by (fast_tac (ZF_cs addEs [succ_in_naturalD,length_type]) 1);
-val drop_length_lemma = result();
+qed "drop_length_lemma";
val drop_length = standard (drop_length_lemma RS bspec);
@@ -347,25 +347,25 @@
val [major] = goal List.thy "xs: list(A) ==> xs@Nil=xs";
by (rtac (major RS list.induct) 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val app_right_Nil = result();
+qed "app_right_Nil";
val prems = goal List.thy "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)";
by (list_ind_tac "xs" prems 1);
by (ALLGOALS (asm_simp_tac list_ss));
-val app_assoc = result();
+qed "app_assoc";
val prems = goal List.thy
"ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)";
by (list_ind_tac "ls" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [app_assoc])));
-val flat_app_distrib = result();
+qed "flat_app_distrib";
(*** theorems about rev ***)
val prems = goal List.thy "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))";
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib])));
-val rev_map_distrib = result();
+qed "rev_map_distrib";
(*Simplifier needs the premises as assumptions because rewriting will not
instantiate the variable ?A in the rules' typing conditions; note that
@@ -375,19 +375,19 @@
"!!xs. [| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)";
by (etac list.induct 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [app_right_Nil,app_assoc])));
-val rev_app_distrib = result();
+qed "rev_app_distrib";
val prems = goal List.thy "l: list(A) ==> rev(rev(l))=l";
by (list_ind_tac "l" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [rev_app_distrib])));
-val rev_rev_ident = result();
+qed "rev_rev_ident";
val prems = goal List.thy
"ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))";
by (list_ind_tac "ls" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps
[map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil])));
-val rev_flat = result();
+qed "rev_flat";
(*** theorems about list_add ***)
@@ -401,21 +401,21 @@
(asm_simp_tac (list_ss addsimps [add_0_right, add_assoc RS sym])));
by (rtac (add_commute RS subst_context) 1);
by (REPEAT (ares_tac [refl, list_add_type] 1));
-val list_add_app = result();
+qed "list_add_app";
val prems = goal List.thy
"l: list(nat) ==> list_add(rev(l)) = list_add(l)";
by (list_ind_tac "l" prems 1);
by (ALLGOALS
(asm_simp_tac (list_ss addsimps [list_add_app, add_0_right])));
-val list_add_rev = result();
+qed "list_add_rev";
val prems = goal List.thy
"ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))";
by (list_ind_tac "ls" prems 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps [list_add_app])));
by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1));
-val list_add_flat = result();
+qed "list_add_flat";
(** New induction rule **)
@@ -427,5 +427,5 @@
by (rtac (major RS rev_rev_ident RS subst) 1);
by (rtac (major RS rev_type RS list.induct) 1);
by (ALLGOALS (asm_simp_tac (list_ss addsimps prems)));
-val list_append_induct = result();
+qed "list_append_induct";
--- a/src/ZF/Nat.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Nat.ML Wed Dec 07 13:12:04 1994 +0100
@@ -13,7 +13,7 @@
by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2));
by (cut_facts_tac [infinity] 1);
by (fast_tac ZF_cs 1);
-val nat_bnd_mono = result();
+qed "nat_bnd_mono";
(* nat = {0} Un {succ(x). x:nat} *)
val nat_unfold = nat_bnd_mono RS (nat_def RS def_lfp_Tarski);
@@ -23,22 +23,22 @@
goal Nat.thy "0 : nat";
by (rtac (nat_unfold RS ssubst) 1);
by (rtac (singletonI RS UnI1) 1);
-val nat_0I = result();
+qed "nat_0I";
val prems = goal Nat.thy "n : nat ==> succ(n) : nat";
by (rtac (nat_unfold RS ssubst) 1);
by (rtac (RepFunI RS UnI2) 1);
by (resolve_tac prems 1);
-val nat_succI = result();
+qed "nat_succI";
goal Nat.thy "1 : nat";
by (rtac (nat_0I RS nat_succI) 1);
-val nat_1I = result();
+qed "nat_1I";
goal Nat.thy "bool <= nat";
by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1
ORELSE eresolve_tac [boolE,ssubst] 1));
-val bool_subset_nat = result();
+qed "bool_subset_nat";
val bool_into_nat = bool_subset_nat RS subsetD;
@@ -50,7 +50,7 @@
"[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)";
by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1);
by (fast_tac (ZF_cs addIs prems) 1);
-val nat_induct = result();
+qed "nat_induct";
(*Perform induction on n, then prove the n:nat subgoal using prems. *)
fun nat_ind_tac a prems i =
@@ -63,12 +63,12 @@
by (rtac (major RS (nat_unfold RS equalityD1 RS subsetD) RS UnE) 1);
by (DEPTH_SOLVE (eresolve_tac [singletonE,RepFunE] 1
ORELSE ares_tac prems 1));
-val natE = result();
+qed "natE";
val prems = goal Nat.thy "n: nat ==> Ord(n)";
by (nat_ind_tac "n" prems 1);
by (REPEAT (ares_tac [Ord_0, Ord_succ] 1));
-val nat_into_Ord = result();
+qed "nat_into_Ord";
(* i: nat ==> 0 le i *)
val nat_0_le = nat_into_Ord RS Ord_0_le;
@@ -79,7 +79,7 @@
by (etac nat_induct 1);
by (fast_tac ZF_cs 1);
by (fast_tac (ZF_cs addIs [nat_0_le]) 1);
-val natE0 = result();
+qed "natE0";
goal Nat.thy "Ord(nat)";
by (rtac OrdI 1);
@@ -88,12 +88,12 @@
by (rtac ballI 1);
by (etac nat_induct 1);
by (REPEAT (ares_tac [empty_subsetI,succ_subsetI] 1));
-val Ord_nat = result();
+qed "Ord_nat";
goalw Nat.thy [Limit_def] "Limit(nat)";
by (safe_tac (ZF_cs addSIs [ltI, nat_0I, nat_1I, nat_succI, Ord_nat]));
by (etac ltD 1);
-val Limit_nat = result();
+qed "Limit_nat";
goal Nat.thy "!!i. Limit(i) ==> nat le i";
by (resolve_tac [subset_imp_le] 1);
@@ -102,7 +102,7 @@
by (fast_tac (ZF_cs addIs [Limit_has_succ RS ltD, ltI, Limit_is_Ord]) 2);
by (REPEAT (ares_tac [Limit_has_0 RS ltD,
Ord_nat, Limit_is_Ord] 1));
-val nat_le_Limit = result();
+qed "nat_le_Limit";
(* succ(i): nat ==> i: nat *)
val succ_natD = [succI1, asm_rl, Ord_nat] MRS Ord_trans;
@@ -114,7 +114,7 @@
by (etac ltE 1);
by (etac (Ord_nat RSN (3,Ord_trans)) 1);
by (assume_tac 1);
-val lt_nat_in_nat = result();
+qed "lt_nat_in_nat";
(** Variations on mathematical induction **)
@@ -130,7 +130,7 @@
by (ALLGOALS
(asm_simp_tac
(ZF_ss addsimps (prems@distrib_rews@[le0_iff, le_succ_iff]))));
-val nat_induct_from_lemma = result();
+qed "nat_induct_from_lemma";
(*Induction starting from m rather than 0*)
val prems = goal Nat.thy
@@ -140,7 +140,7 @@
\ |] ==> P(n)";
by (rtac (nat_induct_from_lemma RS mp RS mp) 1);
by (REPEAT (ares_tac prems 1));
-val nat_induct_from = result();
+qed "nat_induct_from";
(*Induction suitable for subtraction and less-than*)
val prems = goal Nat.thy
@@ -155,7 +155,7 @@
by (rtac ballI 2);
by (nat_ind_tac "x" [] 2);
by (REPEAT (ares_tac (prems@[ballI]) 1 ORELSE etac bspec 1));
-val diff_induct = result();
+qed "diff_induct";
(** Induction principle analogous to trancl_induct **)
@@ -166,7 +166,7 @@
by (ALLGOALS
(EVERY' [rtac (impI RS impI), rtac (nat_induct RS ballI), assume_tac,
fast_tac lt_cs, fast_tac lt_cs]));
-val succ_lt_induct_lemma = result();
+qed "succ_lt_induct_lemma";
val prems = goal Nat.thy
"[| m<n; n: nat; \
@@ -176,17 +176,17 @@
by (res_inst_tac [("P4","P")]
(succ_lt_induct_lemma RS mp RS mp RS bspec RS mp) 1);
by (REPEAT (ares_tac (prems @ [ballI, impI, lt_nat_in_nat]) 1));
-val succ_lt_induct = result();
+qed "succ_lt_induct";
(** nat_case **)
goalw Nat.thy [nat_case_def] "nat_case(a,b,0) = a";
by (fast_tac (ZF_cs addIs [the_equality]) 1);
-val nat_case_0 = result();
+qed "nat_case_0";
goalw Nat.thy [nat_case_def] "nat_case(a,b,succ(m)) = b(m)";
by (fast_tac (ZF_cs addIs [the_equality]) 1);
-val nat_case_succ = result();
+qed "nat_case_succ";
val major::prems = goal Nat.thy
"[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) \
@@ -194,7 +194,7 @@
by (rtac (major RS nat_induct) 1);
by (ALLGOALS
(asm_simp_tac (ZF_ss addsimps (prems @ [nat_case_0, nat_case_succ]))));
-val nat_case_type = result();
+qed "nat_case_type";
(** nat_rec -- used to define eclose and transrec, then obsolete;
@@ -205,23 +205,23 @@
goal Nat.thy "nat_rec(0,a,b) = a";
by (rtac nat_rec_trans 1);
by (rtac nat_case_0 1);
-val nat_rec_0 = result();
+qed "nat_rec_0";
val [prem] = goal Nat.thy
"m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))";
by (rtac nat_rec_trans 1);
by (simp_tac (ZF_ss addsimps [prem, nat_case_succ, nat_succI, Memrel_iff,
vimage_singleton_iff]) 1);
-val nat_rec_succ = result();
+qed "nat_rec_succ";
(** 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";
by (rtac (Un_least_lt RS ltD) 1);
by (REPEAT (ares_tac [ltI, Ord_nat] 1));
-val Un_nat_type = result();
+qed "Un_nat_type";
goal Nat.thy "!!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));
-val Int_nat_type = result();
+qed "Int_nat_type";
--- a/src/ZF/Order.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Order.ML Wed Dec 07 13:12:04 1994 +0100
@@ -23,7 +23,7 @@
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, asym_def]
"!!r. part_ord(A,r) ==> asym(r Int A*A)";
by (fast_tac ZF_cs 1);
-val part_ord_Imp_asym = result();
+qed "part_ord_Imp_asym";
(** Subset properties for the various forms of relation **)
@@ -32,22 +32,22 @@
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def]
"!!A B r. [| part_ord(A,r); B<=A |] ==> part_ord(B,r)";
by (fast_tac ZF_cs 1);
-val part_ord_subset = result();
+qed "part_ord_subset";
goalw Order.thy [linear_def]
"!!A B r. [| linear(A,r); B<=A |] ==> linear(B,r)";
by (fast_tac ZF_cs 1);
-val linear_subset = result();
+qed "linear_subset";
goalw Order.thy [tot_ord_def]
"!!A B r. [| tot_ord(A,r); B<=A |] ==> tot_ord(B,r)";
by (fast_tac (ZF_cs addSEs [part_ord_subset, linear_subset]) 1);
-val tot_ord_subset = result();
+qed "tot_ord_subset";
goalw Order.thy [well_ord_def]
"!!A B r. [| well_ord(A,r); B<=A |] ==> well_ord(B,r)";
by (fast_tac (ZF_cs addSEs [tot_ord_subset, wf_on_subset_A]) 1);
-val well_ord_subset = result();
+qed "well_ord_subset";
(** Order-isomorphisms **)
@@ -55,13 +55,13 @@
goalw Order.thy [ord_iso_def]
"!!f. f: ord_iso(A,r,B,s) ==> f: bij(A,B)";
by (etac CollectD1 1);
-val ord_iso_is_bij = result();
+qed "ord_iso_is_bij";
goalw Order.thy [ord_iso_def]
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: r; x:A; y:A |] ==> \
\ <f`x, f`y> : s";
by (fast_tac ZF_cs 1);
-val ord_iso_apply = result();
+qed "ord_iso_apply";
goalw Order.thy [ord_iso_def]
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: s; x:B; y:B |] ==> \
@@ -71,7 +71,7 @@
by (REPEAT (eresolve_tac [asm_rl,
bij_converse_bij RS bij_is_fun RS apply_type] 1));
by (asm_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
-val ord_iso_converse = result();
+qed "ord_iso_converse";
val major::premx::premy::prems = goalw Order.thy [linear_def]
"[| linear(A,r); x:A; y:A; \
@@ -80,7 +80,7 @@
by (REPEAT_FIRST (eresolve_tac [ballE,disjE]));
by (EVERY1 (map etac prems));
by (ALLGOALS contr_tac);
-val linearE = result();
+qed "linearE";
(*Does the case analysis, deleting linear(A,r) and proving trivial subgoals*)
val linear_case_tac =
@@ -113,7 +113,7 @@
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
by (dres_inst_tac [("t", "op `(f)")] subst_context 1);
by (asm_full_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
-val not_well_ord_iso_pred_lemma = result();
+qed "not_well_ord_iso_pred_lemma";
(*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment
@@ -130,7 +130,7 @@
(*Now we know f`x : pred(A,x,r) *)
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, pred_def]) 1);
by (fast_tac (ZF_cs addSEs [wf_on_not_refl RS notE]) 1);
-val not_well_ord_iso_pred = result();
+qed "not_well_ord_iso_pred";
(*Inductive argument for proving Kunen's Lemma 6.2*)
@@ -157,7 +157,7 @@
by (dres_inst_tac [("t", "op `(converse(g))")] subst_context 1);
by (dres_inst_tac [("t", "op `(f)")] subst_context 1);
by (asm_full_simp_tac bij_inverse_ss 1);
-val well_ord_iso_unique_lemma = result();
+qed "well_ord_iso_unique_lemma";
(*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*)
@@ -169,7 +169,7 @@
by (etac (ord_iso_is_bij RS bij_is_fun) 1);
by (rtac well_ord_iso_unique_lemma 1);
by (REPEAT_SOME assume_tac);
-val well_ord_iso_unique = result();
+qed "well_ord_iso_unique";
goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def,
@@ -182,17 +182,17 @@
by (fast_tac (ZF_cs addSEs [wf_on_asym]) 1);
(*case x<xb*)
by (fast_tac (ZF_cs addSEs [wf_on_chain3]) 1);
-val well_ordI = result();
+qed "well_ordI";
goalw Order.thy [well_ord_def]
"!!r. well_ord(A,r) ==> wf[A](r)";
by (safe_tac ZF_cs);
-val well_ord_is_wf = result();
+qed "well_ord_is_wf";
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
"!!r. well_ord(A,r) ==> trans[A](r)";
by (safe_tac ZF_cs);
-val well_ord_is_trans_on = result();
+qed "well_ord_is_trans_on";
(*** Derived rules for pred(A,x,r) ***)
@@ -200,21 +200,21 @@
"[| y: pred(A,x,r); [| y:A; <y,x>:r |] ==> P |] ==> P";
by (rtac (major RS CollectE) 1);
by (REPEAT (ares_tac [minor] 1));
-val predE = result();
+qed "predE";
goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}";
by (fast_tac ZF_cs 1);
-val pred_subset_under = result();
+qed "pred_subset_under";
goalw Order.thy [pred_def] "pred(A,x,r) <= A";
by (fast_tac ZF_cs 1);
-val pred_subset = result();
+qed "pred_subset";
goalw Order.thy [pred_def] "y : pred(A,x,r) <-> <y,x>:r & y:A";
by (fast_tac ZF_cs 1);
-val pred_iff = result();
+qed "pred_iff";
goalw Order.thy [pred_def]
"pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)";
by (fast_tac eq_cs 1);
-val pred_pred_eq = result();
+qed "pred_pred_eq";
--- a/src/ZF/OrderArith.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/OrderArith.ML Wed Dec 07 13:12:04 1994 +0100
@@ -16,22 +16,22 @@
goalw OrderArith.thy [radd_def]
"<Inl(a), Inr(b)> : radd(A,r,B,s) <-> a:A & b:B";
by (fast_tac sum_cs 1);
-val radd_Inl_Inr_iff = result();
+qed "radd_Inl_Inr_iff";
goalw OrderArith.thy [radd_def]
"<Inl(a'), Inl(a)> : radd(A,r,B,s) <-> <a',a>:r & a':A & a:A";
by (fast_tac sum_cs 1);
-val radd_Inl_iff = result();
+qed "radd_Inl_iff";
goalw OrderArith.thy [radd_def]
"<Inr(b'), Inr(b)> : radd(A,r,B,s) <-> <b',b>:s & b':B & b:B";
by (fast_tac sum_cs 1);
-val radd_Inr_iff = result();
+qed "radd_Inr_iff";
goalw OrderArith.thy [radd_def]
"<Inr(b), Inl(a)> : radd(A,r,B,s) <-> False";
by (fast_tac sum_cs 1);
-val radd_Inr_Inl_iff = result();
+qed "radd_Inr_Inl_iff";
(** Elimination Rule **)
@@ -50,13 +50,13 @@
by (EVERY (map (fn prem =>
EVERY1 [rtac prem, assume_tac, REPEAT o fast_tac sum_cs])
prems));
-val raddE = result();
+qed "raddE";
(** Type checking **)
goalw OrderArith.thy [radd_def] "radd(A,r,B,s) <= (A+B) * (A+B)";
by (rtac Collect_subset 1);
-val radd_type = result();
+qed "radd_type";
val field_radd = standard (radd_type RS field_rel_subset);
@@ -72,7 +72,7 @@
"!!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 radd_ss));
-val linear_radd = result();
+qed "linear_radd";
(** Well-foundedness **)
@@ -95,14 +95,14 @@
by (etac (bspec RS mp) 1);
by (fast_tac sum_cs 1);
by (best_tac (sum_cs addSEs [raddE]) 1);
-val wf_on_radd = result();
+qed "wf_on_radd";
goal OrderArith.thy
"!!r s. [| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))";
by (asm_full_simp_tac (ZF_ss 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));
-val wf_radd = result();
+qed "wf_radd";
goal OrderArith.thy
"!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \
@@ -111,7 +111,7 @@
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, wf_on_radd]) 1);
by (asm_full_simp_tac
(ZF_ss addsimps [well_ord_def, tot_ord_def, linear_radd]) 1);
-val well_ord_radd = result();
+qed "well_ord_radd";
(**** Multiplication of relations -- lexicographic product ****)
@@ -123,7 +123,7 @@
\ (<a',a>: r & a':A & a:A & b': B & b: B) | \
\ (<b',b>: s & a'=a & a:A & b': B & b: B)";
by (fast_tac ZF_cs 1);
-val rmult_iff = result();
+qed "rmult_iff";
val major::prems = goal OrderArith.thy
"[| <<a',b'>, <a,b>> : rmult(A,r,B,s); \
@@ -132,13 +132,13 @@
\ |] ==> Q";
by (rtac (major RS (rmult_iff RS iffD1) RS disjE) 1);
by (DEPTH_SOLVE (eresolve_tac ([asm_rl, conjE] @ prems) 1));
-val rmultE = result();
+qed "rmultE";
(** Type checking **)
goalw OrderArith.thy [rmult_def] "rmult(A,r,B,s) <= (A*B) * (A*B)";
by (rtac Collect_subset 1);
-val rmult_type = result();
+qed "rmult_type";
val field_rmult = standard (rmult_type RS field_rel_subset);
@@ -152,7 +152,7 @@
by (res_inst_tac [("x","xa"), ("y","xb")] (lina RS linearE) 1);
by (res_inst_tac [("x","ya"), ("y","yb")] (linb RS linearE) 4);
by (REPEAT_SOME (fast_tac ZF_cs));
-val linear_rmult = result();
+qed "linear_rmult";
(** Well-foundedness **)
@@ -170,7 +170,7 @@
by (etac (bspec RS mp) 1);
by (fast_tac ZF_cs 1);
by (best_tac (ZF_cs addSEs [rmultE]) 1);
-val wf_on_rmult = result();
+qed "wf_on_rmult";
goal OrderArith.thy
@@ -178,7 +178,7 @@
by (asm_full_simp_tac (ZF_ss 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));
-val wf_rmult = result();
+qed "wf_rmult";
goal OrderArith.thy
"!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \
@@ -187,7 +187,7 @@
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, wf_on_rmult]) 1);
by (asm_full_simp_tac
(ZF_ss addsimps [well_ord_def, tot_ord_def, linear_rmult]) 1);
-val well_ord_rmult = result();
+qed "well_ord_rmult";
(**** Inverse image of a relation ****)
@@ -197,13 +197,13 @@
goalw OrderArith.thy [rvimage_def]
"<a,b> : rvimage(A,f,r) <-> <f`a,f`b>: r & a:A & b:A";
by (fast_tac ZF_cs 1);
-val rvimage_iff = result();
+qed "rvimage_iff";
(** Type checking **)
goalw OrderArith.thy [rvimage_def] "rvimage(A,f,r) <= A*A";
by (rtac Collect_subset 1);
-val rvimage_type = result();
+qed "rvimage_type";
val field_rvimage = standard (rvimage_type RS field_rel_subset);
@@ -218,7 +218,7 @@
by (cut_facts_tac [finj] 1);
by (res_inst_tac [("x","f`x"), ("y","f`y")] (lin RS linearE) 1);
by (REPEAT_SOME (fast_tac (ZF_cs addSIs [apply_type])));
-val linear_rvimage = result();
+qed "linear_rvimage";
(** Well-foundedness **)
@@ -231,7 +231,7 @@
by (eres_inst_tac [("a","f`y")] wf_on_induct 1);
by (fast_tac (ZF_cs addSIs [apply_type]) 1);
by (best_tac (ZF_cs addSIs [apply_type] addSDs [rvimage_iff RS iffD1]) 1);
-val wf_on_rvimage = result();
+qed "wf_on_rvimage";
goal OrderArith.thy
"!!r. [| f: inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))";
@@ -239,4 +239,4 @@
by (rewrite_goals_tac [well_ord_def, tot_ord_def]);
by (fast_tac (ZF_cs addSIs [wf_on_rvimage, inj_is_fun]) 1);
by (fast_tac (ZF_cs addSIs [linear_rvimage]) 1);
-val well_ord_rvimage = result();
+qed "well_ord_rvimage";
--- a/src/ZF/OrderType.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/OrderType.ML Wed Dec 07 13:12:04 1994 +0100
@@ -14,7 +14,7 @@
by (rtac lam_type 1);
by (rtac (lamI RS imageI) 1);
by (REPEAT (assume_tac 1));
-val ordermap_type = result();
+qed "ordermap_type";
(** Unfolding of ordermap **)
@@ -27,7 +27,7 @@
by (assume_tac 1);
by (asm_simp_tac (ZF_ss addsimps [subset_iff, image_lam,
vimage_singleton_iff]) 1);
-val ordermap_eq_image = result();
+qed "ordermap_eq_image";
(*Useful for rewriting PROVIDED pred is not unfolded until later!*)
goal OrderType.thy
@@ -35,7 +35,7 @@
\ ordermap(A,r) ` x = {ordermap(A,r)`y . y : pred(A,x,r)}";
by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, pred_subset,
ordermap_type RS image_fun]) 1);
-val ordermap_pred_unfold = result();
+qed "ordermap_pred_unfold";
(*pred-unfolded version. NOT suitable for rewriting -- loops!*)
val ordermap_unfold = rewrite_rule [pred_def] ordermap_pred_unfold;
@@ -58,7 +58,7 @@
by (safe_tac ZF_cs);
by (ordermap_elim_tac 1);
by (fast_tac (ZF_cs addSEs [trans_onD]) 1);
-val Ord_ordermap = result();
+qed "Ord_ordermap";
goalw OrderType.thy [ordertype_def]
"!!r. well_ord(A,r) ==> Ord(ordertype(A,r))";
@@ -69,7 +69,7 @@
by (safe_tac ZF_cs);
by (ordermap_elim_tac 1);
by (fast_tac ZF_cs 1);
-val Ord_ordertype = result();
+qed "Ord_ordertype";
(** ordermap preserves the orderings in both directions **)
@@ -79,7 +79,7 @@
by (eres_inst_tac [("x1", "x")] (ordermap_unfold RS ssubst) 1);
by (assume_tac 1);
by (fast_tac ZF_cs 1);
-val ordermap_mono = result();
+qed "ordermap_mono";
(*linearity of r is crucial here*)
goalw OrderType.thy [well_ord_def, tot_ord_def]
@@ -92,7 +92,7 @@
by (REPEAT_SOME assume_tac);
by (etac mem_asym 1);
by (assume_tac 1);
-val converse_ordermap_mono = result();
+qed "converse_ordermap_mono";
val ordermap_surj =
(ordermap_type RS surj_image) |>
@@ -109,7 +109,7 @@
by (ALLGOALS (dtac ordermap_mono));
by (REPEAT_SOME assume_tac);
by (ALLGOALS (asm_full_simp_tac (ZF_ss addsimps [mem_not_refl])));
-val ordermap_bij = result();
+qed "ordermap_bij";
goalw OrderType.thy [ord_iso_def]
"!!r. well_ord(A,r) ==> \
@@ -121,7 +121,7 @@
by (rewtac well_ord_def);
by (fast_tac (ZF_cs addSIs [MemrelI, ordermap_mono,
ordermap_type RS apply_type]) 1);
-val ordertype_ord_iso = result();
+qed "ordertype_ord_iso";
(** Unfolding of ordertype **)
@@ -129,7 +129,7 @@
goalw OrderType.thy [ordertype_def]
"ordertype(A,r) = {ordermap(A,r)`y . y : A}";
by (rtac ([ordermap_type, subset_refl] MRS image_fun) 1);
-val ordertype_unfold = result();
+qed "ordertype_unfold";
(** Ordertype of Memrel **)
@@ -141,7 +141,7 @@
by (asm_simp_tac (ZF_ss addsimps [linear_def, Memrel_iff]) 1);
by (REPEAT (resolve_tac [ballI, Ord_linear] 1));;
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1));;
-val well_ord_Memrel = result();
+qed "well_ord_Memrel";
goal OrderType.thy "!!i. [| Ord(i); j:i |] ==> ordermap(i,Memrel(i)) ` j = j";
by (eresolve_tac [Ord_induct] 1);
@@ -152,11 +152,11 @@
by (dresolve_tac [OrdmemD] 1);
by (assume_tac 1);
by (fast_tac eq_cs 1);
-val ordermap_Memrel = result();
+qed "ordermap_Memrel";
goal OrderType.thy "!!i. Ord(i) ==> ordertype(i,Memrel(i)) = i";
by (asm_simp_tac (ZF_ss addsimps [ordertype_unfold, ordermap_Memrel]) 1);
-val ordertype_Memrel = result();
+qed "ordertype_Memrel";
(** Ordertype of rvimage **)
@@ -174,7 +174,7 @@
by (fast_tac bij_apply_cs 1);
by (res_inst_tac [("a", "converse(f)`xb")] RepFun_eqI 1);
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
-val bij_ordermap_vimage = result();
+qed "bij_ordermap_vimage";
goal OrderType.thy
"!!f. [| f: bij(A,B); well_ord(B,r) |] ==> \
@@ -184,7 +184,7 @@
by (fast_tac bij_apply_cs 1);
by (res_inst_tac [("a", "converse(f)`xa")] RepFun_eqI 1);
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
-val bij_ordertype_vimage = result();
+qed "bij_ordertype_vimage";
goal OrderType.thy
@@ -201,7 +201,7 @@
br (refl RSN (2,RepFun_cong)) 1;
bd well_ord_is_trans_on 1;
by (fast_tac (eq_cs addSEs [trans_onD]) 1);
-val ordermap_pred_eq_ordermap = result();
+qed "ordermap_pred_eq_ordermap";
goal OrderType.thy
@@ -215,5 +215,5 @@
[well_ord_is_wf RS ordermap_pred_unfold]) 1);
by (asm_simp_tac (ZF_ss addsimps [ordertype_unfold,
ordermap_pred_eq_ordermap]) 1);
-val ordertype_subset = result();
+qed "ordertype_subset";
--- a/src/ZF/Ordinal.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Ordinal.ML Wed Dec 07 13:12:04 1994 +0100
@@ -14,36 +14,36 @@
goalw Ordinal.thy [Transset_def] "Transset(A) <-> A<=Pow(A)";
by (fast_tac ZF_cs 1);
-val Transset_iff_Pow = result();
+qed "Transset_iff_Pow";
goalw Ordinal.thy [Transset_def] "Transset(A) <-> Union(succ(A)) = A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val Transset_iff_Union_succ = result();
+qed "Transset_iff_Union_succ";
(** Consequences of downwards closure **)
goalw Ordinal.thy [Transset_def]
"!!C a b. [| Transset(C); {a,b}: C |] ==> a:C & b: C";
by (fast_tac ZF_cs 1);
-val Transset_doubleton_D = result();
+qed "Transset_doubleton_D";
val [prem1,prem2] = goalw Ordinal.thy [Pair_def]
"[| Transset(C); <a,b>: C |] ==> a:C & b: C";
by (cut_facts_tac [prem2] 1);
by (fast_tac (ZF_cs addSDs [prem1 RS Transset_doubleton_D]) 1);
-val Transset_Pair_D = result();
+qed "Transset_Pair_D";
val prem1::prems = goal Ordinal.thy
"[| Transset(C); A*B <= C; b: B |] ==> A <= C";
by (cut_facts_tac prems 1);
by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1);
-val Transset_includes_domain = result();
+qed "Transset_includes_domain";
val prem1::prems = goal Ordinal.thy
"[| Transset(C); A*B <= C; a: A |] ==> B <= C";
by (cut_facts_tac prems 1);
by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1);
-val Transset_includes_range = result();
+qed "Transset_includes_range";
val [prem1,prem2] = goalw (merge_theories(Ordinal.thy,Sum.thy)) [sum_def]
"[| Transset(C); A+B <= C |] ==> A <= C & B <= C";
@@ -62,63 +62,63 @@
goalw Ordinal.thy [Transset_def] "Transset(0)";
by (fast_tac ZF_cs 1);
-val Transset_0 = result();
+qed "Transset_0";
goalw Ordinal.thy [Transset_def]
"!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Un j)";
by (fast_tac ZF_cs 1);
-val Transset_Un = result();
+qed "Transset_Un";
goalw Ordinal.thy [Transset_def]
"!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Int j)";
by (fast_tac ZF_cs 1);
-val Transset_Int = result();
+qed "Transset_Int";
goalw Ordinal.thy [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))";
by (fast_tac ZF_cs 1);
-val Transset_succ = result();
+qed "Transset_succ";
goalw Ordinal.thy [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))";
by (fast_tac ZF_cs 1);
-val Transset_Pow = result();
+qed "Transset_Pow";
goalw Ordinal.thy [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))";
by (fast_tac ZF_cs 1);
-val Transset_Union = result();
+qed "Transset_Union";
val [Transprem] = goalw Ordinal.thy [Transset_def]
"[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))";
by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1);
-val Transset_Union_family = result();
+qed "Transset_Union_family";
val [prem,Transprem] = goalw Ordinal.thy [Transset_def]
"[| j:A; !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))";
by (cut_facts_tac [prem] 1);
by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1);
-val Transset_Inter_family = result();
+qed "Transset_Inter_family";
(*** Natural Deduction rules for Ord ***)
val prems = goalw Ordinal.thy [Ord_def]
"[| Transset(i); !!x. x:i ==> Transset(x) |] ==> Ord(i) ";
by (REPEAT (ares_tac (prems@[ballI,conjI]) 1));
-val OrdI = result();
+qed "OrdI";
val [major] = goalw Ordinal.thy [Ord_def]
"Ord(i) ==> Transset(i)";
by (rtac (major RS conjunct1) 1);
-val Ord_is_Transset = result();
+qed "Ord_is_Transset";
val [major,minor] = goalw Ordinal.thy [Ord_def]
"[| Ord(i); j:i |] ==> Transset(j) ";
by (rtac (minor RS (major RS conjunct2 RS bspec)) 1);
-val Ord_contains_Transset = result();
+qed "Ord_contains_Transset";
(*** Lemmas for ordinals ***)
goalw Ordinal.thy [Ord_def,Transset_def] "!!i j.[| Ord(i); j:i |] ==> Ord(j)";
by (fast_tac ZF_cs 1);
-val Ord_in_Ord = result();
+qed "Ord_in_Ord";
(* Ord(succ(j)) ==> Ord(j) *)
val Ord_succD = succI1 RSN (2, Ord_in_Ord);
@@ -126,44 +126,44 @@
goal Ordinal.thy "!!i j. [| Ord(i); Transset(j); j<=i |] ==> Ord(j)";
by (REPEAT (ares_tac [OrdI] 1
ORELSE eresolve_tac [Ord_contains_Transset, subsetD] 1));
-val Ord_subset_Ord = result();
+qed "Ord_subset_Ord";
goalw Ordinal.thy [Ord_def,Transset_def] "!!i j. [| j:i; Ord(i) |] ==> j<=i";
by (fast_tac ZF_cs 1);
-val OrdmemD = result();
+qed "OrdmemD";
goal Ordinal.thy "!!i j k. [| i:j; j:k; Ord(k) |] ==> i:k";
by (REPEAT (ares_tac [OrdmemD RS subsetD] 1));
-val Ord_trans = result();
+qed "Ord_trans";
goal Ordinal.thy "!!i j. [| i:j; Ord(j) |] ==> succ(i) <= j";
by (REPEAT (ares_tac [OrdmemD RSN (2,succ_subsetI)] 1));
-val Ord_succ_subsetI = result();
+qed "Ord_succ_subsetI";
(*** The construction of ordinals: 0, succ, Union ***)
goal Ordinal.thy "Ord(0)";
by (REPEAT (ares_tac [OrdI,Transset_0] 1 ORELSE etac emptyE 1));
-val Ord_0 = result();
+qed "Ord_0";
goal Ordinal.thy "!!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));
-val Ord_succ = result();
+qed "Ord_succ";
goal Ordinal.thy "Ord(succ(i)) <-> Ord(i)";
by (fast_tac (ZF_cs addIs [Ord_succ] addDs [Ord_succD]) 1);
-val Ord_succ_iff = result();
+qed "Ord_succ_iff";
goalw Ordinal.thy [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Un j)";
by (fast_tac (ZF_cs addSIs [Transset_Un]) 1);
-val Ord_Un = result();
+qed "Ord_Un";
goalw Ordinal.thy [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Int j)";
by (fast_tac (ZF_cs addSIs [Transset_Int]) 1);
-val Ord_Int = result();
+qed "Ord_Int";
val nonempty::prems = goal Ordinal.thy
"[| j:A; !!i. i:A ==> Ord(i) |] ==> Ord(Inter(A))";
@@ -171,7 +171,7 @@
by (rtac Ord_is_Transset 1);
by (REPEAT (ares_tac ([Ord_contains_Transset,nonempty]@prems) 1
ORELSE etac InterD 1));
-val Ord_Inter = result();
+qed "Ord_Inter";
val jmemA::prems = goal Ordinal.thy
"[| j:A; !!x. x:A ==> Ord(B(x)) |] ==> Ord(INT x:A. B(x))";
@@ -179,7 +179,7 @@
by (etac RepFunE 1);
by (etac ssubst 1);
by (eresolve_tac prems 1);
-val Ord_INT = result();
+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))";
@@ -192,81 +192,81 @@
by (safe_tac ZF_cs);
by (asm_full_simp_tac ZF_ss 1);
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1));
-val ON_class = result();
+qed "ON_class";
(*** < is 'less than' for ordinals ***)
goalw Ordinal.thy [lt_def] "!!i j. [| i:j; Ord(j) |] ==> i<j";
by (REPEAT (ares_tac [conjI] 1));
-val ltI = result();
+qed "ltI";
val major::prems = goalw Ordinal.thy [lt_def]
"[| i<j; [| i:j; Ord(i); Ord(j) |] ==> P |] ==> P";
by (rtac (major RS conjE) 1);
by (REPEAT (ares_tac (prems@[Ord_in_Ord]) 1));
-val ltE = result();
+qed "ltE";
goal Ordinal.thy "!!i j. i<j ==> i:j";
by (etac ltE 1);
by (assume_tac 1);
-val ltD = result();
+qed "ltD";
goalw Ordinal.thy [lt_def] "~ i<0";
by (fast_tac ZF_cs 1);
-val not_lt0 = result();
+qed "not_lt0";
(* i<0 ==> R *)
val lt0E = standard (not_lt0 RS notE);
goal Ordinal.thy "!!i j k. [| i<j; j<k |] ==> i<k";
by (fast_tac (ZF_cs addSIs [ltI] addSEs [ltE, Ord_trans]) 1);
-val lt_trans = result();
+qed "lt_trans";
goalw Ordinal.thy [lt_def] "!!i j. [| i<j; j<i |] ==> P";
by (REPEAT (eresolve_tac [asm_rl, conjE, mem_asym] 1));
-val lt_asym = result();
+qed "lt_asym";
-val lt_irrefl = prove_goal Ordinal.thy "i<i ==> P"
+qed_goal "lt_irrefl" Ordinal.thy "i<i ==> P"
(fn [major]=> [ (rtac (major RS (major RS lt_asym)) 1) ]);
-val lt_not_refl = prove_goal Ordinal.thy "~ i<i"
+qed_goal "lt_not_refl" Ordinal.thy "~ i<i"
(fn _=> [ (rtac notI 1), (etac lt_irrefl 1) ]);
(** le is less than or equals; recall i le j abbrevs i<succ(j) !! **)
goalw Ordinal.thy [lt_def] "i le j <-> i<j | (i=j & Ord(j))";
by (fast_tac (ZF_cs addSIs [Ord_succ] addSDs [Ord_succD]) 1);
-val le_iff = result();
+qed "le_iff";
goal Ordinal.thy "!!i j. i<j ==> i le j";
by (asm_simp_tac (ZF_ss addsimps [le_iff]) 1);
-val leI = result();
+qed "leI";
goal Ordinal.thy "!!i. [| i=j; Ord(j) |] ==> i le j";
by (asm_simp_tac (ZF_ss addsimps [le_iff]) 1);
-val le_eqI = result();
+qed "le_eqI";
val le_refl = refl RS le_eqI;
val [prem] = goal Ordinal.thy "(~ (i=j & Ord(j)) ==> i<j) ==> i le j";
by (rtac (disjCI RS (le_iff RS iffD2)) 1);
by (etac prem 1);
-val leCI = result();
+qed "leCI";
val major::prems = goal Ordinal.thy
"[| i le j; i<j ==> P; [| i=j; Ord(j) |] ==> P |] ==> P";
by (rtac (major RS (le_iff RS iffD1 RS disjE)) 1);
by (DEPTH_SOLVE (ares_tac prems 1 ORELSE etac conjE 1));
-val leE = result();
+qed "leE";
goal Ordinal.thy "!!i j. [| i le j; j le i |] ==> i=j";
by (asm_full_simp_tac (ZF_ss addsimps [le_iff]) 1);
by (fast_tac (ZF_cs addEs [lt_asym]) 1);
-val le_anti_sym = result();
+qed "le_anti_sym";
goal Ordinal.thy "i le 0 <-> i=0";
by (fast_tac (ZF_cs addSIs [Ord_0 RS le_refl] addSEs [leE, lt0E]) 1);
-val le0_iff = result();
+qed "le0_iff";
val le0D = standard (le0_iff RS iffD1);
@@ -280,11 +280,11 @@
goalw Ordinal.thy [Memrel_def] "<a,b> : Memrel(A) <-> a:b & a:A & b:A";
by (fast_tac ZF_cs 1);
-val Memrel_iff = result();
+qed "Memrel_iff";
val prems = goal Ordinal.thy "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)";
by (REPEAT (resolve_tac (prems@[conjI, Memrel_iff RS iffD2]) 1));
-val MemrelI = result();
+qed "MemrelI";
val [major,minor] = goal Ordinal.thy
"[| <a,b> : Memrel(A); \
@@ -294,7 +294,7 @@
by (etac conjE 1);
by (rtac minor 1);
by (REPEAT (assume_tac 1));
-val MemrelE = result();
+qed "MemrelE";
(*The membership relation (as a set) is well-founded.
Proof idea: show A<=B by applying the foundation axiom to A-B *)
@@ -306,19 +306,19 @@
etac MemrelE,
etac bspec,
REPEAT o assume_tac]);
-val wf_Memrel = result();
+qed "wf_Memrel";
(*Transset(i) does not suffice, though ALL j:i.Transset(j) does*)
goalw Ordinal.thy [Ord_def, Transset_def, trans_def]
"!!i. Ord(i) ==> trans(Memrel(i))";
by (fast_tac (ZF_cs addSIs [MemrelI] addSEs [MemrelE]) 1);
-val trans_Memrel = result();
+qed "trans_Memrel";
(*If Transset(A) then Memrel(A) internalizes the membership relation below A*)
goalw Ordinal.thy [Transset_def]
"!!A. Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A";
by (fast_tac (ZF_cs addSIs [MemrelI] addSEs [MemrelE]) 1);
-val Transset_Memrel_iff = result();
+qed "Transset_Memrel_iff";
(*** Transfinite induction ***)
@@ -334,7 +334,7 @@
by (assume_tac 1);
by (cut_facts_tac prems 1);
by (fast_tac (ZF_cs addIs [MemrelI]) 1);
-val Transset_induct = result();
+qed "Transset_induct";
(*Induction over an ordinal*)
val Ord_induct = Ord_is_Transset RSN (2, Transset_induct);
@@ -348,7 +348,7 @@
by (rtac indhyp 1);
by (rtac (major RS Ord_succ RS Ord_in_Ord) 1);
by (REPEAT (assume_tac 1));
-val trans_induct = result();
+qed "trans_induct";
(*Perform induction on i, then prove the Ord(i) subgoal using prems. *)
fun trans_ind_tac a prems i =
@@ -370,7 +370,7 @@
by (rtac (impI RS allI) 1);
by (trans_ind_tac "j" [] 1);
by (DEPTH_SOLVE (step_tac eq_cs 1 ORELSE Ord_trans_tac 1));
-val Ord_linear_lemma = result();
+qed "Ord_linear_lemma";
val Ord_linear = standard (Ord_linear_lemma RS spec RS mp);
(*The trichotomy law for ordinals!*)
@@ -379,48 +379,48 @@
by (rtac ([ordi,ordj] MRS Ord_linear RS disjE) 1);
by (etac disjE 2);
by (DEPTH_SOLVE (ares_tac ([ordi,ordj,conjI] @ prems) 1));
-val Ord_linear_lt = result();
+qed "Ord_linear_lt";
val prems = goal Ordinal.thy
"[| Ord(i); Ord(j); i<j ==> P; j le i ==> P |] ==> P";
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1);
by (DEPTH_SOLVE (ares_tac ([leI, sym RS le_eqI] @ prems) 1));
-val Ord_linear2 = result();
+qed "Ord_linear2";
val prems = goal Ordinal.thy
"[| Ord(i); Ord(j); i le j ==> P; j le i ==> P |] ==> P";
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1);
by (DEPTH_SOLVE (ares_tac ([leI,le_eqI] @ prems) 1));
-val Ord_linear_le = result();
+qed "Ord_linear_le";
goal Ordinal.thy "!!i j. j le i ==> ~ i<j";
by (fast_tac (lt_cs addEs [lt_asym]) 1);
-val le_imp_not_lt = result();
+qed "le_imp_not_lt";
goal Ordinal.thy "!!i j. [| ~ i<j; Ord(i); Ord(j) |] ==> j le i";
by (res_inst_tac [("i","i"),("j","j")] Ord_linear2 1);
by (REPEAT (SOMEGOAL assume_tac));
by (fast_tac (lt_cs addEs [lt_asym]) 1);
-val not_lt_imp_le = result();
+qed "not_lt_imp_le";
goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i";
by (REPEAT (ares_tac [iffI, le_imp_not_lt, not_lt_imp_le] 1));
-val not_lt_iff_le = result();
+qed "not_lt_iff_le";
goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i";
by (asm_simp_tac (ZF_ss addsimps [not_lt_iff_le RS iff_sym]) 1);
-val not_le_iff_lt = result();
+qed "not_le_iff_lt";
goal Ordinal.thy "!!i. Ord(i) ==> 0 le i";
by (etac (not_lt_iff_le RS iffD1) 1);
by (REPEAT (resolve_tac [Ord_0, not_lt0] 1));
-val Ord_0_le = result();
+qed "Ord_0_le";
goal Ordinal.thy "!!i. [| Ord(i); i~=0 |] ==> 0<i";
by (etac (not_le_iff_lt RS iffD1) 1);
by (rtac Ord_0 1);
by (fast_tac lt_cs 1);
-val Ord_0_lt = result();
+qed "Ord_0_lt";
(*** Results about less-than or equals ***)
@@ -431,72 +431,72 @@
by (assume_tac 1);
by (assume_tac 1);
by (fast_tac (ZF_cs addEs [ltE, mem_irrefl]) 1);
-val subset_imp_le = result();
+qed "subset_imp_le";
goal Ordinal.thy "!!i j. i le j ==> i<=j";
by (etac leE 1);
by (fast_tac ZF_cs 2);
by (fast_tac (subset_cs addIs [OrdmemD] addEs [ltE]) 1);
-val le_imp_subset = result();
+qed "le_imp_subset";
goal Ordinal.thy "j le i <-> j<=i & Ord(i) & Ord(j)";
by (fast_tac (ZF_cs addSEs [subset_imp_le, le_imp_subset]
addEs [ltE, make_elim Ord_succD]) 1);
-val le_subset_iff = result();
+qed "le_subset_iff";
goal Ordinal.thy "i le succ(j) <-> i le j | i=succ(j) & Ord(i)";
by (simp_tac (ZF_ss addsimps [le_iff]) 1);
by (fast_tac (ZF_cs addIs [Ord_succ] addDs [Ord_succD]) 1);
-val le_succ_iff = result();
+qed "le_succ_iff";
(*Just a variant of subset_imp_le*)
val [ordi,ordj,minor] = goal Ordinal.thy
"[| Ord(i); Ord(j); !!x. x<j ==> x<i |] ==> j le i";
by (REPEAT_FIRST (ares_tac [notI RS not_lt_imp_le, ordi, ordj]));
by (etac (minor RS lt_irrefl) 1);
-val all_lt_imp_le = result();
+qed "all_lt_imp_le";
(** Transitive laws **)
goal Ordinal.thy "!!i j. [| i le j; j<k |] ==> i<k";
by (fast_tac (ZF_cs addEs [leE, lt_trans]) 1);
-val lt_trans1 = result();
+qed "lt_trans1";
goal Ordinal.thy "!!i j. [| i<j; j le k |] ==> i<k";
by (fast_tac (ZF_cs addEs [leE, lt_trans]) 1);
-val lt_trans2 = result();
+qed "lt_trans2";
goal Ordinal.thy "!!i j. [| i le j; j le k |] ==> i le k";
by (REPEAT (ares_tac [lt_trans1] 1));
-val le_trans = result();
+qed "le_trans";
goal Ordinal.thy "!!i j. i<j ==> succ(i) le j";
by (rtac (not_lt_iff_le RS iffD1) 1);
by (fast_tac (lt_cs addEs [lt_asym]) 3);
by (ALLGOALS (fast_tac (ZF_cs addEs [ltE] addIs [Ord_succ])));
-val succ_leI = result();
+qed "succ_leI";
goal Ordinal.thy "!!i j. succ(i) le j ==> i<j";
by (rtac (not_le_iff_lt RS iffD1) 1);
by (fast_tac (lt_cs addEs [lt_asym]) 3);
by (ALLGOALS (fast_tac (ZF_cs addEs [ltE, make_elim Ord_succD])));
-val succ_leE = result();
+qed "succ_leE";
goal Ordinal.thy "succ(i) le j <-> i<j";
by (REPEAT (ares_tac [iffI,succ_leI,succ_leE] 1));
-val succ_le_iff = result();
+qed "succ_le_iff";
(** Union and Intersection **)
goal Ordinal.thy "!!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));
-val Un_upper1_le = result();
+qed "Un_upper1_le";
goal Ordinal.thy "!!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));
-val Un_upper2_le = result();
+qed "Un_upper2_le";
(*Replacing k by succ(k') yields the similar rule for le!*)
goal Ordinal.thy "!!i j k. [| i<k; j<k |] ==> i Un j < k";
@@ -505,21 +505,21 @@
by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Un_iff]) 4);
by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Un_iff]) 3);
by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
-val Un_least_lt = result();
+qed "Un_least_lt";
goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k";
by (safe_tac (ZF_cs addSIs [Un_least_lt]));
by (rtac (Un_upper2_le RS lt_trans1) 2);
by (rtac (Un_upper1_le RS lt_trans1) 1);
by (REPEAT_SOME assume_tac);
-val Un_least_lt_iff = result();
+qed "Un_least_lt_iff";
val [ordi,ordj,ordk] = goal Ordinal.thy
"[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k";
by (cut_facts_tac [[ordi,ordj] MRS
read_instantiate [("k","k")] Un_least_lt_iff] 1);
by (asm_full_simp_tac (ZF_ss addsimps [lt_def,ordi,ordj,ordk]) 1);
-val Un_least_mem_iff = result();
+qed "Un_least_mem_iff";
(*Replacing k by succ(k') yields the similar rule for le!*)
goal Ordinal.thy "!!i j k. [| i<k; j<k |] ==> i Int j < k";
@@ -528,7 +528,7 @@
by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Int_iff]) 4);
by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Int_iff]) 3);
by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
-val Int_greatest_lt = result();
+qed "Int_greatest_lt";
(*FIXME: the Intersection duals are missing!*)
@@ -538,7 +538,7 @@
val prems = goal Ordinal.thy "[| !!i. i:A ==> Ord(i) |] ==> Ord(Union(A))";
by (rtac (Ord_is_Transset RS Transset_Union_family RS OrdI) 1);
by (REPEAT (etac UnionE 1 ORELSE ares_tac ([Ord_contains_Transset]@prems) 1));
-val Ord_Union = result();
+qed "Ord_Union";
val prems = goal Ordinal.thy
"[| !!x. x:A ==> Ord(B(x)) |] ==> Ord(UN x:A. B(x))";
@@ -546,56 +546,56 @@
by (etac RepFunE 1);
by (etac ssubst 1);
by (eresolve_tac prems 1);
-val Ord_UN = result();
+qed "Ord_UN";
(* No < version; consider (UN i:nat.i)=nat *)
val [ordi,limit] = goal Ordinal.thy
"[| Ord(i); !!x. x:A ==> b(x) le i |] ==> (UN x:A. b(x)) le i";
by (rtac (le_imp_subset RS UN_least RS subset_imp_le) 1);
by (REPEAT (ares_tac [ordi, Ord_UN, limit] 1 ORELSE etac (limit RS ltE) 1));
-val UN_least_le = result();
+qed "UN_least_le";
val [jlti,limit] = goal Ordinal.thy
"[| j<i; !!x. x:A ==> b(x)<j |] ==> (UN x:A. succ(b(x))) < i";
by (rtac (jlti RS ltE) 1);
by (rtac (UN_least_le RS lt_trans2) 1);
by (REPEAT (ares_tac [jlti, succ_leI, limit] 1));
-val UN_succ_least_lt = result();
+qed "UN_succ_least_lt";
val prems = goal Ordinal.thy
"[| a: A; i le b(a); !!x. x:A ==> Ord(b(x)) |] ==> i le (UN x:A. b(x))";
by (resolve_tac (prems RL [ltE]) 1);
by (rtac (le_imp_subset RS subset_trans RS subset_imp_le) 1);
by (REPEAT (ares_tac (prems @ [UN_upper, Ord_UN]) 1));
-val UN_upper_le = result();
+qed "UN_upper_le";
goal Ordinal.thy "!!i. Ord(i) ==> (UN y:i. succ(y)) = i";
by (fast_tac (eq_cs addEs [Ord_trans]) 1);
-val Ord_equality = result();
+qed "Ord_equality";
(*Holds for all transitive sets, not just ordinals*)
goal Ordinal.thy "!!i. Ord(i) ==> Union(i) <= i";
by (fast_tac (ZF_cs addSEs [Ord_trans]) 1);
-val Ord_Union_subset = result();
+qed "Ord_Union_subset";
(*** Limit ordinals -- general properties ***)
goalw Ordinal.thy [Limit_def] "!!i. Limit(i) ==> Union(i) = i";
by (fast_tac (eq_cs addSIs [ltI] addSEs [ltE] addEs [Ord_trans]) 1);
-val Limit_Union_eq = result();
+qed "Limit_Union_eq";
goalw Ordinal.thy [Limit_def] "!!i. Limit(i) ==> Ord(i)";
by (etac conjunct1 1);
-val Limit_is_Ord = result();
+qed "Limit_is_Ord";
goalw Ordinal.thy [Limit_def] "!!i. Limit(i) ==> 0 < i";
by (etac (conjunct2 RS conjunct1) 1);
-val Limit_has_0 = result();
+qed "Limit_has_0";
goalw Ordinal.thy [Limit_def] "!!i. [| Limit(i); j<i |] ==> succ(j) < i";
by (fast_tac ZF_cs 1);
-val Limit_has_succ = result();
+qed "Limit_has_succ";
goalw Ordinal.thy [Limit_def]
"!!i. [| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)";
@@ -603,17 +603,17 @@
by (rtac (not_le_iff_lt RS iffD1) 2);
by (fast_tac (lt_cs addEs [lt_asym]) 4);
by (REPEAT (eresolve_tac [asm_rl, ltE, Ord_succ] 1));
-val non_succ_LimitI = result();
+qed "non_succ_LimitI";
goal Ordinal.thy "!!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);
-val succ_LimitE = result();
+qed "succ_LimitE";
goal Ordinal.thy "!!i. [| Limit(i); i le succ(j) |] ==> i le j";
by (safe_tac (ZF_cs addSEs [succ_LimitE, leE]));
-val Limit_le_succD = result();
+qed "Limit_le_succD";
--- a/src/ZF/Perm.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Perm.ML Wed Dec 07 13:12:04 1994 +0100
@@ -15,16 +15,16 @@
goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B";
by (etac CollectD1 1);
-val surj_is_fun = result();
+qed "surj_is_fun";
goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))";
by (fast_tac (ZF_cs addIs [apply_equality]
addEs [range_of_fun,domain_type]) 1);
-val fun_is_surj = result();
+qed "fun_is_surj";
goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B";
by (best_tac (ZF_cs addIs [equalityI,apply_Pair] addEs [range_type]) 1);
-val surj_range = result();
+qed "surj_range";
(** A function with a right inverse is a surjection **)
@@ -32,7 +32,7 @@
"[| f: A->B; !!y. y:B ==> d(y): A; !!y. y:B ==> f`d(y) = y \
\ |] ==> f: surj(A,B)";
by (fast_tac (ZF_cs addIs prems) 1);
-val f_imp_surjective = result();
+qed "f_imp_surjective";
val prems = goal Perm.thy
"[| !!x. x:A ==> c(x): B; \
@@ -41,27 +41,27 @@
\ |] ==> (lam x:A.c(x)) : surj(A,B)";
by (res_inst_tac [("d", "d")] f_imp_surjective 1);
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([lam_type]@prems)) ));
-val lam_surjective = result();
+qed "lam_surjective";
(*Cantor's theorem revisited*)
goalw Perm.thy [surj_def] "f ~: surj(A,Pow(A))";
by (safe_tac ZF_cs);
by (cut_facts_tac [cantor] 1);
by (fast_tac subset_cs 1);
-val cantor_surj = result();
+qed "cantor_surj";
(** Injective function space **)
goalw Perm.thy [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B";
by (etac CollectD1 1);
-val inj_is_fun = result();
+qed "inj_is_fun";
goalw Perm.thy [inj_def]
"!!f A B. [| <a,b>:f; <c,b>:f; f: inj(A,B) |] ==> a=c";
by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1));
by (fast_tac ZF_cs 1);
-val inj_equality = result();
+qed "inj_equality";
(** A function with a left inverse is an injection **)
@@ -71,7 +71,7 @@
by (safe_tac ZF_cs);
by (eresolve_tac [subst_context RS box_equals] 1);
by (REPEAT (ares_tac prems 1));
-val f_imp_injective = result();
+qed "f_imp_injective";
val prems = goal Perm.thy
"[| !!x. x:A ==> c(x): B; \
@@ -79,17 +79,17 @@
\ |] ==> (lam x:A.c(x)) : inj(A,B)";
by (res_inst_tac [("d", "d")] f_imp_injective 1);
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([lam_type]@prems)) ));
-val lam_injective = result();
+qed "lam_injective";
(** Bijections **)
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)";
by (etac IntD1 1);
-val bij_is_inj = result();
+qed "bij_is_inj";
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)";
by (etac IntD2 1);
-val bij_is_surj = result();
+qed "bij_is_surj";
(* f: bij(A,B) ==> f: A->B *)
val bij_is_fun = standard (bij_is_inj RS inj_is_fun);
@@ -101,46 +101,46 @@
\ !!y. y:B ==> c(d(y)) = y \
\ |] ==> (lam x:A.c(x)) : bij(A,B)";
by (REPEAT (ares_tac (prems @ [IntI, lam_injective, lam_surjective]) 1));
-val lam_bijective = result();
+qed "lam_bijective";
(** Identity function **)
val [prem] = goalw Perm.thy [id_def] "a:A ==> <a,a> : id(A)";
by (rtac (prem RS lamI) 1);
-val idI = result();
+qed "idI";
val major::prems = goalw Perm.thy [id_def]
"[| p: id(A); !!x.[| x:A; p=<x,x> |] ==> P \
\ |] ==> P";
by (rtac (major RS lamE) 1);
by (REPEAT (ares_tac prems 1));
-val idE = result();
+qed "idE";
goalw Perm.thy [id_def] "id(A) : A->A";
by (rtac lam_type 1);
by (assume_tac 1);
-val id_type = result();
+qed "id_type";
val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)";
by (rtac (prem RS lam_mono) 1);
-val id_mono = result();
+qed "id_mono";
goalw Perm.thy [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 ZF_ss 1);
-val id_subset_inj = result();
+qed "id_subset_inj";
val id_inj = subset_refl RS id_subset_inj;
goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)";
by (fast_tac (ZF_cs addIs [lam_type,beta]) 1);
-val id_surj = result();
+qed "id_surj";
goalw Perm.thy [bij_def] "id(A): bij(A,A)";
by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1);
-val id_bij = result();
+qed "id_bij";
goalw Perm.thy [id_def] "A <= B <-> id(A) : A->B";
by (safe_tac ZF_cs);
@@ -148,7 +148,7 @@
by (dtac apply_type 1);
by (assume_tac 1);
by (asm_full_simp_tac ZF_ss 1);
-val subset_iff_id = result();
+qed "subset_iff_id";
(*** Converse of a function ***)
@@ -160,7 +160,7 @@
by (deepen_tac ZF_cs 0 2);
by (simp_tac (ZF_ss addsimps [function_def, converse_iff]) 1);
by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1);
-val inj_converse_fun = result();
+qed "inj_converse_fun";
(** Equations for converse(f) **)
@@ -168,12 +168,12 @@
val prems = goal Perm.thy
"[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a";
by (fast_tac (ZF_cs addIs (prems@[apply_Pair,apply_equality,converseI])) 1);
-val left_inverse_lemma = result();
+qed "left_inverse_lemma";
goal Perm.thy
"!!f. [| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a";
by (fast_tac (ZF_cs addIs [left_inverse_lemma,inj_converse_fun,inj_is_fun]) 1);
-val left_inverse = result();
+qed "left_inverse";
val left_inverse_bij = bij_is_inj RS left_inverse;
@@ -181,21 +181,21 @@
"[| f: A->B; converse(f): C->A; b: C |] ==> f`(converse(f)`b) = b";
by (rtac (apply_Pair RS (converseD RS apply_equality)) 1);
by (REPEAT (resolve_tac prems 1));
-val right_inverse_lemma = result();
+qed "right_inverse_lemma";
(*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";
by (rtac right_inverse_lemma 1);
by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1));
-val right_inverse = result();
+qed "right_inverse";
goalw Perm.thy [bij_def]
"!!f. [| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b";
by (EVERY1 [etac IntE, etac right_inverse,
etac (surj_range RS ssubst),
assume_tac]);
-val right_inverse_bij = result();
+qed "right_inverse_bij";
(** Converses of injections, surjections, bijections **)
@@ -204,13 +204,13 @@
by (eresolve_tac [inj_converse_fun] 1);
by (resolve_tac [right_inverse] 1);
by (REPEAT (assume_tac 1));
-val inj_converse_inj = result();
+qed "inj_converse_inj";
goal Perm.thy "!!f A B. f: inj(A,B) ==> converse(f): surj(range(f), A)";
by (REPEAT (ares_tac [f_imp_surjective, inj_converse_fun] 1));
by (REPEAT (ares_tac [left_inverse] 2));
by (REPEAT (ares_tac [inj_is_fun, range_of_fun RS apply_type] 1));
-val inj_converse_surj = result();
+qed "inj_converse_surj";
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)";
by (etac IntE 1);
@@ -218,7 +218,7 @@
by (rtac IntI 1);
by (etac inj_converse_inj 1);
by (etac inj_converse_surj 1);
-val bij_converse_bij = result();
+qed "bij_converse_bij";
(** Composition of two relations **)
@@ -227,7 +227,7 @@
goalw Perm.thy [comp_def] "!!r s. [| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
by (fast_tac ZF_cs 1);
-val compI = result();
+qed "compI";
val prems = goalw Perm.thy [comp_def]
"[| xz : r O s; \
@@ -235,7 +235,7 @@
\ |] ==> P";
by (cut_facts_tac prems 1);
by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1));
-val compE = result();
+qed "compE";
val compEpair =
rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac)
@@ -249,56 +249,56 @@
(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*)
goal Perm.thy "range(r O s) <= range(r)";
by (fast_tac comp_cs 1);
-val range_comp = result();
+qed "range_comp";
goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)";
by (rtac (range_comp RS equalityI) 1);
by (fast_tac comp_cs 1);
-val range_comp_eq = result();
+qed "range_comp_eq";
goal Perm.thy "domain(r O s) <= domain(s)";
by (fast_tac comp_cs 1);
-val domain_comp = result();
+qed "domain_comp";
goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)";
by (rtac (domain_comp RS equalityI) 1);
by (fast_tac comp_cs 1);
-val domain_comp_eq = result();
+qed "domain_comp_eq";
goal Perm.thy "(r O s)``A = r``(s``A)";
by (fast_tac (comp_cs addIs [equalityI]) 1);
-val image_comp = result();
+qed "image_comp";
(** Other results **)
goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
by (fast_tac comp_cs 1);
-val comp_mono = result();
+qed "comp_mono";
(*composition preserves relations*)
goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C";
by (fast_tac comp_cs 1);
-val comp_rel = result();
+qed "comp_rel";
(*associative law for composition*)
goal Perm.thy "(r O s) O t = r O (s O t)";
by (fast_tac (comp_cs addIs [equalityI]) 1);
-val comp_assoc = result();
+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";
by (fast_tac (comp_cs addIs [equalityI]) 1);
-val left_comp_id = result();
+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";
by (fast_tac (comp_cs addIs [equalityI]) 1);
-val right_comp_id = result();
+qed "right_comp_id";
(** Composition preserves functions, injections, and surjections **)
@@ -306,7 +306,7 @@
goalw Perm.thy [function_def]
"!!f g. [| function(g); function(f) |] ==> function(f O g)";
by (fast_tac (ZF_cs addIs [compI] addSEs [compE, Pair_inject]) 1);
-val comp_function = result();
+qed "comp_function";
goalw Perm.thy [Pi_def]
"!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C";
@@ -315,12 +315,12 @@
by (rtac (range_rel_subset RS domain_comp_eq RS ssubst) 2 THEN assume_tac 3);
by (fast_tac ZF_cs 2);
by (asm_simp_tac (ZF_ss addsimps [comp_rel]) 1);
-val comp_fun = result();
+qed "comp_fun";
goal Perm.thy "!!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));
-val comp_fun_apply = result();
+qed "comp_fun_apply";
goal Perm.thy "!!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)")]
@@ -328,17 +328,17 @@
by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1));
by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply, left_inverse]
setsolver type_auto_tac [inj_is_fun, apply_type]) 1);
-val comp_inj = result();
+qed "comp_inj";
goalw Perm.thy [surj_def]
"!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)";
by (best_tac (ZF_cs addSIs [comp_fun,comp_fun_apply]) 1);
-val comp_surj = result();
+qed "comp_surj";
goalw Perm.thy [bij_def]
"!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)";
by (fast_tac (ZF_cs addIs [comp_inj,comp_surj]) 1);
-val comp_bij = result();
+qed "comp_bij";
(** Dual properties of inj and surj -- useful for proofs from
@@ -350,7 +350,7 @@
by (safe_tac comp_cs);
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1);
-val comp_mem_injD1 = result();
+qed "comp_mem_injD1";
goalw Perm.thy [inj_def,surj_def]
"!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)";
@@ -362,17 +362,17 @@
by (res_inst_tac [("t", "op `(g)")] subst_context 1);
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1);
-val comp_mem_injD2 = result();
+qed "comp_mem_injD2";
goalw Perm.thy [surj_def]
"!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)";
by (fast_tac (comp_cs addSIs [comp_fun_apply RS sym, apply_type]) 1);
-val comp_mem_surjD1 = result();
+qed "comp_mem_surjD1";
goal Perm.thy
"!!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));
-val comp_fun_applyD = result();
+qed "comp_fun_applyD";
goalw Perm.thy [inj_def,surj_def]
"!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)";
@@ -380,7 +380,7 @@
by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1);
by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_fun_applyD 1));
by (best_tac (comp_cs addSIs [apply_type]) 1);
-val comp_mem_surjD2 = result();
+qed "comp_mem_surjD2";
(** inverses of composition **)
@@ -393,7 +393,7 @@
by (cut_facts_tac [prem RS inj_is_fun] 1);
by (fast_tac (comp_cs addIs [equalityI,apply_Pair]
addEs [domain_type, make_elim injfD]) 1);
-val left_comp_inverse = result();
+qed "left_comp_inverse";
(*right inverse of composition; one inclusion is
f: A->B ==> f O converse(f) <= id(B)
@@ -405,7 +405,7 @@
by (rtac equalityI 1);
by (best_tac (comp_cs addEs [domain_type, range_type, make_elim appfD]) 1);
by (best_tac (comp_cs addIs [apply_Pair]) 1);
-val right_comp_inverse = result();
+qed "right_comp_inverse";
(** Proving that a function is a bijection **)
@@ -418,7 +418,7 @@
by (rtac fun_extension 1);
by (REPEAT (ares_tac [comp_fun, lam_type] 1));
by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1);
-val comp_eq_id_iff = result();
+qed "comp_eq_id_iff";
goalw Perm.thy [bij_def]
"!!f A B. [| f: A->B; g: B->A; f O g = id(B); g O f = id(A) \
@@ -426,16 +426,16 @@
by (asm_full_simp_tac (ZF_ss addsimps [comp_eq_id_iff]) 1);
by (REPEAT (ares_tac [conjI, f_imp_injective, f_imp_surjective] 1
ORELSE eresolve_tac [bspec, apply_type] 1));
-val fg_imp_bijective = result();
+qed "fg_imp_bijective";
goal Perm.thy "!!f A. [| f: A->A; f O f = id(A) |] ==> f : bij(A,A)";
by (REPEAT (ares_tac [fg_imp_bijective] 1));
-val nilpotent_imp_bijective = result();
+qed "nilpotent_imp_bijective";
goal Perm.thy "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)";
by (asm_simp_tac (ZF_ss addsimps [fg_imp_bijective, comp_eq_id_iff,
left_inverse_lemma, right_inverse_lemma]) 1);
-val invertible_imp_bijective = result();
+qed "invertible_imp_bijective";
(** Unions of functions -- cf similar theorems on func.ML **)
@@ -444,7 +444,7 @@
by (DEPTH_SOLVE_1
(resolve_tac [Un_least,converse_mono, Un_upper1,Un_upper2] 2));
by (fast_tac ZF_cs 1);
-val converse_of_Un = result();
+qed "converse_of_Un";
goalw Perm.thy [surj_def]
"!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \
@@ -453,7 +453,7 @@
ORELSE ball_tac 1
ORELSE (rtac trans 1 THEN atac 2)
ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1));
-val surj_disjoint_Un = result();
+qed "surj_disjoint_Un";
(*A simple, high-level proof; the version for injections follows from it,
using f:inj(A,B) <-> f:bij(A,range(f)) *)
@@ -463,7 +463,7 @@
by (rtac invertible_imp_bijective 1);
by (rtac (converse_of_Un RS ssubst) 1);
by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1));
-val bij_disjoint_Un = result();
+qed "bij_disjoint_Un";
(** Restrictions as surjections and bijections *)
@@ -472,7 +472,7 @@
"f: Pi(A,B) ==> f: surj(A, f``A)";
val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]);
by (fast_tac (ZF_cs addIs rls) 1);
-val surj_image = result();
+qed "surj_image";
goal Perm.thy "!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A";
by (rtac equalityI 1);
@@ -480,21 +480,21 @@
by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2
ORELSE ares_tac [subsetI,lamI,imageI] 2));
by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1));
-val restrict_image = result();
+qed "restrict_image";
goalw Perm.thy [inj_def]
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)";
by (safe_tac (ZF_cs addSEs [restrict_type2]));
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD,
box_equals, restrict] 1));
-val restrict_inj = result();
+qed "restrict_inj";
val prems = goal Perm.thy
"[| f: Pi(A,B); C<=A |] ==> restrict(f,C): surj(C, f``C)";
by (rtac (restrict_image RS subst) 1);
by (rtac (restrict_type2 RS surj_image) 3);
by (REPEAT (resolve_tac prems 1));
-val restrict_surj = result();
+qed "restrict_surj";
goalw Perm.thy [inj_def,bij_def]
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)";
@@ -502,14 +502,14 @@
by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD,
box_equals, restrict] 1
ORELSE ares_tac [surj_is_fun,restrict_surj] 1));
-val restrict_bij = result();
+qed "restrict_bij";
(*** Lemmas for Ramsey's Theorem ***)
goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)";
by (fast_tac (ZF_cs addSEs [fun_weaken_type]) 1);
-val inj_weaken_type = result();
+qed "inj_weaken_type";
val [major] = goal Perm.thy
"[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})";
@@ -524,7 +524,7 @@
by (assume_tac 1);
by (res_inst_tac [("a","m")] mem_irrefl 1);
by (fast_tac ZF_cs 1);
-val inj_succ_restrict = result();
+qed "inj_succ_restrict";
goalw Perm.thy [inj_def]
"!!f. [| f: inj(A,B); a~:A; b~:B |] ==> \
@@ -540,4 +540,4 @@
by (ALLGOALS (asm_simp_tac
(FOL_ss addsimps [fun_extend_apply2,fun_extend_apply1])));
by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type])));
-val inj_extend = result();
+qed "inj_extend";
--- a/src/ZF/QPair.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/QPair.ML Wed Dec 07 13:12:04 1994 +0100
@@ -25,17 +25,17 @@
(** Lemmas for showing that <a;b> uniquely determines a and b **)
-val QPair_iff = prove_goalw QPair.thy [QPair_def]
+qed_goalw "QPair_iff" QPair.thy [QPair_def]
"<a;b> = <c;d> <-> a=c & b=d"
(fn _=> [rtac sum_equal_iff 1]);
val QPair_inject = standard (QPair_iff RS iffD1 RS conjE);
-val QPair_inject1 = prove_goal QPair.thy "<a;b> = <c;d> ==> a=c"
+qed_goal "QPair_inject1" QPair.thy "<a;b> = <c;d> ==> a=c"
(fn [major]=>
[ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]);
-val QPair_inject2 = prove_goal QPair.thy "<a;b> = <c;d> ==> b=d"
+qed_goal "QPair_inject2" QPair.thy "<a;b> = <c;d> ==> b=d"
(fn [major]=>
[ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]);
@@ -43,12 +43,12 @@
(*** QSigma: Disjoint union of a family of sets
Generalizes Cartesian product ***)
-val QSigmaI = prove_goalw QPair.thy [QSigma_def]
+qed_goalw "QSigmaI" QPair.thy [QSigma_def]
"[| a:A; b:B(a) |] ==> <a;b> : QSigma(A,B)"
(fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
(*The general elimination rule*)
-val QSigmaE = prove_goalw QPair.thy [QSigma_def]
+qed_goalw "QSigmaE" QPair.thy [QSigma_def]
"[| c: QSigma(A,B); \
\ !!x y.[| x:A; y:B(x); c=<x;y> |] ==> P \
\ |] ==> P"
@@ -63,36 +63,36 @@
THEN prune_params_tac)
(read_instantiate [("c","<a;b>")] QSigmaE);
-val QSigmaD1 = prove_goal QPair.thy "<a;b> : QSigma(A,B) ==> a : A"
+qed_goal "QSigmaD1" QPair.thy "<a;b> : QSigma(A,B) ==> a : A"
(fn [major]=>
[ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]);
-val QSigmaD2 = prove_goal QPair.thy "<a;b> : QSigma(A,B) ==> b : B(a)"
+qed_goal "QSigmaD2" QPair.thy "<a;b> : QSigma(A,B) ==> b : B(a)"
(fn [major]=>
[ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]);
val qpair_cs = ZF_cs addSIs [QSigmaI] addSEs [QSigmaE2, QSigmaE, QPair_inject];
-val QSigma_cong = prove_goalw QPair.thy [QSigma_def]
+qed_goalw "QSigma_cong" QPair.thy [QSigma_def]
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \
\ QSigma(A,B) = QSigma(A',B')"
(fn prems=> [ (simp_tac (ZF_ss addsimps prems) 1) ]);
-val QSigma_empty1 = prove_goal QPair.thy "QSigma(0,B) = 0"
+qed_goal "QSigma_empty1" QPair.thy "QSigma(0,B) = 0"
(fn _ => [ (fast_tac (qpair_cs addIs [equalityI]) 1) ]);
-val QSigma_empty2 = prove_goal QPair.thy "A <*> 0 = 0"
+qed_goal "QSigma_empty2" QPair.thy "A <*> 0 = 0"
(fn _ => [ (fast_tac (qpair_cs addIs [equalityI]) 1) ]);
(*** Eliminator - qsplit ***)
-val qsplit = prove_goalw QPair.thy [qsplit_def]
+qed_goalw "qsplit" QPair.thy [qsplit_def]
"qsplit(%x y.c(x,y), <a;b>) = c(a,b)"
(fn _ => [ (fast_tac (qpair_cs addIs [the_equality]) 1) ]);
-val qsplit_type = prove_goal QPair.thy
+qed_goal "qsplit_type" QPair.thy
"[| p:QSigma(A,B); \
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x;y>) \
\ |] ==> qsplit(%x y.c(x,y), p) : C(p)"
@@ -104,15 +104,15 @@
(*** qconverse ***)
-val qconverseI = prove_goalw QPair.thy [qconverse_def]
+qed_goalw "qconverseI" QPair.thy [qconverse_def]
"!!a b r. <a;b>:r ==> <b;a>:qconverse(r)"
(fn _ => [ (fast_tac qpair_cs 1) ]);
-val qconverseD = prove_goalw QPair.thy [qconverse_def]
+qed_goalw "qconverseD" QPair.thy [qconverse_def]
"!!a b r. <a;b> : qconverse(r) ==> <b;a> : r"
(fn _ => [ (fast_tac qpair_cs 1) ]);
-val qconverseE = prove_goalw QPair.thy [qconverse_def]
+qed_goalw "qconverseE" QPair.thy [qconverse_def]
"[| yx : qconverse(r); \
\ !!x y. [| yx=<y;x>; <x;y>:r |] ==> P \
\ |] ==> P"
@@ -125,18 +125,18 @@
val qconverse_cs = qpair_cs addSIs [qconverseI]
addSEs [qconverseD,qconverseE];
-val qconverse_of_qconverse = prove_goal QPair.thy
+qed_goal "qconverse_of_qconverse" QPair.thy
"!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r"
(fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]);
-val qconverse_type = prove_goal QPair.thy
+qed_goal "qconverse_type" QPair.thy
"!!A B r. r <= A <*> B ==> qconverse(r) <= B <*> A"
(fn _ => [ (fast_tac qconverse_cs 1) ]);
-val qconverse_of_prod = prove_goal QPair.thy "qconverse(A <*> B) = B <*> A"
+qed_goal "qconverse_of_prod" QPair.thy "qconverse(A <*> B) = B <*> A"
(fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]);
-val qconverse_empty = prove_goal QPair.thy "qconverse(0) = 0"
+qed_goal "qconverse_empty" QPair.thy "qconverse(0) = 0"
(fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]);
@@ -144,17 +144,17 @@
goalw QPair.thy [qfsplit_def] "!!R a b. R(a,b) ==> qfsplit(R, <a;b>)";
by (REPEAT (ares_tac [refl,exI,conjI] 1));
-val qfsplitI = result();
+qed "qfsplitI";
val major::prems = goalw QPair.thy [qfsplit_def]
"[| qfsplit(R,z); !!x y. [| z = <x;y>; R(x,y) |] ==> P |] ==> P";
by (cut_facts_tac [major] 1);
by (REPEAT (eresolve_tac (prems@[asm_rl,exE,conjE]) 1));
-val qfsplitE = result();
+qed "qfsplitE";
goal QPair.thy "!!R a b. qfsplit(R,<a;b>) ==> R(a,b)";
by (REPEAT (eresolve_tac [asm_rl,qfsplitE,QPair_inject,ssubst] 1));
-val qfsplitD = result();
+qed "qfsplitD";
(**** The Quine-inspired notion of disjoint sum ****)
@@ -165,11 +165,11 @@
goalw QPair.thy qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B";
by (REPEAT (ares_tac [UnI1,QSigmaI,singletonI] 1));
-val QInlI = result();
+qed "QInlI";
goalw QPair.thy qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B";
by (REPEAT (ares_tac [UnI2,QSigmaI,singletonI] 1));
-val QInrI = result();
+qed "QInrI";
(** Elimination rules **)
@@ -181,25 +181,25 @@
by (rtac (major RS UnE) 1);
by (REPEAT (rtac refl 1
ORELSE eresolve_tac (prems@[QSigmaE,singletonE,ssubst]) 1));
-val qsumE = result();
+qed "qsumE";
(** Injection and freeness equivalences, for rewriting **)
goalw QPair.thy qsum_defs "QInl(a)=QInl(b) <-> a=b";
by (simp_tac (ZF_ss addsimps [QPair_iff]) 1);
-val QInl_iff = result();
+qed "QInl_iff";
goalw QPair.thy qsum_defs "QInr(a)=QInr(b) <-> a=b";
by (simp_tac (ZF_ss addsimps [QPair_iff]) 1);
-val QInr_iff = result();
+qed "QInr_iff";
goalw QPair.thy qsum_defs "QInl(a)=QInr(b) <-> False";
by (simp_tac (ZF_ss addsimps [QPair_iff, one_not_0 RS not_sym]) 1);
-val QInl_QInr_iff = result();
+qed "QInl_QInr_iff";
goalw QPair.thy qsum_defs "QInr(b)=QInl(a) <-> False";
by (simp_tac (ZF_ss addsimps [QPair_iff, one_not_0]) 1);
-val QInr_QInl_iff = result();
+qed "QInr_QInl_iff";
(*Injection and freeness rules*)
@@ -215,39 +215,39 @@
goal QPair.thy "!!A B. QInl(a): A<+>B ==> a: A";
by (fast_tac qsum_cs 1);
-val QInlD = result();
+qed "QInlD";
goal QPair.thy "!!A B. QInr(b): A<+>B ==> b: B";
by (fast_tac qsum_cs 1);
-val QInrD = result();
+qed "QInrD";
(** <+> is itself injective... who cares?? **)
goal QPair.thy
"u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))";
by (fast_tac qsum_cs 1);
-val qsum_iff = result();
+qed "qsum_iff";
goal QPair.thy "A <+> B <= C <+> D <-> A<=C & B<=D";
by (fast_tac qsum_cs 1);
-val qsum_subset_iff = result();
+qed "qsum_subset_iff";
goal QPair.thy "A <+> B = C <+> D <-> A=C & B=D";
by (simp_tac (ZF_ss addsimps [extension,qsum_subset_iff]) 1);
by (fast_tac ZF_cs 1);
-val qsum_equal_iff = result();
+qed "qsum_equal_iff";
(*** Eliminator -- qcase ***)
goalw QPair.thy qsum_defs "qcase(c, d, QInl(a)) = c(a)";
by (rtac (qsplit RS trans) 1);
by (rtac cond_0 1);
-val qcase_QInl = result();
+qed "qcase_QInl";
goalw QPair.thy qsum_defs "qcase(c, d, QInr(b)) = d(b)";
by (rtac (qsplit RS trans) 1);
by (rtac cond_1 1);
-val qcase_QInr = result();
+qed "qcase_QInr";
val major::prems = goal QPair.thy
"[| u: A <+> B; \
@@ -258,22 +258,22 @@
by (ALLGOALS (etac ssubst));
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps
(prems@[qcase_QInl,qcase_QInr]))));
-val qcase_type = result();
+qed "qcase_type";
(** Rules for the Part primitive **)
goal QPair.thy "Part(A <+> B,QInl) = {QInl(x). x: A}";
by (fast_tac (qsum_cs addIs [equalityI]) 1);
-val Part_QInl = result();
+qed "Part_QInl";
goal QPair.thy "Part(A <+> B,QInr) = {QInr(y). y: B}";
by (fast_tac (qsum_cs addIs [equalityI]) 1);
-val Part_QInr = result();
+qed "Part_QInr";
goal QPair.thy "Part(A <+> B, %x.QInr(h(x))) = {QInr(y). y: Part(B,h)}";
by (fast_tac (qsum_cs addIs [equalityI]) 1);
-val Part_QInr2 = result();
+qed "Part_QInr2";
goal QPair.thy "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C";
by (fast_tac (qsum_cs addIs [equalityI]) 1);
-val Part_qsum_equality = result();
+qed "Part_qsum_equality";
--- a/src/ZF/QUniv.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/QUniv.ML Wed Dec 07 13:12:04 1994 +0100
@@ -13,35 +13,35 @@
goalw QUniv.thy [quniv_def]
"!!X A. X <= univ(eclose(A)) ==> X : quniv(A)";
by (etac PowI 1);
-val qunivI = result();
+qed "qunivI";
goalw QUniv.thy [quniv_def]
"!!X A. X : quniv(A) ==> X <= univ(eclose(A))";
by (etac PowD 1);
-val qunivD = result();
+qed "qunivD";
goalw QUniv.thy [quniv_def] "!!A B. A<=B ==> quniv(A) <= quniv(B)";
by (etac (eclose_mono RS univ_mono RS Pow_mono) 1);
-val quniv_mono = result();
+qed "quniv_mono";
(*** Closure properties ***)
goalw QUniv.thy [quniv_def] "univ(eclose(A)) <= quniv(A)";
by (rtac (Transset_iff_Pow RS iffD1) 1);
by (rtac (Transset_eclose RS Transset_univ) 1);
-val univ_eclose_subset_quniv = result();
+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)";
by (rtac (arg_subset_eclose RS univ_mono RS subset_trans) 1);
by (rtac univ_eclose_subset_quniv 1);
-val univ_subset_quniv = result();
+qed "univ_subset_quniv";
val univ_into_quniv = standard (univ_subset_quniv RS subsetD);
goalw QUniv.thy [quniv_def] "Pow(univ(A)) <= quniv(A)";
by (rtac (arg_subset_eclose RS univ_mono RS Pow_mono) 1);
-val Pow_univ_subset_quniv = result();
+qed "Pow_univ_subset_quniv";
val univ_subset_into_quniv = standard
(PowI RS (Pow_univ_subset_quniv RS subsetD));
@@ -61,13 +61,13 @@
goalw QUniv.thy [QPair_def]
"!!A a. [| a <= univ(A); b <= univ(A) |] ==> <a;b> <= univ(A)";
by (REPEAT (ares_tac [sum_subset_univ] 1));
-val QPair_subset_univ = result();
+qed "QPair_subset_univ";
(** Quine disjoint sum **)
goalw QUniv.thy [QInl_def] "!!A a. a <= univ(A) ==> QInl(a) <= univ(A)";
by (etac (empty_subsetI RS QPair_subset_univ) 1);
-val QInl_subset_univ = result();
+qed "QInl_subset_univ";
val naturals_subset_nat =
rewrite_rule [Transset_def] (Ord_nat RS Ord_is_Transset)
@@ -78,7 +78,7 @@
goalw QUniv.thy [QInr_def] "!!A a. a <= univ(A) ==> QInr(a) <= univ(A)";
by (etac (nat_1I RS naturals_subset_univ RS QPair_subset_univ) 1);
-val QInr_subset_univ = result();
+qed "QInr_subset_univ";
(*** Closure for Quine-inspired products and sums ***)
@@ -87,12 +87,12 @@
"!!A a. [| a: quniv(A); b: quniv(A) |] ==> <a;b> : quniv(A)";
by (REPEAT (dtac PowD 1));
by (REPEAT (ares_tac [PowI, sum_subset_univ] 1));
-val QPair_in_quniv = result();
+qed "QPair_in_quniv";
goal QUniv.thy "quniv(A) <*> quniv(A) <= quniv(A)";
by (REPEAT (ares_tac [subsetI, QPair_in_quniv] 1
ORELSE eresolve_tac [QSigmaE, ssubst] 1));
-val QSigma_quniv = result();
+qed "QSigma_quniv";
val QSigma_subset_quniv = standard
(QSigma_mono RS (QSigma_quniv RSN (2,subset_trans)));
@@ -103,30 +103,30 @@
by (etac ([Transset_eclose RS Transset_univ, PowD] MRS
Transset_includes_summands RS conjE) 1);
by (REPEAT (ares_tac [conjI,PowI] 1));
-val quniv_QPair_D = result();
+qed "quniv_QPair_D";
val quniv_QPair_E = standard (quniv_QPair_D RS conjE);
goal QUniv.thy "<a;b> : quniv(A) <-> a: quniv(A) & b: quniv(A)";
by (REPEAT (ares_tac [iffI, QPair_in_quniv, quniv_QPair_D] 1
ORELSE etac conjE 1));
-val quniv_QPair_iff = result();
+qed "quniv_QPair_iff";
(** Quine disjoint sum **)
goalw QUniv.thy [QInl_def] "!!A a. a: quniv(A) ==> QInl(a) : quniv(A)";
by (REPEAT (ares_tac [zero_in_quniv,QPair_in_quniv] 1));
-val QInl_in_quniv = result();
+qed "QInl_in_quniv";
goalw QUniv.thy [QInr_def] "!!A b. b: quniv(A) ==> QInr(b) : quniv(A)";
by (REPEAT (ares_tac [one_in_quniv, QPair_in_quniv] 1));
-val QInr_in_quniv = result();
+qed "QInr_in_quniv";
goal QUniv.thy "quniv(C) <+> quniv(C) <= quniv(C)";
by (REPEAT (ares_tac [subsetI, QInl_in_quniv, QInr_in_quniv] 1
ORELSE eresolve_tac [qsumE, ssubst] 1));
-val qsum_quniv = result();
+qed "qsum_quniv";
val qsum_subset_quniv = standard
(qsum_mono RS (qsum_quniv RSN (2,subset_trans)));
@@ -156,7 +156,7 @@
by (dtac (Transset_Vfrom_succ RS equalityD1 RS subsetD RS PowD) 1);
by (assume_tac 1);
by (fast_tac ZF_cs 1);
-val doubleton_in_Vfrom_D = result();
+qed "doubleton_in_Vfrom_D";
(*This weaker version says a, b exist at the same level*)
val Vfrom_doubleton_D = standard (Transset_Vfrom RS Transset_doubleton_D);
@@ -173,13 +173,13 @@
"!!X. [| <a,b> : Vfrom(X,succ(i)); Transset(X) |] ==> \
\ a: Vfrom(X,i) & b: Vfrom(X,i)";
by (fast_tac (ZF_cs addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1);
-val Pair_in_Vfrom_D = result();
+qed "Pair_in_Vfrom_D";
goal Univ.thy
"!!X. Transset(X) ==> \
\ (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))";
by (fast_tac (ZF_cs addSDs [Pair_in_Vfrom_D]) 1);
-val product_Int_Vfrom_subset = result();
+qed "product_Int_Vfrom_subset";
(*** Intersecting <a;b> with Vfrom... ***)
@@ -190,7 +190,7 @@
by (rtac Un_mono 1);
by (REPEAT (ares_tac [product_Int_Vfrom_subset RS subset_trans,
[Int_lower1, subset_refl] MRS Sigma_mono] 1));
-val QPair_Int_Vfrom_succ_subset = result();
+qed "QPair_Int_Vfrom_succ_subset";
(**** "Take-lemma" rules for proving a=b by coinduction and c: quniv(A) ****)
@@ -200,7 +200,7 @@
"!!X. Transset(X) ==> \
\ <a;b> Int Vfrom(X,i) <= <a Int Vfrom(X,i); b Int Vfrom(X,i)>";
by (etac (Transset_Vfrom RS Transset_sum_Int_subset) 1);
-val QPair_Int_Vfrom_subset = result();
+qed "QPair_Int_Vfrom_subset";
(*[| a Int Vset(i) <= c; b Int Vset(i) <= d |] ==> <a;b> Int Vset(i) <= <c;d>*)
val QPair_Int_Vset_subset_trans = standard
@@ -219,4 +219,4 @@
(*Limit(i) case*)
by (asm_simp_tac (ZF_ss addsimps [Limit_Vfrom_eq, Int_UN_distrib, subset_refl,
UN_mono, QPair_Int_Vset_subset_trans]) 1);
-val QPair_Int_Vset_subset_UN = result();
+qed "QPair_Int_Vset_subset_UN";
--- a/src/ZF/Rel.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Rel.ML Wed Dec 07 13:12:04 1994 +0100
@@ -15,23 +15,23 @@
val prems = goalw Rel.thy [irrefl_def]
"[| !!x. x:A ==> <x,x> ~: r |] ==> irrefl(A,r)";
by (REPEAT (ares_tac (prems @ [ballI]) 1));
-val irreflI = result();
+qed "irreflI";
val prems = goalw Rel.thy [irrefl_def]
"[| irrefl(A,r); x:A |] ==> <x,x> ~: r";
by (rtac (prems MRS bspec) 1);
-val irreflE = result();
+qed "irreflE";
(* symmetry *)
val prems = goalw Rel.thy [sym_def]
"[| !!x y.<x,y>: r ==> <y,x>: r |] ==> sym(r)";
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
-val symI = result();
+qed "symI";
goalw Rel.thy [sym_def] "!!r. [| sym(r); <x,y>: r |] ==> <y,x>: r";
by (fast_tac ZF_cs 1);
-val symE = result();
+qed "symE";
(* antisymmetry *)
@@ -39,22 +39,22 @@
"[| !!x y.[| <x,y>: r; <y,x>: r |] ==> x=y |] ==> \
\ antisym(r)";
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
-val antisymI = result();
+qed "antisymI";
val prems = goalw Rel.thy [antisym_def]
"!!r. [| antisym(r); <x,y>: r; <y,x>: r |] ==> x=y";
by (fast_tac ZF_cs 1);
-val antisymE = result();
+qed "antisymE";
(* transitivity *)
goalw Rel.thy [trans_def]
"!!r. [| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r";
by (fast_tac ZF_cs 1);
-val transD = result();
+qed "transD";
goalw Rel.thy [trans_on_def]
"!!r. [| trans[A](r); <a,b>:r; <b,c>:r; a:A; b:A; c:A |] ==> <a,c>:r";
by (fast_tac ZF_cs 1);
-val trans_onD = result();
+qed "trans_onD";
--- a/src/ZF/Sum.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Sum.ML Wed Dec 07 13:12:04 1994 +0100
@@ -13,12 +13,12 @@
goalw Sum.thy [Part_def]
"a : Part(A,h) <-> a:A & (EX y. a=h(y))";
by (rtac separation 1);
-val Part_iff = result();
+qed "Part_iff";
goalw Sum.thy [Part_def]
"!!a b A h. [| a : A; a=h(b) |] ==> a : Part(A,h)";
by (REPEAT (ares_tac [exI,CollectI] 1));
-val Part_eqI = result();
+qed "Part_eqI";
val PartI = refl RSN (2,Part_eqI);
@@ -28,11 +28,11 @@
by (rtac (major RS CollectE) 1);
by (etac exE 1);
by (REPEAT (ares_tac prems 1));
-val PartE = result();
+qed "PartE";
goalw Sum.thy [Part_def] "Part(A,h) <= A";
by (rtac Collect_subset 1);
-val Part_subset = result();
+qed "Part_subset";
(*** Rules for Disjoint Sums ***)
@@ -41,17 +41,17 @@
goalw Sum.thy (bool_def::sum_defs) "Sigma(bool,C) = C(0) + C(1)";
by (fast_tac eq_cs 1);
-val Sigma_bool = result();
+qed "Sigma_bool";
(** Introduction rules for the injections **)
goalw Sum.thy sum_defs "!!a A B. a : A ==> Inl(a) : A+B";
by (REPEAT (ares_tac [UnI1,SigmaI,singletonI] 1));
-val InlI = result();
+qed "InlI";
goalw Sum.thy sum_defs "!!b A B. b : B ==> Inr(b) : A+B";
by (REPEAT (ares_tac [UnI2,SigmaI,singletonI] 1));
-val InrI = result();
+qed "InrI";
(** Elimination rules **)
@@ -63,25 +63,25 @@
by (rtac (major RS UnE) 1);
by (REPEAT (rtac refl 1
ORELSE eresolve_tac (prems@[SigmaE,singletonE,ssubst]) 1));
-val sumE = result();
+qed "sumE";
(** Injection and freeness equivalences, for rewriting **)
goalw Sum.thy sum_defs "Inl(a)=Inl(b) <-> a=b";
by (simp_tac ZF_ss 1);
-val Inl_iff = result();
+qed "Inl_iff";
goalw Sum.thy sum_defs "Inr(a)=Inr(b) <-> a=b";
by (simp_tac ZF_ss 1);
-val Inr_iff = result();
+qed "Inr_iff";
goalw Sum.thy sum_defs "Inl(a)=Inr(b) <-> False";
by (simp_tac (ZF_ss addsimps [one_not_0 RS not_sym]) 1);
-val Inl_Inr_iff = result();
+qed "Inl_Inr_iff";
goalw Sum.thy sum_defs "Inr(b)=Inl(a) <-> False";
by (simp_tac (ZF_ss addsimps [one_not_0]) 1);
-val Inr_Inl_iff = result();
+qed "Inr_Inl_iff";
(*Injection and freeness rules*)
@@ -99,24 +99,24 @@
goal Sum.thy "!!A B. Inl(a): A+B ==> a: A";
by (fast_tac sum_cs 1);
-val InlD = result();
+qed "InlD";
goal Sum.thy "!!A B. Inr(b): A+B ==> b: B";
by (fast_tac sum_cs 1);
-val InrD = result();
+qed "InrD";
goal Sum.thy "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))";
by (fast_tac sum_cs 1);
-val sum_iff = result();
+qed "sum_iff";
goal Sum.thy "A+B <= C+D <-> A<=C & B<=D";
by (fast_tac sum_cs 1);
-val sum_subset_iff = result();
+qed "sum_subset_iff";
goal Sum.thy "A+B = C+D <-> A=C & B=D";
by (simp_tac (ZF_ss addsimps [extension,sum_subset_iff]) 1);
by (fast_tac ZF_cs 1);
-val sum_equal_iff = result();
+qed "sum_equal_iff";
(*** Eliminator -- case ***)
@@ -124,12 +124,12 @@
goalw Sum.thy sum_defs "case(c, d, Inl(a)) = c(a)";
by (rtac (split RS trans) 1);
by (rtac cond_0 1);
-val case_Inl = result();
+qed "case_Inl";
goalw Sum.thy sum_defs "case(c, d, Inr(b)) = d(b)";
by (rtac (split RS trans) 1);
by (rtac cond_1 1);
-val case_Inr = result();
+qed "case_Inr";
val major::prems = goal Sum.thy
"[| u: A+B; \
@@ -140,7 +140,7 @@
by (ALLGOALS (etac ssubst));
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps
(prems@[case_Inl,case_Inr]))));
-val case_type = result();
+qed "case_type";
goal Sum.thy
"!!u. u: A+B ==> \
@@ -152,40 +152,39 @@
by (fast_tac sum_cs 1);
by (asm_simp_tac (ZF_ss addsimps [case_Inr]) 1);
by (fast_tac sum_cs 1);
-val expand_case = result();
-
+qed "expand_case";
goal Sum.thy "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)";
by (fast_tac sum_cs 1);
-val Part_mono = result();
+qed "Part_mono";
goal Sum.thy "Part(Collect(A,P), h) = Collect(Part(A,h), P)";
by (fast_tac (sum_cs addSIs [equalityI]) 1);
-val Part_Collect = result();
+qed "Part_Collect";
val Part_CollectE =
Part_Collect RS equalityD1 RS subsetD RS CollectE |> standard;
goal Sum.thy "Part(A+B,Inl) = {Inl(x). x: A}";
by (fast_tac (sum_cs addIs [equalityI]) 1);
-val Part_Inl = result();
+qed "Part_Inl";
goal Sum.thy "Part(A+B,Inr) = {Inr(y). y: B}";
by (fast_tac (sum_cs addIs [equalityI]) 1);
-val Part_Inr = result();
+qed "Part_Inr";
goalw Sum.thy [Part_def] "!!a. a : Part(A,h) ==> a : A";
by (etac CollectD1 1);
-val PartD1 = result();
+qed "PartD1";
goal Sum.thy "Part(A,%x.x) = A";
by (fast_tac (sum_cs addIs [equalityI]) 1);
-val Part_id = result();
+qed "Part_id";
goal Sum.thy "Part(A+B, %x.Inr(h(x))) = {Inr(y). y: Part(B,h)}";
by (fast_tac (sum_cs addIs [equalityI]) 1);
-val Part_Inr2 = result();
+qed "Part_Inr2";
goal Sum.thy "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C";
by (fast_tac (sum_cs addIs [equalityI]) 1);
-val Part_sum_equality = result();
+qed "Part_sum_equality";
--- a/src/ZF/Trancl.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Trancl.ML Wed Dec 07 13:12:04 1994 +0100
@@ -12,13 +12,13 @@
by (rtac bnd_monoI 1);
by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2));
by (fast_tac comp_cs 1);
-val rtrancl_bnd_mono = result();
+qed "rtrancl_bnd_mono";
val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*";
by (rtac lfp_mono 1);
by (REPEAT (resolve_tac [rtrancl_bnd_mono, prem, subset_refl, id_mono,
comp_mono, Un_mono, field_mono, Sigma_mono] 1));
-val rtrancl_mono = result();
+qed "rtrancl_mono";
(* r^* = id(field(r)) Un ( r O r^* ) *)
val rtrancl_unfold = rtrancl_bnd_mono RS (rtrancl_def RS def_lfp_Tarski);
@@ -31,7 +31,7 @@
val [prem] = goal Trancl.thy "[| a: field(r) |] ==> <a,a> : r^*";
by (resolve_tac [rtrancl_unfold RS ssubst] 1);
by (rtac (prem RS idI RS UnI1) 1);
-val rtrancl_refl = result();
+qed "rtrancl_refl";
(*Closure under composition with r *)
val prems = goal Trancl.thy
@@ -40,24 +40,24 @@
by (rtac (compI RS UnI2) 1);
by (resolve_tac prems 1);
by (resolve_tac prems 1);
-val rtrancl_into_rtrancl = result();
+qed "rtrancl_into_rtrancl";
(*rtrancl of r contains all pairs in r *)
val prems = goal Trancl.thy "<a,b> : r ==> <a,b> : r^*";
by (resolve_tac [rtrancl_refl RS rtrancl_into_rtrancl] 1);
by (REPEAT (resolve_tac (prems@[fieldI1]) 1));
-val r_into_rtrancl = result();
+qed "r_into_rtrancl";
(*The premise ensures that r consists entirely of pairs*)
val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^*";
by (cut_facts_tac prems 1);
by (fast_tac (ZF_cs addIs [r_into_rtrancl]) 1);
-val r_subset_rtrancl = result();
+qed "r_subset_rtrancl";
goal Trancl.thy "field(r^*) = field(r)";
by (fast_tac (eq_cs addIs [r_into_rtrancl]
addSDs [rtrancl_type RS subsetD]) 1);
-val rtrancl_field = result();
+qed "rtrancl_field";
(** standard induction rule **)
@@ -69,7 +69,7 @@
\ ==> P(<a,b>)";
by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1);
by (fast_tac (ZF_cs addIs prems addSEs [idE,compE]) 1);
-val rtrancl_full_induct = result();
+qed "rtrancl_full_induct";
(*nice induction rule.
Tried adding the typing hypotheses y,z:field(r), but these
@@ -86,14 +86,14 @@
(*now do the induction*)
by (resolve_tac [major RS rtrancl_full_induct] 1);
by (ALLGOALS (fast_tac (ZF_cs addIs prems)));
-val rtrancl_induct = result();
+qed "rtrancl_induct";
(*transitivity of transitive closure!! -- by induction.*)
goalw Trancl.thy [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));
-val trans_rtrancl = result();
+qed "trans_rtrancl";
(*elimination of rtrancl -- by induction on a special formula*)
val major::prems = goal Trancl.thy
@@ -104,7 +104,7 @@
(*see HOL/trancl*)
by (rtac (major RS rtrancl_induct) 2);
by (ALLGOALS (fast_tac (ZF_cs addSEs prems)));
-val rtranclE = result();
+qed "rtranclE";
(**** The relation trancl ****)
@@ -114,31 +114,31 @@
by (safe_tac comp_cs);
by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1);
by (REPEAT (assume_tac 1));
-val trans_trancl = result();
+qed "trans_trancl";
(** Conversions between trancl and rtrancl **)
val [major] = goalw Trancl.thy [trancl_def] "<a,b> : r^+ ==> <a,b> : r^*";
by (resolve_tac [major RS compEpair] 1);
by (REPEAT (ares_tac [rtrancl_into_rtrancl] 1));
-val trancl_into_rtrancl = result();
+qed "trancl_into_rtrancl";
(*r^+ contains all pairs in r *)
val [prem] = goalw Trancl.thy [trancl_def] "<a,b> : r ==> <a,b> : r^+";
by (REPEAT (ares_tac [prem,compI,rtrancl_refl,fieldI1] 1));
-val r_into_trancl = result();
+qed "r_into_trancl";
(*The premise ensures that r consists entirely of pairs*)
val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^+";
by (cut_facts_tac prems 1);
by (fast_tac (ZF_cs addIs [r_into_trancl]) 1);
-val r_subset_trancl = result();
+qed "r_subset_trancl";
(*intro rule by definition: from r^* and r *)
val prems = goalw Trancl.thy [trancl_def]
"[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^+";
by (REPEAT (resolve_tac ([compI]@prems) 1));
-val rtrancl_into_trancl1 = result();
+qed "rtrancl_into_trancl1";
(*intro rule from r and r^* *)
val prems = goal Trancl.thy
@@ -147,7 +147,7 @@
by (resolve_tac (prems RL [r_into_trancl]) 1);
by (etac (trans_trancl RS transD) 1);
by (etac r_into_trancl 1);
-val rtrancl_into_trancl2 = result();
+qed "rtrancl_into_trancl2";
(*Nice induction rule for trancl*)
val major::prems = goal Trancl.thy
@@ -162,7 +162,7 @@
by (fast_tac ZF_cs 1);
by (etac rtrancl_induct 1);
by (ALLGOALS (fast_tac (ZF_cs addIs (rtrancl_into_trancl1::prems))));
-val trancl_induct = result();
+qed "trancl_induct";
(*elimination of r^+ -- NOT an induction rule*)
val major::prems = goal Trancl.thy
@@ -175,13 +175,13 @@
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1);
by (etac rtranclE 1);
by (ALLGOALS (fast_tac (ZF_cs addIs [rtrancl_into_trancl1])));
-val tranclE = result();
+qed "tranclE";
goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)";
by (fast_tac (ZF_cs addEs [compE, rtrancl_type RS subsetD RS SigmaE2]) 1);
-val trancl_type = result();
+qed "trancl_type";
val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+";
by (REPEAT (resolve_tac [prem, comp_mono, rtrancl_mono] 1));
-val trancl_mono = result();
+qed "trancl_mono";
--- a/src/ZF/Univ.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Univ.ML Wed Dec 07 13:12:04 1994 +0100
@@ -12,7 +12,7 @@
goal Univ.thy "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))";
by (rtac (Vfrom_def RS def_transrec RS ssubst) 1);
by (simp_tac ZF_ss 1);
-val Vfrom = result();
+qed "Vfrom";
(** Monotonicity **)
@@ -28,7 +28,7 @@
by (etac (bspec RS spec RS mp) 1);
by (assume_tac 1);
by (rtac subset_refl 1);
-val Vfrom_mono_lemma = result();
+qed "Vfrom_mono_lemma";
(* [| A<=B; i<=x |] ==> Vfrom(A,i) <= Vfrom(B,x) *)
val Vfrom_mono = standard (Vfrom_mono_lemma RS spec RS mp);
@@ -41,7 +41,7 @@
by (rtac (Vfrom RS ssubst) 1);
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac (ZF_cs addSIs [rank_lt RS ltD]) 1);
-val Vfrom_rank_subset1 = result();
+qed "Vfrom_rank_subset1";
goal Univ.thy "Vfrom(A,rank(x)) <= Vfrom(A,x)";
by (eps_ind_tac "x" 1);
@@ -58,13 +58,13 @@
by (rtac (Ord_rank RS Ord_succ) 1);
by (etac bspec 1);
by (assume_tac 1);
-val Vfrom_rank_subset2 = result();
+qed "Vfrom_rank_subset2";
goal Univ.thy "Vfrom(A,rank(x)) = Vfrom(A,x)";
by (rtac equalityI 1);
by (rtac Vfrom_rank_subset2 1);
by (rtac Vfrom_rank_subset1 1);
-val Vfrom_rank_eq = result();
+qed "Vfrom_rank_eq";
(*** Basic closure properties ***)
@@ -72,58 +72,58 @@
goal Univ.thy "!!x y. y:x ==> 0 : Vfrom(A,x)";
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac ZF_cs 1);
-val zero_in_Vfrom = result();
+qed "zero_in_Vfrom";
goal Univ.thy "i <= Vfrom(A,i)";
by (eps_ind_tac "i" 1);
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac ZF_cs 1);
-val i_subset_Vfrom = result();
+qed "i_subset_Vfrom";
goal Univ.thy "A <= Vfrom(A,i)";
by (rtac (Vfrom RS ssubst) 1);
by (rtac Un_upper1 1);
-val A_subset_Vfrom = result();
+qed "A_subset_Vfrom";
val A_into_Vfrom = A_subset_Vfrom RS subsetD |> standard;
goal Univ.thy "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))";
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac ZF_cs 1);
-val subset_mem_Vfrom = result();
+qed "subset_mem_Vfrom";
(** Finite sets and ordered pairs **)
goal Univ.thy "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))";
by (rtac subset_mem_Vfrom 1);
by (safe_tac ZF_cs);
-val singleton_in_Vfrom = result();
+qed "singleton_in_Vfrom";
goal Univ.thy
"!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))";
by (rtac subset_mem_Vfrom 1);
by (safe_tac ZF_cs);
-val doubleton_in_Vfrom = result();
+qed "doubleton_in_Vfrom";
goalw Univ.thy [Pair_def]
"!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> \
\ <a,b> : Vfrom(A,succ(succ(i)))";
by (REPEAT (ares_tac [doubleton_in_Vfrom] 1));
-val Pair_in_Vfrom = result();
+qed "Pair_in_Vfrom";
val [prem] = goal Univ.thy
"a<=Vfrom(A,i) ==> succ(a) : Vfrom(A,succ(succ(i)))";
by (REPEAT (resolve_tac [subset_mem_Vfrom, succ_subsetI] 1));
by (rtac (Vfrom_mono RSN (2,subset_trans)) 2);
by (REPEAT (resolve_tac [prem, subset_refl, subset_succI] 1));
-val succ_in_Vfrom = result();
+qed "succ_in_Vfrom";
(*** 0, successor and limit equations fof Vfrom ***)
goal Univ.thy "Vfrom(A,0) = A";
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac eq_cs 1);
-val Vfrom_0 = result();
+qed "Vfrom_0";
goal Univ.thy "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))";
by (rtac (Vfrom RS trans) 1);
@@ -133,14 +133,14 @@
by (rtac (subset_refl RS Vfrom_mono RS Pow_mono) 1);
by (etac (ltI RS le_imp_subset) 1);
by (etac Ord_succ 1);
-val Vfrom_succ_lemma = result();
+qed "Vfrom_succ_lemma";
goal Univ.thy "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 (rtac (rank_succ RS ssubst) 1);
by (rtac (Ord_rank RS Vfrom_succ_lemma) 1);
-val Vfrom_succ = result();
+qed "Vfrom_succ";
(*The premise distinguishes this from Vfrom(A,0); allowing X=0 forces
the conclusion to be Vfrom(A,Union(X)) = A Un (UN y:X. Vfrom(A,y)) *)
@@ -161,11 +161,11 @@
by (rtac UN_least 1);
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac ZF_cs 1);
-val Vfrom_Union = result();
+qed "Vfrom_Union";
goal Univ.thy "!!i. Ord(i) ==> i=0 | (EX j. i=succ(j)) | Limit(i)";
by (fast_tac (ZF_cs addSIs [non_succ_LimitI, Ord_0_lt]) 1);
-val Ord_cases_lemma = result();
+qed "Ord_cases_lemma";
val major::prems = goal Univ.thy
"[| Ord(i); \
@@ -175,7 +175,7 @@
\ |] ==> P";
by (cut_facts_tac [major RS Ord_cases_lemma] 1);
by (REPEAT (eresolve_tac (prems@[disjE, exE]) 1));
-val Ord_cases = result();
+qed "Ord_cases";
(*** Vfrom applied to Limit ordinals ***)
@@ -187,12 +187,12 @@
by (rtac (limiti RS (Limit_has_0 RS ltD) RS Vfrom_Union RS subst) 1);
by (rtac (limiti RS Limit_Union_eq RS ssubst) 1);
by (rtac refl 1);
-val Limit_Vfrom_eq = result();
+qed "Limit_Vfrom_eq";
goal Univ.thy "!!a. [| a: Vfrom(A,j); Limit(i); j<i |] ==> a : Vfrom(A,i)";
by (rtac (Limit_Vfrom_eq RS equalityD2 RS subsetD) 1);
by (REPEAT (ares_tac [ltD RS UN_I] 1));
-val Limit_VfromI = result();
+qed "Limit_VfromI";
val prems = goal Univ.thy
"[| a: Vfrom(A,i); Limit(i); \
@@ -200,7 +200,7 @@
\ |] ==> R";
by (rtac (Limit_Vfrom_eq RS equalityD1 RS subsetD RS UN_E) 1);
by (REPEAT (ares_tac (prems @ [ltI, Limit_is_Ord]) 1));
-val Limit_VfromE = result();
+qed "Limit_VfromE";
val zero_in_VLimit = Limit_has_0 RS ltD RS zero_in_Vfrom;
@@ -209,7 +209,7 @@
by (rtac ([major,limiti] MRS Limit_VfromE) 1);
by (etac ([singleton_in_Vfrom, limiti] MRS Limit_VfromI) 1);
by (etac (limiti RS Limit_has_succ) 1);
-val singleton_in_VLimit = result();
+qed "singleton_in_VLimit";
val Vfrom_UnI1 = Un_upper1 RS (subset_refl RS Vfrom_mono RS subsetD)
and Vfrom_UnI2 = Un_upper2 RS (subset_refl RS Vfrom_mono RS subsetD);
@@ -224,7 +224,7 @@
by (etac Vfrom_UnI1 1);
by (etac Vfrom_UnI2 1);
by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1));
-val doubleton_in_VLimit = result();
+qed "doubleton_in_VLimit";
val [aprem,bprem,limiti] = goal Univ.thy
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i) |] ==> \
@@ -237,12 +237,12 @@
by (etac Vfrom_UnI1 1);
by (etac Vfrom_UnI2 1);
by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1));
-val Pair_in_VLimit = result();
+qed "Pair_in_VLimit";
goal Univ.thy "!!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));
-val product_VLimit = result();
+qed "product_VLimit";
val Sigma_subset_VLimit =
[Sigma_mono, product_VLimit] MRS subset_trans |> standard;
@@ -253,7 +253,7 @@
goal Univ.thy "!!i. [| n: nat; Limit(i) |] ==> n : Vfrom(A,i)";
by (REPEAT (ares_tac [nat_subset_VLimit RS subsetD] 1));
-val nat_into_VLimit = result();
+qed "nat_into_VLimit";
(** Closure under disjoint union **)
@@ -261,21 +261,21 @@
goal Univ.thy "!!i. Limit(i) ==> 1 : Vfrom(A,i)";
by (REPEAT (ares_tac [nat_into_VLimit, nat_0I, nat_succI] 1));
-val one_in_VLimit = result();
+qed "one_in_VLimit";
goalw Univ.thy [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));
-val Inl_in_VLimit = result();
+qed "Inl_in_VLimit";
goalw Univ.thy [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));
-val Inr_in_VLimit = result();
+qed "Inr_in_VLimit";
goal Univ.thy "!!i. Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)";
by (fast_tac (sum_cs addSIs [Inl_in_VLimit, Inr_in_VLimit]) 1);
-val sum_VLimit = result();
+qed "sum_VLimit";
val sum_subset_VLimit =
[sum_mono, sum_VLimit] MRS subset_trans |> standard;
@@ -289,7 +289,7 @@
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac (ZF_cs addSIs [Transset_Union_family, Transset_Un,
Transset_Pow]) 1);
-val Transset_Vfrom = result();
+qed "Transset_Vfrom";
goal Univ.thy "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))";
by (rtac (Vfrom_succ RS trans) 1);
@@ -297,12 +297,12 @@
by (rtac (subset_refl RSN (2,Un_least)) 1);
by (rtac (A_subset_Vfrom RS subset_trans) 1);
by (etac (Transset_Vfrom RS (Transset_iff_Pow RS iffD1)) 1);
-val Transset_Vfrom_succ = result();
+qed "Transset_Vfrom_succ";
goalw Ordinal.thy [Pair_def,Transset_def]
"!!C. [| <a,b> <= C; Transset(C) |] ==> a: C & b: C";
by (fast_tac ZF_cs 1);
-val Transset_Pair_subset = result();
+qed "Transset_Pair_subset";
goal Univ.thy
"!!a b.[| <a,b> <= Vfrom(A,i); Transset(A); Limit(i) |] ==> \
@@ -310,7 +310,7 @@
by (etac (Transset_Pair_subset RS conjE) 1);
by (etac Transset_Vfrom 1);
by (REPEAT (ares_tac [Pair_in_VLimit] 1));
-val Transset_Pair_subset_VLimit = result();
+qed "Transset_Pair_subset_VLimit";
(*** Closure under product/sum applied to elements -- thus Vfrom(A,i)
@@ -333,7 +333,7 @@
by (etac (Vfrom_UnI1 RS Vfrom_UnI1) 3);
by (rtac (succI1 RS UnI2) 2);
by (REPEAT (ares_tac [limiti, Limit_has_0, Limit_has_succ, Un_least_lt] 1));
-val in_VLimit = result();
+qed "in_VLimit";
(** products **)
@@ -344,7 +344,7 @@
by (rtac subset_mem_Vfrom 1);
by (rewtac Transset_def);
by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1);
-val prod_in_Vfrom = result();
+qed "prod_in_Vfrom";
val [aprem,bprem,limiti,transset] = goal Univ.thy
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \
@@ -352,7 +352,7 @@
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1);
by (REPEAT (ares_tac [exI, conjI, prod_in_Vfrom, transset,
limiti RS Limit_has_succ] 1));
-val prod_in_VLimit = result();
+qed "prod_in_VLimit";
(** Disjoint sums, aka Quine ordered pairs **)
@@ -364,7 +364,7 @@
by (rewtac Transset_def);
by (fast_tac (ZF_cs addIs [zero_in_Vfrom, Pair_in_Vfrom,
i_subset_Vfrom RS subsetD]) 1);
-val sum_in_Vfrom = result();
+qed "sum_in_Vfrom";
val [aprem,bprem,limiti,transset] = goal Univ.thy
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \
@@ -372,7 +372,7 @@
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1);
by (REPEAT (ares_tac [exI, conjI, sum_in_Vfrom, transset,
limiti RS Limit_has_succ] 1));
-val sum_in_VLimit = result();
+qed "sum_in_VLimit";
(** function space! **)
@@ -389,7 +389,7 @@
by (rtac Pow_mono 1);
by (rewtac Transset_def);
by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1);
-val fun_in_Vfrom = result();
+qed "fun_in_Vfrom";
val [aprem,bprem,limiti,transset] = goal Univ.thy
"[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); Transset(A) |] ==> \
@@ -397,7 +397,7 @@
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1);
by (REPEAT (ares_tac [exI, conjI, fun_in_Vfrom, transset,
limiti RS Limit_has_succ] 1));
-val fun_in_VLimit = result();
+qed "fun_in_VLimit";
(*** The set Vset(i) ***)
@@ -405,7 +405,7 @@
goal Univ.thy "Vset(i) = (UN j:i. Pow(Vset(j)))";
by (rtac (Vfrom RS ssubst) 1);
by (fast_tac eq_cs 1);
-val Vset = result();
+qed "Vset";
val Vset_succ = Transset_0 RS Transset_Vfrom_succ;
@@ -421,7 +421,7 @@
by (rtac UN_succ_least_lt 1);
by (fast_tac ZF_cs 2);
by (REPEAT (ares_tac [ltI] 1));
-val Vset_rank_imp1 = result();
+qed "Vset_rank_imp1";
(* [| Ord(i); x : Vset(i) |] ==> rank(x) < i *)
val VsetD = standard (Vset_rank_imp1 RS spec RS mp);
@@ -431,23 +431,23 @@
by (rtac allI 1);
by (rtac (Vset RS ssubst) 1);
by (fast_tac (ZF_cs addSIs [rank_lt RS ltD]) 1);
-val Vset_rank_imp2 = result();
+qed "Vset_rank_imp2";
goal Univ.thy "!!x i. rank(x)<i ==> x : Vset(i)";
by (etac ltE 1);
by (etac (Vset_rank_imp2 RS spec RS mp) 1);
by (assume_tac 1);
-val VsetI = result();
+qed "VsetI";
goal Univ.thy "!!i. Ord(i) ==> b : Vset(i) <-> rank(b) < i";
by (rtac iffI 1);
by (REPEAT (eresolve_tac [asm_rl, VsetD, VsetI] 1));
-val Vset_Ord_rank_iff = result();
+qed "Vset_Ord_rank_iff";
goal Univ.thy "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);
-val Vset_rank_iff = result();
+qed "Vset_rank_iff";
goal Univ.thy "!!i. Ord(i) ==> rank(Vset(i)) = i";
by (rtac (rank RS ssubst) 1);
@@ -459,21 +459,21 @@
assume_tac,
rtac succI1] 3);
by (REPEAT (eresolve_tac [asm_rl, VsetD RS ltD, Ord_trans] 1));
-val rank_Vset = result();
+qed "rank_Vset";
(** Lemmas for reasoning about sets in terms of their elements' ranks **)
goal Univ.thy "a <= Vset(rank(a))";
by (rtac subsetI 1);
by (etac (rank_lt RS VsetI) 1);
-val arg_subset_Vset_rank = result();
+qed "arg_subset_Vset_rank";
val [iprem] = goal Univ.thy
"[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b";
by (rtac ([subset_refl, arg_subset_Vset_rank] MRS
Int_greatest RS subset_trans) 1);
by (rtac (Ord_rank RS iprem) 1);
-val Int_Vset_subset = result();
+qed "Int_Vset_subset";
(** Set up an environment for simplification **)
@@ -491,7 +491,7 @@
by (rtac (transrec RS ssubst) 1);
by (simp_tac (ZF_ss addsimps [Ord_rank, Ord_succ, VsetD RS ltD RS beta,
VsetI RS beta, le_refl]) 1);
-val Vrec = result();
+qed "Vrec";
(*This form avoids giant explosions in proofs. NOTE USE OF == *)
val rew::prems = goal Univ.thy
@@ -499,7 +499,7 @@
\ h(a) = H(a, lam x: Vset(rank(a)). h(x))";
by (rewtac rew);
by (rtac Vrec 1);
-val def_Vrec = result();
+qed "def_Vrec";
(*** univ(A) ***)
@@ -507,22 +507,22 @@
goalw Univ.thy [univ_def] "!!A B. A<=B ==> univ(A) <= univ(B)";
by (etac Vfrom_mono 1);
by (rtac subset_refl 1);
-val univ_mono = result();
+qed "univ_mono";
goalw Univ.thy [univ_def] "!!A. Transset(A) ==> Transset(univ(A))";
by (etac Transset_Vfrom 1);
-val Transset_univ = result();
+qed "Transset_univ";
(** univ(A) as a limit **)
goalw Univ.thy [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))";
by (rtac (Limit_nat RS Limit_Vfrom_eq) 1);
-val univ_eq_UN = result();
+qed "univ_eq_UN";
goal Univ.thy "!!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);
-val subset_univ_eq_Int = result();
+qed "subset_univ_eq_Int";
val [aprem, iprem] = goal Univ.thy
"[| a <= univ(X); \
@@ -531,7 +531,7 @@
by (rtac (aprem RS subset_univ_eq_Int RS ssubst) 1);
by (rtac UN_least 1);
by (etac iprem 1);
-val univ_Int_Vfrom_subset = result();
+qed "univ_Int_Vfrom_subset";
val prems = goal Univ.thy
"[| a <= univ(X); b <= univ(X); \
@@ -542,17 +542,17 @@
(resolve_tac (prems RL [univ_Int_Vfrom_subset]) THEN'
eresolve_tac (prems RL [equalityD1,equalityD2] RL [subset_trans]) THEN'
rtac Int_lower1));
-val univ_Int_Vfrom_eq = result();
+qed "univ_Int_Vfrom_eq";
(** Closure properties **)
goalw Univ.thy [univ_def] "0 : univ(A)";
by (rtac (nat_0I RS zero_in_Vfrom) 1);
-val zero_in_univ = result();
+qed "zero_in_univ";
goalw Univ.thy [univ_def] "A <= univ(A)";
by (rtac A_subset_Vfrom 1);
-val A_subset_univ = result();
+qed "A_subset_univ";
val A_into_univ = A_subset_univ RS subsetD;
@@ -560,28 +560,28 @@
goalw Univ.thy [univ_def] "!!A a. a: univ(A) ==> {a} : univ(A)";
by (REPEAT (ares_tac [singleton_in_VLimit, Limit_nat] 1));
-val singleton_in_univ = result();
+qed "singleton_in_univ";
goalw Univ.thy [univ_def]
"!!A a. [| a: univ(A); b: univ(A) |] ==> {a,b} : univ(A)";
by (REPEAT (ares_tac [doubleton_in_VLimit, Limit_nat] 1));
-val doubleton_in_univ = result();
+qed "doubleton_in_univ";
goalw Univ.thy [univ_def]
"!!A a. [| a: univ(A); b: univ(A) |] ==> <a,b> : univ(A)";
by (REPEAT (ares_tac [Pair_in_VLimit, Limit_nat] 1));
-val Pair_in_univ = result();
+qed "Pair_in_univ";
goalw Univ.thy [univ_def] "univ(A)*univ(A) <= univ(A)";
by (rtac (Limit_nat RS product_VLimit) 1);
-val product_univ = result();
+qed "product_univ";
(** The natural numbers **)
goalw Univ.thy [univ_def] "nat <= univ(A)";
by (rtac i_subset_Vfrom 1);
-val nat_subset_univ = result();
+qed "nat_subset_univ";
(* n:nat ==> n:univ(A) *)
val nat_into_univ = standard (nat_subset_univ RS subsetD);
@@ -590,16 +590,16 @@
goalw Univ.thy [univ_def] "1 : univ(A)";
by (rtac (Limit_nat RS one_in_VLimit) 1);
-val one_in_univ = result();
+qed "one_in_univ";
(*unused!*)
goal Univ.thy "succ(1) : univ(A)";
by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1));
-val two_in_univ = result();
+qed "two_in_univ";
goalw Univ.thy [bool_def] "bool <= univ(A)";
by (fast_tac (ZF_cs addSIs [zero_in_univ,one_in_univ]) 1);
-val bool_subset_univ = result();
+qed "bool_subset_univ";
val bool_into_univ = standard (bool_subset_univ RS subsetD);
@@ -608,15 +608,15 @@
goalw Univ.thy [univ_def] "!!A a. a: univ(A) ==> Inl(a) : univ(A)";
by (etac (Limit_nat RSN (2,Inl_in_VLimit)) 1);
-val Inl_in_univ = result();
+qed "Inl_in_univ";
goalw Univ.thy [univ_def] "!!A b. b: univ(A) ==> Inr(b) : univ(A)";
by (etac (Limit_nat RSN (2,Inr_in_VLimit)) 1);
-val Inr_in_univ = result();
+qed "Inr_in_univ";
goalw Univ.thy [univ_def] "univ(C)+univ(C) <= univ(C)";
by (rtac (Limit_nat RS sum_VLimit) 1);
-val sum_univ = result();
+qed "sum_univ";
val sum_subset_univ = [sum_mono, sum_univ] MRS subset_trans |> standard;
--- a/src/ZF/WF.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/WF.ML Wed Dec 07 13:12:04 1994 +0100
@@ -26,23 +26,23 @@
goalw WF.thy [wf_def, wf_on_def] "!!A r. wf(r) ==> wf[A](r)";
by (fast_tac ZF_cs 1);
-val wf_imp_wf_on = result();
+qed "wf_imp_wf_on";
goalw WF.thy [wf_def, wf_on_def] "!!r. wf[field(r)](r) ==> wf(r)";
by (fast_tac ZF_cs 1);
-val wf_on_field_imp_wf = result();
+qed "wf_on_field_imp_wf";
goal WF.thy "wf(r) <-> wf[field(r)](r)";
by (fast_tac (ZF_cs addSEs [wf_imp_wf_on, wf_on_field_imp_wf]) 1);
-val wf_iff_wf_on_field = result();
+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)";
by (fast_tac ZF_cs 1);
-val wf_on_subset_A = result();
+qed "wf_on_subset_A";
goalw WF.thy [wf_on_def, wf_def] "!!A r s. [| wf[A](r); s<=r |] ==> wf[A](s)";
by (fast_tac ZF_cs 1);
-val wf_on_subset_r = result();
+qed "wf_on_subset_r";
(** Introduction rules for wf_on **)
@@ -53,7 +53,7 @@
by (rtac (equals0I RS disjCI RS allI) 1);
by (res_inst_tac [ ("Z", "Z") ] prem 1);
by (ALLGOALS (fast_tac ZF_cs));
-val wf_onI = result();
+qed "wf_onI";
(*If r allows well-founded induction over A then wf[A](r)
Premise is equivalent to
@@ -67,7 +67,7 @@
by (contr_tac 3);
by (fast_tac ZF_cs 2);
by (fast_tac ZF_cs 1);
-val wf_onI2 = result();
+qed "wf_onI2";
(** Well-founded Induction **)
@@ -84,7 +84,7 @@
by (etac bexE 1);
by (dtac minor 1);
by (fast_tac ZF_cs 1);
-val wf_induct = result();
+qed "wf_induct";
(*Perform induction on i, then prove the wf(r) subgoal using prems. *)
fun wf_ind_tac a prems i =
@@ -102,11 +102,11 @@
by (rtac impI 1);
by (eresolve_tac prems 1);
by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1);
-val wf_induct2 = result();
+qed "wf_induct2";
goal ZF.thy "!!r A. field(r Int A*A) <= A";
by (fast_tac ZF_cs 1);
-val field_Int_square = result();
+qed "field_Int_square";
val wfr::amem::prems = goalw WF.thy [wf_on_def]
"[| wf[A](r); a:A; \
@@ -115,7 +115,7 @@
by (rtac ([wfr, amem, field_Int_square] MRS wf_induct2) 1);
by (REPEAT (ares_tac prems 1));
by (fast_tac ZF_cs 1);
-val wf_on_induct = result();
+qed "wf_on_induct";
fun wf_on_ind_tac a prems i =
EVERY [res_inst_tac [("a",a)] wf_on_induct i,
@@ -130,7 +130,7 @@
\ ==> wf(r)";
by (rtac ([wf_onI2, subs] MRS (wf_on_subset_A RS wf_on_field_imp_wf)) 1);
by (REPEAT (ares_tac [indhyp] 1));
-val wfI = result();
+qed "wfI";
(*** Properties of well-founded relations ***)
@@ -138,26 +138,26 @@
goal WF.thy "!!r. wf(r) ==> <a,a> ~: r";
by (wf_ind_tac "a" [] 1);
by (fast_tac ZF_cs 1);
-val wf_not_refl = result();
+qed "wf_not_refl";
goal WF.thy "!!r. [| wf(r); <a,x>:r; <x,a>:r |] ==> P";
by (subgoal_tac "ALL x. <a,x>:r --> <x,a>:r --> P" 1);
by (wf_ind_tac "a" [] 2);
by (fast_tac ZF_cs 2);
by (fast_tac FOL_cs 1);
-val wf_asym = result();
+qed "wf_asym";
goal WF.thy "!!r. [| wf[A](r); a: A |] ==> <a,a> ~: r";
by (wf_on_ind_tac "a" [] 1);
by (fast_tac ZF_cs 1);
-val wf_on_not_refl = result();
+qed "wf_on_not_refl";
goal WF.thy "!!r. [| wf[A](r); <a,b>:r; <b,a>:r; a:A; b:A |] ==> P";
by (subgoal_tac "ALL y:A. <a,y>:r --> <y,a>:r --> P" 1);
by (wf_on_ind_tac "a" [] 2);
by (fast_tac ZF_cs 2);
by (fast_tac ZF_cs 1);
-val wf_on_asym = result();
+qed "wf_on_asym";
(*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 *)
@@ -168,7 +168,7 @@
by (wf_on_ind_tac "a" [] 2);
by (fast_tac ZF_cs 2);
by (fast_tac ZF_cs 1);
-val wf_on_chain3 = result();
+qed "wf_on_chain3";
(*retains the universal formula for later use!*)
@@ -187,14 +187,14 @@
by (cut_facts_tac [subs] 1);
(*astar_tac is slightly faster*)
by (best_tac ZF_cs 1);
-val wf_on_trancl = result();
+qed "wf_on_trancl";
goal WF.thy "!!r. wf(r) ==> wf(r^+)";
by (asm_full_simp_tac (ZF_ss 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);
by (fast_tac ZF_cs 1);
-val wf_trancl = result();
+qed "wf_trancl";
@@ -210,13 +210,13 @@
by (rtac (major RS ssubst) 1);
by (rtac (lamI RS rangeI RS lam_type) 1);
by (assume_tac 1);
-val is_recfun_type = result();
+qed "is_recfun_type";
val [isrec,rel] = goalw WF.thy [is_recfun_def]
"[| is_recfun(r,a,H,f); <x,a>:r |] ==> f`x = H(x, restrict(f,r-``{x}))";
by (res_inst_tac [("P", "%x.?t(x) = (?u::i)")] (isrec RS ssubst) 1);
by (rtac (rel RS underI RS beta) 1);
-val apply_recfun = result();
+qed "apply_recfun";
(*eresolve_tac transD solves <a,b>:r using transitivity AT MOST ONCE
spec RS mp instantiates induction hypotheses*)
@@ -237,7 +237,7 @@
by (REPEAT (rtac impI 1 ORELSE etac ssubst 1));
by (rewtac restrict_def);
by (asm_simp_tac (wf_super_ss addsimps [vimage_singleton_iff]) 1);
-val is_recfun_equal_lemma = result();
+qed "is_recfun_equal_lemma";
val is_recfun_equal = standard (is_recfun_equal_lemma RS mp RS mp);
val prems as [wfr,transr,recf,recg,_] = goal WF.thy
@@ -250,7 +250,7 @@
by (ALLGOALS
(asm_simp_tac (wf_super_ss addsimps
[ [wfr,transr,recf,recg] MRS is_recfun_equal ])));
-val is_recfun_cut = result();
+qed "is_recfun_cut";
(*** Main Existence Lemma ***)
@@ -260,7 +260,7 @@
by (rtac fun_extension 1);
by (REPEAT (ares_tac [is_recfun_equal] 1
ORELSE eresolve_tac [is_recfun_type,underD] 1));
-val is_recfun_functional = result();
+qed "is_recfun_functional";
(*If some f satisfies is_recfun(r,a,H,-) then so does the_recfun(r,a,H) *)
val prems = goalw WF.thy [the_recfun_def]
@@ -268,7 +268,7 @@
\ ==> is_recfun(r, a, H, the_recfun(r,a,H))";
by (rtac (ex1I RS theI) 1);
by (REPEAT (ares_tac (prems@[is_recfun_functional]) 1));
-val is_the_recfun = result();
+qed "is_the_recfun";
val prems = goal WF.thy
"[| wf(r); trans(r) |] ==> is_recfun(r, a, H, the_recfun(r,a,H))";
@@ -285,7 +285,7 @@
by (ALLGOALS
(asm_simp_tac
(wf_super_ss addsimps [underI RS beta, apply_recfun, is_recfun_cut])));
-val unfold_the_recfun = result();
+qed "unfold_the_recfun";
(*** Unfolding wftrec ***)
@@ -294,7 +294,7 @@
"[| wf(r); trans(r); <b,a>:r |] ==> \
\ restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)";
by (REPEAT (ares_tac (prems @ [is_recfun_cut, unfold_the_recfun]) 1));
-val the_recfun_cut = result();
+qed "the_recfun_cut";
(*NOT SUITABLE FOR REWRITING since it is recursive!*)
goalw WF.thy [wftrec_def]
@@ -303,7 +303,7 @@
by (rtac (rewrite_rule [is_recfun_def] unfold_the_recfun RS ssubst) 1);
by (ALLGOALS (asm_simp_tac
(ZF_ss addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut])));
-val wftrec = result();
+qed "wftrec";
(** Removal of the premise trans(r) **)
@@ -315,7 +315,7 @@
by (rtac (vimage_pair_mono RS restrict_lam_eq RS subst_context) 1);
by (etac r_into_trancl 1);
by (rtac subset_refl 1);
-val wfrec = result();
+qed "wfrec";
(*This form avoids giant explosions in proofs. NOTE USE OF == *)
val rew::prems = goal WF.thy
@@ -323,7 +323,7 @@
\ h(a) = H(a, lam x: r-``{a}. h(x))";
by (rewtac rew);
by (REPEAT (resolve_tac (prems@[wfrec]) 1));
-val def_wfrec = result();
+qed "def_wfrec";
val prems = goal WF.thy
"[| wf(r); a:A; field(r)<=A; \
@@ -333,7 +333,7 @@
by (rtac (wfrec RS ssubst) 4);
by (REPEAT (ares_tac (prems@[lam_type]) 1
ORELSE eresolve_tac [spec RS mp, underD] 1));
-val wfrec_type = result();
+qed "wfrec_type";
goalw WF.thy [wf_on_def, wfrec_on_def]
@@ -341,5 +341,5 @@
\ wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))";
by (etac (wfrec RS trans) 1);
by (asm_simp_tac (ZF_ss addsimps [vimage_Int_square, cons_subset_iff]) 1);
-val wfrec_on = result();
+qed "wfrec_on";
--- a/src/ZF/Zorn.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Zorn.ML Wed Dec 07 13:12:04 1994 +0100
@@ -15,12 +15,12 @@
goal ZF.thy "!!A B C. (ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)";
by (fast_tac ZF_cs 1);
-val Union_lemma0 = result();
+qed "Union_lemma0";
goal ZF.thy
"!!A B C. [| c:C; ALL x:C. A<=x | x<=B |] ==> A<=Inter(C) | Inter(C)<=B";
by (fast_tac ZF_cs 1);
-val Inter_lemma0 = result();
+qed "Inter_lemma0";
(*** Section 2. The Transfinite Construction ***)
@@ -28,13 +28,13 @@
goalw Zorn.thy [increasing_def]
"!!f A. f: increasing(A) ==> f: Pow(A)->Pow(A)";
by (eresolve_tac [CollectD1] 1);
-val increasingD1 = result();
+qed "increasingD1";
goalw Zorn.thy [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);
-val increasingD2 = result();
+qed "increasingD2";
(*Introduction rules*)
val [TFin_nextI, Pow_TFin_UnionI] = TFin.intrs;
@@ -52,7 +52,7 @@
\ |] ==> P(n)";
by (rtac (major RS TFin.induct) 1);
by (ALLGOALS (fast_tac (ZF_cs addIs prems)));
-val TFin_induct = result();
+qed "TFin_induct";
(*Perform induction on n, then prove the major premise using prems. *)
fun TFin_ind_tac a prems i =
@@ -75,7 +75,7 @@
br (major RS TFin_induct) 1;
by (etac Union_lemma0 2); (*or just fast_tac ZF_cs*)
by (fast_tac (subset_cs addIs [increasing_trans]) 1);
-val TFin_linear_lemma1 = result();
+qed "TFin_linear_lemma1";
(*Lemma 2 of section 3.2. Interesting in its own right!
Requires next: increasing(S) in the second induction step. *)
@@ -103,7 +103,7 @@
by (fast_tac subset_cs 1);
br (ninc RS increasingD2 RS subset_trans RS disjI1) 1;
by (REPEAT (ares_tac [TFin_is_subset] 1));
-val TFin_linear_lemma2 = result();
+qed "TFin_linear_lemma2";
(*a more convenient form for Lemma 2*)
goal Zorn.thy
@@ -111,7 +111,7 @@
\ |] ==> n=m | next`n<=m";
br (TFin_linear_lemma2 RS bspec RS mp) 1;
by (REPEAT (assume_tac 1));
-val TFin_subsetD = result();
+qed "TFin_subsetD";
(*Consequences from section 3.3 -- Property 3.2, the ordering is total*)
goal Zorn.thy
@@ -121,7 +121,7 @@
by (REPEAT (assume_tac 1) THEN etac disjI2 1);
by (fast_tac (subset_cs addIs [increasingD2 RS subset_trans,
TFin_is_subset]) 1);
-val TFin_subset_linear = result();
+qed "TFin_subset_linear";
(*Lemma 3 of section 3.3*)
@@ -133,7 +133,7 @@
by (REPEAT (assume_tac 1));
by (fast_tac (ZF_cs addEs [ssubst]) 1);
by (fast_tac (subset_cs addIs [TFin_is_subset]) 1);
-val equal_next_upper = result();
+qed "equal_next_upper";
(*Property 3.3 of section 3.3*)
goal Zorn.thy
@@ -147,7 +147,7 @@
by (rtac (increasingD2 RS equalityI) 1 THEN assume_tac 1);
by (ALLGOALS
(fast_tac (subset_cs addIs [TFin_UnionI, TFin_nextI, TFin_is_subset])));
-val equal_next_Union = result();
+qed "equal_next_Union";
(*** Section 4. Hausdorff's Theorem: every set contains a maximal chain ***)
@@ -157,15 +157,15 @@
goalw Zorn.thy [chain_def] "chain(A) <= Pow(A)";
by (resolve_tac [Collect_subset] 1);
-val chain_subset_Pow = result();
+qed "chain_subset_Pow";
goalw Zorn.thy [super_def] "super(A,c) <= chain(A)";
by (resolve_tac [Collect_subset] 1);
-val super_subset_chain = result();
+qed "super_subset_chain";
goalw Zorn.thy [maxchain_def] "maxchain(A) <= chain(A)";
by (resolve_tac [Collect_subset] 1);
-val maxchain_subset_chain = result();
+qed "maxchain_subset_chain";
goal Zorn.thy
"!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \
@@ -174,7 +174,7 @@
by (eresolve_tac [apply_type] 1);
by (rewrite_goals_tac [super_def, maxchain_def]);
by (fast_tac ZF_cs 1);
-val choice_super = result();
+qed "choice_super";
goal Zorn.thy
"!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \
@@ -185,7 +185,7 @@
by (assume_tac 1);
by (assume_tac 1);
by (asm_full_simp_tac (ZF_ss addsimps [super_def]) 1);
-val choice_not_equals = result();
+qed "choice_not_equals";
(*This justifies Definition 4.4*)
goal Zorn.thy
@@ -213,7 +213,7 @@
by (REPEAT (assume_tac 1));
bw super_def;
by (fast_tac ZF_cs 1);
-val Hausdorff_next_exists = result();
+qed "Hausdorff_next_exists";
(*Lemma 4*)
goal Zorn.thy
@@ -236,7 +236,7 @@
(*fast_tac is just too slow here!*)
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subsetD] 1
ORELSE ball_tac 1 THEN etac (CollectD2 RS bspec RS bspec) 1));
-val TFin_chain_lemma4 = result();
+qed "TFin_chain_lemma4";
goal Zorn.thy "EX c. c : maxchain(S)";
by (rtac (AC_Pi_Pow RS exE) 1);
@@ -258,7 +258,7 @@
setloop split_tac [expand_if]) 1);
by (eresolve_tac [choice_not_equals RS notE] 1);
by (REPEAT (assume_tac 1));
-val Hausdorff = result();
+qed "Hausdorff";
(*** Section 5. Zorn's Lemma: if all chains in S have upper bounds in S
@@ -268,7 +268,7 @@
goalw Zorn.thy [chain_def]
"!!c. [| c: chain(A); z: A; ALL x:c. x<=z |] ==> cons(z,c) : chain(A)";
by (fast_tac ZF_cs 1);
-val chain_extend = result();
+qed "chain_extend";
goal Zorn.thy
"!!S. ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z";
@@ -286,7 +286,7 @@
by (safe_tac eq_cs);
by (fast_tac (ZF_cs addEs [chain_extend]) 1);
by (best_tac (ZF_cs addEs [equalityE]) 1);
-val Zorn = result();
+qed "Zorn";
(*** Section 6. Zermelo's Theorem: every set can be well-ordered ***)
@@ -304,7 +304,7 @@
by (subgoal_tac "x = Inter(Z)" 1);
by (fast_tac ZF_cs 1);
by (fast_tac eq_cs 1);
-val TFin_well_lemma5 = result();
+qed "TFin_well_lemma5";
(*Well-ordering of TFin(S,next)*)
goal Zorn.thy "!!Z. [| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z";
@@ -315,7 +315,7 @@
br (Union_upper RS equalityI) 1;
br (subset_refl RS TFin_UnionI RS TFin_well_lemma5 RS bspec) 2;
by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD]));
-val well_ord_TFin_lemma = result();
+qed "well_ord_TFin_lemma";
(*This theorem just packages the previous result*)
goal Zorn.thy
@@ -337,7 +337,7 @@
by (forward_tac [well_ord_TFin_lemma] 1 THEN assume_tac 1);
by (dres_inst_tac [("x","Inter(Z)")] bspec 1 THEN assume_tac 1);
by (fast_tac eq_cs 1);
-val well_ord_TFin = result();
+qed "well_ord_TFin";
(** Defining the "next" operation for Zermelo's Theorem **)
@@ -346,7 +346,7 @@
\ |] ==> ch ` (S-X) : S-X";
by (eresolve_tac [apply_type] 1);
by (fast_tac (eq_cs addEs [equalityE]) 1);
-val choice_Diff = result();
+qed "choice_Diff";
(*This justifies Definition 6.1*)
goal Zorn.thy
@@ -369,7 +369,7 @@
by (asm_simp_tac (ZF_ss addsimps [Pow_iff, cons_subset_iff, subset_refl]
setloop split_tac [expand_if]) 1);
by (fast_tac (ZF_cs addSIs [choice_Diff RS DiffD1]) 1);
-val Zermelo_next_exists = result();
+qed "Zermelo_next_exists";
(*The construction of the injection*)
@@ -405,7 +405,7 @@
Pow_iff, cons_subset_iff, subset_refl]
setloop split_tac [expand_if]) 1);
by (REPEAT (eresolve_tac [asm_rl, consE, sym, notE] 1));
-val choice_imp_injection = result();
+qed "choice_imp_injection";
(*The wellordering theorem*)
goal Zorn.thy "EX r. well_ord(S,r)";
@@ -417,4 +417,4 @@
by (eresolve_tac [well_ord_TFin] 2);
by (resolve_tac [choice_imp_injection RS inj_weaken_type] 1);
by (REPEAT (ares_tac [Diff_subset] 1));
-val AC_well_ord = result();
+qed "AC_well_ord";
--- a/src/ZF/domrange.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/domrange.ML Wed Dec 07 13:12:04 1994 +0100
@@ -8,19 +8,19 @@
(*** converse ***)
-val converse_iff = prove_goalw ZF.thy [converse_def]
+qed_goalw "converse_iff" ZF.thy [converse_def]
"<a,b>: converse(r) <-> <b,a>:r"
(fn _ => [ (fast_tac pair_cs 1) ]);
-val converseI = prove_goalw ZF.thy [converse_def]
+qed_goalw "converseI" ZF.thy [converse_def]
"!!a b r. <a,b>:r ==> <b,a>:converse(r)"
(fn _ => [ (fast_tac pair_cs 1) ]);
-val converseD = prove_goalw ZF.thy [converse_def]
+qed_goalw "converseD" ZF.thy [converse_def]
"!!a b r. <a,b> : converse(r) ==> <b,a> : r"
(fn _ => [ (fast_tac pair_cs 1) ]);
-val converseE = prove_goalw ZF.thy [converse_def]
+qed_goalw "converseE" ZF.thy [converse_def]
"[| yx : converse(r); \
\ !!x y. [| yx=<y,x>; <x,y>:r |] ==> P \
\ |] ==> P"
@@ -33,51 +33,51 @@
val converse_cs = pair_cs addSIs [converseI]
addSEs [converseD,converseE];
-val converse_of_converse = prove_goal ZF.thy
+qed_goal "converse_of_converse" ZF.thy
"!!A B r. r<=Sigma(A,B) ==> converse(converse(r)) = r"
(fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]);
-val converse_type = prove_goal ZF.thy "!!A B r. r<=A*B ==> converse(r)<=B*A"
+qed_goal "converse_type" ZF.thy "!!A B r. r<=A*B ==> converse(r)<=B*A"
(fn _ => [ (fast_tac converse_cs 1) ]);
-val converse_of_prod = prove_goal ZF.thy "converse(A*B) = B*A"
+qed_goal "converse_of_prod" ZF.thy "converse(A*B) = B*A"
(fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]);
-val converse_empty = prove_goal ZF.thy "converse(0) = 0"
+qed_goal "converse_empty" ZF.thy "converse(0) = 0"
(fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]);
(*** domain ***)
-val domain_iff = prove_goalw ZF.thy [domain_def]
+qed_goalw "domain_iff" ZF.thy [domain_def]
"a: domain(r) <-> (EX y. <a,y>: r)"
(fn _=> [ (fast_tac pair_cs 1) ]);
-val domainI = prove_goal ZF.thy "!!a b r. <a,b>: r ==> a: domain(r)"
+qed_goal "domainI" ZF.thy "!!a b r. <a,b>: r ==> a: domain(r)"
(fn _ => [ (etac (exI RS (domain_iff RS iffD2)) 1) ]);
-val domainE = prove_goal ZF.thy
+qed_goal "domainE" ZF.thy
"[| a : domain(r); !!y. <a,y>: r ==> P |] ==> P"
(fn prems=>
[ (rtac (domain_iff RS iffD1 RS exE) 1),
(REPEAT (ares_tac prems 1)) ]);
-val domain_subset = prove_goal ZF.thy "domain(Sigma(A,B)) <= A"
+qed_goal "domain_subset" ZF.thy "domain(Sigma(A,B)) <= A"
(fn _ => [ (rtac subsetI 1), (etac domainE 1), (etac SigmaD1 1) ]);
(*** range ***)
-val rangeI = prove_goalw ZF.thy [range_def] "!!a b r.<a,b>: r ==> b : range(r)"
+qed_goalw "rangeI" ZF.thy [range_def] "!!a b r.<a,b>: r ==> b : range(r)"
(fn _ => [ (etac (converseI RS domainI) 1) ]);
-val rangeE = prove_goalw ZF.thy [range_def]
+qed_goalw "rangeE" ZF.thy [range_def]
"[| b : range(r); !!x. <x,b>: r ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS domainE) 1),
(resolve_tac prems 1),
(etac converseD 1) ]);
-val range_subset = prove_goalw ZF.thy [range_def] "range(A*B) <= B"
+qed_goalw "range_subset" ZF.thy [range_def] "range(A*B) <= B"
(fn _ =>
[ (rtac (converse_of_prod RS ssubst) 1),
(rtac domain_subset 1) ]);
@@ -85,22 +85,22 @@
(*** field ***)
-val fieldI1 = prove_goalw ZF.thy [field_def] "<a,b>: r ==> a : field(r)"
+qed_goalw "fieldI1" ZF.thy [field_def] "<a,b>: r ==> a : field(r)"
(fn [prem]=>
[ (rtac (prem RS domainI RS UnI1) 1) ]);
-val fieldI2 = prove_goalw ZF.thy [field_def] "<a,b>: r ==> b : field(r)"
+qed_goalw "fieldI2" ZF.thy [field_def] "<a,b>: r ==> b : field(r)"
(fn [prem]=>
[ (rtac (prem RS rangeI RS UnI2) 1) ]);
-val fieldCI = prove_goalw ZF.thy [field_def]
+qed_goalw "fieldCI" ZF.thy [field_def]
"(~ <c,a>:r ==> <a,b>: r) ==> a : field(r)"
(fn [prem]=>
[ (rtac (prem RS domainI RS UnCI) 1),
(swap_res_tac [rangeI] 1),
(etac notnotD 1) ]);
-val fieldE = prove_goalw ZF.thy [field_def]
+qed_goalw "fieldE" ZF.thy [field_def]
"[| a : field(r); \
\ !!x. <a,x>: r ==> P; \
\ !!x. <x,a>: r ==> P |] ==> P"
@@ -108,74 +108,74 @@
[ (rtac (major RS UnE) 1),
(REPEAT (eresolve_tac (prems@[domainE,rangeE]) 1)) ]);
-val field_subset = prove_goal ZF.thy "field(A*B) <= A Un B"
+qed_goal "field_subset" ZF.thy "field(A*B) <= A Un B"
(fn _ => [ (fast_tac (pair_cs addIs [fieldCI] addSEs [fieldE]) 1) ]);
-val domain_subset_field = prove_goalw ZF.thy [field_def]
+qed_goalw "domain_subset_field" ZF.thy [field_def]
"domain(r) <= field(r)"
(fn _ => [ (rtac Un_upper1 1) ]);
-val range_subset_field = prove_goalw ZF.thy [field_def]
+qed_goalw "range_subset_field" ZF.thy [field_def]
"range(r) <= field(r)"
(fn _ => [ (rtac Un_upper2 1) ]);
-val domain_times_range = prove_goal ZF.thy
+qed_goal "domain_times_range" ZF.thy
"!!A B r. r <= Sigma(A,B) ==> r <= domain(r)*range(r)"
(fn _ => [ (fast_tac (pair_cs addIs [domainI,rangeI]) 1) ]);
-val field_times_field = prove_goal ZF.thy
+qed_goal "field_times_field" ZF.thy
"!!A B r. r <= Sigma(A,B) ==> r <= field(r)*field(r)"
(fn _ => [ (fast_tac (pair_cs addIs [fieldI1,fieldI2]) 1) ]);
(*** Image of a set under a function/relation ***)
-val image_iff = prove_goalw ZF.thy [image_def]
+qed_goalw "image_iff" ZF.thy [image_def]
"b : r``A <-> (EX x:A. <x,b>:r)"
(fn _ => [ fast_tac (pair_cs addIs [rangeI]) 1 ]);
-val image_singleton_iff = prove_goal ZF.thy
+qed_goal "image_singleton_iff" ZF.thy
"b : r``{a} <-> <a,b>:r"
(fn _ => [ rtac (image_iff RS iff_trans) 1,
fast_tac pair_cs 1 ]);
-val imageI = prove_goalw ZF.thy [image_def]
+qed_goalw "imageI" ZF.thy [image_def]
"!!a b r. [| <a,b>: r; a:A |] ==> b : r``A"
(fn _ => [ (REPEAT (ares_tac [CollectI,rangeI,bexI] 1)) ]);
-val imageE = prove_goalw ZF.thy [image_def]
+qed_goalw "imageE" ZF.thy [image_def]
"[| b: r``A; !!x.[| <x,b>: r; x:A |] ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS CollectE) 1),
(REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]);
-val image_subset = prove_goal ZF.thy
+qed_goal "image_subset" ZF.thy
"!!A B r. r <= A*B ==> r``C <= B"
(fn _ => [ (fast_tac (pair_cs addSEs [imageE]) 1) ]);
(*** Inverse image of a set under a function/relation ***)
-val vimage_iff = prove_goalw ZF.thy [vimage_def,image_def,converse_def]
+qed_goalw "vimage_iff" ZF.thy [vimage_def,image_def,converse_def]
"a : r-``B <-> (EX y:B. <a,y>:r)"
(fn _ => [ fast_tac (pair_cs addIs [rangeI]) 1 ]);
-val vimage_singleton_iff = prove_goal ZF.thy
+qed_goal "vimage_singleton_iff" ZF.thy
"a : r-``{b} <-> <a,b>:r"
(fn _ => [ rtac (vimage_iff RS iff_trans) 1,
fast_tac pair_cs 1 ]);
-val vimageI = prove_goalw ZF.thy [vimage_def]
+qed_goalw "vimageI" ZF.thy [vimage_def]
"!!A B r. [| <a,b>: r; b:B |] ==> a : r-``B"
(fn _ => [ (REPEAT (ares_tac [converseI RS imageI] 1)) ]);
-val vimageE = prove_goalw ZF.thy [vimage_def]
+qed_goalw "vimageE" ZF.thy [vimage_def]
"[| a: r-``B; !!x.[| <a,x>: r; x:B |] ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS imageE) 1),
(REPEAT (etac converseD 1 ORELSE ares_tac prems 1)) ]);
-val vimage_subset = prove_goalw ZF.thy [vimage_def]
+qed_goalw "vimage_subset" ZF.thy [vimage_def]
"!!A B r. r <= A*B ==> r-``C <= A"
(fn _ => [ (etac (converse_type RS image_subset) 1) ]);
@@ -193,10 +193,10 @@
goal ZF.thy "!!S. (ALL x:S. EX A B. x <= A*B) ==> \
\ Union(S) <= domain(Union(S)) * range(Union(S))";
by (fast_tac ZF_cs 1);
-val rel_Union = result();
+qed "rel_Union";
(** The Union of 2 relations is a relation (Lemma for fun_Un) **)
-val rel_Un = prove_goal ZF.thy
+qed_goal "rel_Un" ZF.thy
"!!r s. [| r <= A*B; s <= C*D |] ==> (r Un s) <= (A Un C) * (B Un D)"
(fn _ => [ (fast_tac ZF_cs 1) ]);
--- a/src/ZF/equalities.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/equalities.ML Wed Dec 07 13:12:04 1994 +0100
@@ -12,23 +12,23 @@
(* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*)
goal ZF.thy "{a} Un B = cons(a,B)";
by (fast_tac eq_cs 1);
-val cons_eq = result();
+qed "cons_eq";
goal ZF.thy "cons(a, cons(b, C)) = cons(b, cons(a, C))";
by (fast_tac eq_cs 1);
-val cons_commute = result();
+qed "cons_commute";
goal ZF.thy "!!B. a: B ==> cons(a,B) = B";
by (fast_tac eq_cs 1);
-val cons_absorb = result();
+qed "cons_absorb";
goal ZF.thy "!!B. a: B ==> cons(a, B-{a}) = B";
by (fast_tac eq_cs 1);
-val cons_Diff = result();
+qed "cons_Diff";
goal ZF.thy "!!C. [| a: C; ALL y:C. y=b |] ==> C = {b}";
by (fast_tac eq_cs 1);
-val equal_singleton_lemma = result();
+qed "equal_singleton_lemma";
val equal_singleton = ballI RSN (2,equal_singleton_lemma);
@@ -36,443 +36,443 @@
goal ZF.thy "0 Int A = 0";
by (fast_tac eq_cs 1);
-val Int_0 = result();
+qed "Int_0";
(*NOT an equality, but it seems to belong here...*)
goal ZF.thy "cons(a,B) Int C <= cons(a, B Int C)";
by (fast_tac eq_cs 1);
-val Int_cons = result();
+qed "Int_cons";
goal ZF.thy "A Int A = A";
by (fast_tac eq_cs 1);
-val Int_absorb = result();
+qed "Int_absorb";
goal ZF.thy "A Int B = B Int A";
by (fast_tac eq_cs 1);
-val Int_commute = result();
+qed "Int_commute";
goal ZF.thy "(A Int B) Int C = A Int (B Int C)";
by (fast_tac eq_cs 1);
-val Int_assoc = result();
+qed "Int_assoc";
goal ZF.thy "(A Un B) Int C = (A Int C) Un (B Int C)";
by (fast_tac eq_cs 1);
-val Int_Un_distrib = result();
+qed "Int_Un_distrib";
goal ZF.thy "A<=B <-> A Int B = A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val subset_Int_iff = result();
+qed "subset_Int_iff";
goal ZF.thy "A<=B <-> B Int A = A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val subset_Int_iff2 = result();
+qed "subset_Int_iff2";
(** Binary Union **)
goal ZF.thy "0 Un A = A";
by (fast_tac eq_cs 1);
-val Un_0 = result();
+qed "Un_0";
goal ZF.thy "cons(a,B) Un C = cons(a, B Un C)";
by (fast_tac eq_cs 1);
-val Un_cons = result();
+qed "Un_cons";
goal ZF.thy "A Un A = A";
by (fast_tac eq_cs 1);
-val Un_absorb = result();
+qed "Un_absorb";
goal ZF.thy "A Un B = B Un A";
by (fast_tac eq_cs 1);
-val Un_commute = result();
+qed "Un_commute";
goal ZF.thy "(A Un B) Un C = A Un (B Un C)";
by (fast_tac eq_cs 1);
-val Un_assoc = result();
+qed "Un_assoc";
goal ZF.thy "(A Int B) Un C = (A Un C) Int (B Un C)";
by (fast_tac eq_cs 1);
-val Un_Int_distrib = result();
+qed "Un_Int_distrib";
goal ZF.thy "A<=B <-> A Un B = B";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val subset_Un_iff = result();
+qed "subset_Un_iff";
goal ZF.thy "A<=B <-> B Un A = B";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val subset_Un_iff2 = result();
+qed "subset_Un_iff2";
(** Simple properties of Diff -- set difference **)
goal ZF.thy "A-A = 0";
by (fast_tac eq_cs 1);
-val Diff_cancel = result();
+qed "Diff_cancel";
goal ZF.thy "0-A = 0";
by (fast_tac eq_cs 1);
-val empty_Diff = result();
+qed "empty_Diff";
goal ZF.thy "A-0 = A";
by (fast_tac eq_cs 1);
-val Diff_0 = result();
+qed "Diff_0";
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
goal ZF.thy "A - cons(a,B) = A - B - {a}";
by (fast_tac eq_cs 1);
-val Diff_cons = result();
+qed "Diff_cons";
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
goal ZF.thy "A - cons(a,B) = A - {a} - B";
by (fast_tac eq_cs 1);
-val Diff_cons2 = result();
+qed "Diff_cons2";
goal ZF.thy "A Int (B-A) = 0";
by (fast_tac eq_cs 1);
-val Diff_disjoint = result();
+qed "Diff_disjoint";
goal ZF.thy "!!A B. A<=B ==> A Un (B-A) = B";
by (fast_tac eq_cs 1);
-val Diff_partition = result();
+qed "Diff_partition";
goal ZF.thy "!!A B. [| A<=B; B<=C |] ==> B-(C-A) = A";
by (fast_tac eq_cs 1);
-val double_complement = result();
+qed "double_complement";
goal ZF.thy "(A Un B) - (B-A) = A";
by (fast_tac eq_cs 1);
-val double_complement_Un = result();
+qed "double_complement_Un";
goal ZF.thy
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)";
by (fast_tac eq_cs 1);
-val Un_Int_crazy = result();
+qed "Un_Int_crazy";
goal ZF.thy "A - (B Un C) = (A-B) Int (A-C)";
by (fast_tac eq_cs 1);
-val Diff_Un = result();
+qed "Diff_Un";
goal ZF.thy "A - (B Int C) = (A-B) Un (A-C)";
by (fast_tac eq_cs 1);
-val Diff_Int = result();
+qed "Diff_Int";
(*Halmos, Naive Set Theory, page 16.*)
goal ZF.thy "(A Int B) Un C = A Int (B Un C) <-> C<=A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val Un_Int_assoc_iff = result();
+qed "Un_Int_assoc_iff";
(** Big Union and Intersection **)
goal ZF.thy "Union(0) = 0";
by (fast_tac eq_cs 1);
-val Union_0 = result();
+qed "Union_0";
goal ZF.thy "Union(cons(a,B)) = a Un Union(B)";
by (fast_tac eq_cs 1);
-val Union_cons = result();
+qed "Union_cons";
goal ZF.thy "Union(A Un B) = Union(A) Un Union(B)";
by (fast_tac eq_cs 1);
-val Union_Un_distrib = result();
+qed "Union_Un_distrib";
goal ZF.thy "Union(A Int B) <= Union(A) Int Union(B)";
by (fast_tac ZF_cs 1);
-val Union_Int_subset = result();
+qed "Union_Int_subset";
goal ZF.thy "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val Union_disjoint = result();
+qed "Union_disjoint";
(* A good challenge: Inter is ill-behaved on the empty set *)
goal ZF.thy "!!A B. [| a:A; b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)";
by (fast_tac eq_cs 1);
-val Inter_Un_distrib = result();
+qed "Inter_Un_distrib";
goal ZF.thy "Union({b}) = b";
by (fast_tac eq_cs 1);
-val Union_singleton = result();
+qed "Union_singleton";
goal ZF.thy "Inter({b}) = b";
by (fast_tac eq_cs 1);
-val Inter_singleton = result();
+qed "Inter_singleton";
(** Unions and Intersections of Families **)
goal ZF.thy "Union(A) = (UN x:A. x)";
by (fast_tac eq_cs 1);
-val Union_eq_UN = result();
+qed "Union_eq_UN";
goalw ZF.thy [Inter_def] "Inter(A) = (INT x:A. x)";
by (fast_tac eq_cs 1);
-val Inter_eq_INT = result();
+qed "Inter_eq_INT";
goal ZF.thy "(UN i:0. A(i)) = 0";
by (fast_tac eq_cs 1);
-val UN_0 = result();
+qed "UN_0";
(*Halmos, Naive Set Theory, page 35.*)
goal ZF.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";
by (fast_tac eq_cs 1);
-val Int_UN_distrib = result();
+qed "Int_UN_distrib";
goal ZF.thy "!!A B. i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";
by (fast_tac eq_cs 1);
-val Un_INT_distrib = result();
+qed "Un_INT_distrib";
goal ZF.thy
"(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))";
by (fast_tac eq_cs 1);
-val Int_UN_distrib2 = result();
+qed "Int_UN_distrib2";
goal ZF.thy
"!!I J. [| i:I; j:J |] ==> \
\ (INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))";
by (fast_tac eq_cs 1);
-val Un_INT_distrib2 = result();
+qed "Un_INT_distrib2";
goal ZF.thy "!!A. a: A ==> (UN y:A. c) = c";
by (fast_tac eq_cs 1);
-val UN_constant = result();
+qed "UN_constant";
goal ZF.thy "!!A. a: A ==> (INT y:A. c) = c";
by (fast_tac eq_cs 1);
-val INT_constant = result();
+qed "INT_constant";
(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:
Union of a family of unions **)
goal ZF.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))";
by (fast_tac eq_cs 1);
-val UN_Un_distrib = result();
+qed "UN_Un_distrib";
goal ZF.thy
"!!A B. i:I ==> \
\ (INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))";
by (fast_tac eq_cs 1);
-val INT_Int_distrib = result();
+qed "INT_Int_distrib";
(** Devlin, page 12, exercise 5: Complements **)
goal ZF.thy "!!A B. i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))";
by (fast_tac eq_cs 1);
-val Diff_UN = result();
+qed "Diff_UN";
goal ZF.thy "!!A B. i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))";
by (fast_tac eq_cs 1);
-val Diff_INT = result();
+qed "Diff_INT";
(** Unions and Intersections with General Sum **)
goal ZF.thy "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)";
by (fast_tac eq_cs 1);
-val Sigma_cons = result();
+qed "Sigma_cons";
goal ZF.thy "(SUM x:(UN y:A. C(y)). B(x)) = (UN y:A. SUM x:C(y). B(x))";
by (fast_tac eq_cs 1);
-val SUM_UN_distrib1 = result();
+qed "SUM_UN_distrib1";
goal ZF.thy "(SUM i:I. UN j:J. C(i,j)) = (UN j:J. SUM i:I. C(i,j))";
by (fast_tac eq_cs 1);
-val SUM_UN_distrib2 = result();
+qed "SUM_UN_distrib2";
goal ZF.thy "(SUM i:I Un J. C(i)) = (SUM i:I. C(i)) Un (SUM j:J. C(j))";
by (fast_tac eq_cs 1);
-val SUM_Un_distrib1 = result();
+qed "SUM_Un_distrib1";
goal ZF.thy "(SUM i:I. A(i) Un B(i)) = (SUM i:I. A(i)) Un (SUM i:I. B(i))";
by (fast_tac eq_cs 1);
-val SUM_Un_distrib2 = result();
+qed "SUM_Un_distrib2";
(*First-order version of the above, for rewriting*)
goal ZF.thy "I * (A Un B) = I*A Un I*B";
by (resolve_tac [SUM_Un_distrib2] 1);
-val prod_Un_distrib2 = result();
+qed "prod_Un_distrib2";
goal ZF.thy "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))";
by (fast_tac eq_cs 1);
-val SUM_Int_distrib1 = result();
+qed "SUM_Int_distrib1";
goal ZF.thy
"(SUM i:I. A(i) Int B(i)) = (SUM i:I. A(i)) Int (SUM i:I. B(i))";
by (fast_tac eq_cs 1);
-val SUM_Int_distrib2 = result();
+qed "SUM_Int_distrib2";
(*First-order version of the above, for rewriting*)
goal ZF.thy "I * (A Int B) = I*A Int I*B";
by (resolve_tac [SUM_Int_distrib2] 1);
-val prod_Int_distrib2 = result();
+qed "prod_Int_distrib2";
(*Cf Aczel, Non-Well-Founded Sets, page 115*)
goal ZF.thy "(SUM i:I. A(i)) = (UN i:I. {i} * A(i))";
by (fast_tac eq_cs 1);
-val SUM_eq_UN = result();
+qed "SUM_eq_UN";
(** Domain **)
-val domain_of_prod = prove_goal ZF.thy "!!A B. b:B ==> domain(A*B) = A"
+qed_goal "domain_of_prod" ZF.thy "!!A B. b:B ==> domain(A*B) = A"
(fn _ => [ fast_tac eq_cs 1 ]);
-val domain_0 = prove_goal ZF.thy "domain(0) = 0"
+qed_goal "domain_0" ZF.thy "domain(0) = 0"
(fn _ => [ fast_tac eq_cs 1 ]);
-val domain_cons = prove_goal ZF.thy
+qed_goal "domain_cons" ZF.thy
"domain(cons(<a,b>,r)) = cons(a, domain(r))"
(fn _ => [ fast_tac eq_cs 1 ]);
goal ZF.thy "domain(A Un B) = domain(A) Un domain(B)";
by (fast_tac eq_cs 1);
-val domain_Un_eq = result();
+qed "domain_Un_eq";
goal ZF.thy "domain(A Int B) <= domain(A) Int domain(B)";
by (fast_tac eq_cs 1);
-val domain_Int_subset = result();
+qed "domain_Int_subset";
goal ZF.thy "domain(A) - domain(B) <= domain(A - B)";
by (fast_tac eq_cs 1);
-val domain_diff_subset = result();
+qed "domain_diff_subset";
goal ZF.thy "domain(converse(r)) = range(r)";
by (fast_tac eq_cs 1);
-val domain_converse = result();
+qed "domain_converse";
(** Range **)
-val range_of_prod = prove_goal ZF.thy
+qed_goal "range_of_prod" ZF.thy
"!!a A B. a:A ==> range(A*B) = B"
(fn _ => [ fast_tac eq_cs 1 ]);
-val range_0 = prove_goal ZF.thy "range(0) = 0"
+qed_goal "range_0" ZF.thy "range(0) = 0"
(fn _ => [ fast_tac eq_cs 1 ]);
-val range_cons = prove_goal ZF.thy
+qed_goal "range_cons" ZF.thy
"range(cons(<a,b>,r)) = cons(b, range(r))"
(fn _ => [ fast_tac eq_cs 1 ]);
goal ZF.thy "range(A Un B) = range(A) Un range(B)";
by (fast_tac eq_cs 1);
-val range_Un_eq = result();
+qed "range_Un_eq";
goal ZF.thy "range(A Int B) <= range(A) Int range(B)";
by (fast_tac ZF_cs 1);
-val range_Int_subset = result();
+qed "range_Int_subset";
goal ZF.thy "range(A) - range(B) <= range(A - B)";
by (fast_tac eq_cs 1);
-val range_diff_subset = result();
+qed "range_diff_subset";
goal ZF.thy "range(converse(r)) = domain(r)";
by (fast_tac eq_cs 1);
-val range_converse = result();
+qed "range_converse";
(** Field **)
-val field_of_prod = prove_goal ZF.thy "field(A*A) = A"
+qed_goal "field_of_prod" ZF.thy "field(A*A) = A"
(fn _ => [ fast_tac eq_cs 1 ]);
-val field_0 = prove_goal ZF.thy "field(0) = 0"
+qed_goal "field_0" ZF.thy "field(0) = 0"
(fn _ => [ fast_tac eq_cs 1 ]);
-val field_cons = prove_goal ZF.thy
+qed_goal "field_cons" ZF.thy
"field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))"
(fn _ => [ rtac equalityI 1, ALLGOALS (fast_tac ZF_cs) ]);
goal ZF.thy "field(A Un B) = field(A) Un field(B)";
by (fast_tac eq_cs 1);
-val field_Un_eq = result();
+qed "field_Un_eq";
goal ZF.thy "field(A Int B) <= field(A) Int field(B)";
by (fast_tac eq_cs 1);
-val field_Int_subset = result();
+qed "field_Int_subset";
goal ZF.thy "field(A) - field(B) <= field(A - B)";
by (fast_tac eq_cs 1);
-val field_diff_subset = result();
+qed "field_diff_subset";
(** Image **)
goal ZF.thy "r``0 = 0";
by (fast_tac eq_cs 1);
-val image_0 = result();
+qed "image_0";
goal ZF.thy "r``(A Un B) = (r``A) Un (r``B)";
by (fast_tac eq_cs 1);
-val image_Un = result();
+qed "image_Un";
goal ZF.thy "r``(A Int B) <= (r``A) Int (r``B)";
by (fast_tac ZF_cs 1);
-val image_Int_subset = result();
+qed "image_Int_subset";
goal ZF.thy "(r Int A*A)``B <= (r``B) Int A";
by (fast_tac ZF_cs 1);
-val image_Int_square_subset = result();
+qed "image_Int_square_subset";
goal ZF.thy "!!r. B<=A ==> (r Int A*A)``B = (r``B) Int A";
by (fast_tac eq_cs 1);
-val image_Int_square = result();
+qed "image_Int_square";
(** Inverse Image **)
goal ZF.thy "r-``0 = 0";
by (fast_tac eq_cs 1);
-val vimage_0 = result();
+qed "vimage_0";
goal ZF.thy "r-``(A Un B) = (r-``A) Un (r-``B)";
by (fast_tac eq_cs 1);
-val vimage_Un = result();
+qed "vimage_Un";
goal ZF.thy "r-``(A Int B) <= (r-``A) Int (r-``B)";
by (fast_tac ZF_cs 1);
-val vimage_Int_subset = result();
+qed "vimage_Int_subset";
goal ZF.thy "(r Int A*A)-``B <= (r-``B) Int A";
by (fast_tac ZF_cs 1);
-val vimage_Int_square_subset = result();
+qed "vimage_Int_square_subset";
goal ZF.thy "!!r. B<=A ==> (r Int A*A)-``B = (r-``B) Int A";
by (fast_tac eq_cs 1);
-val vimage_Int_square = result();
+qed "vimage_Int_square";
(** Converse **)
goal ZF.thy "converse(A Un B) = converse(A) Un converse(B)";
by (fast_tac eq_cs 1);
-val converse_Un = result();
+qed "converse_Un";
goal ZF.thy "converse(A Int B) = converse(A) Int converse(B)";
by (fast_tac eq_cs 1);
-val converse_Int = result();
+qed "converse_Int";
goal ZF.thy "converse(A) - converse(B) = converse(A - B)";
by (fast_tac eq_cs 1);
-val converse_diff = result();
+qed "converse_diff";
(** Pow **)
goal ZF.thy "Pow(A) Un Pow(B) <= Pow(A Un B)";
by (fast_tac upair_cs 1);
-val Un_Pow_subset = result();
+qed "Un_Pow_subset";
goal ZF.thy "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))";
by (fast_tac upair_cs 1);
-val UN_Pow_subset = result();
+qed "UN_Pow_subset";
goal ZF.thy "A <= Pow(Union(A))";
by (fast_tac upair_cs 1);
-val subset_Pow_Union = result();
+qed "subset_Pow_Union";
goal ZF.thy "Union(Pow(A)) = A";
by (fast_tac eq_cs 1);
-val Union_Pow_eq = result();
+qed "Union_Pow_eq";
goal ZF.thy "Pow(A) Int Pow(B) = Pow(A Int B)";
by (fast_tac eq_cs 1);
-val Int_Pow_eq = result();
+qed "Int_Pow_eq";
goal ZF.thy "!!x A. x:A ==> (INT x:A. Pow(B(x))) = Pow(INT x:A. B(x))";
by (fast_tac eq_cs 1);
-val INT_Pow_subset = result();
+qed "INT_Pow_subset";
--- a/src/ZF/ex/BT.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/BT.ML Wed Dec 07 13:12:04 1994 +0100
@@ -46,13 +46,13 @@
goal BT.thy "bt_rec(Lf,c,h) = c";
by (rtac (bt_rec_def RS def_Vrec RS trans) 1);
by (simp_tac (ZF_ss addsimps bt.case_eqns) 1);
-val bt_rec_Lf = result();
+qed "bt_rec_Lf";
goal BT.thy
"bt_rec(Br(a,l,r), c, h) = h(a, l, r, bt_rec(l,c,h), bt_rec(r,c,h))";
by (rtac (bt_rec_def RS def_Vrec RS trans) 1);
by (simp_tac (rank_ss addsimps (bt.case_eqns @ [rank_Br1, rank_Br2])) 1);
-val bt_rec_Br = result();
+qed "bt_rec_Br";
(*Type checking -- proved by induction, as usual*)
val prems = goal BT.thy
@@ -64,7 +64,7 @@
by (bt_ind_tac "t" prems 1);
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps
(prems@[bt_rec_Lf,bt_rec_Br]))));
-val bt_rec_type = result();
+qed "bt_rec_type";
(** Versions for use with definitions **)
@@ -98,7 +98,7 @@
val prems = goalw BT.thy [n_leaves_def]
"xs: bt(A) ==> n_leaves(xs) : nat";
by (REPEAT (ares_tac (prems @ [bt_rec_type, nat_0I, nat_succI, add_type]) 1));
-val n_leaves_type = result();
+qed "n_leaves_type";
(** bt_reflect **)
--- a/src/ZF/ex/Bin.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/Bin.ML Wed Dec 07 13:12:04 1994 +0100
@@ -21,19 +21,19 @@
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac bin.con_defs);
by (simp_tac rank_ss 1);
-val bin_rec_Plus = result();
+qed "bin_rec_Plus";
goal Bin.thy "bin_rec(Minus,a,b,h) = b";
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac bin.con_defs);
by (simp_tac rank_ss 1);
-val bin_rec_Minus = result();
+qed "bin_rec_Minus";
goal Bin.thy "bin_rec(w$$x,a,b,h) = h(w, x, bin_rec(w,a,b,h))";
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac bin.con_defs);
by (simp_tac rank_ss 1);
-val bin_rec_Bcons = result();
+qed "bin_rec_Bcons";
(*Type checking*)
val prems = goal Bin.thy
@@ -83,23 +83,23 @@
goalw Bin.thy [bin_succ_def]
"!!w. w: bin ==> bin_succ(w) : bin";
by (typechk_tac (bin_typechecks0@bool_typechecks));
-val bin_succ_type = result();
+qed "bin_succ_type";
goalw Bin.thy [bin_pred_def]
"!!w. w: bin ==> bin_pred(w) : bin";
by (typechk_tac (bin_typechecks0@bool_typechecks));
-val bin_pred_type = result();
+qed "bin_pred_type";
goalw Bin.thy [bin_minus_def]
"!!w. w: bin ==> bin_minus(w) : bin";
by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks));
-val bin_minus_type = result();
+qed "bin_minus_type";
goalw Bin.thy [bin_add_def]
"!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin";
by (typechk_tac ([bin_succ_type,bin_pred_type]@bin_typechecks0@
bool_typechecks@ZF_typechecks));
-val bin_add_type = result();
+qed "bin_add_type";
goalw Bin.thy [bin_mult_def]
"!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
@@ -134,7 +134,7 @@
"!!z v w. [| z: integ; v: integ; w: integ |] \
\ ==> z $+ (v $+ w) = v $+ (z $+ w)";
by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1));
-val zadd_assoc_swap = result();
+qed "zadd_assoc_swap";
val zadd_cong =
read_instantiate_sg (sign_of Integ.thy) [("t","op $+")] subst_context2;
@@ -149,7 +149,7 @@
"!!z w. [| z: integ; w: integ |] \
\ ==> w $+ (z $+ (w $+ z)) = w $+ (w $+ (z $+ z))";
by (REPEAT (ares_tac [zadd_kill, zadd_assoc_swap] 1));
-val zadd_swap_pairs = result();
+qed "zadd_swap_pairs";
val carry_ss = bin_ss addsimps
@@ -163,7 +163,7 @@
by (etac boolE 1);
by (ALLGOALS (asm_simp_tac (carry_ss addsimps [zadd_assoc])));
by (REPEAT (ares_tac (zadd_swap_pairs::typechecks) 1));
-val integ_of_bin_succ = result();
+qed "integ_of_bin_succ";
goal Bin.thy
"!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)";
@@ -176,7 +176,7 @@
(carry_ss addsimps [zadd_assoc RS sym,
zadd_zminus_inverse, zadd_zminus_inverse2])));
by (REPEAT (ares_tac ([zadd_commute, zadd_cong, refl]@typechecks) 1));
-val integ_of_bin_pred = result();
+qed "integ_of_bin_pred";
(*These two results replace the definitions of bin_succ and bin_pred*)
@@ -202,19 +202,19 @@
goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
by (asm_simp_tac bin_ss 1);
-val bin_add_Plus = result();
+qed "bin_add_Plus";
goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
by (asm_simp_tac bin_ss 1);
-val bin_add_Minus = result();
+qed "bin_add_Minus";
goalw Bin.thy [bin_add_def] "bin_add(v$$x,Plus) = v$$x";
by (simp_tac bin_ss 1);
-val bin_add_Bcons_Plus = result();
+qed "bin_add_Bcons_Plus";
goalw Bin.thy [bin_add_def] "bin_add(v$$x,Minus) = bin_pred(v$$x)";
by (simp_tac bin_ss 1);
-val bin_add_Bcons_Minus = result();
+qed "bin_add_Bcons_Minus";
goalw Bin.thy [bin_add_def]
"!!w y. [| w: bin; y: bool |] ==> \
--- a/src/ZF/ex/CoUnit.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/CoUnit.ML Wed Dec 07 13:12:04 1994 +0100
@@ -48,7 +48,7 @@
by (rtac (qunivI RS singleton_subsetI) 1);
by (rtac ([lfp_subset, empty_subsetI RS univ_mono] MRS subset_trans) 1);
by (fast_tac (ZF_cs addSIs [Con2_bnd_mono RS lfp_Tarski]) 1);
-val lfp_Con2_in_counit2 = result();
+qed "lfp_Con2_in_counit2";
(*Lemma for proving finality. Borrowed from ex/llist_eq.ML!*)
goal CoUnit.thy
--- a/src/ZF/ex/Comb.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/Comb.ML Wed Dec 07 13:12:04 1994 +0100
@@ -33,7 +33,7 @@
goal Comb.thy "field(contract) = comb";
by (fast_tac (ZF_cs addIs [equalityI,contract.K] addSEs [contract_combE2]) 1);
-val field_contract_eq = result();
+qed "field_contract_eq";
val reduction_refl = standard
(field_contract_eq RS equalityD2 RS subsetD RS rtrancl_refl);
@@ -54,7 +54,7 @@
goalw Comb.thy [I_def] "I: comb";
by (REPEAT (ares_tac comb.intrs 1));
-val comb_I = result();
+qed "comb_I";
(** Non-contraction results **)
@@ -218,7 +218,7 @@
(fast_tac
(contract_cs addIs [Ap_reduce1, Ap_reduce2]
addSEs [parcontract_combD1,parcontract_combD2])));
-val parcontract_imp_reduce = result();
+qed "parcontract_imp_reduce";
goal Comb.thy "!!p q. p===>q ==> p--->q";
by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1);
--- a/src/ZF/ex/Integ.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/Integ.ML Wed Dec 07 13:12:04 1994 +0100
@@ -43,13 +43,13 @@
"<<x1,y1>,<x2,y2>>: intrel <-> \
\ x1: nat & y1: nat & x2: nat & y2: nat & x1#+y2 = x2#+y1";
by (fast_tac ZF_cs 1);
-val intrel_iff = result();
+qed "intrel_iff";
goalw Integ.thy [intrel_def]
"!!x1 x2. [| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
\ <<x1,y1>,<x2,y2>>: intrel";
by (fast_tac (ZF_cs addIs prems) 1);
-val intrelI = result();
+qed "intrelI";
(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*)
goalw Integ.thy [intrel_def]
@@ -66,14 +66,14 @@
\ ==> Q";
by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1);
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
-val intrelE = result();
+qed "intrelE";
val intrel_cs = ZF_cs addSIs [intrelI] addSEs [intrelE];
goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def]
"equiv(nat*nat, intrel)";
by (fast_tac (intrel_cs addSEs [sym, integ_trans_lemma]) 1);
-val equiv_intrel = result();
+qed "equiv_intrel";
val intrel_ss =
@@ -88,7 +88,7 @@
goalw Integ.thy [integ_def,quotient_def,znat_def]
"!!m. m : nat ==> $#m : integ";
by (fast_tac (ZF_cs addSIs [nat_0I]) 1);
-val znat_type = result();
+qed "znat_type";
goalw Integ.thy [znat_def]
"!!m n. [| $#m = $#n; n: nat |] ==> m=n";
@@ -113,7 +113,7 @@
"!!z. z : integ ==> $~z : integ";
by (typechk_tac [split_type, SigmaI, zminus_ize UN_equiv_class_type,
quotientI]);
-val zminus_type = result();
+qed "zminus_type";
goalw Integ.thy [integ_def,zminus_def]
"!!z w. [| $~z = $~w; z: integ; w: integ |] ==> z=w";
@@ -127,7 +127,7 @@
goalw Integ.thy [zminus_def]
"!!x y.[| x: nat; y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
by (asm_simp_tac (ZF_ss addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
-val zminus = result();
+qed "zminus";
goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z";
by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1));
@@ -190,7 +190,7 @@
"!!x y. [| x: nat; y: nat |] ==> \
\ zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
by (asm_simp_tac (ZF_ss addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1);
-val zmagnitude = result();
+qed "zmagnitude";
goalw Integ.thy [znat_def]
"!!n. n: nat ==> zmagnitude($# n) = n";
@@ -236,12 +236,12 @@
\ (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = \
\ intrel `` {<x1#+x2, y1#+y2>}";
by (asm_simp_tac (ZF_ss addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1);
-val zadd = result();
+qed "zadd";
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
-val zadd_0 = result();
+qed "zadd_0";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w";
@@ -253,7 +253,7 @@
"!!z w. [| z: integ; w: integ |] ==> z $+ w = w $+ z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
by (asm_simp_tac (intrel_ss addsimps (add_ac @ [zadd])) 1);
-val zadd_commute = result();
+qed "zadd_commute";
goalw Integ.thy [integ_def]
"!!z1 z2 z3. [| z1: integ; z2: integ; z3: integ |] ==> \
@@ -261,7 +261,7 @@
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
(*rewriting is much faster without intrel_iff, etc.*)
by (asm_simp_tac (arith_ss addsimps [zadd, add_assoc]) 1);
-val zadd_assoc = result();
+qed "zadd_assoc";
goalw Integ.thy [znat_def]
"!!m n. [| m: nat; n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
@@ -271,7 +271,7 @@
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
by (asm_simp_tac (intrel_ss addsimps [zminus, zadd, add_commute]) 1);
-val zadd_zminus_inverse = result();
+qed "zadd_zminus_inverse";
goal Integ.thy "!!z. z : integ ==> ($~ z) $+ z = $#0";
by (asm_simp_tac
@@ -281,7 +281,7 @@
goal Integ.thy "!!z. z:integ ==> z $+ $#0 = z";
by (rtac (zadd_commute RS trans) 1);
by (REPEAT (ares_tac [znat_type, nat_0I, zadd_0] 1));
-val zadd_0_right = result();
+qed "zadd_0_right";
(*Need properties of $- ??? Or use $- just as an abbreviation?
@@ -324,7 +324,7 @@
\ (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = \
\ intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
by (asm_simp_tac (ZF_ss addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
-val zmult = result();
+qed "zmult";
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
@@ -335,7 +335,7 @@
"!!z. z : integ ==> $#1 $* z = z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
by (asm_simp_tac (arith_ss addsimps [zmult, add_0_right]) 1);
-val zmult_1 = result();
+qed "zmult_1";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> ($~ z) $* w = $~ (z $* w)";
@@ -371,7 +371,7 @@
by (asm_simp_tac
(intrel_ss addsimps ([zadd, zmult, add_mult_distrib] @
add_ac @ mult_ac)) 1);
-val zadd_zmult_distrib = result();
+qed "zadd_zmult_distrib";
val integ_typechecks =
[znat_type, zminus_type, zmagnitude_type, zadd_type, zmult_type];
--- a/src/ZF/ex/LList.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/LList.ML Wed Dec 07 13:12:04 1994 +0100
@@ -184,7 +184,7 @@
by (asm_simp_tac flip_ss 1);
by (asm_simp_tac flip_ss 1);
by (fast_tac (ZF_cs addSIs [not_type]) 1);
-val flip_type = result();
+qed "flip_type";
val [prem] = goal LList.thy
"l : llist(bool) ==> flip(flip(l)) = l";
--- a/src/ZF/ex/Primrec.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/Primrec.ML Wed Dec 07 13:12:04 1994 +0100
@@ -32,12 +32,12 @@
goalw Primrec.thy [SC_def]
"!!x l. [| x:nat; l: list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)";
by (asm_simp_tac pr_ss 1);
-val SC = result();
+qed "SC";
goalw Primrec.thy [CONST_def]
"!!l. [| l: list(nat) |] ==> CONST(k) ` l = k";
by (asm_simp_tac pr_ss 1);
-val CONST = result();
+qed "CONST";
goalw Primrec.thy [PROJ_def]
"!!l. [| x: nat; l: list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x";
@@ -87,7 +87,7 @@
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j): nat";
by (tc_tac []);
-val ack_type = result();
+qed "ack_type";
(** Ackermann's function cases **)
@@ -140,7 +140,7 @@
by (assume_tac 1);
by (rtac lt_trans 2);
by (REPEAT (ares_tac ([ack_lt_ack_succ2, ack_type] @ pr_typechecks) 1));
-val ack_lt_mono2 = result();
+qed "ack_lt_mono2";
(*PROPERTY A 5', monotonicity for le *)
goal Primrec.thy
@@ -157,14 +157,14 @@
by (rtac ack_le_mono2 1);
by (rtac (lt_ack2 RS succ_leI RS le_trans) 1);
by (REPEAT (ares_tac (ack_typechecks) 1));
-val ack2_le_ack1 = result();
+qed "ack2_le_ack1";
(*PROPERTY A 7-, the single-step lemma*)
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(succ(i),j)";
by (rtac (ack_lt_mono2 RS lt_trans2) 1);
by (rtac ack2_le_ack1 4);
by (REPEAT (ares_tac ([nat_le_refl, ack_type] @ pr_typechecks) 1));
-val ack_lt_ack_succ1 = result();
+qed "ack_lt_ack_succ1";
(*PROPERTY A 7, monotonicity for < *)
goal Primrec.thy "!!i j k. [| i<j; j:nat; k:nat |] ==> ack(i,k) < ack(j,k)";
@@ -173,20 +173,20 @@
by (assume_tac 1);
by (rtac lt_trans 2);
by (REPEAT (ares_tac ([ack_lt_ack_succ1, ack_type] @ pr_typechecks) 1));
-val ack_lt_mono1 = result();
+qed "ack_lt_mono1";
(*PROPERTY A 7', monotonicity for le *)
goal Primrec.thy
"!!i j k. [| i le j; j:nat; k:nat |] ==> ack(i,k) le ack(j,k)";
by (res_inst_tac [("f", "%j.ack(j,k)")] Ord_lt_mono_imp_le_mono 1);
by (REPEAT (ares_tac [ack_lt_mono1, ack_type RS nat_into_Ord] 1));
-val ack_le_mono1 = result();
+qed "ack_le_mono1";
(*PROPERTY A 8*)
goal Primrec.thy "!!j. j:nat ==> ack(1,j) = succ(succ(j))";
by (etac nat_induct 1);
by (ALLGOALS (asm_simp_tac ack_ss));
-val ack_1 = result();
+qed "ack_1";
(*PROPERTY A 9*)
goal Primrec.thy "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
@@ -203,7 +203,7 @@
by (rtac (add_le_self RS ack_le_mono1 RS lt_trans1) 1);
by (rtac (add_le_self2 RS ack_lt_mono1 RS ack_lt_mono2) 5);
by (tc_tac []);
-val ack_nest_bound = result();
+qed "ack_nest_bound";
(*PROPERTY A 11*)
goal Primrec.thy
@@ -216,7 +216,7 @@
by (rtac (add_le_mono RS leI RS leI) 1);
by (REPEAT (ares_tac ([add_le_self, add_le_self2, ack_le_mono1] @
ack_typechecks) 1));
-val ack_add_bound = result();
+qed "ack_add_bound";
(*PROPERTY A 12. Article uses existential quantifier but the ALF proof
used k#+4. Quantified version must be nested EX k'. ALL i,j... *)
@@ -247,7 +247,7 @@
by (asm_simp_tac (ack_ss addsimps [nat_0_le]) 1);
by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1);
by (tc_tac []);
-val lt_ack1 = result();
+qed "lt_ack1";
goalw Primrec.thy [CONST_def]
"!!l. [| l: list(nat); k: nat |] ==> CONST(k) ` l < ack(k, list_add(l))";
--- a/src/ZF/ex/PropLog.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/PropLog.ML Wed Dec 07 13:12:04 1994 +0100
@@ -53,7 +53,7 @@
goalw PropLog.thy [is_true_def]
"is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))";
by (simp_tac (prop_rec_ss setloop (split_tac [expand_if])) 1);
-val is_true_Imp = result();
+qed "is_true_Imp";
(** The function hyps **)
@@ -90,7 +90,7 @@
goal PropLog.thy "!!p q H. [| H |- p=>q; H |- p |] ==> H |- q";
by (rtac thms.MP 1);
by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1));
-val thms_MP = result();
+qed "thms_MP";
(*Rule is called I for Identity Combinator, not for Introduction*)
goal PropLog.thy "!!p H. p:prop ==> H |- p=>p";
@@ -98,7 +98,7 @@
by (rtac thms.K 5);
by (rtac thms.K 4);
by (REPEAT (ares_tac prop.intrs 1));
-val thms_I = result();
+qed "thms_I";
(** Weakening, left and right **)
@@ -114,7 +114,7 @@
goal PropLog.thy "!!H p q. [| H |- q; p:prop |] ==> H |- p=>q";
by (rtac (thms.K RS thms_MP) 1);
by (REPEAT (ares_tac [thms_in_pl] 1));
-val weaken_right = result();
+qed "weaken_right";
(*The deduction theorem*)
goal PropLog.thy "!!p q H. [| cons(p,H) |- q; p:prop |] ==> H |- p=>q";
@@ -124,7 +124,7 @@
by (fast_tac (ZF_cs addIs [thms.S RS weaken_right]) 1);
by (fast_tac (ZF_cs addIs [thms.DN RS weaken_right]) 1);
by (fast_tac (ZF_cs addIs [thms.S RS thms_MP RS thms_MP]) 1);
-val deduction = result();
+qed "deduction";
(*The cut rule*)
@@ -157,7 +157,7 @@
by (rtac deduction 1);
by (rtac (premf RS weaken_left_cons RS thms_notE) 1);
by (REPEAT (ares_tac [premq, consI1, thms.H] 1));
-val Fls_Imp = result();
+qed "Fls_Imp";
val [premp,premq] = goal PropLog.thy
"[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls";
@@ -208,7 +208,7 @@
"[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p: prop |] ==> H |- q";
by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1);
by (REPEAT (resolve_tac (prems@prop.intrs@[deduction,thms_in_pl]) 1));
-val thms_excluded_middle_rule = result();
+qed "thms_excluded_middle_rule";
(*** Completeness -- lemmas for reducing the set of assumptions ***)
@@ -240,11 +240,11 @@
goal ZF.thy "B-C <= cons(a, B-cons(a,C))";
by (fast_tac ZF_cs 1);
-val cons_Diff_same = result();
+qed "cons_Diff_same";
goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
by (fast_tac ZF_cs 1);
-val cons_Diff_subset2 = result();
+qed "cons_Diff_subset2";
(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls;
could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*)
--- a/src/ZF/ex/Ramsey.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/Ramsey.ML Wed Dec 07 13:12:04 1994 +0100
@@ -44,7 +44,7 @@
goalw Ramsey.thy [Atleast_def, inj_def, Pi_def, function_def] "Atleast(0,A)";
by (fast_tac ZF_cs 1);
-val Atleast0 = result();
+qed "Atleast0";
val [major] = goalw Ramsey.thy [Atleast_def]
"Atleast(succ(m),A) ==> EX x:A. Atleast(m, A-{x})";
@@ -54,7 +54,7 @@
by (rtac succI1 2);
by (rtac exI 1);
by (etac inj_succ_restrict 1);
-val Atleast_succD = result();
+qed "Atleast_succD";
val major::prems = goalw Ramsey.thy [Atleast_def]
"[| Atleast(n,A); A<=B |] ==> Atleast(n,B)";
@@ -62,7 +62,7 @@
by (rtac exI 1);
by (etac inj_weaken_type 1);
by (resolve_tac prems 1);
-val Atleast_superset = result();
+qed "Atleast_superset";
val prems = goalw Ramsey.thy [Atleast_def,succ_def]
"[| Atleast(m,B); b~: B |] ==> Atleast(succ(m), cons(b,B))";
@@ -72,7 +72,7 @@
by (etac inj_extend 1);
by (rtac mem_not_refl 1);
by (assume_tac 1);
-val Atleast_succI = result();
+qed "Atleast_succI";
val prems = goal Ramsey.thy
"[| Atleast(m, B-{x}); x: B |] ==> Atleast(succ(m), B)";
@@ -80,7 +80,7 @@
by (etac (Atleast_succI RS Atleast_superset) 1);
by (fast_tac ZF_cs 1);
by (fast_tac ZF_cs 1);
-val Atleast_Diff_succI = result();
+qed "Atleast_Diff_succI";
(*** Main Cardinality Lemma ***)
--- a/src/ZF/ex/TF.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/TF.ML Wed Dec 07 13:12:04 1994 +0100
@@ -41,7 +41,7 @@
by (fast_tac (sum_cs addSIs (equalityI :: (map rew intrs RL [PartD1]))
addEs [rew elim]) 1)
end;
-val tree_forest_unfold = result();
+qed "tree_forest_unfold";
val tree_forest_unfold' = rewrite_rule [tree_def, forest_def]
tree_forest_unfold;
@@ -67,20 +67,20 @@
by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac tree_forest.con_defs);
by (simp_tac rank_ss 1);
-val TF_rec_Tcons = result();
+qed "TF_rec_Tcons";
goal TF.thy "TF_rec(Fnil, b, c, d) = c";
by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac tree_forest.con_defs);
by (simp_tac rank_ss 1);
-val TF_rec_Fnil = result();
+qed "TF_rec_Fnil";
goal TF.thy "TF_rec(Fcons(t,f), b, c, d) = \
\ d(t, f, TF_rec(t, b, c, d), TF_rec(f, b, c, d))";
by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac tree_forest.con_defs);
by (simp_tac rank_ss 1);
-val TF_rec_Fcons = result();
+qed "TF_rec_Fcons";
(*list_ss includes list operations as well as arith_ss*)
val TF_rec_ss = list_ss addsimps
@@ -98,7 +98,7 @@
\ |] ==> TF_rec(z,b,c,d) : C(z)";
by (rtac (major RS tree_forest.induct) 1);
by (ALLGOALS (asm_simp_tac (TF_rec_ss addsimps prems)));
-val TF_rec_type = result();
+qed "TF_rec_type";
(*Mutually recursive version*)
val prems = goal TF.thy
--- a/src/ZF/ex/Term.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/ex/Term.ML Wed Dec 07 13:12:04 1994 +0100
@@ -38,7 +38,7 @@
by (rtac (major RS term.induct) 1);
by (resolve_tac prems 1);
by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1));
-val term_induct_eqn = result();
+qed "term_induct_eqn";
(** Lemmas to justify using "term" in other recursive type definitions **)
@@ -86,7 +86,7 @@
by (rewrite_goals_tac term.con_defs);
val term_rec_ss = ZF_ss addsimps [Ord_rank, rank_pair2, prem RS map_lemma];
by (simp_tac term_rec_ss 1);
-val term_rec = result();
+qed "term_rec";
(*Slightly odd typing condition on r in the second premise!*)
val major::prems = goal Term.thy
@@ -110,7 +110,7 @@
\ j(Apply(a,ts)) = d(a, ts, map(%Z.j(Z), ts))";
by (rewtac rew);
by (rtac (tslist RS term_rec) 1);
-val def_term_rec = result();
+qed "def_term_rec";
val prems = goal Term.thy
"[| t: term(A); \
@@ -119,7 +119,7 @@
\ |] ==> term_rec(t,d) : C";
by (REPEAT (ares_tac (term_rec_type::prems) 1));
by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1);
-val term_rec_simple_type = result();
+qed "term_rec_simple_type";
(** term_map **)
@@ -129,13 +129,13 @@
val prems = goalw Term.thy [term_map_def]
"[| t: term(A); !!x. x: A ==> f(x): B |] ==> term_map(f,t) : term(B)";
by (REPEAT (ares_tac ([term_rec_simple_type, term.Apply_I] @ prems) 1));
-val term_map_type = result();
+qed "term_map_type";
val [major] = goal Term.thy
"t: term(A) ==> term_map(f,t) : term({f(u). u:A})";
by (rtac (major RS term_map_type) 1);
by (etac RepFunI 1);
-val term_map_type2 = result();
+qed "term_map_type2";
(** term_size **)
--- a/src/ZF/func.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/func.ML Wed Dec 07 13:12:04 1994 +0100
@@ -11,7 +11,7 @@
goalw ZF.thy [Pi_def]
"f: Pi(A,B) <-> function(f) & f<=Sigma(A,B) & A<=domain(f)";
by (fast_tac ZF_cs 1);
-val Pi_iff = result();
+qed "Pi_iff";
(*For upward compatibility with the former definition*)
goalw ZF.thy [Pi_def, function_def]
@@ -21,99 +21,99 @@
by (best_tac ZF_cs 1);
by (set_mp_tac 1);
by (deepen_tac ZF_cs 3 1);
-val Pi_iff_old = result();
+qed "Pi_iff_old";
goal ZF.thy "!!f. f: Pi(A,B) ==> function(f)";
by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff]) 1);
-val fun_is_function = result();
+qed "fun_is_function";
(**Two "destruct" rules for Pi **)
val [major] = goalw ZF.thy [Pi_def] "f: Pi(A,B) ==> f <= Sigma(A,B)";
by (rtac (major RS CollectD1 RS PowD) 1);
-val fun_is_rel = result();
+qed "fun_is_rel";
goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> EX! y. <a,y>: f";
by (fast_tac (ZF_cs addSDs [Pi_iff_old RS iffD1]) 1);
-val fun_unique_Pair = result();
+qed "fun_unique_Pair";
val prems = goalw ZF.thy [Pi_def]
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')";
by (simp_tac (FOL_ss addsimps prems addcongs [Sigma_cong]) 1);
-val Pi_cong = result();
+qed "Pi_cong";
(*Weakening one function type to another; see also Pi_type*)
goalw ZF.thy [Pi_def] "!!f. [| f: A->B; B<=D |] ==> f: A->D";
by (best_tac ZF_cs 1);
-val fun_weaken_type = result();
+qed "fun_weaken_type";
(*Empty function spaces*)
goalw ZF.thy [Pi_def, function_def] "Pi(0,A) = {0}";
by (fast_tac eq_cs 1);
-val Pi_empty1 = result();
+qed "Pi_empty1";
goalw ZF.thy [Pi_def] "!!A a. a:A ==> A->0 = 0";
by (fast_tac eq_cs 1);
-val Pi_empty2 = result();
+qed "Pi_empty2";
(*The empty function*)
goalw ZF.thy [Pi_def, function_def] "0: 0->A";
by (fast_tac ZF_cs 1);
-val empty_fun = result();
+qed "empty_fun";
(*The singleton function*)
goalw ZF.thy [Pi_def, function_def] "{<a,b>} : {a} -> {b}";
by (fast_tac eq_cs 1);
-val single_fun = result();
+qed "single_fun";
(*** Function Application ***)
goalw ZF.thy [Pi_def, function_def]
"!!a b f. [| <a,b>: f; <a,c>: f; f: Pi(A,B) |] ==> b=c";
by (deepen_tac ZF_cs 3 1);
-val apply_equality2 = result();
+qed "apply_equality2";
goalw ZF.thy [apply_def] "!!a b f. [| <a,b>: f; f: Pi(A,B) |] ==> f`a = b";
by (rtac the_equality 1);
by (rtac apply_equality2 2);
by (REPEAT (assume_tac 1));
-val apply_equality = result();
+qed "apply_equality";
(*Applying a function outside its domain yields 0*)
goalw ZF.thy [apply_def]
"!!a b f. [| a ~: domain(f); f: Pi(A,B) |] ==> f`a = 0";
by (rtac the_0 1);
by (fast_tac ZF_cs 1);
-val apply_0 = result();
+qed "apply_0";
goal ZF.thy "!!f. [| f: Pi(A,B); c: f |] ==> EX x:A. c = <x,f`x>";
by (forward_tac [fun_is_rel] 1);
by (fast_tac (ZF_cs addDs [apply_equality]) 1);
-val Pi_memberD = result();
+qed "Pi_memberD";
goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> <a,f`a>: f";
by (rtac (fun_unique_Pair RS ex1E) 1);
by (resolve_tac [apply_equality RS ssubst] 3);
by (REPEAT (assume_tac 1));
-val apply_Pair = result();
+qed "apply_Pair";
(*Conclusion is flexible -- use res_inst_tac or else apply_funtype below!*)
goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> f`a : B(a)";
by (rtac (fun_is_rel RS subsetD RS SigmaE2) 1);
by (REPEAT (ares_tac [apply_Pair] 1));
-val apply_type = result();
+qed "apply_type";
(*This version is acceptable to the simplifier*)
goal ZF.thy "!!f. [| f: A->B; a:A |] ==> f`a : B";
by (REPEAT (ares_tac [apply_type] 1));
-val apply_funtype = result();
+qed "apply_funtype";
val [major] = goal ZF.thy
"f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b";
by (cut_facts_tac [major RS fun_is_rel] 1);
by (fast_tac (ZF_cs addSIs [major RS apply_Pair,
major RSN (2,apply_equality)]) 1);
-val apply_iff = result();
+qed "apply_iff";
(*Refining one Pi type to another*)
val pi_prem::prems = goal ZF.thy
@@ -121,19 +121,19 @@
by (cut_facts_tac [pi_prem] 1);
by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff]) 1);
by (fast_tac (ZF_cs addIs prems addSDs [pi_prem RS Pi_memberD]) 1);
-val Pi_type = result();
+qed "Pi_type";
(** Elimination of membership in a function **)
goal ZF.thy "!!a A. [| <a,b> : f; f: Pi(A,B) |] ==> a : A";
by (REPEAT (ares_tac [fun_is_rel RS subsetD RS SigmaD1] 1));
-val domain_type = result();
+qed "domain_type";
goal ZF.thy "!!b B a. [| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)";
by (etac (fun_is_rel RS subsetD RS SigmaD2) 1);
by (assume_tac 1);
-val range_type = result();
+qed "range_type";
val prems = goal ZF.thy
"[| <a,b>: f; f: Pi(A,B); \
@@ -142,43 +142,43 @@
by (cut_facts_tac prems 1);
by (resolve_tac prems 1);
by (REPEAT (eresolve_tac [asm_rl,domain_type,range_type,apply_equality] 1));
-val Pair_mem_PiE = result();
+qed "Pair_mem_PiE";
(*** Lambda Abstraction ***)
goalw ZF.thy [lam_def] "!!A b. a:A ==> <a,b(a)> : (lam x:A. b(x))";
by (etac RepFunI 1);
-val lamI = result();
+qed "lamI";
val major::prems = goalw ZF.thy [lam_def]
"[| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P \
\ |] ==> P";
by (rtac (major RS RepFunE) 1);
by (REPEAT (ares_tac prems 1));
-val lamE = result();
+qed "lamE";
goal ZF.thy "!!a b c. [| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)";
by (REPEAT (eresolve_tac [asm_rl,lamE,Pair_inject,ssubst] 1));
-val lamD = result();
+qed "lamD";
val prems = goalw ZF.thy [lam_def, Pi_def, function_def]
"[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A.b(x)) : Pi(A,B)";
by (fast_tac (ZF_cs addIs prems) 1);
-val lam_type = result();
+qed "lam_type";
goal ZF.thy "(lam x:A.b(x)) : A -> {b(x). x:A}";
by (REPEAT (ares_tac [refl,lam_type,RepFunI] 1));
-val lam_funtype = result();
+qed "lam_funtype";
goal ZF.thy "!!a A. a : A ==> (lam x:A.b(x)) ` a = b(a)";
by (REPEAT (ares_tac [apply_equality,lam_funtype,lamI] 1));
-val beta = result();
+qed "beta";
(*congruence rule for lambda abstraction*)
val prems = goalw ZF.thy [lam_def]
"[| A=A'; !!x. x:A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')";
by (simp_tac (FOL_ss addsimps prems addcongs [RepFun_cong]) 1);
-val lam_cong = result();
+qed "lam_cong";
val [major] = goal ZF.thy
"(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)";
@@ -187,7 +187,7 @@
by (rtac (beta RS ssubst) 1);
by (assume_tac 1);
by (etac (major RS theI) 1);
-val lam_theI = result();
+qed "lam_theI";
(** Extensionality **)
@@ -201,19 +201,19 @@
by (etac ssubst 1);
by (resolve_tac (prems RL [ssubst]) 1);
by (REPEAT (ares_tac (prems@[apply_Pair,subsetD]) 1));
-val fun_subset = result();
+qed "fun_subset";
val prems = goal ZF.thy
"[| f : Pi(A,B); g: Pi(A,D); \
\ !!x. x:A ==> f`x = g`x |] ==> f=g";
by (REPEAT (ares_tac (prems @ (prems RL [sym]) @
[subset_refl,equalityI,fun_subset]) 1));
-val fun_extension = result();
+qed "fun_extension";
goal ZF.thy "!!f A B. f : Pi(A,B) ==> (lam x:A. f`x) = f";
by (rtac fun_extension 1);
by (REPEAT (ares_tac [lam_type,apply_type,beta] 1));
-val eta = result();
+qed "eta";
(*Every element of Pi(A,B) may be expressed as a lambda abstraction!*)
val prems = goal ZF.thy
@@ -223,20 +223,20 @@
by (resolve_tac prems 1);
by (rtac (eta RS sym) 2);
by (REPEAT (ares_tac (prems@[ballI,apply_type]) 1));
-val Pi_lamE = result();
+qed "Pi_lamE";
(** Images of functions **)
goalw ZF.thy [lam_def] "!!C A. C <= A ==> (lam x:A.b(x)) `` C = {b(x). x:C}";
by (fast_tac eq_cs 1);
-val image_lam = result();
+qed "image_lam";
goal ZF.thy "!!C A. [| f : Pi(A,B); C <= A |] ==> f``C = {f`x. x:C}";
by (etac (eta RS subst) 1);
by (asm_full_simp_tac (FOL_ss addsimps [beta, image_lam, subset_iff]
addcongs [RepFun_cong]) 1);
-val image_fun = result();
+qed "image_fun";
(*** properties of "restrict" ***)
@@ -244,39 +244,39 @@
goalw ZF.thy [restrict_def,lam_def]
"!!f A. [| f: Pi(C,B); A<=C |] ==> restrict(f,A) <= f";
by (fast_tac (ZF_cs addIs [apply_Pair]) 1);
-val restrict_subset = result();
+qed "restrict_subset";
val prems = goalw ZF.thy [restrict_def]
"[| !!x. x:A ==> f`x: B(x) |] ==> restrict(f,A) : Pi(A,B)";
by (rtac lam_type 1);
by (eresolve_tac prems 1);
-val restrict_type = result();
+qed "restrict_type";
val [pi,subs] = goal ZF.thy
"[| f: Pi(C,B); A<=C |] ==> restrict(f,A) : Pi(A,B)";
by (rtac (pi RS apply_type RS restrict_type) 1);
by (etac (subs RS subsetD) 1);
-val restrict_type2 = result();
+qed "restrict_type2";
goalw ZF.thy [restrict_def] "!!a A. a : A ==> restrict(f,A) ` a = f`a";
by (etac beta 1);
-val restrict = result();
+qed "restrict";
(*NOT SAFE as a congruence rule for the simplifier! Can cause it to fail!*)
val prems = goalw ZF.thy [restrict_def]
"[| A=B; !!x. x:B ==> f`x=g`x |] ==> restrict(f,A) = restrict(g,B)";
by (REPEAT (ares_tac (prems@[lam_cong]) 1));
-val restrict_eqI = result();
+qed "restrict_eqI";
goalw ZF.thy [restrict_def, lam_def] "domain(restrict(f,C)) = C";
by (fast_tac eq_cs 1);
-val domain_restrict = result();
+qed "domain_restrict";
val [prem] = goalw ZF.thy [restrict_def]
"A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))";
by (rtac (refl RS lam_cong) 1);
by (etac (prem RS subsetD RS beta) 1); (*easier than calling simp_tac*)
-val restrict_lam_eq = result();
+qed "restrict_lam_eq";
@@ -289,14 +289,14 @@
\ ALL x:S. ALL y:S. x<=y | y<=x |] ==> \
\ function(Union(S))";
by (fast_tac (ZF_cs addSDs [bspec RS bspec]) 1);
-val function_Union = result();
+qed "function_Union";
goalw ZF.thy [Pi_def]
"!!S. [| ALL f:S. EX C D. f:C->D; \
\ ALL f:S. ALL y:S. f<=y | y<=f |] ==> \
\ Union(S) : domain(Union(S)) -> range(Union(S))";
by (fast_tac (subset_cs addSIs [rel_Union, function_Union]) 1);
-val fun_Union = result();
+qed "fun_Union";
(** The Union of 2 disjoint functions is a function **)
@@ -319,37 +319,37 @@
by (REPEAT_FIRST (dtac (spec RS spec)));
by (deepen_tac ZF_cs 1 1);
by (deepen_tac ZF_cs 1 1);
-val fun_disjoint_Un = result();
+qed "fun_disjoint_Un";
goal ZF.thy
"!!f g a. [| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \
\ (f Un g)`a = f`a";
by (rtac (apply_Pair RS UnI1 RS apply_equality) 1);
by (REPEAT (ares_tac [fun_disjoint_Un] 1));
-val fun_disjoint_apply1 = result();
+qed "fun_disjoint_apply1";
goal ZF.thy
"!!f g c. [| c:C; f: A->B; g: C->D; A Int C = 0 |] ==> \
\ (f Un g)`c = g`c";
by (rtac (apply_Pair RS UnI2 RS apply_equality) 1);
by (REPEAT (ares_tac [fun_disjoint_Un] 1));
-val fun_disjoint_apply2 = result();
+qed "fun_disjoint_apply2";
(** Domain and range of a function/relation **)
goalw ZF.thy [Pi_def] "!!f. f : Pi(A,B) ==> domain(f)=A";
by (fast_tac eq_cs 1);
-val domain_of_fun = result();
+qed "domain_of_fun";
goal ZF.thy "!!f. [| f : Pi(A,B); a: A |] ==> f`a : range(f)";
by (etac (apply_Pair RS rangeI) 1);
by (assume_tac 1);
-val apply_rangeI = result();
+qed "apply_rangeI";
val [major] = goal ZF.thy "f : Pi(A,B) ==> f : A->range(f)";
by (rtac (major RS Pi_type) 1);
by (etac (major RS apply_rangeI) 1);
-val range_of_fun = result();
+qed "range_of_fun";
(*** Extensions of functions ***)
@@ -358,24 +358,24 @@
by (forward_tac [single_fun RS fun_disjoint_Un] 1);
by (asm_full_simp_tac (FOL_ss addsimps [cons_eq]) 2);
by (fast_tac eq_cs 1);
-val fun_extend = result();
+qed "fun_extend";
goal ZF.thy
"!!f A B. [| f: A->B; c~:A; b: B |] ==> cons(<c,b>,f) : cons(c,A) -> B";
by (fast_tac (ZF_cs addEs [fun_extend RS fun_weaken_type]) 1);
-val fun_extend3 = result();
+qed "fun_extend3";
goal ZF.thy "!!f A B. [| f: A->B; a:A; c~:A |] ==> cons(<c,b>,f)`a = f`a";
by (rtac (apply_Pair RS consI2 RS apply_equality) 1);
by (rtac fun_extend 3);
by (REPEAT (assume_tac 1));
-val fun_extend_apply1 = result();
+qed "fun_extend_apply1";
goal ZF.thy "!!f A B. [| f: A->B; c~:A |] ==> cons(<c,b>,f)`c = b";
by (rtac (consI1 RS apply_equality) 1);
by (rtac fun_extend 1);
by (REPEAT (assume_tac 1));
-val fun_extend_apply2 = result();
+qed "fun_extend_apply2";
(*For Finite.ML. Inclusion of right into left is easy*)
goal ZF.thy
@@ -393,5 +393,5 @@
by (ALLGOALS
(asm_simp_tac (FOL_ss addsimps [restrict, fun_extend_apply1,
fun_extend_apply2])));
-val cons_fun_eq = result();
+qed "cons_fun_eq";
--- a/src/ZF/mono.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/mono.ML Wed Dec 07 13:12:04 1994 +0100
@@ -12,121 +12,121 @@
would have to be single-valued*)
goal ZF.thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)";
by (fast_tac (ZF_cs addSEs [ReplaceE]) 1);
-val Replace_mono = result();
+qed "Replace_mono";
goal ZF.thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}";
by (fast_tac ZF_cs 1);
-val RepFun_mono = result();
+qed "RepFun_mono";
goal ZF.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)";
by (fast_tac ZF_cs 1);
-val Pow_mono = result();
+qed "Pow_mono";
goal ZF.thy "!!A B. A<=B ==> Union(A) <= Union(B)";
by (fast_tac ZF_cs 1);
-val Union_mono = result();
+qed "Union_mono";
val prems = goal ZF.thy
"[| A<=C; !!x. x:A ==> B(x)<=D(x) \
\ |] ==> (UN x:A. B(x)) <= (UN x:C. D(x))";
by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1);
-val UN_mono = result();
+qed "UN_mono";
(*Intersection is ANTI-monotonic. There are TWO premises! *)
goal ZF.thy "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)";
by (fast_tac ZF_cs 1);
-val Inter_anti_mono = result();
+qed "Inter_anti_mono";
goal ZF.thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)";
by (fast_tac ZF_cs 1);
-val cons_mono = result();
+qed "cons_mono";
goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D";
by (fast_tac ZF_cs 1);
-val Un_mono = result();
+qed "Un_mono";
goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D";
by (fast_tac ZF_cs 1);
-val Int_mono = result();
+qed "Int_mono";
goal ZF.thy "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D";
by (fast_tac ZF_cs 1);
-val Diff_mono = result();
+qed "Diff_mono";
(** Standard products, sums and function spaces **)
goal ZF.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \
\ Sigma(A,B) <= Sigma(C,D)";
by (fast_tac ZF_cs 1);
-val Sigma_mono_lemma = result();
+qed "Sigma_mono_lemma";
val Sigma_mono = ballI RSN (2,Sigma_mono_lemma);
goalw Sum.thy 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));
-val sum_mono = result();
+qed "sum_mono";
(*Note that B->A and C->A are typically disjoint!*)
goal ZF.thy "!!A B C. B<=C ==> A->B <= A->C";
by (fast_tac (ZF_cs addIs [lam_type] addEs [Pi_lamE]) 1);
-val Pi_mono = result();
+qed "Pi_mono";
goalw ZF.thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)";
by (etac RepFun_mono 1);
-val lam_mono = result();
+qed "lam_mono";
(** Quine-inspired ordered pairs, products, injections and sums **)
goalw QPair.thy [QPair_def] "!!a b c d. [| a<=c; b<=d |] ==> <a;b> <= <c;d>";
by (REPEAT (ares_tac [sum_mono] 1));
-val QPair_mono = result();
+qed "QPair_mono";
goal QPair.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \
\ QSigma(A,B) <= QSigma(C,D)";
by (fast_tac qpair_cs 1);
-val QSigma_mono_lemma = result();
+qed "QSigma_mono_lemma";
val QSigma_mono = ballI RSN (2,QSigma_mono_lemma);
goalw QPair.thy [QInl_def] "!!a b. a<=b ==> QInl(a) <= QInl(b)";
by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1));
-val QInl_mono = result();
+qed "QInl_mono";
goalw QPair.thy [QInr_def] "!!a b. a<=b ==> QInr(a) <= QInr(b)";
by (REPEAT (ares_tac [subset_refl RS QPair_mono] 1));
-val QInr_mono = result();
+qed "QInr_mono";
goal QPair.thy "!!A B C D. [| A<=C; B<=D |] ==> A <+> B <= C <+> D";
by (fast_tac qsum_cs 1);
-val qsum_mono = result();
+qed "qsum_mono";
(** Converse, domain, range, field **)
goal ZF.thy "!!r s. r<=s ==> converse(r) <= converse(s)";
by (fast_tac ZF_cs 1);
-val converse_mono = result();
+qed "converse_mono";
goal ZF.thy "!!r s. r<=s ==> domain(r)<=domain(s)";
by (fast_tac ZF_cs 1);
-val domain_mono = result();
+qed "domain_mono";
val domain_rel_subset =
[domain_mono, domain_subset] MRS subset_trans |> standard;
goal ZF.thy "!!r s. r<=s ==> range(r)<=range(s)";
by (fast_tac ZF_cs 1);
-val range_mono = result();
+qed "range_mono";
val range_rel_subset =
[range_mono, range_subset] MRS subset_trans |> standard;
goal ZF.thy "!!r s. r<=s ==> field(r)<=field(s)";
by (fast_tac ZF_cs 1);
-val field_mono = result();
+qed "field_mono";
goal ZF.thy "!!r A. r <= A*A ==> field(r) <= A";
by (etac (field_mono RS subset_trans) 1);
by (fast_tac ZF_cs 1);
-val field_rel_subset = result();
+qed "field_rel_subset";
(** Images **)
@@ -134,58 +134,58 @@
val [prem1,prem2] = goal ZF.thy
"[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A <= s``B";
by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1);
-val image_pair_mono = result();
+qed "image_pair_mono";
val [prem1,prem2] = goal ZF.thy
"[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A <= s-``B";
by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1);
-val vimage_pair_mono = result();
+qed "vimage_pair_mono";
goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B";
by (fast_tac ZF_cs 1);
-val image_mono = result();
+qed "image_mono";
goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B";
by (fast_tac ZF_cs 1);
-val vimage_mono = result();
+qed "vimage_mono";
val [sub,PQimp] = goal ZF.thy
"[| A<=B; !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)";
by (fast_tac (ZF_cs addIs [sub RS subsetD, PQimp RS mp]) 1);
-val Collect_mono = result();
+qed "Collect_mono";
(** Monotonicity of implications -- some could go to FOL **)
goal ZF.thy "!!A B x. A<=B ==> x:A --> x:B";
by (fast_tac ZF_cs 1);
-val in_mono = result();
+qed "in_mono";
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)";
by (Int.fast_tac 1);
-val conj_mono = result();
+qed "conj_mono";
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)";
by (Int.fast_tac 1);
-val disj_mono = result();
+qed "disj_mono";
goal IFOL.thy "!!P1 P2 Q1 Q2.[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)";
by (Int.fast_tac 1);
-val imp_mono = result();
+qed "imp_mono";
goal IFOL.thy "P-->P";
by (rtac impI 1);
by (assume_tac 1);
-val imp_refl = result();
+qed "imp_refl";
val [PQimp] = goal IFOL.thy
"[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))";
by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1);
-val ex_mono = result();
+qed "ex_mono";
val [PQimp] = goal IFOL.thy
"[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))";
by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1);
-val all_mono = result();
+qed "all_mono";
(*Used in intr_elim.ML and in individual datatype definitions*)
val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono,
--- a/src/ZF/pair.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/pair.ML Wed Dec 07 13:12:04 1994 +0100
@@ -8,38 +8,38 @@
(** Lemmas for showing that <a,b> uniquely determines a and b **)
-val doubleton_iff = prove_goal ZF.thy
+qed_goal "doubleton_iff" ZF.thy
"{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)"
(fn _=> [ (resolve_tac [extension RS iff_trans] 1),
(fast_tac upair_cs 1) ]);
-val Pair_iff = prove_goalw ZF.thy [Pair_def]
+qed_goalw "Pair_iff" ZF.thy [Pair_def]
"<a,b> = <c,d> <-> a=c & b=d"
(fn _=> [ (simp_tac (FOL_ss addsimps [doubleton_iff]) 1),
(fast_tac FOL_cs 1) ]);
val Pair_inject = standard (Pair_iff RS iffD1 RS conjE);
-val Pair_inject1 = prove_goal ZF.thy "<a,b> = <c,d> ==> a=c"
+qed_goal "Pair_inject1" ZF.thy "<a,b> = <c,d> ==> a=c"
(fn [major]=>
[ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]);
-val Pair_inject2 = prove_goal ZF.thy "<a,b> = <c,d> ==> b=d"
+qed_goal "Pair_inject2" ZF.thy "<a,b> = <c,d> ==> b=d"
(fn [major]=>
[ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]);
-val Pair_neq_0 = prove_goalw ZF.thy [Pair_def] "<a,b>=0 ==> P"
+qed_goalw "Pair_neq_0" ZF.thy [Pair_def] "<a,b>=0 ==> P"
(fn [major]=>
[ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1),
(rtac consI1 1) ]);
-val Pair_neq_fst = prove_goalw ZF.thy [Pair_def] "<a,b>=a ==> P"
+qed_goalw "Pair_neq_fst" ZF.thy [Pair_def] "<a,b>=a ==> P"
(fn [major]=>
[ (rtac (consI1 RS mem_asym RS FalseE) 1),
(rtac (major RS subst) 1),
(rtac consI1 1) ]);
-val Pair_neq_snd = prove_goalw ZF.thy [Pair_def] "<a,b>=b ==> P"
+qed_goalw "Pair_neq_snd" ZF.thy [Pair_def] "<a,b>=b ==> P"
(fn [major]=>
[ (rtac (consI1 RS consI2 RS mem_asym RS FalseE) 1),
(rtac (major RS subst) 1),
@@ -49,12 +49,12 @@
(*** Sigma: Disjoint union of a family of sets
Generalizes Cartesian product ***)
-val SigmaI = prove_goalw ZF.thy [Sigma_def]
+qed_goalw "SigmaI" ZF.thy [Sigma_def]
"[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"
(fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
(*The general elimination rule*)
-val SigmaE = prove_goalw ZF.thy [Sigma_def]
+qed_goalw "SigmaE" ZF.thy [Sigma_def]
"[| c: Sigma(A,B); \
\ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \
\ |] ==> P"
@@ -63,12 +63,12 @@
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);
(** Elimination of <a,b>:A*B -- introduces no eigenvariables **)
-val SigmaD1 = prove_goal ZF.thy "<a,b> : Sigma(A,B) ==> a : A"
+qed_goal "SigmaD1" ZF.thy "<a,b> : Sigma(A,B) ==> a : A"
(fn [major]=>
[ (rtac (major RS SigmaE) 1),
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
-val SigmaD2 = prove_goal ZF.thy "<a,b> : Sigma(A,B) ==> b : B(a)"
+qed_goal "SigmaD2" ZF.thy "<a,b> : Sigma(A,B) ==> b : B(a)"
(fn [major]=>
[ (rtac (major RS SigmaE) 1),
(REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
@@ -77,7 +77,7 @@
rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac)
THEN prune_params_tac)
(read_instantiate [("c","<a,b>")] SigmaE); *)
-val SigmaE2 = prove_goal ZF.thy
+qed_goal "SigmaE2" ZF.thy
"[| <a,b> : Sigma(A,B); \
\ [| a:A; b:B(a) |] ==> P \
\ |] ==> P"
@@ -86,26 +86,26 @@
(rtac (major RS SigmaD1) 1),
(rtac (major RS SigmaD2) 1) ]);
-val Sigma_cong = prove_goalw ZF.thy [Sigma_def]
+qed_goalw "Sigma_cong" ZF.thy [Sigma_def]
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \
\ Sigma(A,B) = Sigma(A',B')"
(fn prems=> [ (simp_tac (FOL_ss addsimps prems addcongs [RepFun_cong]) 1) ]);
-val Sigma_empty1 = prove_goal ZF.thy "Sigma(0,B) = 0"
+qed_goal "Sigma_empty1" ZF.thy "Sigma(0,B) = 0"
(fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]);
-val Sigma_empty2 = prove_goal ZF.thy "A*0 = 0"
+qed_goal "Sigma_empty2" ZF.thy "A*0 = 0"
(fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]);
(*** Eliminator - split ***)
-val split = prove_goalw ZF.thy [split_def]
+qed_goalw "split" ZF.thy [split_def]
"split(%x y.c(x,y), <a,b>) = c(a,b)"
(fn _ =>
[ (fast_tac (upair_cs addIs [the_equality] addSEs [Pair_inject]) 1) ]);
-val split_type = prove_goal ZF.thy
+qed_goal "split_type" ZF.thy
"[| p:Sigma(A,B); \
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \
\ |] ==> split(%x y.c(x,y), p) : C(p)"
@@ -121,15 +121,15 @@
by (etac SigmaE 1);
by (asm_simp_tac (FOL_ss addsimps [split]) 1);
by (fast_tac (upair_cs addSEs [Pair_inject]) 1);
-val expand_split = result();
+qed "expand_split";
(*** conversions for fst and snd ***)
-val fst_conv = prove_goalw ZF.thy [fst_def] "fst(<a,b>) = a"
+qed_goalw "fst_conv" ZF.thy [fst_def] "fst(<a,b>) = a"
(fn _=> [ (rtac split 1) ]);
-val snd_conv = prove_goalw ZF.thy [snd_def] "snd(<a,b>) = b"
+qed_goalw "snd_conv" ZF.thy [snd_def] "snd(<a,b>) = b"
(fn _=> [ (rtac split 1) ]);
@@ -137,17 +137,17 @@
goalw ZF.thy [fsplit_def] "!!R a b. R(a,b) ==> fsplit(R, <a,b>)";
by (REPEAT (ares_tac [refl,exI,conjI] 1));
-val fsplitI = result();
+qed "fsplitI";
val major::prems = goalw ZF.thy [fsplit_def]
"[| fsplit(R,z); !!x y. [| z = <x,y>; R(x,y) |] ==> P |] ==> P";
by (cut_facts_tac [major] 1);
by (REPEAT (eresolve_tac (prems@[asm_rl,exE,conjE]) 1));
-val fsplitE = result();
+qed "fsplitE";
goal ZF.thy "!!R a b. fsplit(R,<a,b>) ==> R(a,b)";
by (REPEAT (eresolve_tac [asm_rl,fsplitE,Pair_inject,ssubst] 1));
-val fsplitD = result();
+qed "fsplitD";
val pair_cs = upair_cs
addSIs [SigmaI]
@@ -172,7 +172,7 @@
goal ZF.thy "{x.x:A} = A";
by (fast_tac (pair_cs addSIs [equalityI]) 1);
-val triv_RepFun = result();
+qed "triv_RepFun";
(*INCLUDED: should be??*)
val bquant_simps = map prove_fun
--- a/src/ZF/subset.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/subset.ML Wed Dec 07 13:12:04 1994 +0100
@@ -9,43 +9,43 @@
(*** cons ***)
-val cons_subsetI = prove_goal ZF.thy "[| a:C; B<=C |] ==> cons(a,B) <= C"
+qed_goal "cons_subsetI" ZF.thy "[| a:C; B<=C |] ==> cons(a,B) <= C"
(fn prems=>
[ (cut_facts_tac prems 1),
(REPEAT (ares_tac [subsetI] 1
ORELSE eresolve_tac [consE,ssubst,subsetD] 1)) ]);
-val subset_consI = prove_goal ZF.thy "B <= cons(a,B)"
+qed_goal "subset_consI" ZF.thy "B <= cons(a,B)"
(fn _=> [ (rtac subsetI 1), (etac consI2 1) ]);
(*Useful for rewriting!*)
-val cons_subset_iff = prove_goal ZF.thy "cons(a,B)<=C <-> a:C & B<=C"
+qed_goal "cons_subset_iff" ZF.thy "cons(a,B)<=C <-> a:C & B<=C"
(fn _=> [ (fast_tac upair_cs 1) ]);
(*A safe special case of subset elimination, adding no new variables
[| cons(a,B) <= C; [| a : C; B <= C |] ==> R |] ==> R *)
val cons_subsetE = standard (cons_subset_iff RS iffD1 RS conjE);
-val subset_empty_iff = prove_goal ZF.thy "A<=0 <-> A=0"
+qed_goal "subset_empty_iff" ZF.thy "A<=0 <-> A=0"
(fn _=> [ (fast_tac (upair_cs addIs [equalityI]) 1) ]);
-val subset_cons_iff = prove_goal ZF.thy
+qed_goal "subset_cons_iff" ZF.thy
"C<=cons(a,B) <-> C<=B | (a:C & C-{a} <= B)"
(fn _=> [ (fast_tac upair_cs 1) ]);
(*** succ ***)
-val subset_succI = prove_goal ZF.thy "i <= succ(i)"
+qed_goal "subset_succI" ZF.thy "i <= succ(i)"
(fn _=> [ (rtac subsetI 1), (etac succI2 1) ]);
(*But if j is an ordinal or is transitive, then i:j implies i<=j!
See ordinal/Ord_succ_subsetI*)
-val succ_subsetI = prove_goalw ZF.thy [succ_def]
+qed_goalw "succ_subsetI" ZF.thy [succ_def]
"[| i:j; i<=j |] ==> succ(i)<=j"
(fn prems=>
[ (REPEAT (ares_tac (prems@[cons_subsetI]) 1)) ]);
-val succ_subsetE = prove_goalw ZF.thy [succ_def]
+qed_goalw "succ_subsetE" ZF.thy [succ_def]
"[| succ(i) <= j; [| i:j; i<=j |] ==> P \
\ |] ==> P"
(fn major::prems=>
@@ -54,25 +54,25 @@
(*** singletons ***)
-val singleton_subsetI = prove_goal ZF.thy
+qed_goal "singleton_subsetI" ZF.thy
"a:C ==> {a} <= C"
(fn prems=>
[ (REPEAT (resolve_tac (prems@[cons_subsetI,empty_subsetI]) 1)) ]);
-val singleton_subsetD = prove_goal ZF.thy
+qed_goal "singleton_subsetD" ZF.thy
"{a} <= C ==> a:C"
(fn prems=> [ (REPEAT (ares_tac (prems@[cons_subsetE]) 1)) ]);
(*** Big Union -- least upper bound of a set ***)
-val Union_subset_iff = prove_goal ZF.thy "Union(A) <= C <-> (ALL x:A. x <= C)"
+qed_goal "Union_subset_iff" ZF.thy "Union(A) <= C <-> (ALL x:A. x <= C)"
(fn _ => [ fast_tac upair_cs 1 ]);
-val Union_upper = prove_goal ZF.thy
+qed_goal "Union_upper" ZF.thy
"B:A ==> B <= Union(A)"
(fn prems=> [ (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1)) ]);
-val Union_least = prove_goal ZF.thy
+qed_goal "Union_least" ZF.thy
"[| !!x. x:A ==> x<=C |] ==> Union(A) <= C"
(fn [prem]=>
[ (rtac (ballI RS (Union_subset_iff RS iffD2)) 1),
@@ -82,17 +82,17 @@
goal ZF.thy "A <= (UN i:I. B(i)) <-> A = (UN i:I. A Int B(i))";
by (fast_tac (upair_cs addSIs [equalityI] addSEs [equalityE]) 1);
-val subset_UN_iff_eq = result();
+qed "subset_UN_iff_eq";
-val UN_subset_iff = prove_goal ZF.thy
+qed_goal "UN_subset_iff" ZF.thy
"(UN x:A.B(x)) <= C <-> (ALL x:A. B(x) <= C)"
(fn _ => [ fast_tac upair_cs 1 ]);
-val UN_upper = prove_goal ZF.thy
+qed_goal "UN_upper" ZF.thy
"!!x A. x:A ==> B(x) <= (UN x:A.B(x))"
(fn _ => [ etac (RepFunI RS Union_upper) 1 ]);
-val UN_least = prove_goal ZF.thy
+qed_goal "UN_least" ZF.thy
"[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A.B(x)) <= C"
(fn [prem]=>
[ (rtac (ballI RS (UN_subset_iff RS iffD2)) 1),
@@ -101,16 +101,16 @@
(*** Big Intersection -- greatest lower bound of a nonempty set ***)
-val Inter_subset_iff = prove_goal ZF.thy
+qed_goal "Inter_subset_iff" ZF.thy
"!!a A. a: A ==> C <= Inter(A) <-> (ALL x:A. C <= x)"
(fn _ => [ fast_tac upair_cs 1 ]);
-val Inter_lower = prove_goal ZF.thy "B:A ==> Inter(A) <= B"
+qed_goal "Inter_lower" ZF.thy "B:A ==> Inter(A) <= B"
(fn prems=>
[ (REPEAT (resolve_tac (prems@[subsetI]) 1
ORELSE etac InterD 1)) ]);
-val Inter_greatest = prove_goal ZF.thy
+qed_goal "Inter_greatest" ZF.thy
"[| a:A; !!x. x:A ==> C<=x |] ==> C <= Inter(A)"
(fn [prem1,prem2]=>
[ (rtac ([prem1, ballI] MRS (Inter_subset_iff RS iffD2)) 1),
@@ -118,12 +118,12 @@
(*** Intersection of a family of sets ***)
-val INT_lower = prove_goal ZF.thy
+qed_goal "INT_lower" ZF.thy
"x:A ==> (INT x:A.B(x)) <= B(x)"
(fn [prem] =>
[ rtac (prem RS RepFunI RS Inter_lower) 1 ]);
-val INT_greatest = prove_goal ZF.thy
+qed_goal "INT_greatest" ZF.thy
"[| a:A; !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A.B(x))"
(fn [nonempty,prem] =>
[ rtac (nonempty RS RepFunI RS Inter_greatest) 1,
@@ -132,32 +132,32 @@
(*** Finite Union -- the least upper bound of 2 sets ***)
-val Un_subset_iff = prove_goal ZF.thy "A Un B <= C <-> A <= C & B <= C"
+qed_goal "Un_subset_iff" ZF.thy "A Un B <= C <-> A <= C & B <= C"
(fn _ => [ fast_tac upair_cs 1 ]);
-val Un_upper1 = prove_goal ZF.thy "A <= A Un B"
+qed_goal "Un_upper1" ZF.thy "A <= A Un B"
(fn _ => [ (REPEAT (ares_tac [subsetI,UnI1] 1)) ]);
-val Un_upper2 = prove_goal ZF.thy "B <= A Un B"
+qed_goal "Un_upper2" ZF.thy "B <= A Un B"
(fn _ => [ (REPEAT (ares_tac [subsetI,UnI2] 1)) ]);
-val Un_least = prove_goal ZF.thy "!!A B C. [| A<=C; B<=C |] ==> A Un B <= C"
+qed_goal "Un_least" ZF.thy "!!A B C. [| A<=C; B<=C |] ==> A Un B <= C"
(fn _ =>
[ (rtac (Un_subset_iff RS iffD2) 1),
(REPEAT (ares_tac [conjI] 1)) ]);
(*** Finite Intersection -- the greatest lower bound of 2 sets *)
-val Int_subset_iff = prove_goal ZF.thy "C <= A Int B <-> C <= A & C <= B"
+qed_goal "Int_subset_iff" ZF.thy "C <= A Int B <-> C <= A & C <= B"
(fn _ => [ fast_tac upair_cs 1 ]);
-val Int_lower1 = prove_goal ZF.thy "A Int B <= A"
+qed_goal "Int_lower1" ZF.thy "A Int B <= A"
(fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)) ]);
-val Int_lower2 = prove_goal ZF.thy "A Int B <= B"
+qed_goal "Int_lower2" ZF.thy "A Int B <= B"
(fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)) ]);
-val Int_greatest = prove_goal ZF.thy
+qed_goal "Int_greatest" ZF.thy
"!!A B C. [| C<=A; C<=B |] ==> C <= A Int B"
(fn prems=>
[ (rtac (Int_subset_iff RS iffD2) 1),
@@ -165,10 +165,10 @@
(*** Set difference *)
-val Diff_subset = prove_goal ZF.thy "A-B <= A"
+qed_goal "Diff_subset" ZF.thy "A-B <= A"
(fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac DiffE 1)) ]);
-val Diff_contains = prove_goal ZF.thy
+qed_goal "Diff_contains" ZF.thy
"[| C<=A; C Int B = 0 |] ==> C <= A-B"
(fn prems=>
[ (cut_facts_tac prems 1),
@@ -178,7 +178,7 @@
(** Collect **)
-val Collect_subset = prove_goal ZF.thy "Collect(A,P) <= A"
+qed_goal "Collect_subset" ZF.thy "Collect(A,P) <= A"
(fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac CollectD1 1)) ]);
(** RepFun **)
@@ -188,7 +188,7 @@
by (etac RepFunE 1);
by (etac ssubst 1);
by (eresolve_tac prems 1);
-val RepFun_subset = result();
+qed "RepFun_subset";
(*A more powerful claset for subset reasoning*)
val subset_cs = subset0_cs
--- a/src/ZF/upair.ML Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/upair.ML Wed Dec 07 13:12:04 1994 +0100
@@ -21,17 +21,17 @@
(*** Unordered pairs - Upair ***)
-val pairing = prove_goalw ZF.thy [Upair_def]
+qed_goalw "pairing" ZF.thy [Upair_def]
"c : Upair(a,b) <-> (c=a | c=b)"
(fn _ => [ (fast_tac (lemmas_cs addEs [Pow_neq_0, sym RS Pow_neq_0]) 1) ]);
-val UpairI1 = prove_goal ZF.thy "a : Upair(a,b)"
+qed_goal "UpairI1" ZF.thy "a : Upair(a,b)"
(fn _ => [ (rtac (refl RS disjI1 RS (pairing RS iffD2)) 1) ]);
-val UpairI2 = prove_goal ZF.thy "b : Upair(a,b)"
+qed_goal "UpairI2" ZF.thy "b : Upair(a,b)"
(fn _ => [ (rtac (refl RS disjI2 RS (pairing RS iffD2)) 1) ]);
-val UpairE = prove_goal ZF.thy
+qed_goal "UpairE" ZF.thy
"[| a : Upair(b,c); a=b ==> P; a=c ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS (pairing RS iffD1 RS disjE)) 1),
@@ -39,13 +39,13 @@
(*** Rules for binary union -- Un -- defined via Upair ***)
-val UnI1 = prove_goalw ZF.thy [Un_def] "c : A ==> c : A Un B"
+qed_goalw "UnI1" ZF.thy [Un_def] "c : A ==> c : A Un B"
(fn [prem]=> [ (rtac (prem RS (UpairI1 RS UnionI)) 1) ]);
-val UnI2 = prove_goalw ZF.thy [Un_def] "c : B ==> c : A Un B"
+qed_goalw "UnI2" ZF.thy [Un_def] "c : B ==> c : A Un B"
(fn [prem]=> [ (rtac (prem RS (UpairI2 RS UnionI)) 1) ]);
-val UnE = prove_goalw ZF.thy [Un_def]
+qed_goalw "UnE" ZF.thy [Un_def]
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS UnionE) 1),
@@ -53,7 +53,7 @@
(REPEAT (EVERY1 [resolve_tac prems, etac subst, assume_tac])) ]);
(*Stronger version of the rule above*)
-val UnE' = prove_goal ZF.thy
+qed_goal "UnE'" ZF.thy
"[| c : A Un B; c:A ==> P; [| c:B; c~:A |] ==> P |] ==> P"
(fn major::prems =>
[(rtac (major RS UnE) 1),
@@ -63,11 +63,11 @@
(swap_res_tac prems 1),
(etac notnotD 1)]);
-val Un_iff = prove_goal ZF.thy "c : A Un B <-> (c:A | c:B)"
+qed_goal "Un_iff" ZF.thy "c : A Un B <-> (c:A | c:B)"
(fn _ => [ (fast_tac (lemmas_cs addIs [UnI1,UnI2] addSEs [UnE]) 1) ]);
(*Classical introduction rule: no commitment to A vs B*)
-val UnCI = prove_goal ZF.thy "(c ~: B ==> c : A) ==> c : A Un B"
+qed_goal "UnCI" ZF.thy "(c ~: B ==> c : A) ==> c : A Un B"
(fn [prem]=>
[ (rtac (disjCI RS (Un_iff RS iffD2)) 1),
(etac prem 1) ]);
@@ -75,69 +75,69 @@
(*** Rules for small intersection -- Int -- defined via Upair ***)
-val IntI = prove_goalw ZF.thy [Int_def]
+qed_goalw "IntI" ZF.thy [Int_def]
"[| c : A; c : B |] ==> c : A Int B"
(fn prems=>
[ (REPEAT (resolve_tac (prems @ [UpairI1,InterI]) 1
ORELSE eresolve_tac [UpairE, ssubst] 1)) ]);
-val IntD1 = prove_goalw ZF.thy [Int_def] "c : A Int B ==> c : A"
+qed_goalw "IntD1" ZF.thy [Int_def] "c : A Int B ==> c : A"
(fn [major]=>
[ (rtac (UpairI1 RS (major RS InterD)) 1) ]);
-val IntD2 = prove_goalw ZF.thy [Int_def] "c : A Int B ==> c : B"
+qed_goalw "IntD2" ZF.thy [Int_def] "c : A Int B ==> c : B"
(fn [major]=>
[ (rtac (UpairI2 RS (major RS InterD)) 1) ]);
-val IntE = prove_goal ZF.thy
+qed_goal "IntE" ZF.thy
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"
(fn prems=>
[ (resolve_tac prems 1),
(REPEAT (resolve_tac (prems RL [IntD1,IntD2]) 1)) ]);
-val Int_iff = prove_goal ZF.thy "c : A Int B <-> (c:A & c:B)"
+qed_goal "Int_iff" ZF.thy "c : A Int B <-> (c:A & c:B)"
(fn _ => [ (fast_tac (lemmas_cs addSIs [IntI] addSEs [IntE]) 1) ]);
(*** Rules for set difference -- defined via Upair ***)
-val DiffI = prove_goalw ZF.thy [Diff_def]
+qed_goalw "DiffI" ZF.thy [Diff_def]
"[| c : A; c ~: B |] ==> c : A - B"
(fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI]) 1)) ]);
-val DiffD1 = prove_goalw ZF.thy [Diff_def]
+qed_goalw "DiffD1" ZF.thy [Diff_def]
"c : A - B ==> c : A"
(fn [major]=> [ (rtac (major RS CollectD1) 1) ]);
-val DiffD2 = prove_goalw ZF.thy [Diff_def]
+qed_goalw "DiffD2" ZF.thy [Diff_def]
"c : A - B ==> c ~: B"
(fn [major]=> [ (rtac (major RS CollectD2) 1) ]);
-val DiffE = prove_goal ZF.thy
+qed_goal "DiffE" ZF.thy
"[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"
(fn prems=>
[ (resolve_tac prems 1),
(REPEAT (ares_tac (prems RL [DiffD1, DiffD2]) 1)) ]);
-val Diff_iff = prove_goal ZF.thy "c : A-B <-> (c:A & c~:B)"
+qed_goal "Diff_iff" ZF.thy "c : A-B <-> (c:A & c~:B)"
(fn _ => [ (fast_tac (lemmas_cs addSIs [DiffI] addSEs [DiffE]) 1) ]);
(*** Rules for cons -- defined via Un and Upair ***)
-val consI1 = prove_goalw ZF.thy [cons_def] "a : cons(a,B)"
+qed_goalw "consI1" ZF.thy [cons_def] "a : cons(a,B)"
(fn _ => [ (rtac (UpairI1 RS UnI1) 1) ]);
-val consI2 = prove_goalw ZF.thy [cons_def] "a : B ==> a : cons(b,B)"
+qed_goalw "consI2" ZF.thy [cons_def] "a : B ==> a : cons(b,B)"
(fn [prem]=> [ (rtac (prem RS UnI2) 1) ]);
-val consE = prove_goalw ZF.thy [cons_def]
+qed_goalw "consE" ZF.thy [cons_def]
"[| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS UnE) 1),
(REPEAT (eresolve_tac (prems @ [UpairE]) 1)) ]);
(*Stronger version of the rule above*)
-val consE' = prove_goal ZF.thy
+qed_goal "consE'" ZF.thy
"[| a : cons(b,A); a=b ==> P; [| a:A; a~=b |] ==> P |] ==> P"
(fn major::prems =>
[(rtac (major RS consE) 1),
@@ -147,21 +147,21 @@
(swap_res_tac prems 1),
(etac notnotD 1)]);
-val cons_iff = prove_goal ZF.thy "a : cons(b,A) <-> (a=b | a:A)"
+qed_goal "cons_iff" ZF.thy "a : cons(b,A) <-> (a=b | a:A)"
(fn _ => [ (fast_tac (lemmas_cs addIs [consI1,consI2] addSEs [consE]) 1) ]);
(*Classical introduction rule*)
-val consCI = prove_goal ZF.thy "(a~:B ==> a=b) ==> a: cons(b,B)"
+qed_goal "consCI" ZF.thy "(a~:B ==> a=b) ==> a: cons(b,B)"
(fn [prem]=>
[ (rtac (disjCI RS (cons_iff RS iffD2)) 1),
(etac prem 1) ]);
(*** Singletons - using cons ***)
-val singletonI = prove_goal ZF.thy "a : {a}"
+qed_goal "singletonI" ZF.thy "a : {a}"
(fn _=> [ (rtac consI1 1) ]);
-val singletonE = prove_goal ZF.thy "[| a: {b}; a=b ==> P |] ==> P"
+qed_goal "singletonE" ZF.thy "[| a: {b}; a=b ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS consE) 1),
(REPEAT (eresolve_tac (prems @ [emptyE]) 1)) ]);
@@ -169,26 +169,26 @@
(*** Rules for Descriptions ***)
-val the_equality = prove_goalw ZF.thy [the_def]
+qed_goalw "the_equality" ZF.thy [the_def]
"[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a"
(fn [pa,eq] =>
[ (fast_tac (lemmas_cs addSIs [singletonI,pa] addIs [equalityI]
addEs [eq RS subst]) 1) ]);
(* Only use this if you already know EX!x. P(x) *)
-val the_equality2 = prove_goal ZF.thy
+qed_goal "the_equality2" ZF.thy
"!!P. [| EX! x. P(x); P(a) |] ==> (THE x. P(x)) = a"
(fn _ =>
[ (deepen_tac (lemmas_cs addSIs [the_equality]) 1 1) ]);
-val theI = prove_goal ZF.thy "EX! x. P(x) ==> P(THE x. P(x))"
+qed_goal "theI" ZF.thy "EX! x. P(x) ==> P(THE x. P(x))"
(fn [major]=>
[ (rtac (major RS ex1E) 1),
(resolve_tac [major RS the_equality2 RS ssubst] 1),
(REPEAT (assume_tac 1)) ]);
(*Easier to apply than theI: conclusion has only one occurrence of P*)
-val theI2 = prove_goal ZF.thy
+qed_goal "theI2" ZF.thy
"[| EX! x. P(x); !!x. P(x) ==> Q(x) |] ==> Q(THE x.P(x))"
(fn prems => [ resolve_tac prems 1,
rtac theI 1,
@@ -198,7 +198,7 @@
(THE x.P(x)) rewrites to (THE x. Q(x)) *)
(*If it's "undefined", it's zero!*)
-val the_0 = prove_goalw ZF.thy [the_def]
+qed_goalw "the_0" ZF.thy [the_def]
"!!P. ~ (EX! x. P(x)) ==> (THE x. P(x))=0"
(fn _ =>
[ (fast_tac (lemmas_cs addIs [equalityI] addSEs [ReplaceE]) 1) ]);
@@ -208,31 +208,31 @@
goalw ZF.thy [if_def] "if(True,a,b) = a";
by (fast_tac (lemmas_cs addIs [the_equality]) 1);
-val if_true = result();
+qed "if_true";
goalw ZF.thy [if_def] "if(False,a,b) = b";
by (fast_tac (lemmas_cs addIs [the_equality]) 1);
-val if_false = result();
+qed "if_false";
(*Never use with case splitting, or if P is known to be true or false*)
val prems = goalw ZF.thy [if_def]
"[| P<->Q; Q ==> a=c; ~Q ==> b=d |] ==> if(P,a,b) = if(Q,c,d)";
by (simp_tac (FOL_ss addsimps prems addcongs [conj_cong]) 1);
-val if_cong = result();
+qed "if_cong";
(*Not needed for rewriting, since P would rewrite to True anyway*)
goalw ZF.thy [if_def] "!!P. P ==> if(P,a,b) = a";
by (fast_tac (lemmas_cs addSIs [the_equality]) 1);
-val if_P = result();
+qed "if_P";
(*Not needed for rewriting, since P would rewrite to False anyway*)
goalw ZF.thy [if_def] "!!P. ~P ==> if(P,a,b) = b";
by (fast_tac (lemmas_cs addSIs [the_equality]) 1);
-val if_not_P = result();
+qed "if_not_P";
val if_ss = FOL_ss addsimps [if_true,if_false];
-val expand_if = prove_goal ZF.thy
+qed_goal "expand_if" ZF.thy
"P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))"
(fn _=> [ (excluded_middle_tac "Q" 1),
(asm_simp_tac if_ss 1),
@@ -242,13 +242,13 @@
"[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A";
by (excluded_middle_tac "P" 1);
by (ALLGOALS (asm_simp_tac (if_ss addsimps prems)));
-val if_type = result();
+qed "if_type";
(*** Foundation lemmas ***)
(*was called mem_anti_sym*)
-val mem_asym = prove_goal ZF.thy "!!P. [| a:b; b:a |] ==> P"
+qed_goal "mem_asym" ZF.thy "!!P. [| a:b; b:a |] ==> P"
(fn _=>
[ (res_inst_tac [("A1","{a,b}")] (foundation RS disjE) 1),
(etac equals0D 1),
@@ -257,58 +257,58 @@
addSEs [consE,equalityE]) 1) ]);
(*was called mem_anti_refl*)
-val mem_irrefl = prove_goal ZF.thy "a:a ==> P"
+qed_goal "mem_irrefl" ZF.thy "a:a ==> P"
(fn [major]=> [ (rtac (major RS (major RS mem_asym)) 1) ]);
-val mem_not_refl = prove_goal ZF.thy "a ~: a"
+qed_goal "mem_not_refl" ZF.thy "a ~: a"
(K [ (rtac notI 1), (etac mem_irrefl 1) ]);
(*Good for proving inequalities by rewriting*)
-val mem_imp_not_eq = prove_goal ZF.thy "!!a A. a:A ==> a ~= A"
+qed_goal "mem_imp_not_eq" ZF.thy "!!a A. a:A ==> a ~= A"
(fn _=> [ fast_tac (lemmas_cs addSEs [mem_irrefl]) 1 ]);
(*** Rules for succ ***)
-val succI1 = prove_goalw ZF.thy [succ_def] "i : succ(i)"
+qed_goalw "succI1" ZF.thy [succ_def] "i : succ(i)"
(fn _=> [ (rtac consI1 1) ]);
-val succI2 = prove_goalw ZF.thy [succ_def]
+qed_goalw "succI2" ZF.thy [succ_def]
"i : j ==> i : succ(j)"
(fn [prem]=> [ (rtac (prem RS consI2) 1) ]);
-val succE = prove_goalw ZF.thy [succ_def]
+qed_goalw "succE" ZF.thy [succ_def]
"[| i : succ(j); i=j ==> P; i:j ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS consE) 1),
(REPEAT (eresolve_tac prems 1)) ]);
-val succ_iff = prove_goal ZF.thy "i : succ(j) <-> i=j | i:j"
+qed_goal "succ_iff" ZF.thy "i : succ(j) <-> i=j | i:j"
(fn _ => [ (fast_tac (lemmas_cs addIs [succI1,succI2] addSEs [succE]) 1) ]);
(*Classical introduction rule*)
-val succCI = prove_goal ZF.thy "(i~:j ==> i=j) ==> i: succ(j)"
+qed_goal "succCI" ZF.thy "(i~:j ==> i=j) ==> i: succ(j)"
(fn [prem]=>
[ (rtac (disjCI RS (succ_iff RS iffD2)) 1),
(etac prem 1) ]);
-val succ_neq_0 = prove_goal ZF.thy "succ(n)=0 ==> P"
+qed_goal "succ_neq_0" ZF.thy "succ(n)=0 ==> P"
(fn [major]=>
[ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1),
(rtac succI1 1) ]);
(*Useful for rewriting*)
-val succ_not_0 = prove_goal ZF.thy "succ(n) ~= 0"
+qed_goal "succ_not_0" ZF.thy "succ(n) ~= 0"
(fn _=> [ (rtac notI 1), (etac succ_neq_0 1) ]);
(* succ(c) <= B ==> c : B *)
val succ_subsetD = succI1 RSN (2,subsetD);
-val succ_inject = prove_goal ZF.thy "!!m n. succ(m) = succ(n) ==> m=n"
+qed_goal "succ_inject" ZF.thy "!!m n. succ(m) = succ(n) ==> m=n"
(fn _ =>
[ (fast_tac (lemmas_cs addSEs [succE, equalityE, make_elim succ_subsetD]
addEs [mem_asym]) 1) ]);
-val succ_inject_iff = prove_goal ZF.thy "succ(m) = succ(n) <-> m=n"
+qed_goal "succ_inject_iff" ZF.thy "succ(m) = succ(n) <-> m=n"
(fn _=> [ (fast_tac (FOL_cs addSEs [succ_inject]) 1) ]);
(*UpairI1/2 should become UpairCI; mem_irrefl as a hazE? *)