src/ZF/Nat.ML
changeset 760 f0200e91b272
parent 484 70b789956bd3
child 829 ba28811c3496
--- 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";