replaced 'val ... = result()' by 'qed "..."'
--- a/Arith.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Arith.ML Mon Nov 21 17:50:34 1994 +0100
@@ -116,7 +116,7 @@
by (rtac (prem RS 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";
(*** Remainder ***)
@@ -125,7 +125,7 @@
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (etac less_SucE 3);
by (ALLGOALS(asm_simp_tac arith_ss));
-val diff_less_Suc = result();
+qed "diff_less_Suc";
(*In ordinary notation: if 0<n and n<=m then m-n < m *)
goal Arith.thy "!!m. [| 0<n; ~ m<n |] ==> m - n < m";
@@ -133,23 +133,23 @@
by (fast_tac HOL_cs 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS(asm_simp_tac(arith_ss addsimps [diff_less_Suc])));
-val div_termination = result();
+qed "div_termination";
val wf_less_trans = wf_pred_nat RS wf_trancl RSN (2, def_wfrec RS trans);
goalw Nat.thy [less_def] "<m,n> : pred_nat^+ = (m<n)";
by (rtac refl 1);
-val less_eq = result();
+qed "less_eq";
goal Arith.thy "!!m. m<n ==> m mod n = m";
by (rtac (mod_def RS wf_less_trans) 1);
by(asm_simp_tac HOL_ss 1);
-val mod_less = result();
+qed "mod_less";
goal Arith.thy "!!m. [| 0<n; ~m<n |] ==> m mod n = (m-n) mod n";
by (rtac (mod_def RS wf_less_trans) 1);
by(asm_simp_tac (nat_ss addsimps [div_termination, cut_apply, less_eq]) 1);
-val mod_geq = result();
+qed "mod_geq";
(*** Quotient ***)
@@ -157,12 +157,12 @@
goal Arith.thy "!!m. m<n ==> m div n = 0";
by (rtac (div_def RS wf_less_trans) 1);
by(asm_simp_tac nat_ss 1);
-val div_less = result();
+qed "div_less";
goal Arith.thy "!!M. [| 0<n; ~m<n |] ==> m div n = Suc((m-n) div n)";
by (rtac (div_def RS wf_less_trans) 1);
by(asm_simp_tac (nat_ss addsimps [div_termination, cut_apply, less_eq]) 1);
-val div_geq = result();
+qed "div_geq";
(*Main Result about quotient and remainder.*)
goal Arith.thy "!!m. 0<n ==> (m div n)*n + m mod n = m";
@@ -172,7 +172,7 @@
by (ALLGOALS (asm_simp_tac(arith_ss addsimps (add_ac @
[mod_less, mod_geq, div_less, div_geq,
add_diff_inverse, div_termination]))));
-val mod_div_equality = result();
+qed "mod_div_equality";
(*** More results about difference ***)
@@ -181,12 +181,12 @@
by (rtac (prem RS rev_mp) 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS (asm_simp_tac arith_ss));
-val less_imp_diff_is_0 = result();
+qed "less_imp_diff_is_0";
val prems = goal Arith.thy "m-n = 0 --> n-m = 0 --> m=n";
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (REPEAT(simp_tac arith_ss 1 THEN TRY(atac 1)));
-val diffs0_imp_equal_lemma = result();
+qed "diffs0_imp_equal_lemma";
(* [| m-n = 0; n-m = 0 |] ==> m=n *)
val diffs0_imp_equal = standard (diffs0_imp_equal_lemma RS mp RS mp);
@@ -195,29 +195,29 @@
by (rtac (prem RS rev_mp) 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS(asm_simp_tac arith_ss));
-val less_imp_diff_positive = result();
+qed "less_imp_diff_positive";
val [prem] = goal Arith.thy "n < Suc(m) ==> Suc(m)-n = Suc(m-n)";
by (rtac (prem RS rev_mp) 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS(asm_simp_tac arith_ss));
-val Suc_diff_n = result();
+qed "Suc_diff_n";
goal Arith.thy "Suc(m)-n = if(m<n, 0, Suc(m-n))";
by(simp_tac (nat_ss addsimps [less_imp_diff_is_0, not_less_eq, Suc_diff_n]
setloop (split_tac [expand_if])) 1);
-val if_Suc_diff_n = result();
+qed "if_Suc_diff_n";
goal Arith.thy "P(k) --> (!n. P(Suc(n))--> P(n)) --> P(k-i)";
by (res_inst_tac [("m","k"),("n","i")] diff_induct 1);
by (ALLGOALS (strip_tac THEN' simp_tac arith_ss THEN' TRY o fast_tac HOL_cs));
-val zero_induct_lemma = result();
+qed "zero_induct_lemma";
val prems = goal Arith.thy "[| P(k); !!n. P(Suc(n)) ==> P(n) |] ==> P(0)";
by (rtac (diff_self_eq_0 RS subst) 1);
by (rtac (zero_induct_lemma RS mp RS mp) 1);
by (REPEAT (ares_tac ([impI,allI]@prems) 1));
-val zero_induct = result();
+qed "zero_induct";
(*13 July 1992: loaded in 105.7s*)
@@ -228,12 +228,12 @@
by (ALLGOALS(simp_tac arith_ss));
by (etac le_trans 1);
by (rtac (lessI RS less_imp_le) 1);
-val le_add2 = result();
+qed "le_add2";
goal Arith.thy "n <= ((n + m)::nat)";
by (simp_tac (arith_ss addsimps add_ac) 1);
by (rtac le_add2 1);
-val le_add1 = result();
+qed "le_add1";
val less_add_Suc1 = standard (lessI RS (le_add1 RS le_less_trans));
val less_add_Suc2 = standard (lessI RS (le_add2 RS le_less_trans));
@@ -242,5 +242,5 @@
by (nat_ind_tac "k" 1);
by (ALLGOALS (asm_simp_tac arith_ss));
by (fast_tac (HOL_cs addDs [Suc_leD]) 1);
-val plus_leD1_lemma = result();
+qed "plus_leD1_lemma";
val plus_leD1 = plus_leD1_lemma RS mp;
--- a/Finite.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Finite.ML Mon Nov 21 17:50:34 1994 +0100
@@ -11,11 +11,11 @@
goalw Finite.thy Fin.defs "!!A B. A<=B ==> Fin(A) <= Fin(B)";
br lfp_mono 1;
by (REPEAT (ares_tac basic_monos 1));
-val Fin_mono = result();
+qed "Fin_mono";
goalw Finite.thy Fin.defs "Fin(A) <= Pow(A)";
by (fast_tac (set_cs addSIs [lfp_lowerbound]) 1);
-val Fin_subset_Pow = result();
+qed "Fin_subset_Pow";
(* A : Fin(B) ==> A <= B *)
val FinD = Fin_subset_Pow RS subsetD RS PowD;
@@ -29,7 +29,7 @@
by (excluded_middle_tac "a:b" 2);
by (etac (insert_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 **)
@@ -40,7 +40,7 @@
"[| F: Fin(A); G: Fin(A) |] ==> F Un G : Fin(A)";
by (rtac (major RS Fin_induct) 1);
by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Un_insert_left]))));
-val Fin_UnI = result();
+qed "Fin_UnI";
(*Every subset of a finite set is finite*)
val [subs,fin] = goal Finite.thy "[| A<=B; B: Fin(M) |] ==> A: Fin(M)";
@@ -52,7 +52,7 @@
by (safe_tac (set_cs addSDs [subset_insert_iff RS iffD1]));
by (eres_inst_tac [("t","C")] (insert_Diff RS subst) 2);
by (ALLGOALS (asm_simp_tac Fin_ss));
-val Fin_subset = result();
+qed "Fin_subset";
(*The image of a finite set is finite*)
val major::_ = goal Finite.thy
@@ -60,7 +60,7 @@
by (rtac (major RS Fin_induct) 1);
by (simp_tac Fin_ss 1);
by (asm_simp_tac (set_ss addsimps [image_eqI RS Fin.insertI, image_insert]) 1);
-val Fin_imageI = result();
+qed "Fin_imageI";
val major::prems = goal Finite.thy
"[| c: Fin(A); b: Fin(A); \
@@ -71,7 +71,7 @@
by (rtac (Diff_insert RS ssubst) 2);
by (ALLGOALS (asm_simp_tac
(Fin_ss addsimps (prems@[Diff_subset RS Fin_subset]))));
-val Fin_empty_induct_lemma = result();
+qed "Fin_empty_induct_lemma";
val prems = goal Finite.thy
"[| b: Fin(A); \
@@ -81,4 +81,4 @@
by (rtac (Diff_cancel RS subst) 1);
by (rtac (Fin_empty_induct_lemma RS mp) 1);
by (REPEAT (ares_tac (subset_refl::prems) 1));
-val Fin_empty_induct = result();
+qed "Fin_empty_induct";
--- a/Fun.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Fun.ML Mon Nov 21 17:50:34 1994 +0100
@@ -10,13 +10,13 @@
by (rtac iffI 1);
by(asm_simp_tac HOL_ss 1);
by(rtac ext 1 THEN asm_simp_tac HOL_ss 1);
-val expand_fun_eq = result();
+qed "expand_fun_eq";
val prems = goal Fun.thy
"[| f(x)=u; !!x. P(x) ==> g(f(x)) = x; P(x) |] ==> x=g(u)";
by (rtac (arg_cong RS box_equals) 1);
by (REPEAT (resolve_tac (prems@[refl]) 1));
-val apply_inverse = result();
+qed "apply_inverse";
(*** Range of a function ***)
@@ -24,7 +24,7 @@
(*Frequently b does not have the syntactic form of f(x).*)
val [prem] = goalw Fun.thy [range_def] "b=f(x) ==> b : range(f)";
by (EVERY1 [rtac CollectI, rtac exI, rtac prem]);
-val range_eqI = result();
+qed "range_eqI";
val rangeI = refl RS range_eqI;
@@ -32,13 +32,13 @@
"[| b : range(%x.f(x)); !!x. b=f(x) ==> P |] ==> P";
by (rtac (major RS CollectD RS exE) 1);
by (etac minor 1);
-val rangeE = result();
+qed "rangeE";
(*** Image of a set under a function ***)
val prems = goalw Fun.thy [image_def] "[| b=f(x); x:A |] ==> b : f``A";
by (REPEAT (resolve_tac (prems @ [CollectI,bexI,prem]) 1));
-val image_eqI = result();
+qed "image_eqI";
val imageI = refl RS image_eqI;
@@ -47,62 +47,62 @@
"[| b : (%x.f(x))``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P";
by (rtac (major RS CollectD RS bexE) 1);
by (REPEAT (ares_tac prems 1));
-val imageE = result();
+qed "imageE";
goalw Fun.thy [o_def] "(f o g)``r = f``(g``r)";
by (rtac set_ext 1);
by (fast_tac (HOL_cs addIs [imageI] addSEs [imageE]) 1);
-val image_compose = result();
+qed "image_compose";
goal Fun.thy "f``(A Un B) = f``A Un f``B";
by (rtac set_ext 1);
by (fast_tac (HOL_cs addIs [imageI,UnCI] addSEs [imageE,UnE]) 1);
-val image_Un = result();
+qed "image_Un";
(*** inj(f): f is a one-to-one function ***)
val prems = goalw Fun.thy [inj_def]
"[| !! x y. f(x) = f(y) ==> x=y |] ==> inj(f)";
by (fast_tac (HOL_cs addIs prems) 1);
-val injI = result();
+qed "injI";
val [major] = goal Fun.thy "(!!x. g(f(x)) = x) ==> inj(f)";
by (rtac injI 1);
by (etac (arg_cong RS box_equals) 1);
by (rtac major 1);
by (rtac major 1);
-val inj_inverseI = result();
+qed "inj_inverseI";
val [major,minor] = goalw Fun.thy [inj_def]
"[| inj(f); f(x) = f(y) |] ==> x=y";
by (rtac (major RS spec RS spec RS mp) 1);
by (rtac minor 1);
-val injD = result();
+qed "injD";
(*Useful with the simplifier*)
val [major] = goal Fun.thy "inj(f) ==> (f(x) = f(y)) = (x=y)";
by (rtac iffI 1);
by (etac (major RS injD) 1);
by (etac arg_cong 1);
-val inj_eq = result();
+qed "inj_eq";
val [major] = goal Fun.thy "inj(f) ==> (@x.f(x)=f(y)) = y";
by (rtac (major RS injD) 1);
by (rtac selectI 1);
by (rtac refl 1);
-val inj_select = result();
+qed "inj_select";
(*A one-to-one function has an inverse (given using select).*)
val [major] = goalw Fun.thy [Inv_def] "inj(f) ==> Inv(f,f(x)) = x";
by (EVERY1 [rtac (major RS inj_select)]);
-val Inv_f_f = result();
+qed "Inv_f_f";
(* Useful??? *)
val [oneone,minor] = goal Fun.thy
"[| inj(f); !!y. y: range(f) ==> P(Inv(f,y)) |] ==> P(x)";
by (res_inst_tac [("t", "x")] (oneone RS (Inv_f_f RS subst)) 1);
by (rtac (rangeI RS minor) 1);
-val inj_transfer = result();
+qed "inj_transfer";
(*** inj_onto(f,A): f is one-to-one over A ***)
@@ -110,27 +110,27 @@
val prems = goalw Fun.thy [inj_onto_def]
"(!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> inj_onto(f,A)";
by (fast_tac (HOL_cs addIs prems addSIs [ballI]) 1);
-val inj_ontoI = result();
+qed "inj_ontoI";
val [major] = goal Fun.thy
"(!!x. x:A ==> g(f(x)) = x) ==> inj_onto(f,A)";
by (rtac inj_ontoI 1);
by (etac (apply_inverse RS trans) 1);
by (REPEAT (eresolve_tac [asm_rl,major] 1));
-val inj_onto_inverseI = result();
+qed "inj_onto_inverseI";
val major::prems = goalw Fun.thy [inj_onto_def]
"[| inj_onto(f,A); f(x)=f(y); x:A; y:A |] ==> x=y";
by (rtac (major RS bspec RS bspec RS mp) 1);
by (REPEAT (resolve_tac prems 1));
-val inj_ontoD = result();
+qed "inj_ontoD";
val major::prems = goal Fun.thy
"[| inj_onto(f,A); ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)";
by (rtac contrapos 1);
by (etac (major RS inj_ontoD) 2);
by (REPEAT (resolve_tac prems 1));
-val inj_onto_contraD = result();
+qed "inj_onto_contraD";
(*** Lemmas about inj ***)
@@ -140,28 +140,28 @@
by (cut_facts_tac prems 1);
by (fast_tac (HOL_cs addIs [injI,rangeI]
addEs [injD,inj_ontoD]) 1);
-val comp_inj = result();
+qed "comp_inj";
val [prem] = goal Fun.thy "inj(f) ==> inj_onto(f,A)";
by (fast_tac (HOL_cs addIs [prem RS injD, inj_ontoI]) 1);
-val inj_imp = result();
+qed "inj_imp";
val [prem] = goalw Fun.thy [Inv_def] "y : range(f) ==> f(Inv(f,y)) = y";
by (EVERY1 [rtac (prem RS rangeE), rtac selectI, etac sym]);
-val f_Inv_f = result();
+qed "f_Inv_f";
val prems = goal Fun.thy
"[| Inv(f,x)=Inv(f,y); x: range(f); y: range(f) |] ==> x=y";
by (rtac (arg_cong RS box_equals) 1);
by (REPEAT (resolve_tac (prems @ [f_Inv_f]) 1));
-val Inv_injective = result();
+qed "Inv_injective";
val prems = goal Fun.thy
"[| inj(f); A<=range(f) |] ==> inj_onto(Inv(f), A)";
by (cut_facts_tac prems 1);
by (fast_tac (HOL_cs addIs [inj_ontoI]
addEs [Inv_injective,injD,subsetD]) 1);
-val inj_onto_Inv = result();
+qed "inj_onto_Inv";
(*** Set reasoning tools ***)
--- a/Gfp.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Gfp.ML Mon Nov 21 17:50:34 1994 +0100
@@ -15,27 +15,27 @@
val prems = goalw Gfp.thy [gfp_def] "[| X <= f(X) |] ==> X <= gfp(f)";
by (rtac (CollectI RS Union_upper) 1);
by (resolve_tac prems 1);
-val gfp_upperbound = result();
+qed "gfp_upperbound";
val prems = goalw Gfp.thy [gfp_def]
"[| !!u. u <= f(u) ==> u<=X |] ==> gfp(f) <= X";
by (REPEAT (ares_tac ([Union_least]@prems) 1));
by (etac CollectD 1);
-val gfp_least = result();
+qed "gfp_least";
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))";
by (EVERY1 [rtac gfp_least, rtac subset_trans, atac,
rtac (mono RS monoD), rtac gfp_upperbound, atac]);
-val gfp_lemma2 = result();
+qed "gfp_lemma2";
val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)";
by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD),
rtac gfp_lemma2, rtac mono]);
-val gfp_lemma3 = result();
+qed "gfp_lemma3";
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))";
by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1));
-val gfp_Tarski = result();
+qed "gfp_Tarski";
(*** Coinduction rules for greatest fixed points ***)
@@ -44,7 +44,7 @@
"[| a: X; X <= f(X) |] ==> a : gfp(f)";
by (rtac (gfp_upperbound RS subsetD) 1);
by (REPEAT (ares_tac prems 1));
-val weak_coinduct = result();
+qed "weak_coinduct";
val [prem,mono] = goal Gfp.thy
"[| X <= f(X Un gfp(f)); mono(f) |] ==> \
@@ -53,21 +53,21 @@
by (rtac (mono RS gfp_lemma2 RS subset_trans) 1);
by (rtac (Un_upper2 RS subset_trans) 1);
by (rtac (mono RS mono_Un) 1);
-val coinduct_lemma = result();
+qed "coinduct_lemma";
(*strong version, thanks to Coen & Frost*)
goal Gfp.thy
"!!X. [| mono(f); a: X; X <= f(X Un gfp(f)) |] ==> a : gfp(f)";
by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1);
by (REPEAT (ares_tac [UnI1, Un_least] 1));
-val coinduct = result();
+qed "coinduct";
val [mono,prem] = goal Gfp.thy
"[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))";
br (mono RS mono_Un RS subsetD) 1;
br (mono RS gfp_lemma2 RS subsetD RS UnI2) 1;
by (rtac prem 1);
-val gfp_fun_UnI2 = result();
+qed "gfp_fun_UnI2";
(*** Even Stronger version of coinduct [by Martin Coen]
- instead of the condition X <= f(X)
@@ -75,7 +75,7 @@
val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un X Un B)";
by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1));
-val coinduct3_mono_lemma= result();
+qed "coinduct3_mono_lemma";
val [prem,mono] = goal Gfp.thy
"[| X <= f(lfp(%x.f(x) Un X Un gfp(f))); mono(f) |] ==> \
@@ -89,7 +89,7 @@
by (rtac (mono RS monoD) 1);
by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1);
by (rtac Un_upper2 1);
-val coinduct3_lemma = result();
+qed "coinduct3_lemma";
val prems = goal Gfp.thy
"[| mono(f); a:X; X <= f(lfp(%x.f(x) Un X Un gfp(f))) |] ==> a : gfp(f)";
@@ -97,7 +97,7 @@
by (resolve_tac (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1);
by (rtac (UnI2 RS UnI1) 1);
by (REPEAT (resolve_tac prems 1));
-val coinduct3 = result();
+qed "coinduct3";
(** Definition forms of gfp_Tarski and coinduct, to control unfolding **)
@@ -105,13 +105,13 @@
val [rew,mono] = goal Gfp.thy "[| A==gfp(f); mono(f) |] ==> A = f(A)";
by (rewtac rew);
by (rtac (mono RS gfp_Tarski) 1);
-val def_gfp_Tarski = result();
+qed "def_gfp_Tarski";
val rew::prems = goal Gfp.thy
"[| A==gfp(f); mono(f); a:X; X <= f(X Un A) |] ==> a: A";
by (rewtac rew);
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct]) 1));
-val def_coinduct = result();
+qed "def_coinduct";
(*The version used in the induction/coinduction package*)
val prems = goal Gfp.thy
@@ -120,13 +120,13 @@
\ a : A";
by (rtac def_coinduct 1);
by (REPEAT (ares_tac (prems @ [subsetI,CollectI]) 1));
-val def_Collect_coinduct = result();
+qed "def_Collect_coinduct";
val rew::prems = goal Gfp.thy
"[| A==gfp(f); mono(f); a:X; X <= f(lfp(%x.f(x) Un X Un A)) |] ==> a: A";
by (rewtac rew);
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1));
-val def_coinduct3 = result();
+qed "def_coinduct3";
(*Monotonicity of gfp!*)
val prems = goal Gfp.thy
@@ -142,4 +142,4 @@
val [prem] = goal Gfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
br (gfp_upperbound RS gfp_least) 1;
be (prem RSN (2,subset_trans)) 1;
-val gfp_mono = result();
+qed "gfp_mono";
--- a/IMP/Equiv.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IMP/Equiv.ML Mon Nov 21 17:50:34 1994 +0100
@@ -9,18 +9,18 @@
by (ALLGOALS(simp_tac (HOL_ss addsimps A_simps))); (* rewr. Den. *)
by (TRYALL (fast_tac (set_cs addSIs (evala.intrs@prems)
addSEs evala_elim_cases)));
-val aexp_iff = result() RS spec;
+val aexp_iff = store_thm("aexp_iff", result() RS spec);
goal HOL.thy "(? x. x=t & P) = P";
by(fast_tac HOL_cs 1);
-val elim_ex_conv = result();
+qed "elim_ex_conv";
goal Equiv.thy "!w. (<b,s> -b-> w) = (w = B(b,s))";
by (bexp.induct_tac "b" 1);
by (ALLGOALS(asm_simp_tac (HOL_ss addcongs [conj_cong]
addsimps (aexp_iff::elim_ex_conv::B_simps@evalb_simps))));
-val bexp_iff = result() RS spec;
+val bexp_iff = store_thm("bexp_iff", result() RS spec);
val bexp1 = bexp_iff RS iffD1;
val bexp2 = bexp_iff RS iffD2;
@@ -51,7 +51,7 @@
by (rtac (rewrite_rule [Gamma_def] (Gamma_mono RS lfp_Tarski RS ssubst)) 1);
by (fast_tac (comp_cs addSIs [bexp1,BfstI] addSDs [BfstD,bexp1]) 1);
-val com1 = result();
+qed "com1";
val com_ss = prod_ss addsimps (aexp_iff::bexp_iff::evalc.intrs);
@@ -73,7 +73,7 @@
by (EVERY1 [dtac bspec, atac]);
by (ALLGOALS (asm_full_simp_tac com_ss));
-val com2 = result();
+qed "com2";
(**** Proof of Equivalence ****)
@@ -87,4 +87,4 @@
(* <= *)
by (REPEAT (step_tac com_iff_cs 1));
by (asm_full_simp_tac (prod_ss addsimps [surjective_pairing RS sym])1);
-val com_equivalence = result();
+qed "com_equivalence";
--- a/IOA/example/Correctness.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/example/Correctness.ML Mon Nov 21 17:50:34 1994 +0100
@@ -42,7 +42,7 @@
by(Action.action.induct_tac "a" 1);
by(ALLGOALS(simp_tac (action_ss addsimps
(actions_def :: asig_projections @ set_lemmas))));
-val externals_lemma = result();
+qed "externals_lemma";
val sels = [Sender.sbit_def, Sender.sq_def, Sender.ssending_def,
--- a/IOA/example/Impl.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/example/Impl.ML Mon Nov 21 17:50:34 1994 +0100
@@ -47,7 +47,7 @@
(* Instantiation of a tautology? *)
goal Packet.thy "!x. x = packet --> hdr(x) = hdr(packet)";
by (simp_tac (HOL_ss addsimps [Packet.hdr_def]) 1);
-val eq_packet_imp_eq_hdr = result();
+qed "eq_packet_imp_eq_hdr";
(* INVARIANT 1 *)
@@ -124,7 +124,7 @@
by (EVERY1[tac, tac, tac, tac, tac, tac]);
-val inv1 = result();
+qed "inv1";
@@ -197,7 +197,7 @@
by (dres_inst_tac [("k","hdr_sum(srch(s), sbit(sen(s)))")]
(standard(leq_add_leq RS mp)) 1);
by (asm_full_simp_tac HOL_ss 1);
-val inv2 = result();
+qed "inv2";
(* INVARIANT 3 *)
@@ -297,7 +297,7 @@
by (dtac mp 1);
by (assume_tac 1);
by (assume_tac 1);
-val inv3 = result();
+qed "inv3";
@@ -380,7 +380,7 @@
by (dtac (left_add_leq RS mp) 1);
by (asm_full_simp_tac list_ss 1);
by (asm_full_simp_tac arith_ss 1);
-val inv4 = result();
+qed "inv4";
--- a/IOA/example/Lemmas.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/example/Lemmas.ML Mon Nov 21 17:50:34 1994 +0100
@@ -11,37 +11,37 @@
(* Logic *)
val prems = goal HOL.thy "(P ==> Q-->R) ==> P&Q --> R";
by(fast_tac (HOL_cs addDs prems) 1);
-val imp_conj_lemma = result();
+qed "imp_conj_lemma";
goal HOL.thy "(P --> (? x. Q(x))) = (? x. P --> Q(x))";
by(fast_tac HOL_cs 1);
-val imp_ex_equiv = result();
+qed "imp_ex_equiv";
goal HOL.thy "(A --> B & C) = ((A --> B) & (A --> C))";
by (fast_tac HOL_cs 1);
-val fork_lemma = result();
+qed "fork_lemma";
goal HOL.thy "((A --> B) & (C --> B)) = ((A | C) --> B)";
by (fast_tac HOL_cs 1);
-val imp_or_lem = result();
+qed "imp_or_lem";
goal HOL.thy "(X = (~ Y)) = ((~X) = Y)";
by (fast_tac HOL_cs 1);
-val neg_flip = result();
+qed "neg_flip";
goal HOL.thy "P --> Q(M) --> Q(if(P,M,N))";
by (rtac impI 1);
by (rtac impI 1);
by (rtac (expand_if RS iffD2) 1);
by (fast_tac HOL_cs 1);
-val imp_true_decompose = result();
+qed "imp_true_decompose";
goal HOL.thy "(~P) --> Q(N) --> Q(if(P,M,N))";
by (rtac impI 1);
by (rtac impI 1);
by (rtac (expand_if RS iffD2) 1);
by (fast_tac HOL_cs 1);
-val imp_false_decompose = result();
+qed "imp_false_decompose";
(* Sets *)
@@ -57,7 +57,7 @@
goal Arith.thy "n ~= 0 --> Suc(m+pred(n)) = m+n";
by (nat_ind_tac "n" 1);
by (REPEAT(simp_tac arith_ss 1));
-val Suc_pred_lemma = result() RS mp;
+val Suc_pred_lemma = store_thm("Suc_pred_lemma", result() RS mp);
goal Arith.thy "x <= y --> x <= Suc(y)";
by (rtac impI 1);
@@ -67,43 +67,43 @@
by (etac disjE 1);
by (etac less_SucI 1);
by (asm_simp_tac nat_ss 1);
-val leq_imp_leq_suc = result() RS mp;
+val leq_imp_leq_suc = store_thm("leq_imp_leq_suc", result() RS mp);
(* Same as previous! *)
goal Arith.thy "(x::nat)<=y --> x<=Suc(y)";
by (simp_tac (arith_ss addsimps [le_eq_less_or_eq]) 1);
-val leq_suc = result();
+qed "leq_suc";
goal Arith.thy "((m::nat) + n = m + p) = (n = p)";
by (nat_ind_tac "m" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
-val left_plus_cancel = result();
+qed "left_plus_cancel";
goal Arith.thy "((x::nat) + y = Suc(x + z)) = (y = Suc(z))";
by (nat_ind_tac "x" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
-val left_plus_cancel_inside_succ = result();
+qed "left_plus_cancel_inside_succ";
goal Arith.thy "(x ~= 0) = (? y. x = Suc(y))";
by (nat_ind_tac "x" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
by (fast_tac HOL_cs 1);
-val nonzero_is_succ = result();
+qed "nonzero_is_succ";
goal Arith.thy "(m::nat) < n --> m + p < n + p";
by (nat_ind_tac "p" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
-val less_add_same_less = result();
+qed "less_add_same_less";
goal Arith.thy "(x::nat)<= y --> x<=y+k";
by (nat_ind_tac "k" 1);
by (simp_tac arith_ss 1);
by (asm_full_simp_tac (arith_ss addsimps [leq_suc]) 1);
-val leq_add_leq = result();
+qed "leq_add_leq";
goal Arith.thy "(x::nat) + y <= z --> x <= z";
by (nat_ind_tac "y" 1);
@@ -112,7 +112,7 @@
by (rtac impI 1);
by (dtac Suc_leD 1);
by (fast_tac HOL_cs 1);
-val left_add_leq = result();
+qed "left_add_leq";
goal Arith.thy "(A::nat) < B --> C < D --> A + C < B + D";
by (rtac impI 1);
@@ -124,12 +124,12 @@
by (res_inst_tac [("m1","B")] (add_commute RS ssubst) 1);
by (rtac (less_add_same_less RS mp) 1);
by (assume_tac 1);
-val less_add_cong = result();
+qed "less_add_cong";
goal Nat.thy "!!i. [| i < j; j <= k |] ==> i < (k::nat)";
by (dtac le_imp_less_or_eq 1);
by (fast_tac (HOL_cs addIs [less_trans]) 1);
-val less_leq_less = result();
+qed "less_leq_less";
goal Arith.thy "(A::nat) <= B --> C <= D --> A + C <= B + D";
by (rtac impI 1);
@@ -145,7 +145,7 @@
by (res_inst_tac [("m1","B")] (add_commute RS ssubst) 1);
by (rtac (less_add_same_less RS mp) 1);
by (assume_tac 1);
-val less_eq_add_cong = result();
+qed "less_eq_add_cong";
goal Arith.thy "(w <= y) --> ((x::nat) + y <= z) --> (x + w <= z)";
by (rtac impI 1);
@@ -156,7 +156,7 @@
by (rtac impI 1);
by (etac le_trans 1);
by (assume_tac 1);
-val leq_add_left_cong = result();
+qed "leq_add_left_cong";
goal Arith.thy "(? x. y = Suc(x)) = (~(y = 0))";
by (nat_ind_tac "y" 1);
@@ -164,7 +164,7 @@
by (rtac iffI 1);
by (asm_full_simp_tac arith_ss 1);
by (fast_tac HOL_cs 1);
-val suc_not_zero = result();
+qed "suc_not_zero";
goal Arith.thy "Suc(x) <= y --> (? z. y = Suc(z))";
by (rtac impI 1);
@@ -176,7 +176,7 @@
by (asm_full_simp_tac (arith_ss addsimps [suc_not_zero]) 1);
by (hyp_subst_tac 1);
by (asm_full_simp_tac arith_ss 1);
-val suc_leq_suc = result();
+qed "suc_leq_suc";
goal Arith.thy "~0<n --> n = 0";
by (nat_ind_tac "n" 1);
@@ -184,43 +184,43 @@
by (safe_tac HOL_cs);
by (asm_full_simp_tac arith_ss 1);
by (asm_full_simp_tac arith_ss 1);
-val zero_eq = result();
+qed "zero_eq";
goal Arith.thy "x < Suc(y) --> x<=y";
by (nat_ind_tac "n" 1);
by (asm_simp_tac arith_ss 1);
by (safe_tac HOL_cs);
by (etac less_imp_le 1);
-val less_suc_imp_leq = result();
+qed "less_suc_imp_leq";
goal Arith.thy "0<x --> Suc(pred(x)) = x";
by (nat_ind_tac "x" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
-val suc_pred_id = result();
+qed "suc_pred_id";
goal Arith.thy "0<x --> (pred(x) = y) = (x = Suc(y))";
by (nat_ind_tac "x" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
-val pred_suc = result();
+qed "pred_suc";
goal Arith.thy "(x ~= 0) = (0<x)";
by (nat_ind_tac "x" 1);
by (simp_tac arith_ss 1);
by (asm_simp_tac arith_ss 1);
-val unzero_less = result();
+qed "unzero_less";
(* Duplicate of earlier lemma! *)
goal Arith.thy "x<(y::nat) --> y<=z --> x<(z::nat)";
by (rtac impI 1); by (rtac impI 1);
by (dtac le_imp_less_or_eq 1);
by (fast_tac (HOL_cs addIs [less_trans]) 1);
-val less_leq_less = result();
+qed "less_leq_less";
goal Arith.thy "(Suc(n) <= Suc(m)) = (n <= m)";
by (simp_tac (arith_ss addsimps [le_eq_less_or_eq]) 1);
-val succ_leq_mono = result();
+qed "succ_leq_mono";
(* Odd proof. Why do induction? *)
goal Arith.thy "((x::nat) = y + z) --> (y <= x)";
@@ -228,7 +228,7 @@
by (simp_tac arith_ss 1);
by (simp_tac (arith_ss addsimps
[succ_leq_mono, le_refl RS (leq_add_leq RS mp)]) 1);
-val plus_leq_lem = result();
+qed "plus_leq_lem";
(* Lists *)
@@ -236,20 +236,20 @@
by (list_ind_tac "L" 1);
by (simp_tac list_ss 1);
by (asm_simp_tac list_ss 1);
-val append_cons = result();
+qed "append_cons";
goal List.thy "(X ~= hd(L@M)) = (X ~= if(L = [], hd(M), hd(L)))";
by (list_ind_tac "L" 1);
by (simp_tac list_ss 1);
by (asm_full_simp_tac list_ss 1);
-val not_hd_append = result();
+qed "not_hd_append";
goal List.thy "(L = (x#rst)) --> (L = []) --> P";
by (simp_tac list_ss 1);
-val list_cases = result();
+qed "list_cases";
goal List.thy "(? L2. L1 = x#L2) --> (L1 ~= [])";
by (strip_tac 1);
by (etac exE 1);
by (asm_simp_tac list_ss 1);
-val cons_imp_not_null = result();
+qed "cons_imp_not_null";
--- a/IOA/example/Multiset.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/example/Multiset.ML Mon Nov 21 17:50:34 1994 +0100
@@ -10,21 +10,21 @@
goalw Multiset.thy [Multiset.count_def, Multiset.countm_empty_def]
"count({|},x) = 0";
by (rtac refl 1);
-val count_empty = result();
+qed "count_empty";
goal Multiset.thy
"count(addm(M,x),y) = if(y=x,Suc(count(M,y)),count(M,y))";
by (asm_simp_tac (arith_ss addsimps
[Multiset.count_def,Multiset.countm_nonempty_def]
setloop (split_tac [expand_if])) 1);
-val count_addm_simp = result();
+qed "count_addm_simp";
goal Multiset.thy "count(M,y) <= count(addm(M,x),y)";
by (simp_tac (arith_ss addsimps [count_addm_simp]
setloop (split_tac [expand_if])) 1);
by (rtac impI 1);
by (rtac (le_refl RS (leq_suc RS mp)) 1);
-val count_leq_addm = result();
+qed "count_leq_addm";
goalw Multiset.thy [Multiset.count_def]
"count(delm(M,x),y) = if(y=x,pred(count(M,y)),count(M,y))";
@@ -38,7 +38,7 @@
setloop (split_tac [expand_if])) 1);
by (safe_tac HOL_cs);
by (asm_full_simp_tac HOL_ss 1);
-val count_delm_simp = result();
+qed "count_delm_simp";
goal Multiset.thy "!!M. (!x. P(x) --> Q(x)) ==> (countm(M,P) <= countm(M,Q))";
by (res_inst_tac [("M","M")] Multiset.induction 1);
@@ -47,7 +47,7 @@
by (etac (less_eq_add_cong RS mp RS mp) 1);
by (asm_full_simp_tac (arith_ss addsimps [le_eq_less_or_eq]
setloop (split_tac [expand_if])) 1);
-val countm_props = result();
+qed "countm_props";
goal Multiset.thy "!!P. ~P(obj) ==> countm(M,P) = countm(delm(M,obj),P)";
by (res_inst_tac [("M","M")] Multiset.induction 1);
@@ -56,7 +56,7 @@
by (asm_simp_tac (arith_ss addsimps[Multiset.countm_nonempty_def,
Multiset.delm_nonempty_def]
setloop (split_tac [expand_if])) 1);
-val countm_spurious_delm = result();
+qed "countm_spurious_delm";
goal Multiset.thy "!!P. P(x) ==> 0<count(M,x) --> 0<countm(M,P)";
@@ -68,7 +68,7 @@
[Multiset.count_def,Multiset.delm_nonempty_def,
Multiset.countm_nonempty_def]
setloop (split_tac [expand_if])) 1);
-val pos_count_imp_pos_countm = standard(result() RS mp);
+val pos_count_imp_pos_countm = store_thm("pos_count_imp_pos_countm", standard(result() RS mp));
goal Multiset.thy
"!!P. P(x) ==> 0<count(M,x) --> countm(delm(M,x),P) = pred(countm(M,P))";
@@ -81,4 +81,4 @@
Multiset.countm_nonempty_def,pos_count_imp_pos_countm,
suc_pred_id]
setloop (split_tac [expand_if])) 1);
-val countm_done_delm = result();
+qed "countm_done_delm";
--- a/IOA/example/Receiver.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/example/Receiver.ML Mon Nov 21 17:50:34 1994 +0100
@@ -18,7 +18,7 @@
by(simp_tac (action_ss addsimps
(Receiver.receiver_asig_def :: actions_def ::
asig_projections @ set_lemmas)) 1);
-val in_receiver_asig = result();
+qed "in_receiver_asig";
val receiver_projections =
[Receiver.rq_def,
--- a/IOA/example/Sender.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/example/Sender.ML Mon Nov 21 17:50:34 1994 +0100
@@ -18,7 +18,7 @@
by(simp_tac (action_ss addsimps
(Sender.sender_asig_def :: actions_def ::
asig_projections @ set_lemmas)) 1);
-val in_sender_asig = result();
+qed "in_sender_asig";
val sender_projections =
[Sender.sq_def,Sender.ssent_def,Sender.srcvd_def,
--- a/IOA/meta_theory/IOA.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/meta_theory/IOA.ML Mon Nov 21 17:50:34 1994 +0100
@@ -15,14 +15,14 @@
goal IOA.thy
"asig_of(<x,y,z>) = x & starts_of(<x,y,z>) = y & trans_of(<x,y,z>) = z";
by (simp_tac (SS addsimps ioa_projections) 1);
- val ioa_triple_proj = result();
+ qed "ioa_triple_proj";
goalw IOA.thy [ioa_def,state_trans_def,actions_def, is_asig_def]
"!!A. [| IOA(A); <s1,a,s2>:trans_of(A) |] ==> a:actions(asig_of(A))";
by (REPEAT(etac conjE 1));
by (EVERY1[etac allE, etac impE, atac]);
by (asm_full_simp_tac SS 1);
-val trans_in_actions = result();
+qed "trans_in_actions";
goal IOA.thy "filter_oseq(p,filter_oseq(p,s)) = filter_oseq(p,s)";
@@ -31,7 +31,7 @@
by (Option.option.induct_tac "s(i)" 1);
by (simp_tac SS 1);
by (simp_tac (SS setloop (split_tac [expand_if])) 1);
-val filter_oseq_idemp = result();
+qed "filter_oseq_idemp";
goalw IOA.thy [mk_behaviour_def,filter_oseq_def]
"(mk_behaviour(A, s, n) = None) = \
@@ -42,13 +42,13 @@
by (Option.option.induct_tac "s(n)" 1);
by (ALLGOALS (simp_tac (SS setloop (split_tac [expand_if]))));
by (fast_tac HOL_cs 1);
-val mk_behaviour_thm = result();
+qed "mk_behaviour_thm";
goalw IOA.thy [reachable_def] "!!A. s:starts_of(A) ==> reachable(A,s)";
by (res_inst_tac [("x","<(%i.None),(%i.s)>::('action,'a)execution")] bexI 1);
by (simp_tac SS 1);
by (asm_simp_tac (SS addsimps exec_rws) 1);
-val reachable_0 = result();
+qed "reachable_0";
goalw IOA.thy (reachable_def::exec_rws)
"!!A. [| reachable(A,s); <s,a,t> : trans_of(A) |] ==> reachable(A,t)";
@@ -68,7 +68,7 @@
by(fast_tac HOL_cs 1);
by(forward_tac [less_not_sym] 1);
by(asm_simp_tac (SS addsimps [less_not_refl2]) 1);
-val reachable_n = result();
+qed "reachable_n";
val [p1,p2] = goalw IOA.thy [invariant_def]
"[| !!s. s:starts_of(A) ==> P(s); \
@@ -87,37 +87,37 @@
by (etac (p2 RS mp) 1);
by (ALLGOALS(fast_tac(set_cs addDs [hd Option.option.inject RS iffD1,
reachable_n])));
-val invariantI = result();
+qed "invariantI";
val [p1,p2] = goal IOA.thy
"[| !!s. s : starts_of(A) ==> P(s); \
\ !!s t a. reachable(A, s) ==> P(s) --> <s,a,t>:trans_of(A) --> P(t) \
\ |] ==> invariant(A, P)";
by (fast_tac (HOL_cs addSIs [invariantI] addSDs [p1,p2]) 1);
-val invariantI1 = result();
+qed "invariantI1";
val [p1,p2] = goalw IOA.thy [invariant_def]
"[| invariant(A,P); reachable(A,s) |] ==> P(s)";
br(p2 RS (p1 RS spec RS mp))1;
-val invariantE = result();
+qed "invariantE";
goal IOA.thy
"actions(asig_comp(a,b)) = actions(a) Un actions(b)";
by(simp_tac (pair_ss addsimps
([actions_def,asig_comp_def]@asig_projections)) 1);
by(fast_tac eq_cs 1);
-val actions_asig_comp = result();
+qed "actions_asig_comp";
goal IOA.thy
"starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
by(simp_tac (SS addsimps (par_def::ioa_projections)) 1);
-val starts_of_par = result();
+qed "starts_of_par";
(* Every state in an execution is reachable *)
goalw IOA.thy [reachable_def]
"!!A. ex:executions(A) ==> !n. reachable(A, snd(ex,n))";
by (fast_tac set_cs 1);
-val states_of_exec_reachable = result();
+qed "states_of_exec_reachable";
goal IOA.thy
@@ -136,15 +136,15 @@
by(simp_tac (SS addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]@
ioa_projections)
setloop (split_tac [expand_if])) 1);
-val trans_of_par4 = result();
+qed "trans_of_par4";
goal IOA.thy "starts_of(restrict(ioa,acts)) = starts_of(ioa) & \
\ trans_of(restrict(ioa,acts)) = trans_of(ioa) & \
\ reachable(restrict(ioa,acts),s) = reachable(ioa,s)";
by(simp_tac (SS addsimps ([is_execution_fragment_def,executions_def,
reachable_def,restrict_def]@ioa_projections)) 1);
-val cancel_restrict = result();
+qed "cancel_restrict";
goal IOA.thy "asig_of(A || B) = asig_comp(asig_of(A),asig_of(B))";
by(simp_tac (SS addsimps (par_def::ioa_projections)) 1);
-val asig_of_par = result();
+qed "asig_of_par";
--- a/IOA/meta_theory/Option.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/meta_theory/Option.ML Mon Nov 21 17:50:34 1994 +0100
@@ -13,4 +13,4 @@
br (prem RS rev_mp) 1;
by (Option.option.induct_tac "opt" 1);
by (ALLGOALS(fast_tac HOL_cs));
-val optE = standard(result() RS disjE);
+val optE = store_thm("optE", standard(result() RS disjE));
--- a/IOA/meta_theory/Solve.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/IOA/meta_theory/Solve.ML Mon Nov 21 17:50:34 1994 +0100
@@ -50,4 +50,4 @@
by (eres_inst_tac [("x","snd(ex,Suc(n))")] allE 1);
by (eres_inst_tac [("x","a")] allE 1);
by (asm_full_simp_tac SS 1);
-val trace_inclusion = result();
+qed "trace_inclusion";
--- a/Lfp.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Lfp.ML Mon Nov 21 17:50:34 1994 +0100
@@ -15,27 +15,27 @@
val prems = goalw Lfp.thy [lfp_def] "[| f(A) <= A |] ==> lfp(f) <= A";
by (rtac (CollectI RS Inter_lower) 1);
by (resolve_tac prems 1);
-val lfp_lowerbound = result();
+qed "lfp_lowerbound";
val prems = goalw Lfp.thy [lfp_def]
"[| !!u. f(u) <= u ==> A<=u |] ==> A <= lfp(f)";
by (REPEAT (ares_tac ([Inter_greatest]@prems) 1));
by (etac CollectD 1);
-val lfp_greatest = result();
+qed "lfp_greatest";
val [mono] = goal Lfp.thy "mono(f) ==> f(lfp(f)) <= lfp(f)";
by (EVERY1 [rtac lfp_greatest, rtac subset_trans,
rtac (mono RS monoD), rtac lfp_lowerbound, atac, atac]);
-val lfp_lemma2 = result();
+qed "lfp_lemma2";
val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) <= f(lfp(f))";
by (EVERY1 [rtac lfp_lowerbound, rtac (mono RS monoD),
rtac lfp_lemma2, rtac mono]);
-val lfp_lemma3 = result();
+qed "lfp_lemma3";
val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))";
by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1));
-val lfp_Tarski = result();
+qed "lfp_Tarski";
(*** General induction rule for least fixed points ***)
@@ -50,14 +50,14 @@
rtac (Int_lower1 RS (mono RS monoD)),
rtac (mono RS lfp_lemma2),
rtac (CollectI RS subsetI), rtac indhyp, atac]);
-val induct = result();
+qed "induct";
(** Definition forms of lfp_Tarski and induct, to control unfolding **)
val [rew,mono] = goal Lfp.thy "[| h==lfp(f); mono(f) |] ==> h = f(h)";
by (rewtac rew);
by (rtac (mono RS lfp_Tarski) 1);
-val def_lfp_Tarski = result();
+qed "def_lfp_Tarski";
val rew::prems = goal Lfp.thy
"[| A == lfp(f); mono(f); a:A; \
@@ -65,10 +65,10 @@
\ |] ==> P(a)";
by (EVERY1 [rtac induct, (*backtracking to force correct induction*)
REPEAT1 o (ares_tac (map (rewrite_rule [rew]) prems))]);
-val def_induct = result();
+qed "def_induct";
(*Monotonicity of lfp!*)
val [prem] = goal Lfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> lfp(f) <= lfp(g)";
br (lfp_lowerbound RS lfp_greatest) 1;
be (prem RS subset_trans) 1;
-val lfp_mono = result();
+qed "lfp_mono";
--- a/List.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/List.ML Mon Nov 21 17:50:34 1994 +0100
@@ -15,19 +15,19 @@
by (fast_tac (univ_cs addSIs (equalityI :: map rew list.intrs)
addEs [rew list.elim]) 1)
end;
-val list_unfold = result();
+qed "list_unfold";
(*This justifies using list in other recursive type definitions*)
goalw List.thy list.defs "!!A B. A<=B ==> list(A) <= list(B)";
by (rtac lfp_mono 1);
by (REPEAT (ares_tac basic_monos 1));
-val list_mono = result();
+qed "list_mono";
(*Type checking -- list creates well-founded sets*)
goalw List.thy (list_con_defs @ list.defs) "list(sexp) <= sexp";
by (rtac lfp_lowerbound 1);
by (fast_tac (univ_cs addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1);
-val list_sexp = result();
+qed "list_sexp";
(* A <= sexp ==> list(A) <= sexp *)
val list_subset_sexp = standard ([list_mono, list_sexp] MRS subset_trans);
@@ -40,7 +40,7 @@
by (rtac (Rep_list RS list.induct) 1);
by (REPEAT (ares_tac prems 1
ORELSE eresolve_tac [rangeE, ssubst, Abs_list_inverse RS subst] 1));
-val list_induct = result();
+qed "list_induct";
(*Perform induction on xs. *)
fun list_ind_tac a M =
@@ -52,18 +52,18 @@
goal List.thy "inj(Rep_list)";
by (rtac inj_inverseI 1);
by (rtac Rep_list_inverse 1);
-val inj_Rep_list = result();
+qed "inj_Rep_list";
goal List.thy "inj_onto(Abs_list,list(range(Leaf)))";
by (rtac inj_onto_inverseI 1);
by (etac Abs_list_inverse 1);
-val inj_onto_Abs_list = result();
+qed "inj_onto_Abs_list";
(** Distinctness of constructors **)
goalw List.thy list_con_defs "CONS(M,N) ~= NIL";
by (rtac In1_not_In0 1);
-val CONS_not_NIL = result();
+qed "CONS_not_NIL";
val NIL_not_CONS = standard (CONS_not_NIL RS not_sym);
val CONS_neq_NIL = standard (CONS_not_NIL RS notE);
@@ -72,7 +72,7 @@
goalw List.thy [Nil_def,Cons_def] "x # xs ~= Nil";
by (rtac (CONS_not_NIL RS (inj_onto_Abs_list RS inj_onto_contraD)) 1);
by (REPEAT (resolve_tac (list.intrs @ [rangeI, Rep_list]) 1));
-val Cons_not_Nil = result();
+qed "Cons_not_Nil";
val Nil_not_Cons = standard (Cons_not_Nil RS not_sym);
@@ -83,7 +83,7 @@
goalw List.thy [CONS_def] "(CONS(K,M)=CONS(L,N)) = (K=L & M=N)";
by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1);
-val CONS_CONS_eq = result();
+qed "CONS_CONS_eq";
val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE);
@@ -95,20 +95,20 @@
goalw List.thy [Cons_def] "(x#xs=y#ys) = (x=y & xs=ys)";
by (fast_tac list_cs 1);
-val Cons_Cons_eq = result();
+qed "Cons_Cons_eq";
val Cons_inject = standard (Cons_Cons_eq RS iffD1 RS conjE);
val [major] = goal List.thy "CONS(M,N): list(A) ==> M: A & N: list(A)";
by (rtac (major RS setup_induction) 1);
by (etac list.induct 1);
by (ALLGOALS (fast_tac list_cs));
-val CONS_D = result();
+qed "CONS_D";
val prems = goalw List.thy [CONS_def,In1_def]
"CONS(M,N): sexp ==> M: sexp & N: sexp";
by (cut_facts_tac prems 1);
by (fast_tac (set_cs addSDs [Scons_D]) 1);
-val sexp_CONS_D = result();
+qed "sexp_CONS_D";
(*Basic ss with constructors and their freeness*)
@@ -120,12 +120,12 @@
goal List.thy "!!N. N: list(A) ==> !M. N ~= CONS(M,N)";
by (etac list.induct 1);
by (ALLGOALS (asm_simp_tac list_free_ss));
-val not_CONS_self = result();
+qed "not_CONS_self";
goal List.thy "!x. l ~= x#l";
by (list_ind_tac "l" 1);
by (ALLGOALS (asm_simp_tac list_free_ss));
-val not_Cons_self = result();
+qed "not_Cons_self";
goal List.thy "(xs ~= []) = (? y ys. xs = y#ys)";
@@ -133,17 +133,17 @@
by(simp_tac list_free_ss 1);
by(asm_simp_tac list_free_ss 1);
by(REPEAT(resolve_tac [exI,refl,conjI] 1));
-val neq_Nil_conv = result();
+qed "neq_Nil_conv";
(** Conversion rules for List_case: case analysis operator **)
goalw List.thy [List_case_def,NIL_def] "List_case(c, h, NIL) = c";
by (rtac Case_In0 1);
-val List_case_NIL = result();
+qed "List_case_NIL";
goalw List.thy [List_case_def,CONS_def] "List_case(c, h, CONS(M,N)) = h(M,N)";
by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1);
-val List_case_CONS = result();
+qed "List_case_CONS";
(*** List_rec -- by wf recursion on pred_sexp ***)
@@ -158,12 +158,12 @@
goalw List.thy [CONS_def,In1_def]
"!!M. [| M: sexp; N: sexp |] ==> <M, CONS(M,N)> : pred_sexp^+";
by (asm_simp_tac pred_sexp_ss 1);
-val pred_sexp_CONS_I1 = result();
+qed "pred_sexp_CONS_I1";
goalw List.thy [CONS_def,In1_def]
"!!M. [| M: sexp; N: sexp |] ==> <N, CONS(M,N)> : pred_sexp^+";
by (asm_simp_tac pred_sexp_ss 1);
-val pred_sexp_CONS_I2 = result();
+qed "pred_sexp_CONS_I2";
val [prem] = goal List.thy
"<CONS(M1,M2), N> : pred_sexp^+ ==> \
@@ -173,14 +173,14 @@
by (etac (sexp_CONS_D RS conjE) 1);
by (REPEAT (ares_tac [conjI, pred_sexp_CONS_I1, pred_sexp_CONS_I2,
prem RSN (2, trans_trancl RS transD)] 1));
-val pred_sexp_CONS_D = result();
+qed "pred_sexp_CONS_D";
(** Conversion rules for List_rec **)
goal List.thy "List_rec(NIL,c,h) = c";
by (rtac (List_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [List_case_NIL]) 1);
-val List_rec_NIL = result();
+qed "List_rec_NIL";
goal List.thy "!!M. [| M: sexp; N: sexp |] ==> \
\ List_rec(CONS(M,N), c, h) = h(M, N, List_rec(N,c,h))";
@@ -188,7 +188,7 @@
by (asm_simp_tac
(HOL_ss addsimps [List_case_CONS, list.CONS_I, pred_sexp_CONS_I2,
cut_apply])1);
-val List_rec_CONS = result();
+qed "List_rec_CONS";
(*** list_rec -- by List_rec ***)
@@ -227,46 +227,46 @@
val sexp_A_I = A_subset_sexp RS subsetD;
by (rtac (major RS list.induct) 1);
by (ALLGOALS(asm_simp_tac (list_ss addsimps ([sexp_A_I,sexp_ListA_I]@prems))));
-val List_rec_type = result();
+qed "List_rec_type";
(** Generalized map functionals **)
goalw List.thy [Rep_map_def] "Rep_map(f,Nil) = NIL";
by (rtac list_rec_Nil 1);
-val Rep_map_Nil = result();
+qed "Rep_map_Nil";
goalw List.thy [Rep_map_def]
"Rep_map(f, x#xs) = CONS(f(x), Rep_map(f,xs))";
by (rtac list_rec_Cons 1);
-val Rep_map_Cons = result();
+qed "Rep_map_Cons";
goalw List.thy [Rep_map_def] "!!f. (!!x. f(x): A) ==> Rep_map(f,xs): list(A)";
by (rtac list_induct 1);
by(ALLGOALS(asm_simp_tac list_ss));
-val Rep_map_type = result();
+qed "Rep_map_type";
goalw List.thy [Abs_map_def] "Abs_map(g,NIL) = Nil";
by (rtac List_rec_NIL 1);
-val Abs_map_NIL = result();
+qed "Abs_map_NIL";
val prems = goalw List.thy [Abs_map_def]
"[| M: sexp; N: sexp |] ==> \
\ Abs_map(g, CONS(M,N)) = g(M) # Abs_map(g,N)";
by (REPEAT (resolve_tac (List_rec_CONS::prems) 1));
-val Abs_map_CONS = result();
+qed "Abs_map_CONS";
(*These 2 rules ease the use of primitive recursion. NOTE USE OF == *)
val [rew] = goal List.thy
"[| !!xs. f(xs) == list_rec(xs,c,h) |] ==> f([]) = 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
"[| !!xs. f(xs) == list_rec(xs,c,h) |] ==> f(x#xs) = h(x,xs,f(xs))";
by (rewtac rew);
by (rtac list_rec_Cons 1);
-val def_list_rec_Cons = result();
+qed "def_list_rec_Cons";
fun list_recs def =
[standard (def RS def_list_rec_Nil),
@@ -302,42 +302,42 @@
goal List.thy "(xs@ys)@zs = xs@(ys@zs)";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac list_ss));
-val append_assoc = result();
+qed "append_assoc";
goal List.thy "xs @ [] = xs";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac list_ss));
-val append_Nil2 = result();
+qed "append_Nil2";
(** mem **)
goal List.thy "x mem (xs@ys) = (x mem xs | x mem ys)";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if]))));
-val mem_append = result();
+qed "mem_append";
goal List.thy "x mem [x:xs.P(x)] = (x mem xs & P(x))";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if]))));
-val mem_filter = result();
+qed "mem_filter";
(** list_all **)
goal List.thy "(Alls x:xs.True) = True";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac list_ss));
-val list_all_True = result();
+qed "list_all_True";
goal List.thy "list_all(p,xs@ys) = (list_all(p,xs) & list_all(p,ys))";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac list_ss));
-val list_all_conj = result();
+qed "list_all_conj";
goal List.thy "(Alls x:xs.P(x)) = (!x. x mem xs --> P(x))";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac (list_ss setloop (split_tac [expand_if]))));
by(fast_tac HOL_cs 1);
-val list_all_mem_conv = result();
+qed "list_all_mem_conv";
(** The functional "map" **)
@@ -352,7 +352,7 @@
\ ==> Rep_map(f, Abs_map(g,M)) = M";
by (rtac (major RS list.induct) 1);
by (ALLGOALS (asm_simp_tac(map_ss addsimps [sexp_A_I,sexp_ListA_I,minor])));
-val Abs_map_inverse = result();
+qed "Abs_map_inverse";
(*Rep_map_inverse is obtained via Abs_Rep_map and map_ident*)
@@ -364,7 +364,7 @@
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac list_ss));
by(fast_tac HOL_cs 1);
-val expand_list_case = result();
+qed "expand_list_case";
(** Additional mapping lemmas **)
@@ -372,24 +372,24 @@
goal List.thy "map(%x.x, xs) = xs";
by (list_ind_tac "xs" 1);
by (ALLGOALS (asm_simp_tac map_ss));
-val map_ident = result();
+qed "map_ident";
goal List.thy "map(f, xs@ys) = map(f,xs) @ map(f,ys)";
by (list_ind_tac "xs" 1);
by (ALLGOALS (asm_simp_tac (map_ss addsimps [append_Nil,append_Cons])));
-val map_append = result();
+qed "map_append";
goalw List.thy [o_def] "map(f o g, xs) = map(f, map(g, xs))";
by (list_ind_tac "xs" 1);
by (ALLGOALS (asm_simp_tac map_ss));
-val map_compose = result();
+qed "map_compose";
goal List.thy "!!f. (!!x. f(x): sexp) ==> \
\ Abs_map(g, Rep_map(f,xs)) = map(%t. g(f(t)), xs)";
by (list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac(map_ss addsimps
[Rep_map_type,list_sexp RS subsetD])));
-val Abs_Rep_map = result();
+qed "Abs_Rep_map";
val list_ss = list_ss addsimps
[mem_append, mem_filter, append_assoc, append_Nil2, map_ident,
--- a/Nat.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Nat.ML Mon Nov 21 17:50:34 1994 +0100
@@ -10,7 +10,7 @@
goal Nat.thy "mono(%X. {Zero_Rep} Un (Suc_Rep``X))";
by (REPEAT (ares_tac [monoI, subset_refl, image_mono, Un_mono] 1));
-val Nat_fun_mono = result();
+qed "Nat_fun_mono";
val Nat_unfold = Nat_fun_mono RS (Nat_def RS def_lfp_Tarski);
@@ -18,13 +18,13 @@
goal Nat.thy "Zero_Rep: Nat";
by (rtac (Nat_unfold RS ssubst) 1);
by (rtac (singletonI RS UnI1) 1);
-val Zero_RepI = result();
+qed "Zero_RepI";
val prems = goal Nat.thy "i: Nat ==> Suc_Rep(i) : Nat";
by (rtac (Nat_unfold RS ssubst) 1);
by (rtac (imageI RS UnI2) 1);
by (resolve_tac prems 1);
-val Suc_RepI = result();
+qed "Suc_RepI";
(*** Induction ***)
@@ -33,7 +33,7 @@
\ !!j. [| j: Nat; P(j) |] ==> P(Suc_Rep(j)) |] ==> P(i)";
by (rtac ([Nat_def, Nat_fun_mono, major] MRS def_induct) 1);
by (fast_tac (set_cs addIs prems) 1);
-val Nat_induct = result();
+qed "Nat_induct";
val prems = goalw Nat.thy [Zero_def,Suc_def]
"[| P(0); \
@@ -42,7 +42,7 @@
by (rtac (Rep_Nat RS Nat_induct) 1);
by (REPEAT (ares_tac prems 1
ORELSE eresolve_tac [Abs_Nat_inverse RS subst] 1));
-val nat_induct = result();
+qed "nat_induct";
(*Perform induction on n. *)
fun nat_ind_tac a i =
@@ -60,7 +60,7 @@
by (rtac allI 2);
by (nat_ind_tac "x" 2);
by (REPEAT (ares_tac (prems@[allI]) 1 ORELSE etac spec 1));
-val diff_induct = result();
+qed "diff_induct";
(*Case analysis on the natural numbers*)
val prems = goal Nat.thy
@@ -70,7 +70,7 @@
by (nat_ind_tac "n" 1);
by (rtac (refl RS disjI1) 1);
by (fast_tac HOL_cs 1);
-val natE = result();
+qed "natE";
(*** Isomorphisms: Abs_Nat and Rep_Nat ***)
@@ -80,12 +80,12 @@
goal Nat.thy "inj(Rep_Nat)";
by (rtac inj_inverseI 1);
by (rtac Rep_Nat_inverse 1);
-val inj_Rep_Nat = result();
+qed "inj_Rep_Nat";
goal Nat.thy "inj_onto(Abs_Nat,Nat)";
by (rtac inj_onto_inverseI 1);
by (etac Abs_Nat_inverse 1);
-val inj_onto_Abs_Nat = result();
+qed "inj_onto_Abs_Nat";
(*** Distinctness of constructors ***)
@@ -93,7 +93,7 @@
by (rtac (inj_onto_Abs_Nat RS inj_onto_contraD) 1);
by (rtac Suc_Rep_not_Zero_Rep 1);
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI, Zero_RepI] 1));
-val Suc_not_Zero = result();
+qed "Suc_not_Zero";
val Zero_not_Suc = standard (Suc_not_Zero RS not_sym);
@@ -108,18 +108,18 @@
by (REPEAT (resolve_tac [Rep_Nat, Suc_RepI] 1));
by (dtac (inj_Suc_Rep RS injD) 1);
by (etac (inj_Rep_Nat RS injD) 1);
-val inj_Suc = result();
+qed "inj_Suc";
val Suc_inject = inj_Suc RS injD;;
goal Nat.thy "(Suc(m)=Suc(n)) = (m=n)";
by (EVERY1 [rtac iffI, etac Suc_inject, etac arg_cong]);
-val Suc_Suc_eq = result();
+qed "Suc_Suc_eq";
goal Nat.thy "n ~= Suc(n)";
by (nat_ind_tac "n" 1);
by (ALLGOALS(asm_simp_tac (HOL_ss addsimps [Zero_not_Suc,Suc_Suc_eq])));
-val n_not_Suc_n = result();
+qed "n_not_Suc_n";
val Suc_n_not_n = n_not_Suc_n RS not_sym;
@@ -127,25 +127,25 @@
goalw Nat.thy [nat_case_def] "nat_case(a, f, 0) = a";
by (fast_tac (set_cs addIs [select_equality] addEs [Zero_neq_Suc]) 1);
-val nat_case_0 = result();
+qed "nat_case_0";
goalw Nat.thy [nat_case_def] "nat_case(a, f, Suc(k)) = f(k)";
by (fast_tac (set_cs addIs [select_equality]
addEs [make_elim Suc_inject, Suc_neq_Zero]) 1);
-val nat_case_Suc = result();
+qed "nat_case_Suc";
(** Introduction rules for 'pred_nat' **)
goalw Nat.thy [pred_nat_def] "<n, Suc(n)> : pred_nat";
by (fast_tac set_cs 1);
-val pred_natI = result();
+qed "pred_natI";
val major::prems = goalw Nat.thy [pred_nat_def]
"[| p : pred_nat; !!x n. [| p = <n, Suc(n)> |] ==> R \
\ |] ==> R";
by (rtac (major RS CollectE) 1);
by (REPEAT (eresolve_tac ([asm_rl,exE]@prems) 1));
-val pred_natE = result();
+qed "pred_natE";
goalw Nat.thy [wf_def] "wf(pred_nat)";
by (strip_tac 1);
@@ -153,7 +153,7 @@
by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject,
make_elim Suc_inject]) 2);
by (fast_tac (HOL_cs addSEs [mp, pred_natE, Pair_inject, Zero_neq_Suc]) 1);
-val wf_pred_nat = result();
+qed "wf_pred_nat";
(*** nat_rec -- by wf recursion on pred_nat ***)
@@ -165,25 +165,25 @@
goal Nat.thy "nat_rec(0,c,h) = c";
by (rtac (nat_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [nat_case_0]) 1);
-val nat_rec_0 = result();
+qed "nat_rec_0";
goal Nat.thy "nat_rec(Suc(n), c, h) = h(n, nat_rec(n,c,h))";
by (rtac (nat_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [nat_case_Suc, pred_natI, cut_apply]) 1);
-val nat_rec_Suc = result();
+qed "nat_rec_Suc";
(*These 2 rules ease the use of primitive recursion. NOTE USE OF == *)
val [rew] = goal Nat.thy
"[| !!n. f(n) == nat_rec(n,c,h) |] ==> f(0) = c";
by (rewtac rew);
by (rtac nat_rec_0 1);
-val def_nat_rec_0 = result();
+qed "def_nat_rec_0";
val [rew] = goal Nat.thy
"[| !!n. f(n) == nat_rec(n,c,h) |] ==> f(Suc(n)) = h(n,f(n))";
by (rewtac rew);
by (rtac nat_rec_Suc 1);
-val def_nat_rec_Suc = result();
+qed "def_nat_rec_Suc";
fun nat_recs def =
[standard (def RS def_nat_rec_0),
@@ -198,11 +198,11 @@
by (rtac (trans_trancl RS transD) 1);
by (resolve_tac prems 1);
by (resolve_tac prems 1);
-val less_trans = result();
+qed "less_trans";
goalw Nat.thy [less_def] "n < Suc(n)";
by (rtac (pred_natI RS r_into_trancl) 1);
-val lessI = result();
+qed "lessI";
(* i<j ==> i<Suc(j) *)
val less_SucI = lessI RSN (2, less_trans);
@@ -212,13 +212,13 @@
by (rtac lessI 1);
by (etac less_trans 1);
by (rtac lessI 1);
-val zero_less_Suc = result();
+qed "zero_less_Suc";
(** Elimination properties **)
val prems = goalw Nat.thy [less_def] "n<m ==> ~ m<(n::nat)";
by(fast_tac (HOL_cs addIs ([wf_pred_nat, wf_trancl RS wf_asym]@prems))1);
-val less_not_sym = result();
+qed "less_not_sym";
(* [| n<m; m<n |] ==> R *)
val less_asym = standard (less_not_sym RS notE);
@@ -226,14 +226,14 @@
goalw Nat.thy [less_def] "~ n<(n::nat)";
by (rtac notI 1);
by (etac (wf_pred_nat RS wf_trancl RS wf_anti_refl) 1);
-val less_not_refl = result();
+qed "less_not_refl";
(* n<n ==> R *)
val less_anti_refl = standard (less_not_refl RS notE);
goal Nat.thy "!!m. n<m ==> m ~= (n::nat)";
by(fast_tac (HOL_cs addEs [less_anti_refl]) 1);
-val less_not_refl2 = result();
+qed "less_not_refl2";
val major::prems = goalw Nat.thy [less_def]
@@ -242,14 +242,14 @@
by (rtac (major RS tranclE) 1);
by (fast_tac (HOL_cs addSEs (prems@[pred_natE, Pair_inject])) 1);
by (fast_tac (HOL_cs addSEs (prems@[pred_natE, Pair_inject])) 1);
-val lessE = result();
+qed "lessE";
goal Nat.thy "~ n<0";
by (rtac notI 1);
by (etac lessE 1);
by (etac Zero_neq_Suc 1);
by (etac Zero_neq_Suc 1);
-val not_less0 = result();
+qed "not_less0";
(* n<0 ==> R *)
val less_zeroE = standard (not_less0 RS notE);
@@ -261,12 +261,12 @@
by (fast_tac (HOL_cs addSDs [Suc_inject]) 1);
by (rtac less 1);
by (fast_tac (HOL_cs addSDs [Suc_inject]) 1);
-val less_SucE = result();
+qed "less_SucE";
goal Nat.thy "(m < Suc(n)) = (m < n | m = n)";
by (fast_tac (HOL_cs addSIs [lessI]
addEs [less_trans, less_SucE]) 1);
-val less_Suc_eq = result();
+qed "less_Suc_eq";
(** Inductive (?) properties **)
@@ -279,7 +279,7 @@
by (fast_tac (HOL_cs addSIs [lessI RS less_SucI]
addSDs [Suc_inject]
addEs [less_trans, lessE]) 1);
-val Suc_lessD = result();
+qed "Suc_lessD";
val [major,minor] = goal Nat.thy
"[| Suc(i)<k; !!j. [| i<j; k=Suc(j) |] ==> P \
@@ -288,13 +288,13 @@
by (etac (lessI RS minor) 1);
by (etac (Suc_lessD RS minor) 1);
by (assume_tac 1);
-val Suc_lessE = result();
+qed "Suc_lessE";
val [major] = goal Nat.thy "Suc(m) < Suc(n) ==> m<n";
by (rtac (major RS lessE) 1);
by (REPEAT (rtac lessI 1
ORELSE eresolve_tac [make_elim Suc_inject, ssubst, Suc_lessD] 1));
-val Suc_less_SucD = result();
+qed "Suc_less_SucD";
val prems = goal Nat.thy "m<n ==> Suc(m) < Suc(n)";
by (subgoal_tac "m<n --> Suc(m) < Suc(n)" 1);
@@ -305,15 +305,15 @@
by (fast_tac (HOL_cs addSIs [lessI]
addSDs [Suc_inject]
addEs [less_trans, lessE]) 1);
-val Suc_mono = result();
+qed "Suc_mono";
goal Nat.thy "(Suc(m) < Suc(n)) = (m<n)";
by (EVERY1 [rtac iffI, etac Suc_less_SucD, etac Suc_mono]);
-val Suc_less_eq = result();
+qed "Suc_less_eq";
goal Nat.thy "~(Suc(n) < n)";
by(fast_tac (HOL_cs addEs [Suc_lessD RS less_anti_refl]) 1);
-val not_Suc_n_less_n = result();
+qed "not_Suc_n_less_n";
(*"Less than" is a linear ordering*)
goal Nat.thy "m<n | m=n | n<(m::nat)";
@@ -322,28 +322,28 @@
by (rtac (refl RS disjI1 RS disjI2) 1);
by (rtac (zero_less_Suc RS disjI1) 1);
by (fast_tac (HOL_cs addIs [lessI, Suc_mono, less_SucI] addEs [lessE]) 1);
-val less_linear = result();
+qed "less_linear";
(*Can be used with less_Suc_eq to get n=m | n<m *)
goal Nat.thy "(~ m < n) = (n < Suc(m))";
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by(ALLGOALS(asm_simp_tac (HOL_ss addsimps
[not_less0,zero_less_Suc,Suc_less_eq])));
-val not_less_eq = result();
+qed "not_less_eq";
(*Complete induction, aka course-of-values induction*)
val prems = goalw Nat.thy [less_def]
"[| !!n. [| ! m::nat. m<n --> P(m) |] ==> P(n) |] ==> P(n)";
by (wf_ind_tac "n" [wf_pred_nat RS wf_trancl] 1);
by (eresolve_tac prems 1);
-val less_induct = result();
+qed "less_induct";
(*** Properties of <= ***)
goalw Nat.thy [le_def] "0 <= n";
by (rtac not_less0 1);
-val le0 = result();
+qed "le0";
val nat_simps = [not_less0, less_not_refl, zero_less_Suc, lessI,
Suc_less_eq, less_Suc_eq, le0, not_Suc_n_less_n,
@@ -364,64 +364,64 @@
val prems = goalw Nat.thy [le_def] "~(n<m) ==> m<=(n::nat)";
by (resolve_tac prems 1);
-val leI = result();
+qed "leI";
val prems = goalw Nat.thy [le_def] "m<=n ==> ~(n<(m::nat))";
by (resolve_tac prems 1);
-val leD = result();
+qed "leD";
val leE = make_elim leD;
goalw Nat.thy [le_def] "!!m. ~ m <= n ==> n<(m::nat)";
by (fast_tac HOL_cs 1);
-val not_leE = result();
+qed "not_leE";
goalw Nat.thy [le_def] "!!m. m < n ==> Suc(m) <= n";
by(simp_tac nat_ss0 1);
by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1);
-val lessD = result();
+qed "lessD";
goalw Nat.thy [le_def] "!!m. Suc(m) <= n ==> m <= n";
by(asm_full_simp_tac nat_ss0 1);
by(fast_tac HOL_cs 1);
-val Suc_leD = result();
+qed "Suc_leD";
goalw Nat.thy [le_def] "!!m. m < n ==> m <= (n::nat)";
by (fast_tac (HOL_cs addEs [less_asym]) 1);
-val less_imp_le = result();
+qed "less_imp_le";
goalw Nat.thy [le_def] "!!m. m <= n ==> m < n | m=(n::nat)";
by (cut_facts_tac [less_linear] 1);
by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1);
-val le_imp_less_or_eq = result();
+qed "le_imp_less_or_eq";
goalw Nat.thy [le_def] "!!m. m<n | m=n ==> m <=(n::nat)";
by (cut_facts_tac [less_linear] 1);
by (fast_tac (HOL_cs addEs [less_anti_refl,less_asym]) 1);
by (flexflex_tac);
-val less_or_eq_imp_le = result();
+qed "less_or_eq_imp_le";
goal Nat.thy "(x <= (y::nat)) = (x < y | x=y)";
by (REPEAT(ares_tac [iffI,less_or_eq_imp_le,le_imp_less_or_eq] 1));
-val le_eq_less_or_eq = result();
+qed "le_eq_less_or_eq";
goal Nat.thy "n <= (n::nat)";
by(simp_tac (HOL_ss addsimps [le_eq_less_or_eq]) 1);
-val le_refl = result();
+qed "le_refl";
val prems = goal Nat.thy "!!i. [| i <= j; j < k |] ==> i < (k::nat)";
by (dtac le_imp_less_or_eq 1);
by (fast_tac (HOL_cs addIs [less_trans]) 1);
-val le_less_trans = result();
+qed "le_less_trans";
goal Nat.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::nat)";
by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
rtac less_or_eq_imp_le, fast_tac (HOL_cs addIs [less_trans])]);
-val le_trans = result();
+qed "le_trans";
val prems = goal Nat.thy "!!m. [| m <= n; n <= m |] ==> m = (n::nat)";
by (EVERY1[dtac le_imp_less_or_eq, dtac le_imp_less_or_eq,
fast_tac (HOL_cs addEs [less_anti_refl,less_asym])]);
-val le_anti_sym = result();
+qed "le_anti_sym";
val nat_ss = nat_ss0 addsimps [le_refl];
--- a/Ord.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Ord.ML Mon Nov 21 17:50:34 1994 +0100
@@ -11,11 +11,11 @@
val [prem] = goalw Ord.thy [mono_def]
"[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)";
by (REPEAT (ares_tac [allI, impI, prem] 1));
-val monoI = result();
+qed "monoI";
val [major,minor] = goalw Ord.thy [mono_def]
"[| mono(f); A <= B |] ==> f(A) <= f(B)";
by (rtac (major RS spec RS spec RS mp) 1);
by (rtac minor 1);
-val monoD = result();
+qed "monoD";
--- a/Prod.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Prod.ML Mon Nov 21 17:50:34 1994 +0100
@@ -11,52 +11,52 @@
(*This counts as a non-emptiness result for admitting 'a * 'b as a type*)
goalw Prod.thy [Prod_def] "Pair_Rep(a,b) : Prod";
by (EVERY1 [rtac CollectI, rtac exI, rtac exI, rtac refl]);
-val ProdI = result();
+qed "ProdI";
val [major] = goalw Prod.thy [Pair_Rep_def]
"Pair_Rep(a, b) = Pair_Rep(a',b') ==> a=a' & b=b'";
by (EVERY1 [rtac (major RS fun_cong RS fun_cong RS subst),
rtac conjI, rtac refl, rtac refl]);
-val Pair_Rep_inject = result();
+qed "Pair_Rep_inject";
goal Prod.thy "inj_onto(Abs_Prod,Prod)";
by (rtac inj_onto_inverseI 1);
by (etac Abs_Prod_inverse 1);
-val inj_onto_Abs_Prod = result();
+qed "inj_onto_Abs_Prod";
val prems = goalw Prod.thy [Pair_def]
"[| <a, b> = <a',b'>; [| a=a'; b=b' |] ==> R |] ==> R";
by (rtac (inj_onto_Abs_Prod RS inj_ontoD RS Pair_Rep_inject RS conjE) 1);
by (REPEAT (ares_tac (prems@[ProdI]) 1));
-val Pair_inject = result();
+qed "Pair_inject";
goal Prod.thy "(<a,b> = <a',b'>) = (a=a' & b=b')";
by (fast_tac (set_cs addIs [Pair_inject]) 1);
-val Pair_eq = result();
+qed "Pair_eq";
goalw Prod.thy [fst_def] "fst(<a,b>) = a";
by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1);
-val fst_conv = result();
+qed "fst_conv";
goalw Prod.thy [snd_def] "snd(<a,b>) = b";
by (fast_tac (set_cs addIs [select_equality] addSEs [Pair_inject]) 1);
-val snd_conv = result();
+qed "snd_conv";
goalw Prod.thy [Pair_def] "? x y. p = <x,y>";
by (rtac (rewrite_rule [Prod_def] Rep_Prod RS CollectE) 1);
by (EVERY1[etac exE, etac exE, rtac exI, rtac exI,
rtac (Rep_Prod_inverse RS sym RS trans), etac arg_cong]);
-val PairE_lemma = result();
+qed "PairE_lemma";
val [prem] = goal Prod.thy "[| !!x y. p = <x,y> ==> Q |] ==> Q";
by (rtac (PairE_lemma RS exE) 1);
by (REPEAT (eresolve_tac [prem,exE] 1));
-val PairE = result();
+qed "PairE";
goalw Prod.thy [split_def] "split(c, <a,b>) = c(a,b)";
by (sstac [fst_conv, snd_conv] 1);
by (rtac refl 1);
-val split = result();
+qed "split";
val pair_ss = set_ss addsimps [fst_conv, snd_conv, split, Pair_eq];
@@ -64,7 +64,7 @@
by (res_inst_tac[("p","s")] PairE 1);
by (res_inst_tac[("p","t")] PairE 1);
by (asm_simp_tac pair_ss 1);
-val Pair_fst_snd_eq = result();
+qed "Pair_fst_snd_eq";
(*Prevents simplification of c: much faster*)
val split_weak_cong = prove_goal Prod.thy
@@ -74,19 +74,19 @@
goal Prod.thy "p = <fst(p),snd(p)>";
by (res_inst_tac [("p","p")] PairE 1);
by (asm_simp_tac pair_ss 1);
-val surjective_pairing = result();
+qed "surjective_pairing";
goal Prod.thy "p = split(%x y.<x,y>, p)";
by (res_inst_tac [("p","p")] PairE 1);
by (asm_simp_tac pair_ss 1);
-val surjective_pairing2 = result();
+qed "surjective_pairing2";
(*For use with split_tac and the simplifier*)
goal Prod.thy "R(split(c,p)) = (! x y. p = <x,y> --> R(c(x,y)))";
by (stac surjective_pairing 1);
by (stac split 1);
by (fast_tac (HOL_cs addSEs [Pair_inject]) 1);
-val expand_split = result();
+qed "expand_split";
(** split used as a logical connective or set former **)
@@ -95,50 +95,50 @@
goal Prod.thy "!!a b c. c(a,b) ==> split(c, <a,b>)";
by (asm_simp_tac pair_ss 1);
-val splitI = result();
+qed "splitI";
val prems = goalw Prod.thy [split_def]
"[| split(c,p); !!x y. [| p = <x,y>; c(x,y) |] ==> Q |] ==> Q";
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
-val splitE = result();
+qed "splitE";
goal Prod.thy "!!R a b. split(R,<a,b>) ==> R(a,b)";
by (etac (split RS iffD1) 1);
-val splitD = result();
+qed "splitD";
goal Prod.thy "!!a b c. z: c(a,b) ==> z: split(c, <a,b>)";
by (asm_simp_tac pair_ss 1);
-val mem_splitI = result();
+qed "mem_splitI";
val prems = goalw Prod.thy [split_def]
"[| z: split(c,p); !!x y. [| p = <x,y>; z: c(x,y) |] ==> Q |] ==> Q";
by (REPEAT (resolve_tac (prems@[surjective_pairing]) 1));
-val mem_splitE = result();
+qed "mem_splitE";
(*** prod_fun -- action of the product functor upon functions ***)
goalw Prod.thy [prod_fun_def] "prod_fun(f,g,<a,b>) = <f(a),g(b)>";
by (rtac split 1);
-val prod_fun = result();
+qed "prod_fun";
goal Prod.thy
"prod_fun(f1 o f2, g1 o g2) = (prod_fun(f1,g1) o prod_fun(f2,g2))";
by (rtac ext 1);
by (res_inst_tac [("p","x")] PairE 1);
by (asm_simp_tac (pair_ss addsimps [prod_fun,o_def]) 1);
-val prod_fun_compose = result();
+qed "prod_fun_compose";
goal Prod.thy "prod_fun(%x.x, %y.y) = (%z.z)";
by (rtac ext 1);
by (res_inst_tac [("p","z")] PairE 1);
by (asm_simp_tac (pair_ss addsimps [prod_fun]) 1);
-val prod_fun_ident = result();
+qed "prod_fun_ident";
val prems = goal Prod.thy "<a,b>:r ==> <f(a),g(b)> : prod_fun(f,g)``r";
by (rtac image_eqI 1);
by (rtac (prod_fun RS sym) 1);
by (resolve_tac prems 1);
-val prod_fun_imageI = result();
+qed "prod_fun_imageI";
val major::prems = goal Prod.thy
"[| c: prod_fun(f,g)``r; !!x y. [| c=<f(x),g(y)>; <x,y>:r |] ==> P \
@@ -148,7 +148,7 @@
by (resolve_tac prems 1);
by (fast_tac HOL_cs 2);
by (fast_tac (HOL_cs addIs [prod_fun]) 1);
-val prod_fun_imageE = result();
+qed "prod_fun_imageE";
(*** Disjoint union of a family of sets - Sigma ***)
@@ -192,7 +192,7 @@
by (rtac bexI 1);
by (rtac (fst_conv RS sym) 1);
by (resolve_tac prems 1);
-val fst_imageI = result();
+qed "fst_imageI";
val major::prems = goal Prod.thy
"[| a : fst``r; !!y.[| <a,y> : r |] ==> P |] ==> P";
@@ -201,7 +201,7 @@
by (etac ssubst 1);
by (rtac (surjective_pairing RS subst) 1);
by (assume_tac 1);
-val fst_imageE = result();
+qed "fst_imageE";
(*** Range of a relation ***)
@@ -210,7 +210,7 @@
by (rtac bexI 1);
by (rtac (snd_conv RS sym) 1);
by (resolve_tac prems 1);
-val snd_imageI = result();
+qed "snd_imageI";
val major::prems = goal Prod.thy
"[| a : snd``r; !!y.[| <y,a> : r |] ==> P |] ==> P";
@@ -219,7 +219,7 @@
by (etac ssubst 1);
by (rtac (surjective_pairing RS subst) 1);
by (assume_tac 1);
-val snd_imageE = result();
+qed "snd_imageE";
(** Exhaustion rule for unit -- a degenerate form of induction **)
@@ -227,7 +227,7 @@
"u = Unity";
by (stac (rewrite_rule [Unit_def] Rep_Unit RS CollectD RS sym) 1);
by (rtac (Rep_Unit_inverse RS sym) 1);
-val unit_eq = result();
+qed "unit_eq";
val prod_cs = set_cs addSIs [SigmaI, mem_splitI]
addIs [fst_imageI, snd_imageI, prod_fun_imageI]
--- a/Set.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Set.ML Mon Nov 21 17:50:34 1994 +0100
@@ -11,21 +11,21 @@
val [prem] = goal Set.thy "[| P(a) |] ==> a : {x.P(x)}";
by (rtac (mem_Collect_eq RS ssubst) 1);
by (rtac prem 1);
-val CollectI = result();
+qed "CollectI";
val prems = goal Set.thy "[| a : {x.P(x)} |] ==> P(a)";
by (resolve_tac (prems RL [mem_Collect_eq RS subst]) 1);
-val CollectD = result();
+qed "CollectD";
val [prem] = goal Set.thy "[| !!x. (x:A) = (x:B) |] ==> A = B";
by (rtac (prem RS ext RS arg_cong RS box_equals) 1);
by (rtac Collect_mem_eq 1);
by (rtac Collect_mem_eq 1);
-val set_ext = result();
+qed "set_ext";
val [prem] = goal Set.thy "[| !!x. P(x)=Q(x) |] ==> {x. P(x)} = {x. Q(x)}";
by (rtac (prem RS ext RS arg_cong) 1);
-val Collect_cong = result();
+qed "Collect_cong";
val CollectE = make_elim CollectD;
@@ -34,18 +34,18 @@
val prems = goalw Set.thy [Ball_def]
"[| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)";
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));
-val ballI = result();
+qed "ballI";
val [major,minor] = goalw Set.thy [Ball_def]
"[| ! x:A. P(x); x:A |] ==> P(x)";
by (rtac (minor RS (major RS spec RS mp)) 1);
-val bspec = result();
+qed "bspec";
val major::prems = goalw Set.thy [Ball_def]
"[| ! x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q";
by (rtac (major RS spec RS impCE) 1);
by (REPEAT (eresolve_tac prems 1));
-val ballE = result();
+qed "ballE";
(*Takes assumptions ! x:A.P(x) and a:A; creates assumption P(a)*)
fun ball_tac i = etac ballE i THEN contr_tac (i+1);
@@ -53,7 +53,7 @@
val prems = goalw Set.thy [Bex_def]
"[| P(x); x:A |] ==> ? x:A. P(x)";
by (REPEAT (ares_tac (prems @ [exI,conjI]) 1));
-val bexI = result();
+qed "bexI";
val bexCI = prove_goal Set.thy
"[| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A.P(x)"
@@ -65,13 +65,13 @@
"[| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q";
by (rtac (major RS exE) 1);
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));
-val bexE = result();
+qed "bexE";
(*Trival rewrite rule; (! x:A.P)=P holds only if A is nonempty!*)
val prems = goal Set.thy
"(! x:A. True) = True";
by (REPEAT (ares_tac [TrueI,ballI,iffI] 1));
-val ball_rew = result();
+qed "ball_rew";
(** Congruence rules **)
@@ -81,7 +81,7 @@
by (resolve_tac (prems RL [ssubst]) 1);
by (REPEAT (ares_tac [ballI,iffI] 1
ORELSE eresolve_tac ([make_elim bspec, mp] @ (prems RL [iffE])) 1));
-val ball_cong = result();
+qed "ball_cong";
val prems = goal Set.thy
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \
@@ -89,19 +89,19 @@
by (resolve_tac (prems RL [ssubst]) 1);
by (REPEAT (etac bexE 1
ORELSE ares_tac ([bexI,iffI] @ (prems RL [iffD1,iffD2])) 1));
-val bex_cong = result();
+qed "bex_cong";
(*** Subsets ***)
val prems = goalw Set.thy [subset_def] "(!!x.x:A ==> x:B) ==> A <= B";
by (REPEAT (ares_tac (prems @ [ballI]) 1));
-val subsetI = result();
+qed "subsetI";
(*Rule in Modus Ponens style*)
val major::prems = goalw Set.thy [subset_def] "[| A <= B; c:A |] ==> c:B";
by (rtac (major RS bspec) 1);
by (resolve_tac prems 1);
-val subsetD = result();
+qed "subsetD";
(*The same, with reversed premises for use with etac -- cf rev_mp*)
val rev_subsetD = prove_goal Set.thy "[| c:A; A <= B |] ==> c:B"
@@ -112,7 +112,7 @@
"[| A <= B; c~:A ==> P; c:B ==> P |] ==> P";
by (rtac (major RS ballE) 1);
by (REPEAT (eresolve_tac prems 1));
-val subsetCE = result();
+qed "subsetCE";
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)
fun set_mp_tac i = etac subsetCE i THEN mp_tac i;
@@ -123,7 +123,7 @@
val prems = goal Set.thy "[| A<=B; B<=C |] ==> A<=(C::'a set)";
by (cut_facts_tac prems 1);
by (REPEAT (ares_tac [subsetI] 1 ORELSE set_mp_tac 1));
-val subset_trans = result();
+qed "subset_trans";
(*** Equality ***)
@@ -132,31 +132,31 @@
val prems = goal Set.thy "[| A <= B; B <= A |] ==> A = (B::'a set)";
by (rtac (iffI RS set_ext) 1);
by (REPEAT (ares_tac (prems RL [subsetD]) 1));
-val subset_antisym = result();
+qed "subset_antisym";
val equalityI = subset_antisym;
(* Equality rules from ZF set theory -- are they appropriate here? *)
val prems = goal Set.thy "A = B ==> A<=(B::'a set)";
by (resolve_tac (prems RL [subst]) 1);
by (rtac subset_refl 1);
-val equalityD1 = result();
+qed "equalityD1";
val prems = goal Set.thy "A = B ==> B<=(A::'a set)";
by (resolve_tac (prems RL [subst]) 1);
by (rtac subset_refl 1);
-val equalityD2 = result();
+qed "equalityD2";
val prems = goal Set.thy
"[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P";
by (resolve_tac prems 1);
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));
-val equalityE = result();
+qed "equalityE";
val major::prems = goal Set.thy
"[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P";
by (rtac (major RS equalityE) 1);
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));
-val equalityCE = result();
+qed "equalityCE";
(*Lemma for creating induction formulae -- for "pattern matching" on p
To make the induction hypotheses usable, apply "spec" or "bspec" to
@@ -165,7 +165,7 @@
"[| p:A; !!z. z:A ==> p=z --> R |] ==> R";
by (rtac mp 1);
by (REPEAT (resolve_tac (refl::prems) 1));
-val setup_induction = result();
+qed "setup_induction";
(*** Set complement -- Compl ***)
@@ -173,7 +173,7 @@
val prems = goalw Set.thy [Compl_def]
"[| c:A ==> False |] ==> c : Compl(A)";
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));
-val ComplI = result();
+qed "ComplI";
(*This form, with negated conclusion, works well with the Classical prover.
Negated assumptions behave like formulae on the right side of the notional
@@ -181,7 +181,7 @@
val major::prems = goalw Set.thy [Compl_def]
"[| c : Compl(A) |] ==> c~:A";
by (rtac (major RS CollectD) 1);
-val ComplD = result();
+qed "ComplD";
val ComplE = make_elim ComplD;
@@ -190,11 +190,11 @@
val prems = goalw Set.thy [Un_def] "c:A ==> c : A Un B";
by (REPEAT (resolve_tac (prems @ [CollectI,disjI1]) 1));
-val UnI1 = result();
+qed "UnI1";
val prems = goalw Set.thy [Un_def] "c:B ==> c : A Un B";
by (REPEAT (resolve_tac (prems @ [CollectI,disjI2]) 1));
-val UnI2 = result();
+qed "UnI2";
(*Classical introduction rule: no commitment to A vs B*)
val UnCI = prove_goal Set.thy "(c~:B ==> c:A) ==> c : A Un B"
@@ -207,7 +207,7 @@
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P";
by (rtac (major RS CollectD RS disjE) 1);
by (REPEAT (eresolve_tac prems 1));
-val UnE = result();
+qed "UnE";
(*** Binary intersection -- Int ***)
@@ -215,22 +215,22 @@
val prems = goalw Set.thy [Int_def]
"[| c:A; c:B |] ==> c : A Int B";
by (REPEAT (resolve_tac (prems @ [CollectI,conjI]) 1));
-val IntI = result();
+qed "IntI";
val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:A";
by (rtac (major RS CollectD RS conjunct1) 1);
-val IntD1 = result();
+qed "IntD1";
val [major] = goalw Set.thy [Int_def] "c : A Int B ==> c:B";
by (rtac (major RS CollectD RS conjunct2) 1);
-val IntD2 = result();
+qed "IntD2";
val [major,minor] = goal Set.thy
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P";
by (rtac minor 1);
by (rtac (major RS IntD1) 1);
by (rtac (major RS IntD2) 1);
-val IntE = result();
+qed "IntE";
(*** Set difference ***)
@@ -311,14 +311,14 @@
goalw Set.thy [insert_def] "!!a. b : {a} ==> b=a";
by(fast_tac (HOL_cs addSEs [emptyE,CollectE,UnE]) 1);
-val singletonD = result();
+qed "singletonD";
val singletonE = make_elim singletonD;
val [major] = goal Set.thy "{a}={b} ==> a=b";
by (rtac (major RS equalityD1 RS subsetD RS singletonD) 1);
by (rtac singletonI 1);
-val singleton_inject = result();
+qed "singleton_inject";
(*** Unions of families -- UNION x:A. B(x) is Union(B``A) ***)
@@ -326,13 +326,13 @@
val prems = goalw Set.thy [UNION_def]
"[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))";
by (REPEAT (resolve_tac (prems @ [bexI,CollectI]) 1));
-val UN_I = result();
+qed "UN_I";
val major::prems = goalw Set.thy [UNION_def]
"[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R";
by (rtac (major RS CollectD RS bexE) 1);
by (REPEAT (ares_tac prems 1));
-val UN_E = result();
+qed "UN_E";
val prems = goal Set.thy
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \
@@ -340,7 +340,7 @@
by (REPEAT (etac UN_E 1
ORELSE ares_tac ([UN_I,equalityI,subsetI] @
(prems RL [equalityD1,equalityD2] RL [subsetD])) 1));
-val UN_cong = result();
+qed "UN_cong";
(*** Intersections of families -- INTER x:A. B(x) is Inter(B``A) *)
@@ -348,20 +348,20 @@
val prems = goalw Set.thy [INTER_def]
"(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));
-val INT_I = result();
+qed "INT_I";
val major::prems = goalw Set.thy [INTER_def]
"[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)";
by (rtac (major RS CollectD RS bspec) 1);
by (resolve_tac prems 1);
-val INT_D = result();
+qed "INT_D";
(*"Classical" elimination -- by the Excluded Middle on a:A *)
val major::prems = goalw Set.thy [INTER_def]
"[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R";
by (rtac (major RS CollectD RS ballE) 1);
by (REPEAT (eresolve_tac prems 1));
-val INT_E = result();
+qed "INT_E";
val prems = goal Set.thy
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \
@@ -369,7 +369,7 @@
by (REPEAT_FIRST (resolve_tac [INT_I,equalityI,subsetI]));
by (REPEAT (dtac INT_D 1
ORELSE ares_tac (prems RL [equalityD1,equalityD2] RL [subsetD]) 1));
-val INT_cong = result();
+qed "INT_cong";
(*** Unions over a type; UNION1(B) = Union(range(B)) ***)
@@ -378,13 +378,13 @@
val prems = goalw Set.thy [UNION1_def]
"b: B(x) ==> b: (UN x. B(x))";
by (REPEAT (resolve_tac (prems @ [TrueI, CollectI RS UN_I]) 1));
-val UN1_I = result();
+qed "UN1_I";
val major::prems = goalw Set.thy [UNION1_def]
"[| b : (UN x. B(x)); !!x. b: B(x) ==> R |] ==> R";
by (rtac (major RS UN_E) 1);
by (REPEAT (ares_tac prems 1));
-val UN1_E = result();
+qed "UN1_E";
(*** Intersections over a type; INTER1(B) = Inter(range(B)) *)
@@ -392,12 +392,12 @@
val prems = goalw Set.thy [INTER1_def]
"(!!x. b: B(x)) ==> b : (INT x. B(x))";
by (REPEAT (ares_tac (INT_I::prems) 1));
-val INT1_I = result();
+qed "INT1_I";
val [major] = goalw Set.thy [INTER1_def]
"b : (INT x. B(x)) ==> b: B(a)";
by (rtac (TrueI RS (CollectI RS (major RS INT_D))) 1);
-val INT1_D = result();
+qed "INT1_D";
(*** Unions ***)
@@ -405,20 +405,20 @@
val prems = goalw Set.thy [Union_def]
"[| X:C; A:X |] ==> A : Union(C)";
by (REPEAT (resolve_tac (prems @ [UN_I]) 1));
-val UnionI = result();
+qed "UnionI";
val major::prems = goalw Set.thy [Union_def]
"[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R";
by (rtac (major RS UN_E) 1);
by (REPEAT (ares_tac prems 1));
-val UnionE = result();
+qed "UnionE";
(*** Inter ***)
val prems = goalw Set.thy [Inter_def]
"[| !!X. X:C ==> A:X |] ==> A : Inter(C)";
by (REPEAT (ares_tac ([INT_I] @ prems) 1));
-val InterI = result();
+qed "InterI";
(*A "destruct" rule -- every X in C contains A as an element, but
A:X can hold when X:C does not! This rule is analogous to "spec". *)
@@ -426,14 +426,14 @@
"[| A : Inter(C); X:C |] ==> A:X";
by (rtac (major RS INT_D) 1);
by (resolve_tac prems 1);
-val InterD = result();
+qed "InterD";
(*"Classical" elimination rule -- does not require proving X:C *)
val major::prems = goalw Set.thy [Inter_def]
"[| A : Inter(C); A:X ==> R; X~:C ==> R |] ==> R";
by (rtac (major RS INT_E) 1);
by (REPEAT (eresolve_tac prems 1));
-val InterE = result();
+qed "InterE";
(*** Powerset ***)
--- a/Sexp.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Sexp.ML Mon Nov 21 17:50:34 1994 +0100
@@ -19,17 +19,17 @@
goalw Sexp.thy [sexp_case_def] "sexp_case(c, d, e, Leaf(a)) = c(a)";
by (resolve_tac [select_equality] 1);
by (ALLGOALS (fast_tac sexp_free_cs));
-val sexp_case_Leaf = result();
+qed "sexp_case_Leaf";
goalw Sexp.thy [sexp_case_def] "sexp_case(c, d, e, Numb(k)) = d(k)";
by (resolve_tac [select_equality] 1);
by (ALLGOALS (fast_tac sexp_free_cs));
-val sexp_case_Numb = result();
+qed "sexp_case_Numb";
goalw Sexp.thy [sexp_case_def] "sexp_case(c, d, e, M$N) = e(M,N)";
by (resolve_tac [select_equality] 1);
by (ALLGOALS (fast_tac sexp_free_cs));
-val sexp_case_Scons = result();
+qed "sexp_case_Scons";
(** Introduction rules for sexp constructors **)
@@ -37,31 +37,31 @@
val [prem] = goalw Sexp.thy [In0_def]
"M: sexp ==> In0(M) : sexp";
by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1);
-val sexp_In0I = result();
+qed "sexp_In0I";
val [prem] = goalw Sexp.thy [In1_def]
"M: sexp ==> In1(M) : sexp";
by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1);
-val sexp_In1I = result();
+qed "sexp_In1I";
val sexp_cs = set_cs addIs sexp.intrs@[SigmaI, uprodI];
goal Sexp.thy "range(Leaf) <= sexp";
by (fast_tac sexp_cs 1);
-val range_Leaf_subset_sexp = result();
+qed "range_Leaf_subset_sexp";
val [major] = goal Sexp.thy "M$N : sexp ==> M: sexp & N: sexp";
by (rtac (major RS setup_induction) 1);
by (etac sexp.induct 1);
by (ALLGOALS
(fast_tac (set_cs addSEs [Scons_neq_Leaf,Scons_neq_Numb,Scons_inject])));
-val Scons_D = result();
+qed "Scons_D";
(** Introduction rules for 'pred_sexp' **)
goalw Sexp.thy [pred_sexp_def] "pred_sexp <= Sigma(sexp, %u.sexp)";
by (fast_tac sexp_cs 1);
-val pred_sexp_subset_Sigma = result();
+qed "pred_sexp_subset_Sigma";
(* <a,b> : pred_sexp^+ ==> a : sexp *)
val trancl_pred_sexpD1 =
@@ -72,12 +72,12 @@
val prems = goalw Sexp.thy [pred_sexp_def]
"[| M: sexp; N: sexp |] ==> <M, M$N> : pred_sexp";
by (fast_tac (set_cs addIs prems) 1);
-val pred_sexpI1 = result();
+qed "pred_sexpI1";
val prems = goalw Sexp.thy [pred_sexp_def]
"[| M: sexp; N: sexp |] ==> <N, M$N> : pred_sexp";
by (fast_tac (set_cs addIs prems) 1);
-val pred_sexpI2 = result();
+qed "pred_sexpI2";
(*Combinations involving transitivity and the rules above*)
val pred_sexp_t1 = pred_sexpI1 RS r_into_trancl
@@ -100,7 +100,7 @@
\ |] ==> R";
by (cut_facts_tac [major] 1);
by (REPEAT (eresolve_tac ([asm_rl,emptyE,insertE,UN_E]@prems) 1));
-val pred_sexpE = result();
+qed "pred_sexpE";
goal Sexp.thy "wf(pred_sexp)";
by (rtac (pred_sexp_subset_Sigma RS wfI) 1);
@@ -108,7 +108,7 @@
by (fast_tac (HOL_cs addSEs [mp, pred_sexpE, Pair_inject, Scons_inject]) 3);
by (fast_tac (HOL_cs addSEs [mp, pred_sexpE, Pair_inject, Numb_neq_Scons]) 2);
by (fast_tac (HOL_cs addSEs [mp, pred_sexpE, Pair_inject, Leaf_neq_Scons]) 1);
-val wf_pred_sexp = result();
+qed "wf_pred_sexp";
(*** sexp_rec -- by wf recursion on pred_sexp ***)
@@ -120,16 +120,16 @@
goal Sexp.thy "sexp_rec(Leaf(a), c, d, h) = c(a)";
by (stac sexp_rec_unfold 1);
by (rtac sexp_case_Leaf 1);
-val sexp_rec_Leaf = result();
+qed "sexp_rec_Leaf";
goal Sexp.thy "sexp_rec(Numb(k), c, d, h) = d(k)";
by (stac sexp_rec_unfold 1);
by (rtac sexp_case_Numb 1);
-val sexp_rec_Numb = result();
+qed "sexp_rec_Numb";
goal Sexp.thy "!!M. [| M: sexp; N: sexp |] ==> \
\ sexp_rec(M$N, c, d, h) = h(M, N, sexp_rec(M,c,d,h), sexp_rec(N,c,d,h))";
by (rtac (sexp_rec_unfold RS trans) 1);
by (asm_simp_tac(HOL_ss addsimps
[sexp_case_Scons,pred_sexpI1,pred_sexpI2,cut_apply])1);
-val sexp_rec_Scons = result();
+qed "sexp_rec_Scons";
--- a/Subst/AList.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Subst/AList.ML Mon Nov 21 17:50:34 1994 +0100
@@ -34,7 +34,7 @@
by (etac ssubst 1);
by (resolve_tac prems 1);
by (assume_tac 1);
-val alist_induct = result();
+qed "alist_induct";
(*Perform induction on xs. *)
fun alist_ind_tac a M =
@@ -46,7 +46,7 @@
"[| P(Nil); !! x y xs. P(xs) --> P(<x,y>#xs) |] ==> P(a)";
by (alist_ind_tac "a" 1);
by (ALLGOALS (cut_facts_tac prems THEN' fast_tac HOL_cs));
-val alist_induct2 = result();
+qed "alist_induct2";
add_inds alist_induct2;
*)
--- a/Subst/Setplus.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Subst/Setplus.ML Mon Nov 21 17:50:34 1994 +0100
@@ -14,39 +14,39 @@
goal Set.thy "A <= B = (! t.t:A --> t:B)";
by (fast_tac set_cs 1);
-val subset_iff = result();
+qed "subset_iff";
goalw Setplus.thy [ssubset_def] "A < B = ((A <= B) & ~(A=B))";
by (rtac refl 1);
-val ssubset_iff = result();
+qed "ssubset_iff";
goal Setplus.thy "((A::'a set) <= B) = ((A < B) | (A=B))";
by (simp_tac (set_ss addsimps [ssubset_iff]) 1);
by (fast_tac set_cs 1);
-val subseteq_iff_subset_eq = result();
+qed "subseteq_iff_subset_eq";
(*Rule in Modus Ponens style*)
goal Setplus.thy "A < B --> c:A --> c:B";
by (simp_tac (set_ss addsimps [ssubset_iff]) 1);
by (fast_tac set_cs 1);
-val ssubsetD = result();
+qed "ssubsetD";
(*********)
goalw Setplus.thy [empty_def] "~ a : {}";
by (fast_tac set_cs 1);
-val not_in_empty = result();
+qed "not_in_empty";
goalw Setplus.thy [empty_def] "(A = {}) = (ALL a.~ a:A)";
by (fast_tac (set_cs addIs [set_ext]) 1);
-val empty_iff = result();
+qed "empty_iff";
(*********)
goal Set.thy "(~A=B) = ((? x.x:A & ~x:B) | (? x.~x:A & x:B))";
by (fast_tac (set_cs addIs [set_ext]) 1);
-val not_equal_iff = result();
+qed "not_equal_iff";
(*********)
@@ -59,6 +59,6 @@
by (rtac (excluded_middle RS disjE) 1);
by (etac (prem2 RS mp) 1);
by (etac (prem1 RS mp) 1);
-val imp_excluded_middle = result();
+qed "imp_excluded_middle";
fun imp_excluded_middle_tac s = res_inst_tac [("P",s)] imp_excluded_middle;
--- a/Subst/Subst.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Subst/Subst.ML Mon Nov 21 17:50:34 1994 +0100
@@ -34,12 +34,12 @@
goal Subst.thy "t <| Nil = t";
by (uterm_ind_tac "t" 1);
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst])));
-val subst_Nil = result();
+qed "subst_Nil";
goal Subst.thy "t <: u --> t <| s <: u <| s";
by (uterm_ind_tac "u" 1);
by (ALLGOALS (asm_simp_tac subst_ss));
-val subst_mono = result() RS mp;
+val subst_mono = store_thm("subst_mono", result() RS mp);
goal Subst.thy "~ (Var(v) <: t) --> t <| <v,t <| s>#s = t <| s";
by (imp_excluded_middle_tac "t = Var(v)" 1);
@@ -48,7 +48,7 @@
uterm_induct 2);
by (ALLGOALS (simp_tac (subst_ss addsimps [Var_subst])));
by (fast_tac HOL_cs 1);
-val Var_not_occs = result() RS mp;
+val Var_not_occs = store_thm("Var_not_occs", result() RS mp);
goal Subst.thy
"(t <|r = t <|s) = (! v.v : vars_of(t) --> Var(v) <|r = Var(v) <|s)";
@@ -56,24 +56,24 @@
by (REPEAT (etac rev_mp 3));
by (ALLGOALS (asm_simp_tac subst_ss));
by (ALLGOALS (fast_tac HOL_cs));
-val agreement = result();
+qed "agreement";
goal Subst.thy "~ v: vars_of(t) --> t <| <v,u>#s = t <| s";
by(simp_tac(subst_ss addsimps [agreement,Var_subst]
setloop (split_tac [expand_if])) 1);
-val repl_invariance = result() RS mp;
+val repl_invariance = store_thm("repl_invariance", result() RS mp);
val asms = goal Subst.thy
"v : vars_of(t) --> w : vars_of(t <| <v,Var(w)>#s)";
by (uterm_ind_tac "t" 1);
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst])));
-val Var_in_subst = result() RS mp;
+val Var_in_subst = store_thm("Var_in_subst", result() RS mp);
(**** Equality between Substitutions ****)
goalw Subst.thy [subst_eq_def] "r =s= s = (! t.t <| r = t <| s)";
by (simp_tac subst_ss 1);
-val subst_eq_iff = result();
+qed "subst_eq_iff";
local fun mk_thm s = prove_goal Subst.thy s
(fn prems => [cut_facts_tac prems 1,
@@ -89,7 +89,7 @@
"[| r =s= s; P(t <| r,u <| r) |] ==> P(t <| s,u <| s)";
by (resolve_tac [eq RS spec RS subst] 1);
by (resolve_tac (prems RL [eq RS spec RS subst]) 1);
-val subst_subst2 = result();
+qed "subst_subst2";
val ssubst_subst2 = subst_sym RS subst_subst2;
@@ -98,7 +98,7 @@
goal Subst.thy "s <> Nil = s";
by (alist_ind_tac "s" 1);
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [subst_Nil])));
-val comp_Nil = result();
+qed "comp_Nil";
goal Subst.thy "(t <| r <> s) = (t <| r <| s)";
by (uterm_ind_tac "t" 1);
@@ -106,11 +106,11 @@
by (alist_ind_tac "r" 1);
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst,subst_Nil]
setloop (split_tac [expand_if]))));
-val subst_comp = result();
+qed "subst_comp";
goal Subst.thy "q <> r <> s =s= q <> (r <> s)";
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1);
-val comp_assoc = result();
+qed "comp_assoc";
goal Subst.thy "<w,Var(w) <| s>#s =s= s";
by (rtac (allI RS (subst_eq_iff RS iffD2)) 1);
@@ -118,12 +118,12 @@
by (REPEAT (etac rev_mp 3));
by (ALLGOALS (simp_tac (subst_ss addsimps[Var_subst]
setloop (split_tac [expand_if]))));
-val Cons_trivial = result();
+qed "Cons_trivial";
val [prem] = goal Subst.thy "q <> r =s= s ==> t <| q <| r = t <| s";
by (simp_tac (subst_ss addsimps [prem RS (subst_eq_iff RS iffD1),
subst_comp RS sym]) 1);
-val comp_subst_subst = result();
+qed "comp_subst_subst";
(**** Domain and range of Substitutions ****)
@@ -132,31 +132,31 @@
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]
setloop (split_tac[expand_if]))));
by (fast_tac HOL_cs 1);
-val sdom_iff = result();
+qed "sdom_iff";
goalw Subst.thy [srange_def]
"v : srange(s) = (? w.w : sdom(s) & v : vars_of(Var(w) <| s))";
by (fast_tac set_cs 1);
-val srange_iff = result();
+qed "srange_iff";
goal Subst.thy "(t <| s = t) = (sdom(s) Int vars_of(t) = {})";
by (uterm_ind_tac "t" 1);
by (REPEAT (etac rev_mp 3));
by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,Var_subst])));
by (ALLGOALS (fast_tac set_cs));
-val invariance = result();
+qed "invariance";
goal Subst.thy "v : sdom(s) --> ~v : srange(s) --> ~v : vars_of(t <| s)";
by (uterm_ind_tac "t" 1);
by (imp_excluded_middle_tac "x : sdom(s)" 1);
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [sdom_iff,srange_iff])));
by (ALLGOALS (fast_tac set_cs));
-val Var_elim = result() RS mp RS mp;
+val Var_elim = store_thm("Var_elim", result() RS mp RS mp);
val asms = goal Subst.thy
"[| v : sdom(s); v : vars_of(t <| s) |] ==> v : srange(s)";
by (REPEAT (ares_tac (asms @ [Var_elim RS swap RS classical]) 1));
-val Var_elim2 = result();
+qed "Var_elim2";
goal Subst.thy "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)";
by (uterm_ind_tac "t" 1);
@@ -166,20 +166,20 @@
by (etac notE 1);
by (etac subst 1);
by (ALLGOALS (fast_tac set_cs));
-val Var_intro = result() RS mp;
+val Var_intro = store_thm("Var_intro", result() RS mp);
goal Subst.thy
"v : srange(s) --> (? w.w : sdom(s) & v : vars_of(Var(w) <| s))";
by (simp_tac (subst_ss addsimps [srange_iff]) 1);
-val srangeE = make_elim (result() RS mp);
+val srangeE = store_thm("srangeE", make_elim (result() RS mp));
val asms = goal Subst.thy
"sdom(s) Int srange(s) = {} = (! t.sdom(s) Int vars_of(t <| s) = {})";
by (simp_tac subst_ss 1);
by (fast_tac (set_cs addIs [Var_elim2] addEs [srangeE]) 1);
-val dom_range_disjoint = result();
+qed "dom_range_disjoint";
val asms = goal Subst.thy "~ u <| s = u --> (? x.x : sdom(s))";
by (simp_tac (subst_ss addsimps [invariance]) 1);
by (fast_tac set_cs 1);
-val subst_not_empty = result() RS mp;
+val subst_not_empty = store_thm("subst_not_empty", result() RS mp);
--- a/Subst/UTLemmas.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Subst/UTLemmas.ML Mon Nov 21 17:50:34 1994 +0100
@@ -30,23 +30,23 @@
by (uterm_ind_tac "v" 1);
by (ALLGOALS (simp_tac utlemmas_ss));
by (fast_tac HOL_cs 1);
-val occs_trans = conjI RS (result() RS mp);
+val occs_trans = store_thm("occs_trans", conjI RS (result() RS mp));
goal UTLemmas.thy "~ t <: v --> ~ t <: u | ~ u <: v";
by (fast_tac (HOL_cs addIs [occs_trans]) 1);
-val contr_occs_trans = result() RS mp;
+val contr_occs_trans = store_thm("contr_occs_trans", result() RS mp);
goal UTLemmas.thy "t <: Comb(t,u)";
by (simp_tac utlemmas_ss 1);
-val occs_Comb1 = result();
+qed "occs_Comb1";
goal UTLemmas.thy "u <: Comb(t,u)";
by (simp_tac utlemmas_ss 1);
-val occs_Comb2 = result();
+qed "occs_Comb2";
goal HOL.thy "(~(P|Q)) = (~P & ~Q)";
by (fast_tac HOL_cs 1);
-val demorgan_disj = result();
+qed "demorgan_disj";
goal UTLemmas.thy "~ t <: t";
by (uterm_ind_tac "t" 1);
@@ -56,11 +56,11 @@
resolve_tac [occs_Comb1,occs_Comb2] 1) ORELSE
(etac (contr_occs_trans RS disjE) 1 THEN assume_tac 2) ORELSE
eresolve_tac ([occs_Comb1,occs_Comb2] RLN(2,[notE])) 1));
-val occs_irrefl = result();
+qed "occs_irrefl";
goal UTLemmas.thy "t <: u --> ~t=u";
by (fast_tac (HOL_cs addEs [occs_irrefl RS notE]) 1);
-val occs_irrefl2 = result() RS mp;
+val occs_irrefl2 = store_thm("occs_irrefl2", result() RS mp);
(**** vars_of lemmas ****)
@@ -68,10 +68,10 @@
goal UTLemmas.thy "(v : vars_of(Var(w))) = (w=v)";
by (simp_tac utlemmas_ss 1);
by (fast_tac HOL_cs 1);
-val vars_var_iff = result();
+qed "vars_var_iff";
goal UTLemmas.thy "(x : vars_of(t)) = (Var(x) <: t | Var(x) = t)";
by (uterm_ind_tac "t" 1);
by (ALLGOALS (simp_tac utlemmas_ss));
by (fast_tac HOL_cs 1);
-val vars_iff_occseq = result();
+qed "vars_iff_occseq";
--- a/Subst/UTerm.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Subst/UTerm.ML Mon Nov 21 17:50:34 1994 +0100
@@ -15,21 +15,21 @@
by (fast_tac (univ_cs addSIs (equalityI :: map rew uterm.intrs)
addEs [rew uterm.elim]) 1)
end;
-val uterm_unfold = result();
+qed "uterm_unfold";
(** the uterm functional **)
(*This justifies using uterm in other recursive type definitions*)
goalw UTerm.thy uterm.defs "!!A B. A<=B ==> uterm(A) <= uterm(B)";
by (REPEAT (ares_tac (lfp_mono::basic_monos) 1));
-val uterm_mono = result();
+qed "uterm_mono";
(** Type checking rules -- uterm creates well-founded sets **)
goalw UTerm.thy (uterm_con_defs @ uterm.defs) "uterm(sexp) <= sexp";
by (rtac lfp_lowerbound 1);
by (fast_tac (univ_cs addIs sexp.intrs@[sexp_In0I,sexp_In1I]) 1);
-val uterm_sexp = result();
+qed "uterm_sexp";
(* A <= sexp ==> uterm(A) <= sexp *)
val uterm_subset_sexp = standard ([uterm_mono, uterm_sexp] MRS subset_trans);
@@ -44,7 +44,7 @@
by (rtac (Rep_uterm RS uterm.induct) 1);
by (REPEAT (ares_tac prems 1
ORELSE eresolve_tac [rangeE, ssubst, Abs_uterm_inverse RS subst] 1));
-val uterm_induct = result();
+qed "uterm_induct";
(*Perform induction on xs. *)
fun uterm_ind_tac a M =
@@ -57,33 +57,33 @@
goal UTerm.thy "inj(Rep_uterm)";
by (rtac inj_inverseI 1);
by (rtac Rep_uterm_inverse 1);
-val inj_Rep_uterm = result();
+qed "inj_Rep_uterm";
goal UTerm.thy "inj_onto(Abs_uterm,uterm(range(Leaf)))";
by (rtac inj_onto_inverseI 1);
by (etac Abs_uterm_inverse 1);
-val inj_onto_Abs_uterm = result();
+qed "inj_onto_Abs_uterm";
(** Distinctness of constructors **)
goalw UTerm.thy uterm_con_defs "~ CONST(c) = COMB(u,v)";
by (rtac notI 1);
by (etac (In1_inject RS (In0_not_In1 RS notE)) 1);
-val CONST_not_COMB = result();
+qed "CONST_not_COMB";
val COMB_not_CONST = standard (CONST_not_COMB RS not_sym);
val CONST_neq_COMB = standard (CONST_not_COMB RS notE);
val COMB_neq_CONST = sym RS CONST_neq_COMB;
goalw UTerm.thy uterm_con_defs "~ COMB(u,v) = VAR(x)";
by (rtac In1_not_In0 1);
-val COMB_not_VAR = result();
+qed "COMB_not_VAR";
val VAR_not_COMB = standard (COMB_not_VAR RS not_sym);
val COMB_neq_VAR = standard (COMB_not_VAR RS notE);
val VAR_neq_COMB = sym RS COMB_neq_VAR;
goalw UTerm.thy uterm_con_defs "~ VAR(x) = CONST(c)";
by (rtac In0_not_In1 1);
-val VAR_not_CONST = result();
+qed "VAR_not_CONST";
val CONST_not_VAR = standard (VAR_not_CONST RS not_sym);
val VAR_neq_CONST = standard (VAR_not_CONST RS notE);
val CONST_neq_VAR = sym RS VAR_neq_CONST;
@@ -92,7 +92,7 @@
goalw UTerm.thy [Const_def,Comb_def] "~ Const(c) = Comb(u,v)";
by (rtac (CONST_not_COMB RS (inj_onto_Abs_uterm RS inj_onto_contraD)) 1);
by (REPEAT (resolve_tac (uterm.intrs @ [rangeI, Rep_uterm]) 1));
-val Const_not_Comb = result();
+qed "Const_not_Comb";
val Comb_not_Const = standard (Const_not_Comb RS not_sym);
val Const_neq_Comb = standard (Const_not_Comb RS notE);
val Comb_neq_Const = sym RS Const_neq_Comb;
@@ -100,7 +100,7 @@
goalw UTerm.thy [Comb_def,Var_def] "~ Comb(u,v) = Var(x)";
by (rtac (COMB_not_VAR RS (inj_onto_Abs_uterm RS inj_onto_contraD)) 1);
by (REPEAT (resolve_tac (uterm.intrs @ [rangeI, Rep_uterm]) 1));
-val Comb_not_Var = result();
+qed "Comb_not_Var";
val Var_not_Comb = standard (Comb_not_Var RS not_sym);
val Comb_neq_Var = standard (Comb_not_Var RS notE);
val Var_neq_Comb = sym RS Comb_neq_Var;
@@ -108,7 +108,7 @@
goalw UTerm.thy [Var_def,Const_def] "~ Var(x) = Const(c)";
by (rtac (VAR_not_CONST RS (inj_onto_Abs_uterm RS inj_onto_contraD)) 1);
by (REPEAT (resolve_tac (uterm.intrs @ [rangeI, Rep_uterm]) 1));
-val Var_not_Const = result();
+qed "Var_not_Const";
val Const_not_Var = standard (Var_not_Const RS not_sym);
val Var_neq_Const = standard (Var_not_Const RS notE);
val Const_neq_Var = sym RS Var_neq_Const;
@@ -121,15 +121,15 @@
goalw UTerm.thy [VAR_def] "(VAR(M)=VAR(N)) = (M=N)";
by (fast_tac inject_cs 1);
-val VAR_VAR_eq = result();
+qed "VAR_VAR_eq";
goalw UTerm.thy [CONST_def] "(CONST(M)=CONST(N)) = (M=N)";
by (fast_tac inject_cs 1);
-val CONST_CONST_eq = result();
+qed "CONST_CONST_eq";
goalw UTerm.thy [COMB_def] "(COMB(K,L)=COMB(M,N)) = (K=M & L=N)";
by (fast_tac inject_cs 1);
-val COMB_COMB_eq = result();
+qed "COMB_COMB_eq";
val VAR_inject = standard (VAR_VAR_eq RS iffD1);
val CONST_inject = standard (CONST_CONST_eq RS iffD1);
@@ -147,37 +147,37 @@
goalw UTerm.thy [Var_def] "(Var(x)=Var(y)) = (x=y)";
by (fast_tac uterm_cs 1);
-val Var_Var_eq = result();
+qed "Var_Var_eq";
val Var_inject = standard (Var_Var_eq RS iffD1);
goalw UTerm.thy [Const_def] "(Const(x)=Const(y)) = (x=y)";
by (fast_tac uterm_cs 1);
-val Const_Const_eq = result();
+qed "Const_Const_eq";
val Const_inject = standard (Const_Const_eq RS iffD1);
goalw UTerm.thy [Comb_def] "(Comb(u,v)=Comb(x,y)) = (u=x & v=y)";
by (fast_tac uterm_cs 1);
-val Comb_Comb_eq = result();
+qed "Comb_Comb_eq";
val Comb_inject = standard (Comb_Comb_eq RS iffD1 RS conjE);
val [major] = goal UTerm.thy "VAR(M): uterm(A) ==> M : A";
by (rtac (major RS setup_induction) 1);
by (etac uterm.induct 1);
by (ALLGOALS (fast_tac uterm_cs));
-val VAR_D = result();
+qed "VAR_D";
val [major] = goal UTerm.thy "CONST(M): uterm(A) ==> M : A";
by (rtac (major RS setup_induction) 1);
by (etac uterm.induct 1);
by (ALLGOALS (fast_tac uterm_cs));
-val CONST_D = result();
+qed "CONST_D";
val [major] = goal UTerm.thy
"COMB(M,N): uterm(A) ==> M: uterm(A) & N: uterm(A)";
by (rtac (major RS setup_induction) 1);
by (etac uterm.induct 1);
by (ALLGOALS (fast_tac uterm_cs));
-val COMB_D = result();
+qed "COMB_D";
(*Basic ss with constructors and their freeness*)
val uterm_free_simps = uterm.intrs @
@@ -194,14 +194,14 @@
by (rtac (Var_not_Comb RS allI) 1);
by (rtac (Const_not_Comb RS allI) 1);
by (asm_simp_tac uterm_free_ss 1);
-val t_not_Comb_t = result();
+qed "t_not_Comb_t";
goal UTerm.thy "!t. u~=Comb(t,u)";
by (uterm_ind_tac "u" 1);
by (rtac (Var_not_Comb RS allI) 1);
by (rtac (Const_not_Comb RS allI) 1);
by (asm_simp_tac uterm_free_ss 1);
-val u_not_Comb_u = result();
+qed "u_not_Comb_u";
(*** UTerm_rec -- by wf recursion on pred_sexp ***)
@@ -214,12 +214,12 @@
goalw UTerm.thy [VAR_def] "UTerm_rec(VAR(x),b,c,d) = b(x)";
by (rtac (UTerm_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [Case_In0]) 1);
-val UTerm_rec_VAR = result();
+qed "UTerm_rec_VAR";
goalw UTerm.thy [CONST_def] "UTerm_rec(CONST(x),b,c,d) = c(x)";
by (rtac (UTerm_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [Case_In0,Case_In1]) 1);
-val UTerm_rec_CONST = result();
+qed "UTerm_rec_CONST";
goalw UTerm.thy [COMB_def]
"!!M N. [| M: sexp; N: sexp |] ==> \
@@ -228,7 +228,7 @@
by (rtac (UTerm_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [Split,Case_In1]) 1);
by (asm_simp_tac (pred_sexp_ss addsimps [In1_def]) 1);
-val UTerm_rec_COMB = result();
+qed "UTerm_rec_COMB";
(*** uterm_rec -- by UTerm_rec ***)
@@ -244,16 +244,16 @@
goalw UTerm.thy [uterm_rec_def, Var_def] "uterm_rec(Var(x),b,c,d) = b(x)";
by (simp_tac uterm_rec_ss 1);
-val uterm_rec_Var = result();
+qed "uterm_rec_Var";
goalw UTerm.thy [uterm_rec_def, Const_def] "uterm_rec(Const(x),b,c,d) = c(x)";
by (simp_tac uterm_rec_ss 1);
-val uterm_rec_Const = result();
+qed "uterm_rec_Const";
goalw UTerm.thy [uterm_rec_def, Comb_def]
"uterm_rec(Comb(u,v),b,c,d) = d(u,v,uterm_rec(u,b,c,d),uterm_rec(v,b,c,d))";
by (simp_tac uterm_rec_ss 1);
-val uterm_rec_Comb = result();
+qed "uterm_rec_Comb";
val uterm_simps = [UTerm_rec_VAR, UTerm_rec_CONST, UTerm_rec_COMB,
uterm_rec_Var, uterm_rec_Const, uterm_rec_Comb];
--- a/Subst/Unifier.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Subst/Unifier.ML Mon Nov 21 17:50:34 1994 +0100
@@ -16,29 +16,29 @@
goalw Unifier.thy [Unifier_def] "Unifier(s,t,u) = (t <| s = u <| s)";
by (rtac refl 1);
-val Unifier_iff = result();
+qed "Unifier_iff";
goal Unifier.thy
"Unifier(s,Comb(t,u),Comb(v,w)) --> Unifier(s,t,v) & Unifier(s,u,w)";
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1);
-val Unifier_Comb = result() RS mp RS conjE;
+val Unifier_Comb = store_thm("Unifier_Comb", result() RS mp RS conjE);
goal Unifier.thy
"~v : vars_of(t) --> ~v : vars_of(u) -->Unifier(s,t,u) --> \
\ Unifier(<v,r>#s,t,u)";
by (simp_tac (subst_ss addsimps [Unifier_iff,repl_invariance]) 1);
-val Cons_Unifier = result() RS mp RS mp RS mp;
+val Cons_Unifier = store_thm("Cons_Unifier", result() RS mp RS mp RS mp);
(**** Most General Unifiers ****)
goalw Unifier.thy [MoreGeneral_def] "r >> s = (EX q. s =s= r <> q)";
by (rtac refl 1);
-val MoreGen_iff = result();
+qed "MoreGen_iff";
goal Unifier.thy "Nil >> s";
by (simp_tac (subst_ss addsimps [MoreGen_iff]) 1);
by (fast_tac (set_cs addIs [refl RS subst_refl]) 1);
-val MoreGen_Nil = result();
+qed "MoreGen_Nil";
goalw Unifier.thy unify_defs
"MGUnifier(s,t,u) = (ALL r.Unifier(r,t,u) = s >> r)";
@@ -46,7 +46,7 @@
eresolve_tac [conjE,allE,mp,exE,ssubst_subst2] 1));
by (asm_simp_tac (subst_ss addsimps [subst_comp]) 1);
by (fast_tac (set_cs addIs [comp_Nil RS sym RS subst_refl]) 1);
-val MGU_iff = result();
+qed "MGU_iff";
val [prem] = goal Unifier.thy
"~ Var(v) <: t ==> MGUnifier(<v,t>#Nil,Var(v),t)";
@@ -56,52 +56,52 @@
by (etac ssubst_subst2 2);
by (rtac (Cons_trivial RS subst_sym) 1);
by (simp_tac (subst_ss addsimps [prem RS Var_not_occs,Var_subst]) 1);
-val MGUnifier_Var = result();
+qed "MGUnifier_Var";
(**** Most General Idempotent Unifiers ****)
goal Unifier.thy "r <> r =s= r --> s =s= r <> q --> r <> s =s= s";
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1);
-val MGIU_iff_lemma = result() RS mp RS mp;
+val MGIU_iff_lemma = store_thm("MGIU_iff_lemma", result() RS mp RS mp);
goalw Unifier.thy unify_defs
"MGIUnifier(s,t,u) = \
\ (Idem(s) & Unifier(s,t,u) & (ALL r.Unifier(r,t,u) --> s<>r=s=r))";
by (fast_tac (set_cs addEs [subst_sym,MGIU_iff_lemma]) 1);
-val MGIU_iff = result();
+qed "MGIU_iff";
(**** Idempotence ****)
goalw Unifier.thy unify_defs "Idem(s) = (s <> s =s= s)";
by (rtac refl 1);
-val raw_Idem_iff = result();
+qed "raw_Idem_iff";
goal Unifier.thy "Idem(s) = (sdom(s) Int srange(s) = {})";
by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp,
invariance,dom_range_disjoint])1);
-val Idem_iff = result();
+qed "Idem_iff";
goal Unifier.thy "Idem(Nil)";
by (simp_tac (subst_ss addsimps [raw_Idem_iff,refl RS subst_refl]) 1);
-val Idem_Nil = result();
+qed "Idem_Nil";
goal Unifier.thy "~ (Var(v) <: t) --> Idem(<v,t>#Nil)";
by (simp_tac (subst_ss addsimps [Var_subst,vars_iff_occseq,Idem_iff,srange_iff]
setloop (split_tac [expand_if])) 1);
by (fast_tac set_cs 1);
-val Var_Idem = result() RS mp;
+val Var_Idem = store_thm("Var_Idem", result() RS mp);
val [prem] = goalw Unifier.thy [Idem_def]
"Idem(r) ==> Unifier(s,t <| r,u <| r) --> Unifier(r <> s,t <| r,u <| r)";
by (simp_tac (subst_ss addsimps
[Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1);
-val Unifier_Idem_subst = result() RS mp;
+val Unifier_Idem_subst = store_thm("Unifier_Idem_subst", result() RS mp);
val [prem] = goal Unifier.thy
"r <> s =s= s ==> Unifier(s,t,u) --> Unifier(s,t <| r,u <| r)";
by (simp_tac (subst_ss addsimps
[Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1);
-val Unifier_comp_subst = result() RS mp;
+val Unifier_comp_subst = store_thm("Unifier_comp_subst", result() RS mp);
(*** The domain of a MGIU is a subset of the variables in the terms ***)
(*** NB this and one for range are only needed for termination ***)
@@ -110,7 +110,7 @@
"~ vars_of(Var(x) <| r) = vars_of(Var(x) <| s) ==> ~r =s= s";
by (rtac (prem RS contrapos) 1);
by (fast_tac (set_cs addEs [subst_subst2]) 1);
-val lemma_lemma = result();
+qed "lemma_lemma";
val prems = goal Unifier.thy
"x : sdom(s) --> ~x : srange(s) --> \
@@ -122,7 +122,7 @@
br conjI 1;
by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1);
by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1);
-val MGIU_sdom_lemma = result() RS mp RS mp RS lemma_lemma RS notE;;
+val MGIU_sdom_lemma = store_thm("MGIU_sdom_lemma", result() RS mp RS mp RS lemma_lemma RS notE);
goal Unifier.thy "MGIUnifier(s,t,u) --> sdom(s) <= vars_of(t) Un vars_of(u)";
by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1);
@@ -131,13 +131,13 @@
by (eresolve_tac ([spec] RL [impE]) 1);
by (rtac Cons_Unifier 1);
by (ALLGOALS (fast_tac (set_cs addIs [Cons_Unifier,MGIU_sdom_lemma])));
-val MGIU_sdom = result() RS mp;
+val MGIU_sdom = store_thm("MGIU_sdom", result() RS mp);
(*** The range of a MGIU is a subset of the variables in the terms ***)
val prems = goal HOL.thy "P = Q ==> (~P) = (~Q)";
by (simp_tac (set_ss addsimps prems) 1);
-val not_cong = result();
+qed "not_cong";
val prems = goal Unifier.thy
"~w=x --> x : vars_of(Var(w) <| s) --> w : sdom(s) --> ~w : srange(s) --> \
@@ -149,7 +149,7 @@
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,
vars_var_iff RS not_cong RS iffD2 RS repl_invariance]) ));
by (fast_tac (set_cs addIs [Var_in_subst]) 1);
-val MGIU_srange_lemma =result() RS mp RS mp RS mp RS mp RS lemma_lemma RS notE;
+val MGIU_srange_lemma = store_thm("MGIU_srange_lemma", result() RS mp RS mp RS mp RS mp RS lemma_lemma RS notE);
goal Unifier.thy "MGIUnifier(s,t,u) --> srange(s) <= vars_of(t) Un vars_of(u)";
by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1);
@@ -161,7 +161,7 @@
by (imp_excluded_middle_tac "w=ta" 4);
by (fast_tac (set_cs addEs [MGIU_srange_lemma]) 5);
by (ALLGOALS (fast_tac (set_cs addIs [Var_elim2])));
-val MGIU_srange = result() RS mp;
+val MGIU_srange = store_thm("MGIU_srange", result() RS mp);
(*************** Correctness of a simple unification algorithm ***************)
(* *)
@@ -179,41 +179,41 @@
goalw Unifier.thy unify_defs "MGIUnifier(s,t,u) = MGIUnifier(s,u,t)";
by (safe_tac (HOL_cs addSEs ([sym]@([spec] RL [mp]))));
-val Unify_comm = result();
+qed "Unify_comm";
goal Unifier.thy "MGIUnifier(Nil,Const(n),Const(n))";
by (simp_tac (subst_ss addsimps
[MGIU_iff,MGU_iff,Unifier_iff,subst_eq_iff,Idem_Nil]) 1);
-val Unify1 = result();
+qed "Unify1";
goal Unifier.thy "~m=n --> (ALL l.~Unifier(l,Const(m),Const(n)))";
by (simp_tac (subst_ss addsimps[Unifier_iff]) 1);
-val Unify2 = result() RS mp;
+val Unify2 = store_thm("Unify2", result() RS mp);
val [prem] = goalw Unifier.thy [MGIUnifier_def]
"~Var(v) <: t ==> MGIUnifier(<v,t>#Nil,Var(v),t)";
by (fast_tac (HOL_cs addSIs [prem RS MGUnifier_Var,prem RS Var_Idem]) 1);
-val Unify3 = result();
+qed "Unify3";
val [prem] = goal Unifier.thy "Var(v) <: t ==> (ALL l.~Unifier(l,Var(v),t))";
by (simp_tac (subst_ss addsimps
[Unifier_iff,prem RS subst_mono RS occs_irrefl2]) 1);
-val Unify4 = result();
+qed "Unify4";
goal Unifier.thy "ALL l.~Unifier(l,Const(m),Comb(t,u))";
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1);
-val Unify5 = result();
+qed "Unify5";
goal Unifier.thy
"(ALL l.~Unifier(l,t,v)) --> (ALL l.~Unifier(l,Comb(t,u),Comb(v,w)))";
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1);
-val Unify6 = result() RS mp;
+val Unify6 = store_thm("Unify6", result() RS mp);
goal Unifier.thy "MGIUnifier(s,t,v) --> (ALL l.~Unifier(l,u <| s,w <| s)) --> \
\ (ALL l.~Unifier(l,Comb(t,u),Comb(v,w)))";
by (simp_tac (subst_ss addsimps [MGIU_iff]) 1);
by (fast_tac (set_cs addIs [Unifier_comp_subst] addSEs [Unifier_Comb]) 1);
-val Unify7 = result() RS mp RS mp;
+val Unify7 = store_thm("Unify7", result() RS mp RS mp);
val [p1,p2,p3] = goal Unifier.thy
"[| Idem(r); Unifier(s,t <| r,u <| r); \
@@ -223,7 +223,7 @@
p2 RS (p1 RS Unifier_Idem_subst RS (p3 RS spec RS mp))] 1);
by (REPEAT_SOME (etac rev_mp));
by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp]) 1);
-val Unify8_lemma1 = result();
+qed "Unify8_lemma1";
val [p1,p2,p3,p4] = goal Unifier.thy
"[| Unifier(q,t,v); Unifier(q,u,w); (! q.Unifier(q,t,v) --> r <> q =s= q); \
@@ -234,7 +234,7 @@
p2 RS (pp RS Unifier_comp_subst) RS (p4 RS spec RS mp)] 1);
by (REPEAT_SOME (etac rev_mp));
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1);
-val Unify8_lemma2 = result();
+qed "Unify8_lemma2";
goal Unifier.thy "MGIUnifier(r,t,v) --> MGIUnifier(s,u <| r,w <| r) --> \
\ MGIUnifier(r <> s,Comb(t,u),Comb(v,w))";
@@ -245,7 +245,7 @@
[Unifier_iff,MGIU_iff,subst_comp,comp_assoc]) 2);
by (ALLGOALS (fast_tac (set_cs addEs
[Unifier_Comb,Unify8_lemma1,Unify8_lemma2])));
-val Unify8 = result();
+qed "Unify8";
(********************** Termination of the algorithm *************************)
@@ -257,7 +257,7 @@
goalw Unifier.thy [UWFD_def] "UWFD(t,t',Comb(t,u),Comb(t',u'))";
by (simp_tac subst_ss 1);
by (fast_tac set_cs 1);
-val UnifyWFD1 = result();
+qed "UnifyWFD1";
val [prem] = goal Unifier.thy
"MGIUnifier(s,t,t') ==> vars_of(u <| s) Un vars_of(u' <| s) <= \
@@ -268,7 +268,7 @@
by (ALLGOALS (simp_tac (subst_ss addsimps [Var_intro,subset_iff])));
by (ALLGOALS (fast_tac (set_cs addDs
[Var_intro,prem RS MGIU_srange RS subsetD])));
-val UWFD2_lemma1 = result();
+qed "UWFD2_lemma1";
val [major,minor] = goal Unifier.thy
"[| MGIUnifier(s,t,t'); ~ u <| s = u |] ==> \
@@ -282,7 +282,7 @@
by (REPEAT (etac rev_mp 1));
by (asm_simp_tac subst_ss 1);
by (fast_tac (set_cs addIs [Var_elim2]) 1);
-val UWFD2_lemma2 = result();
+qed "UWFD2_lemma2";
val [prem] = goalw Unifier.thy [UWFD_def]
"MGIUnifier(s,t,t') ==> UWFD(u <| s,u' <| s,Comb(t,u),Comb(t',u'))";
@@ -296,4 +296,4 @@
by (asm_simp_tac (set_ss addsimps [subseteq_iff_subset_eq]) 1);
by (asm_simp_tac subst_ss 1);
by (fast_tac (set_cs addDs [prem RS UWFD2_lemma2]) 1);
-val UnifyWFD2 = result();
+qed "UnifyWFD2";
--- a/Sum.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Sum.ML Mon Nov 21 17:50:34 1994 +0100
@@ -13,16 +13,16 @@
(*This counts as a non-emptiness result for admitting 'a+'b as a type*)
goalw Sum.thy [Sum_def] "Inl_Rep(a) : Sum";
by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]);
-val Inl_RepI = result();
+qed "Inl_RepI";
goalw Sum.thy [Sum_def] "Inr_Rep(b) : Sum";
by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]);
-val Inr_RepI = result();
+qed "Inr_RepI";
goal Sum.thy "inj_onto(Abs_Sum,Sum)";
by (rtac inj_onto_inverseI 1);
by (etac Abs_Sum_inverse 1);
-val inj_onto_Abs_Sum = result();
+qed "inj_onto_Abs_Sum";
(** Distinctness of Inl and Inr **)
@@ -31,25 +31,25 @@
etac (fun_cong RS fun_cong RS fun_cong RS iffE),
rtac (notE RS ccontr), etac (mp RS conjunct2),
REPEAT o (ares_tac [refl,conjI]) ]);
-val Inl_Rep_not_Inr_Rep = result();
+qed "Inl_Rep_not_Inr_Rep";
goalw Sum.thy [Inl_def,Inr_def] "Inl(a) ~= Inr(b)";
by (rtac (inj_onto_Abs_Sum RS inj_onto_contraD) 1);
by (rtac Inl_Rep_not_Inr_Rep 1);
by (rtac Inl_RepI 1);
by (rtac Inr_RepI 1);
-val Inl_not_Inr = result();
+qed "Inl_not_Inr";
val Inl_neq_Inr = standard (Inl_not_Inr RS notE);
val Inr_neq_Inl = sym RS Inl_neq_Inr;
goal Sum.thy "(Inl(a)=Inr(b)) = False";
by (simp_tac (HOL_ss addsimps [Inl_not_Inr]) 1);
-val Inl_Inr_eq = result();
+qed "Inl_Inr_eq";
goal Sum.thy "(Inr(b)=Inl(a)) = False";
by (simp_tac (HOL_ss addsimps [Inl_not_Inr RS not_sym]) 1);
-val Inr_Inl_eq = result();
+qed "Inr_Inl_eq";
(** Injectiveness of Inl and Inr **)
@@ -57,19 +57,19 @@
val [major] = goalw Sum.thy [Inl_Rep_def] "Inl_Rep(a) = Inl_Rep(c) ==> a=c";
by (rtac (major RS fun_cong RS fun_cong RS fun_cong RS iffE) 1);
by (fast_tac HOL_cs 1);
-val Inl_Rep_inject = result();
+qed "Inl_Rep_inject";
val [major] = goalw Sum.thy [Inr_Rep_def] "Inr_Rep(b) = Inr_Rep(d) ==> b=d";
by (rtac (major RS fun_cong RS fun_cong RS fun_cong RS iffE) 1);
by (fast_tac HOL_cs 1);
-val Inr_Rep_inject = result();
+qed "Inr_Rep_inject";
goalw Sum.thy [Inl_def] "inj(Inl)";
by (rtac injI 1);
by (etac (inj_onto_Abs_Sum RS inj_ontoD RS Inl_Rep_inject) 1);
by (rtac Inl_RepI 1);
by (rtac Inl_RepI 1);
-val inj_Inl = result();
+qed "inj_Inl";
val Inl_inject = inj_Inl RS injD;
goalw Sum.thy [Inr_def] "inj(Inr)";
@@ -77,16 +77,16 @@
by (etac (inj_onto_Abs_Sum RS inj_ontoD RS Inr_Rep_inject) 1);
by (rtac Inr_RepI 1);
by (rtac Inr_RepI 1);
-val inj_Inr = result();
+qed "inj_Inr";
val Inr_inject = inj_Inr RS injD;
goal Sum.thy "(Inl(x)=Inl(y)) = (x=y)";
by (fast_tac (HOL_cs addSEs [Inl_inject]) 1);
-val Inl_eq = result();
+qed "Inl_eq";
goal Sum.thy "(Inr(x)=Inr(y)) = (x=y)";
by (fast_tac (HOL_cs addSEs [Inr_inject]) 1);
-val Inr_eq = result();
+qed "Inr_eq";
(*** Rules for the disjoint sum of two SETS ***)
@@ -94,11 +94,11 @@
goalw Sum.thy [sum_def] "!!a A B. a : A ==> Inl(a) : A plus B";
by (REPEAT (ares_tac [UnI1,imageI] 1));
-val InlI = result();
+qed "InlI";
goalw Sum.thy [sum_def] "!!b A B. b : B ==> Inr(b) : A plus B";
by (REPEAT (ares_tac [UnI2,imageI] 1));
-val InrI = result();
+qed "InrI";
(** Elimination rules **)
@@ -110,7 +110,7 @@
by (rtac (major RS UnE) 1);
by (REPEAT (rtac refl 1
ORELSE eresolve_tac (prems@[imageE,ssubst]) 1));
-val plusE = result();
+qed "plusE";
val sum_cs = set_cs addSIs [InlI, InrI]
@@ -122,11 +122,11 @@
goalw Sum.thy [sum_case_def] "sum_case(f, g, Inl(x)) = f(x)";
by (fast_tac (sum_cs addIs [select_equality]) 1);
-val sum_case_Inl = result();
+qed "sum_case_Inl";
goalw Sum.thy [sum_case_def] "sum_case(f, g, Inr(x)) = g(x)";
by (fast_tac (sum_cs addIs [select_equality]) 1);
-val sum_case_Inr = result();
+qed "sum_case_Inr";
(** Exhaustion rule for sums -- a degenerate form of induction **)
@@ -138,13 +138,13 @@
ORELSE EVERY1 [resolve_tac prems,
etac subst,
rtac (Rep_Sum_inverse RS sym)]));
-val sumE = result();
+qed "sumE";
goal Sum.thy "sum_case(%x::'a. f(Inl(x)), %y::'b. f(Inr(y)), s) = f(s)";
by (EVERY1 [res_inst_tac [("s","s")] sumE,
etac ssubst, rtac sum_case_Inl,
etac ssubst, rtac sum_case_Inr]);
-val surjective_sum = result();
+qed "surjective_sum";
goal Sum.thy "R(sum_case(f,g,s)) = \
\ ((! x. s = Inl(x) --> R(f(x))) & (! y. s = Inr(y) --> R(g(y))))";
@@ -155,7 +155,7 @@
by (etac ssubst 1);
by (stac sum_case_Inr 1);
by (fast_tac (set_cs addSEs [make_elim Inr_inject, Inr_neq_Inl]) 1);
-val expand_sum_case = result();
+qed "expand_sum_case";
val sum_ss = prod_ss addsimps [Inl_eq, Inr_eq, Inl_Inr_eq, Inr_Inl_eq,
sum_case_Inl, sum_case_Inr];
@@ -173,7 +173,7 @@
goalw Sum.thy [Part_def]
"!!a b A h. [| a : A; a=h(b) |] ==> a : Part(A,h)";
by (fast_tac set_cs 1);
-val Part_eqI = result();
+qed "Part_eqI";
val PartI = refl RSN (2,Part_eqI);
@@ -184,21 +184,21 @@
by (etac 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 Int_lower1 1);
-val Part_subset = result();
+qed "Part_subset";
goal Sum.thy "!!A B. A<=B ==> Part(A,h) <= Part(B,h)";
by (fast_tac (set_cs addSIs [PartI] addSEs [PartE]) 1);
-val Part_mono = result();
+qed "Part_mono";
goalw Sum.thy [Part_def] "!!a. a : Part(A,h) ==> a : A";
by (etac IntD1 1);
-val PartD1 = result();
+qed "PartD1";
goal Sum.thy "Part(A,%x.x) = A";
by (fast_tac (set_cs addIs [PartI,equalityI] addSEs [PartE]) 1);
-val Part_id = result();
+qed "Part_id";
--- a/Trancl.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Trancl.ML Mon Nov 21 17:50:34 1994 +0100
@@ -13,13 +13,13 @@
val prems = goalw Trancl.thy [trans_def]
"(!! x y z. [| <x,y>:r; <y,z>:r |] ==> <x,z>:r) ==> trans(r)";
by (REPEAT (ares_tac (prems@[allI,impI]) 1));
-val transI = result();
+qed "transI";
val major::prems = goalw Trancl.thy [trans_def]
"[| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r";
by (cut_facts_tac [major] 1);
by (fast_tac (HOL_cs addIs prems) 1);
-val transD = result();
+qed "transD";
(** Identity relation **)
@@ -27,7 +27,7 @@
by (rtac CollectI 1);
by (rtac exI 1);
by (rtac refl 1);
-val idI = result();
+qed "idI";
val major::prems = goalw Trancl.thy [id_def]
"[| p: id; !!x.[| p = <x,x> |] ==> P \
@@ -35,14 +35,14 @@
by (rtac (major RS CollectE) 1);
by (etac exE 1);
by (eresolve_tac prems 1);
-val idE = result();
+qed "idE";
(** Composition of two relations **)
val prems = goalw Trancl.thy [comp_def]
"[| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
by (fast_tac (set_cs addIs prems) 1);
-val compI = result();
+qed "compI";
(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*)
val prems = goalw Trancl.thy [comp_def]
@@ -51,7 +51,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 prems = goal Trancl.thy
"[| <a,c> : r O s; \
@@ -59,19 +59,19 @@
\ |] ==> P";
by (rtac compE 1);
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Pair_inject,ssubst] 1));
-val compEpair = result();
+qed "compEpair";
val comp_cs = prod_cs addIs [compI, idI] addSEs [compE, idE];
goal Trancl.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";
goal Trancl.thy
"!!r s. [| s <= Sigma(A,%x.B); r <= Sigma(B,%x.C) |] ==> \
\ (r O s) <= Sigma(A,%x.C)";
by (fast_tac comp_cs 1);
-val comp_subset_Sigma = result();
+qed "comp_subset_Sigma";
(** The relation rtrancl **)
@@ -79,7 +79,7 @@
goal Trancl.thy "mono(%s. id Un (r O s))";
by (rtac monoI 1);
by (REPEAT (ares_tac [monoI, subset_refl, comp_mono, Un_mono] 1));
-val rtrancl_fun_mono = result();
+qed "rtrancl_fun_mono";
val rtrancl_unfold = rtrancl_fun_mono RS (rtrancl_def RS def_lfp_Tarski);
@@ -87,25 +87,25 @@
goal Trancl.thy "<a,a> : r^*";
by (stac rtrancl_unfold 1);
by (fast_tac comp_cs 1);
-val rtrancl_refl = result();
+qed "rtrancl_refl";
(*Closure under composition with r*)
val prems = goal Trancl.thy
"[| <a,b> : r^*; <b,c> : r |] ==> <a,c> : r^*";
by (stac rtrancl_unfold 1);
by (fast_tac (comp_cs addIs prems) 1);
-val rtrancl_into_rtrancl = result();
+qed "rtrancl_into_rtrancl";
(*rtrancl of r contains r*)
val [prem] = goal Trancl.thy "[| <a,b> : r |] ==> <a,b> : r^*";
by (rtac (rtrancl_refl RS rtrancl_into_rtrancl) 1);
by (rtac prem 1);
-val r_into_rtrancl = result();
+qed "r_into_rtrancl";
(*monotonicity of rtrancl*)
goalw Trancl.thy [rtrancl_def] "!!r s. r <= s ==> r^* <= s^*";
by(REPEAT(ares_tac [lfp_mono,Un_mono,comp_mono,subset_refl] 1));
-val rtrancl_mono = result();
+qed "rtrancl_mono";
(** standard induction rule **)
@@ -116,7 +116,7 @@
\ ==> P(<a,b>)";
by (rtac ([rtrancl_def, rtrancl_fun_mono, major] MRS def_induct) 1);
by (fast_tac (comp_cs addIs prems) 1);
-val rtrancl_full_induct = result();
+qed "rtrancl_full_induct";
(*nice induction rule*)
val major::prems = goal Trancl.thy
@@ -132,14 +132,14 @@
by (resolve_tac [major RS rtrancl_full_induct] 1);
by (fast_tac (comp_cs addIs prems) 1);
by (fast_tac (comp_cs addIs prems) 1);
-val rtrancl_induct = result();
+qed "rtrancl_induct";
(*transitivity of transitive closure!! -- by induction.*)
goal Trancl.thy "trans(r^*)";
by (rtac transI 1);
by (res_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
@@ -151,7 +151,7 @@
by (fast_tac (set_cs addIs prems) 2);
by (fast_tac (set_cs addIs prems) 2);
by (REPEAT (eresolve_tac ([asm_rl,exE,disjE,conjE]@prems) 1));
-val rtranclE = result();
+qed "rtranclE";
(**** The relation trancl ****)
@@ -162,19 +162,19 @@
"<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 r*)
val [prem] = goalw Trancl.thy [trancl_def]
"[| <a,b> : r |] ==> <a,b> : r^+";
by (REPEAT (ares_tac [prem,compI,rtrancl_refl] 1));
-val r_into_trancl = result();
+qed "r_into_trancl";
(*intro rule by definition: from rtrancl 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 rtrancl*)
val prems = goal Trancl.thy
@@ -184,7 +184,7 @@
by (resolve_tac (prems RL [r_into_trancl]) 1);
by (rtac (trans_rtrancl RS transD RS rtrancl_into_trancl1) 1);
by (REPEAT (ares_tac (prems@[r_into_rtrancl]) 1));
-val rtrancl_into_trancl2 = result();
+qed "rtrancl_into_trancl2";
(*elimination of r^+ -- NOT an induction rule*)
val major::prems = goal Trancl.thy
@@ -198,7 +198,7 @@
by (etac rtranclE 1);
by (fast_tac comp_cs 1);
by (fast_tac (comp_cs addSIs [rtrancl_into_trancl1]) 1);
-val tranclE = result();
+qed "tranclE";
(*Transitivity of r^+.
Proved by unfolding since it uses transitivity of rtrancl. *)
@@ -207,14 +207,14 @@
by (REPEAT (etac compEpair 1));
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";
val prems = goal Trancl.thy
"[| <a,b> : r; <b,c> : r^+ |] ==> <a,c> : r^+";
by (rtac (r_into_trancl RS (trans_trancl RS transD)) 1);
by (resolve_tac prems 1);
by (resolve_tac prems 1);
-val trancl_into_trancl2 = result();
+qed "trancl_into_trancl2";
val major::prems = goal Trancl.thy
@@ -223,10 +223,10 @@
by (rtac (major RS rtrancl_induct) 1);
by (rtac (refl RS disjI1) 1);
by (fast_tac (comp_cs addSEs [SigmaE2]) 1);
-val trancl_subset_Sigma_lemma = result();
+qed "trancl_subset_Sigma_lemma";
goalw Trancl.thy [trancl_def]
"!!r. r <= Sigma(A,%x.A) ==> trancl(r) <= Sigma(A,%x.A)";
by (fast_tac (comp_cs addSDs [trancl_subset_Sigma_lemma]) 1);
-val trancl_subset_Sigma = result();
+qed "trancl_subset_Sigma";
--- a/Univ.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/Univ.ML Mon Nov 21 17:50:34 1994 +0100
@@ -17,7 +17,7 @@
by (fast_tac (HOL_cs addSIs [prem1,prem2]) 1);
by (cut_facts_tac [less_linear] 1);
by (fast_tac (HOL_cs addSIs [prem1] addSDs [prem2]) 1);
-val Least_equality = result();
+qed "Least_equality";
val [prem] = goal Univ.thy "P(k) ==> P(LEAST x.P(x))";
by (rtac (prem RS rev_mp) 1);
@@ -28,7 +28,7 @@
by (assume_tac 1);
by (assume_tac 2);
by (fast_tac HOL_cs 1);
-val LeastI = result();
+qed "LeastI";
(*Proof is almost identical to the one above!*)
val [prem] = goal Univ.thy "P(k) ==> (LEAST x.P(x)) <= k";
@@ -40,20 +40,20 @@
by (assume_tac 1);
by (rtac le_refl 2);
by (fast_tac (HOL_cs addIs [less_imp_le,le_trans]) 1);
-val Least_le = result();
+qed "Least_le";
val [prem] = goal Univ.thy "k < (LEAST x.P(x)) ==> ~P(k)";
by (rtac notI 1);
by (etac (rewrite_rule [le_def] Least_le RS notE) 1);
by (rtac prem 1);
-val not_less_Least = result();
+qed "not_less_Least";
(** apfst -- can be used in similar type definitions **)
goalw Univ.thy [apfst_def] "apfst(f,<a,b>) = <f(a),b>";
by (rtac split 1);
-val apfst = result();
+qed "apfst";
val [major,minor] = goal Univ.thy
"[| q = apfst(f,p); !!x y. [| p = <x,y>; q = <f(x),y> |] ==> R \
@@ -64,7 +64,7 @@
by (rtac (major RS trans) 1);
by (etac ssubst 1);
by (rtac apfst 1);
-val apfstE = result();
+qed "apfstE";
(** Push -- an injection, analogous to Cons on lists **)
@@ -72,37 +72,37 @@
by (rtac (major RS fun_cong RS box_equals RS Suc_inject) 1);
by (rtac nat_case_0 1);
by (rtac nat_case_0 1);
-val Push_inject1 = result();
+qed "Push_inject1";
val [major] = goalw Univ.thy [Push_def] "Push(i,f)=Push(j,g) ==> f=g";
by (rtac (major RS fun_cong RS ext RS box_equals) 1);
by (rtac (nat_case_Suc RS ext) 1);
by (rtac (nat_case_Suc RS ext) 1);
-val Push_inject2 = result();
+qed "Push_inject2";
val [major,minor] = goal Univ.thy
"[| Push(i,f)=Push(j,g); [| i=j; f=g |] ==> P \
\ |] ==> P";
by (rtac ((major RS Push_inject2) RS ((major RS Push_inject1) RS minor)) 1);
-val Push_inject = result();
+qed "Push_inject";
val [major] = goalw Univ.thy [Push_def] "Push(k,f)=(%z.0) ==> P";
by (rtac (major RS fun_cong RS box_equals RS Suc_neq_Zero) 1);
by (rtac nat_case_0 1);
by (rtac refl 1);
-val Push_neq_K0 = result();
+qed "Push_neq_K0";
(*** Isomorphisms ***)
goal Univ.thy "inj(Rep_Node)";
by (rtac inj_inverseI 1); (*cannot combine by RS: multiple unifiers*)
by (rtac Rep_Node_inverse 1);
-val inj_Rep_Node = result();
+qed "inj_Rep_Node";
goal Univ.thy "inj_onto(Abs_Node,Node)";
by (rtac inj_onto_inverseI 1);
by (etac Abs_Node_inverse 1);
-val inj_onto_Abs_Node = result();
+qed "inj_onto_Abs_Node";
val Abs_Node_inject = inj_onto_Abs_Node RS inj_ontoD;
@@ -111,12 +111,12 @@
goalw Univ.thy [Node_def] "<%k. 0,a> : Node";
by (fast_tac set_cs 1);
-val Node_K0_I = result();
+qed "Node_K0_I";
goalw Univ.thy [Node_def,Push_def]
"!!p. p: Node ==> apfst(Push(i), p) : Node";
by (fast_tac (set_cs addSIs [apfst, nat_case_Suc RS trans]) 1);
-val Node_Push_I = result();
+qed "Node_Push_I";
(*** Distinctness of constructors ***)
@@ -130,7 +130,7 @@
by (REPEAT (eresolve_tac [imageE, Abs_Node_inject RS apfstE,
Pair_inject, sym RS Push_neq_K0] 1
ORELSE resolve_tac [Node_K0_I, Rep_Node RS Node_Push_I] 1));
-val Scons_not_Atom = result();
+qed "Scons_not_Atom";
val Atom_not_Scons = standard (Scons_not_Atom RS not_sym);
val Scons_neq_Atom = standard (Scons_not_Atom RS notE);
@@ -144,20 +144,20 @@
by (rtac injI 1);
by (etac (singleton_inject RS Abs_Node_inject RS Pair_inject) 1);
by (REPEAT (ares_tac [Node_K0_I] 1));
-val inj_Atom = result();
+qed "inj_Atom";
val Atom_inject = inj_Atom RS injD;
goalw Univ.thy [Leaf_def,o_def] "inj(Leaf)";
by (rtac injI 1);
by (etac (Atom_inject RS Inl_inject) 1);
-val inj_Leaf = result();
+qed "inj_Leaf";
val Leaf_inject = inj_Leaf RS injD;
goalw Univ.thy [Numb_def,o_def] "inj(Numb)";
by (rtac injI 1);
by (etac (Atom_inject RS Inr_inject) 1);
-val inj_Numb = result();
+qed "inj_Numb";
val Numb_inject = inj_Numb RS injD;
@@ -175,7 +175,7 @@
by (rtac (inj_Rep_Node RS injD) 1);
by (etac trans 1);
by (safe_tac (HOL_cs addSEs [Pair_inject,Push_inject,sym]));
-val Push_Node_inject = result();
+qed "Push_Node_inject";
(** Injectiveness of Scons **)
@@ -184,38 +184,38 @@
by (cut_facts_tac [major] 1);
by (fast_tac (set_cs addSDs [Suc_inject]
addSEs [Push_Node_inject, Zero_neq_Suc]) 1);
-val Scons_inject_lemma1 = result();
+qed "Scons_inject_lemma1";
val [major] = goalw Univ.thy [Scons_def] "M$N <= M'$N' ==> N<=N'";
by (cut_facts_tac [major] 1);
by (fast_tac (set_cs addSDs [Suc_inject]
addSEs [Push_Node_inject, Suc_neq_Zero]) 1);
-val Scons_inject_lemma2 = result();
+qed "Scons_inject_lemma2";
val [major] = goal Univ.thy "M$N = M'$N' ==> M=M'";
by (rtac (major RS equalityE) 1);
by (REPEAT (ares_tac [equalityI, Scons_inject_lemma1] 1));
-val Scons_inject1 = result();
+qed "Scons_inject1";
val [major] = goal Univ.thy "M$N = M'$N' ==> N=N'";
by (rtac (major RS equalityE) 1);
by (REPEAT (ares_tac [equalityI, Scons_inject_lemma2] 1));
-val Scons_inject2 = result();
+qed "Scons_inject2";
val [major,minor] = goal Univ.thy
"[| M$N = M'$N'; [| M=M'; N=N' |] ==> P \
\ |] ==> P";
by (rtac ((major RS Scons_inject2) RS ((major RS Scons_inject1) RS minor)) 1);
-val Scons_inject = result();
+qed "Scons_inject";
(*rewrite rules*)
goal Univ.thy "(Atom(a)=Atom(b)) = (a=b)";
by (fast_tac (HOL_cs addSEs [Atom_inject]) 1);
-val Atom_Atom_eq = result();
+qed "Atom_Atom_eq";
goal Univ.thy "(M$N = M'$N') = (M=M' & N=N')";
by (fast_tac (HOL_cs addSEs [Scons_inject]) 1);
-val Scons_Scons_eq = result();
+qed "Scons_Scons_eq";
(*** Distinctness involving Leaf and Numb ***)
@@ -223,7 +223,7 @@
goalw Univ.thy [Leaf_def,o_def] "(M$N) ~= Leaf(a)";
by (rtac Scons_not_Atom 1);
-val Scons_not_Leaf = result();
+qed "Scons_not_Leaf";
val Leaf_not_Scons = standard (Scons_not_Leaf RS not_sym);
val Scons_neq_Leaf = standard (Scons_not_Leaf RS notE);
@@ -233,7 +233,7 @@
goalw Univ.thy [Numb_def,o_def] "(M$N) ~= Numb(k)";
by (rtac Scons_not_Atom 1);
-val Scons_not_Numb = result();
+qed "Scons_not_Numb";
val Numb_not_Scons = standard (Scons_not_Numb RS not_sym);
val Scons_neq_Numb = standard (Scons_not_Numb RS notE);
@@ -243,7 +243,7 @@
goalw Univ.thy [Leaf_def,Numb_def] "Leaf(a) ~= Numb(k)";
by (simp_tac (HOL_ss addsimps [Atom_Atom_eq,Inl_not_Inr]) 1);
-val Leaf_not_Numb = result();
+qed "Leaf_not_Numb";
val Numb_not_Leaf = standard (Leaf_not_Numb RS not_sym);
val Leaf_neq_Numb = standard (Leaf_not_Numb RS notE);
@@ -261,14 +261,14 @@
by (rtac Least_equality 1);
by (rtac refl 1);
by (etac less_zeroE 1);
-val ndepth_K0 = result();
+qed "ndepth_K0";
goal Univ.thy "k < Suc(LEAST x. f(x)=0) --> nat_case(Suc(i), f, k) ~= 0";
by (nat_ind_tac "k" 1);
by (ALLGOALS (simp_tac nat_ss));
by (rtac impI 1);
by (etac not_less_Least 1);
-val ndepth_Push_lemma = result();
+qed "ndepth_Push_lemma";
goalw Univ.thy [ndepth_def,Push_Node_def]
"ndepth (Push_Node(i,n)) = Suc(ndepth(n))";
@@ -282,28 +282,28 @@
by (rtac (nat_case_Suc RS trans) 1);
by (etac LeastI 1);
by (etac (ndepth_Push_lemma RS mp) 1);
-val ndepth_Push_Node = result();
+qed "ndepth_Push_Node";
(*** ntrunc applied to the various node sets ***)
goalw Univ.thy [ntrunc_def] "ntrunc(0, M) = {}";
by (safe_tac (set_cs addSIs [equalityI] addSEs [less_zeroE]));
-val ntrunc_0 = result();
+qed "ntrunc_0";
goalw Univ.thy [Atom_def,ntrunc_def] "ntrunc(Suc(k), Atom(a)) = Atom(a)";
by (safe_tac (set_cs addSIs [equalityI]));
by (stac ndepth_K0 1);
by (rtac zero_less_Suc 1);
-val ntrunc_Atom = result();
+qed "ntrunc_Atom";
goalw Univ.thy [Leaf_def,o_def] "ntrunc(Suc(k), Leaf(a)) = Leaf(a)";
by (rtac ntrunc_Atom 1);
-val ntrunc_Leaf = result();
+qed "ntrunc_Leaf";
goalw Univ.thy [Numb_def,o_def] "ntrunc(Suc(k), Numb(i)) = Numb(i)";
by (rtac ntrunc_Atom 1);
-val ntrunc_Numb = result();
+qed "ntrunc_Numb";
goalw Univ.thy [Scons_def,ntrunc_def]
"ntrunc(Suc(k), M$N) = ntrunc(k,M) $ ntrunc(k,N)";
@@ -312,7 +312,7 @@
by (REPEAT (rtac Suc_less_SucD 1 THEN
rtac (ndepth_Push_Node RS subst) 1 THEN
assume_tac 1));
-val ntrunc_Scons = result();
+qed "ntrunc_Scons";
(** Injection nodes **)
@@ -320,30 +320,30 @@
by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_0]) 1);
by (rewtac Scons_def);
by (safe_tac (set_cs addSIs [equalityI]));
-val ntrunc_one_In0 = result();
+qed "ntrunc_one_In0";
goalw Univ.thy [In0_def]
"ntrunc(Suc(Suc(k)), In0(M)) = In0 (ntrunc(Suc(k),M))";
by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_Numb]) 1);
-val ntrunc_In0 = result();
+qed "ntrunc_In0";
goalw Univ.thy [In1_def] "ntrunc(Suc(0), In1(M)) = {}";
by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_0]) 1);
by (rewtac Scons_def);
by (safe_tac (set_cs addSIs [equalityI]));
-val ntrunc_one_In1 = result();
+qed "ntrunc_one_In1";
goalw Univ.thy [In1_def]
"ntrunc(Suc(Suc(k)), In1(M)) = In1 (ntrunc(Suc(k),M))";
by (simp_tac (univ_ss addsimps [ntrunc_Scons,ntrunc_Numb]) 1);
-val ntrunc_In1 = result();
+qed "ntrunc_In1";
(*** Cartesian Product ***)
goalw Univ.thy [uprod_def] "!!M N. [| M:A; N:B |] ==> (M$N) : A<*>B";
by (REPEAT (ares_tac [singletonI,UN_I] 1));
-val uprodI = result();
+qed "uprodI";
(*The general elimination rule*)
val major::prems = goalw Univ.thy [uprod_def]
@@ -353,7 +353,7 @@
by (cut_facts_tac [major] 1);
by (REPEAT (eresolve_tac [asm_rl,singletonE,UN_E] 1
ORELSE resolve_tac prems 1));
-val uprodE = result();
+qed "uprodE";
(*Elimination of a pair -- introduces no eigenvariables*)
val prems = goal Univ.thy
@@ -361,18 +361,18 @@
\ |] ==> P";
by (rtac uprodE 1);
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Scons_inject,ssubst] 1));
-val uprodE2 = result();
+qed "uprodE2";
(*** Disjoint Sum ***)
goalw Univ.thy [usum_def] "!!M. M:A ==> In0(M) : A<+>B";
by (fast_tac set_cs 1);
-val usum_In0I = result();
+qed "usum_In0I";
goalw Univ.thy [usum_def] "!!N. N:B ==> In1(N) : A<+>B";
by (fast_tac set_cs 1);
-val usum_In1I = result();
+qed "usum_In1I";
val major::prems = goalw Univ.thy [usum_def]
"[| u : A<+>B; \
@@ -382,7 +382,7 @@
by (rtac (major RS UnE) 1);
by (REPEAT (rtac refl 1
ORELSE eresolve_tac (prems@[imageE,ssubst]) 1));
-val usumE = result();
+qed "usumE";
(** Injection **)
@@ -390,7 +390,7 @@
goalw Univ.thy [In0_def,In1_def] "In0(M) ~= In1(N)";
by (rtac notI 1);
by (etac (Scons_inject1 RS Numb_inject RS Zero_neq_Suc) 1);
-val In0_not_In1 = result();
+qed "In0_not_In1";
val In1_not_In0 = standard (In0_not_In1 RS not_sym);
val In0_neq_In1 = standard (In0_not_In1 RS notE);
@@ -398,24 +398,24 @@
val [major] = goalw Univ.thy [In0_def] "In0(M) = In0(N) ==> M=N";
by (rtac (major RS Scons_inject2) 1);
-val In0_inject = result();
+qed "In0_inject";
val [major] = goalw Univ.thy [In1_def] "In1(M) = In1(N) ==> M=N";
by (rtac (major RS Scons_inject2) 1);
-val In1_inject = result();
+qed "In1_inject";
(*** proving equality of sets and functions using ntrunc ***)
goalw Univ.thy [ntrunc_def] "ntrunc(k,M) <= M";
by (fast_tac set_cs 1);
-val ntrunc_subsetI = result();
+qed "ntrunc_subsetI";
val [major] = goalw Univ.thy [ntrunc_def]
"(!!k. ntrunc(k,M) <= N) ==> M<=N";
by (fast_tac (set_cs addIs [less_add_Suc1, less_add_Suc2,
major RS subsetD]) 1);
-val ntrunc_subsetD = result();
+qed "ntrunc_subsetD";
(*A generalized form of the take-lemma*)
val [major] = goal Univ.thy "(!!k. ntrunc(k,M) = ntrunc(k,N)) ==> M=N";
@@ -424,81 +424,81 @@
by (ALLGOALS (rtac (ntrunc_subsetI RSN (2, subset_trans))));
by (rtac (major RS equalityD1) 1);
by (rtac (major RS equalityD2) 1);
-val ntrunc_equality = result();
+qed "ntrunc_equality";
val [major] = goalw Univ.thy [o_def]
"[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2";
by (rtac (ntrunc_equality RS ext) 1);
by (rtac (major RS fun_cong) 1);
-val ntrunc_o_equality = result();
+qed "ntrunc_o_equality";
(*** Monotonicity ***)
goalw Univ.thy [uprod_def] "!!A B. [| A<=A'; B<=B' |] ==> A<*>B <= A'<*>B'";
by (fast_tac set_cs 1);
-val uprod_mono = result();
+qed "uprod_mono";
goalw Univ.thy [usum_def] "!!A B. [| A<=A'; B<=B' |] ==> A<+>B <= A'<+>B'";
by (fast_tac set_cs 1);
-val usum_mono = result();
+qed "usum_mono";
goalw Univ.thy [Scons_def] "!!M N. [| M<=M'; N<=N' |] ==> M$N <= M'$N'";
by (fast_tac set_cs 1);
-val Scons_mono = result();
+qed "Scons_mono";
goalw Univ.thy [In0_def] "!!M N. M<=N ==> In0(M) <= In0(N)";
by (REPEAT (ares_tac [subset_refl,Scons_mono] 1));
-val In0_mono = result();
+qed "In0_mono";
goalw Univ.thy [In1_def] "!!M N. M<=N ==> In1(M) <= In1(N)";
by (REPEAT (ares_tac [subset_refl,Scons_mono] 1));
-val In1_mono = result();
+qed "In1_mono";
(*** Split and Case ***)
goalw Univ.thy [Split_def] "Split(c, M$N) = c(M,N)";
by (fast_tac (set_cs addIs [select_equality] addEs [Scons_inject]) 1);
-val Split = result();
+qed "Split";
goalw Univ.thy [Case_def] "Case(c, d, In0(M)) = c(M)";
by (fast_tac (set_cs addIs [select_equality]
addEs [make_elim In0_inject, In0_neq_In1]) 1);
-val Case_In0 = result();
+qed "Case_In0";
goalw Univ.thy [Case_def] "Case(c, d, In1(N)) = d(N)";
by (fast_tac (set_cs addIs [select_equality]
addEs [make_elim In1_inject, In1_neq_In0]) 1);
-val Case_In1 = result();
+qed "Case_In1";
(**** UN x. B(x) rules ****)
goalw Univ.thy [ntrunc_def] "ntrunc(k, UN x.f(x)) = (UN x. ntrunc(k, f(x)))";
by (fast_tac (set_cs addIs [equalityI]) 1);
-val ntrunc_UN1 = result();
+qed "ntrunc_UN1";
goalw Univ.thy [Scons_def] "(UN x.f(x)) $ M = (UN x. f(x) $ M)";
by (fast_tac (set_cs addIs [equalityI]) 1);
-val Scons_UN1_x = result();
+qed "Scons_UN1_x";
goalw Univ.thy [Scons_def] "M $ (UN x.f(x)) = (UN x. M $ f(x))";
by (fast_tac (set_cs addIs [equalityI]) 1);
-val Scons_UN1_y = result();
+qed "Scons_UN1_y";
goalw Univ.thy [In0_def] "In0(UN x.f(x)) = (UN x. In0(f(x)))";
br Scons_UN1_y 1;
-val In0_UN1 = result();
+qed "In0_UN1";
goalw Univ.thy [In1_def] "In1(UN x.f(x)) = (UN x. In1(f(x)))";
br Scons_UN1_y 1;
-val In1_UN1 = result();
+qed "In1_UN1";
(*** Equality : the diagonal relation ***)
goalw Univ.thy [diag_def] "!!a A. [| a=b; a:A |] ==> <a,b> : diag(A)";
by (fast_tac set_cs 1);
-val diag_eqI = result();
+qed "diag_eqI";
val diagI = refl RS diag_eqI |> standard;
@@ -509,14 +509,14 @@
\ |] ==> P";
by (rtac (major RS UN_E) 1);
by (REPEAT (eresolve_tac [asm_rl,singletonE] 1 ORELSE resolve_tac prems 1));
-val diagE = result();
+qed "diagE";
(*** Equality for Cartesian Product ***)
goalw Univ.thy [dprod_def]
"!!r s. [| <M,M'>:r; <N,N'>:s |] ==> <M$N, M'$N'> : r<**>s";
by (fast_tac prod_cs 1);
-val dprodI = result();
+qed "dprodI";
(*The general elimination rule*)
val major::prems = goalw Univ.thy [dprod_def]
@@ -526,18 +526,18 @@
by (cut_facts_tac [major] 1);
by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, mem_splitE, singletonE]));
by (REPEAT (ares_tac prems 1 ORELSE hyp_subst_tac 1));
-val dprodE = result();
+qed "dprodE";
(*** Equality for Disjoint Sum ***)
goalw Univ.thy [dsum_def] "!!r. <M,M'>:r ==> <In0(M), In0(M')> : r<++>s";
by (fast_tac prod_cs 1);
-val dsum_In0I = result();
+qed "dsum_In0I";
goalw Univ.thy [dsum_def] "!!r. <N,N'>:s ==> <In1(N), In1(N')> : r<++>s";
by (fast_tac prod_cs 1);
-val dsum_In1I = result();
+qed "dsum_In1I";
val major::prems = goalw Univ.thy [dsum_def]
"[| w : r<++>s; \
@@ -547,7 +547,7 @@
by (cut_facts_tac [major] 1);
by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, UnE, mem_splitE, singletonE]));
by (DEPTH_SOLVE (ares_tac prems 1 ORELSE hyp_subst_tac 1));
-val dsumE = result();
+qed "dsumE";
val univ_cs =
@@ -560,22 +560,22 @@
goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<**>s <= r'<**>s'";
by (fast_tac univ_cs 1);
-val dprod_mono = result();
+qed "dprod_mono";
goal Univ.thy "!!r s. [| r<=r'; s<=s' |] ==> r<++>s <= r'<++>s'";
by (fast_tac univ_cs 1);
-val dsum_mono = result();
+qed "dsum_mono";
(*** Bounding theorems ***)
goal Univ.thy "diag(A) <= Sigma(A,%x.A)";
by (fast_tac univ_cs 1);
-val diag_subset_Sigma = result();
+qed "diag_subset_Sigma";
goal Univ.thy "(Sigma(A,%x.B) <**> Sigma(C,%x.D)) <= Sigma(A<*>C, %z. B<*>D)";
by (fast_tac univ_cs 1);
-val dprod_Sigma = result();
+qed "dprod_Sigma";
val dprod_subset_Sigma = [dprod_mono, dprod_Sigma] MRS subset_trans |>standard;
@@ -585,11 +585,11 @@
by (safe_tac univ_cs);
by (stac Split 1);
by (fast_tac univ_cs 1);
-val dprod_subset_Sigma2 = result();
+qed "dprod_subset_Sigma2";
goal Univ.thy "(Sigma(A,%x.B) <++> Sigma(C,%x.D)) <= Sigma(A<+>C, %z. B<+>D)";
by (fast_tac univ_cs 1);
-val dsum_Sigma = result();
+qed "dsum_Sigma";
val dsum_subset_Sigma = [dsum_mono, dsum_Sigma] MRS subset_trans |> standard;
@@ -598,18 +598,18 @@
goal Univ.thy "fst `` diag(A) = A";
by (fast_tac (prod_cs addIs [equalityI, diagI] addSEs [diagE]) 1);
-val fst_image_diag = result();
+qed "fst_image_diag";
goal Univ.thy "fst `` (r<**>s) = (fst``r) <*> (fst``s)";
by (fast_tac (prod_cs addIs [equalityI, uprodI, dprodI]
addSEs [uprodE, dprodE]) 1);
-val fst_image_dprod = result();
+qed "fst_image_dprod";
goal Univ.thy "fst `` (r<++>s) = (fst``r) <+> (fst``s)";
by (fast_tac (prod_cs addIs [equalityI, usum_In0I, usum_In1I,
dsum_In0I, dsum_In1I]
addSEs [usumE, dsumE]) 1);
-val fst_image_dsum = result();
+qed "fst_image_dsum";
val fst_image_simps = [fst_image_diag, fst_image_dprod, fst_image_dsum];
val fst_image_ss = univ_ss addsimps fst_image_simps;
--- a/WF.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/WF.ML Mon Nov 21 17:50:34 1994 +0100
@@ -21,7 +21,7 @@
by (rtac allE 1);
by (assume_tac 1);
by (best_tac (HOL_cs addSEs [prem1 RS subsetD RS SigmaE2] addIs [prem2]) 1);
-val wfI = result();
+qed "wfI";
val major::prems = goalw WF.thy [wf_def]
"[| wf(r); \
@@ -29,7 +29,7 @@
\ |] ==> P(a)";
by (rtac (major RS spec RS mp RS spec) 1);
by (fast_tac (HOL_cs addEs prems) 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 =
@@ -42,12 +42,12 @@
by (fast_tac (HOL_cs addIs prems) 1);
by (wf_ind_tac "a" prems 1);
by (fast_tac set_cs 1);
-val wf_asym = result();
+qed "wf_asym";
val prems = goal WF.thy "[| wf(r); <a,a>: r |] ==> P";
by (rtac wf_asym 1);
by (REPEAT (resolve_tac prems 1));
-val wf_anti_refl = result();
+qed "wf_anti_refl";
(*transitive closure of a WF relation is WF!*)
val [prem] = goal WF.thy "wf(r) ==> wf(r^+)";
@@ -61,7 +61,7 @@
by (etac tranclE 1);
by (fast_tac HOL_cs 1);
by (fast_tac HOL_cs 1);
-val wf_trancl = result();
+qed "wf_trancl";
(** cut **)
@@ -72,11 +72,11 @@
"(cut(f,r,x) = cut(g,r,x)) = (!y. <y,x>:r --> f(y)=g(y))";
by(simp_tac (HOL_ss addsimps [expand_fun_eq]
setloop (split_tac [expand_if])) 1);
-val cut_cut_eq = result();
+qed "cut_cut_eq";
goalw WF.thy [cut_def] "!!x. <x,a>:r ==> cut(f,r,a)(x) = f(x)";
by(asm_simp_tac HOL_ss 1);
-val cut_apply = result();
+qed "cut_apply";
(*** is_recfun ***)
@@ -85,7 +85,7 @@
"!!f. [| is_recfun(r,a,H,f); ~<b,a>:r |] ==> f(b) = (@z.True)";
by (etac ssubst 1);
by(asm_simp_tac HOL_ss 1);
-val is_recfun_undef = result();
+qed "is_recfun_undef";
(*eresolve_tac transD solves <a,b>:r using transitivity AT MOST ONCE
mp amd allE instantiate induction hypotheses*)
@@ -110,7 +110,7 @@
by (etac wf_induct 1);
by (REPEAT (rtac impI 1 ORELSE etac ssubst 1));
by (asm_simp_tac (wf_super_ss addcongs [if_cong]) 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);
@@ -124,7 +124,7 @@
by (rtac ext 1);
by (asm_simp_tac (wf_super_ss addsimps [gundef,fisg]
setloop (split_tac [expand_if])) 1);
-val is_recfun_cut = result();
+qed "is_recfun_cut";
(*** Main Existence Lemma -- Basic Properties of the_recfun ***)
@@ -132,7 +132,7 @@
"is_recfun(r,a,H,f) ==> is_recfun(r, a, H, the_recfun(r,a,H))";
by (res_inst_tac [("P", "is_recfun(r,a,H)")] selectI 1);
by (resolve_tac prems 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))";
@@ -146,7 +146,7 @@
etac (mp RS ssubst), atac]);
by (fold_tac [is_recfun_def]);
by (asm_simp_tac (wf_super_ss addsimps[cut_apply,is_recfun_cut,cut_cut_eq]) 1);
-val unfold_the_recfun = result();
+qed "unfold_the_recfun";
(*Beware incompleteness of unification!*)
@@ -154,13 +154,13 @@
"[| wf(r); trans(r); <c,a>:r; <c,b>:r |] \
\ ==> the_recfun(r,a,H,c) = the_recfun(r,b,H,c)";
by (DEPTH_SOLVE (ares_tac (prems@[is_recfun_equal,unfold_the_recfun]) 1));
-val the_recfun_equal = result();
+qed "the_recfun_equal";
val prems = goal WF.thy
"[| wf(r); trans(r); <b,a>:r |] \
\ ==> cut(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";
(*** Unfolding wftrec ***)
@@ -170,7 +170,7 @@
by (EVERY1 [stac (rewrite_rule [is_recfun_def] unfold_the_recfun),
REPEAT o atac, rtac H_cong1]);
by (asm_simp_tac (HOL_ss addsimps [cut_cut_eq,the_recfun_cut]) 1);
-val wftrec = result();
+qed "wftrec";
(*Unused but perhaps interesting*)
val prems = goal WF.thy
@@ -178,7 +178,7 @@
\ wftrec(r,a,H) = H(a, %x.wftrec(r,x,H))";
by (rtac (wftrec RS trans) 1);
by (REPEAT (resolve_tac prems 1));
-val wftrec2 = result();
+qed "wftrec2";
(** Removal of the premise trans(r) **)
@@ -188,11 +188,11 @@
by (rtac trans_trancl 1);
by (rtac (refl RS H_cong) 1); (*expose the equality of cuts*)
by (simp_tac (HOL_ss addsimps [cut_cut_eq, cut_apply, r_into_trancl]) 1);
-val wfrec = result();
+qed "wfrec";
(*This form avoids giant explosions in proofs. NOTE USE OF == *)
val rew::prems = goal WF.thy
"[| !!x. f(x)==wfrec(r,x,H); wf(r) |] ==> f(a) = H(a, cut(%x.f(x),r,a))";
by (rewtac rew);
by (REPEAT (resolve_tac (prems@[wfrec]) 1));
-val def_wfrec = result();
+qed "def_wfrec";
--- a/equalities.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/equalities.ML Mon Nov 21 17:50:34 1994 +0100
@@ -14,168 +14,168 @@
goal Set.thy "x ~: {}";
by(fast_tac set_cs 1);
-val in_empty = result();
+qed "in_empty";
goal Set.thy "x : insert(y,A) = (x=y | x:A)";
by(fast_tac set_cs 1);
-val in_insert = result();
+qed "in_insert";
(** insert **)
goal Set.thy "!!a. a:A ==> insert(a,A) = A";
by (fast_tac eq_cs 1);
-val insert_absorb = result();
+qed "insert_absorb";
goal Set.thy "(insert(x,A) <= B) = (x:B & A <= B)";
by (fast_tac set_cs 1);
-val insert_subset = result();
+qed "insert_subset";
(** Image **)
goal Set.thy "f``{} = {}";
by (fast_tac eq_cs 1);
-val image_empty = result();
+qed "image_empty";
goal Set.thy "f``insert(a,B) = insert(f(a), f``B)";
by (fast_tac eq_cs 1);
-val image_insert = result();
+qed "image_insert";
(** Binary Intersection **)
goal Set.thy "A Int A = A";
by (fast_tac eq_cs 1);
-val Int_absorb = result();
+qed "Int_absorb";
goal Set.thy "A Int B = B Int A";
by (fast_tac eq_cs 1);
-val Int_commute = result();
+qed "Int_commute";
goal Set.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 Set.thy "{} Int B = {}";
by (fast_tac eq_cs 1);
-val Int_empty_left = result();
+qed "Int_empty_left";
goal Set.thy "A Int {} = {}";
by (fast_tac eq_cs 1);
-val Int_empty_right = result();
+qed "Int_empty_right";
goal Set.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 Set.thy "(A<=B) = (A Int B = A)";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val subset_Int_eq = result();
+qed "subset_Int_eq";
(** Binary Union **)
goal Set.thy "A Un A = A";
by (fast_tac eq_cs 1);
-val Un_absorb = result();
+qed "Un_absorb";
goal Set.thy "A Un B = B Un A";
by (fast_tac eq_cs 1);
-val Un_commute = result();
+qed "Un_commute";
goal Set.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 Set.thy "{} Un B = B";
by(fast_tac eq_cs 1);
-val Un_empty_left = result();
+qed "Un_empty_left";
goal Set.thy "A Un {} = A";
by(fast_tac eq_cs 1);
-val Un_empty_right = result();
+qed "Un_empty_right";
goal Set.thy "insert(a,B) Un C = insert(a,B Un C)";
by(fast_tac eq_cs 1);
-val Un_insert_left = result();
+qed "Un_insert_left";
goal Set.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 Set.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 Set.thy "(A<=B) = (A Un B = B)";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val subset_Un_eq = result();
+qed "subset_Un_eq";
goal Set.thy "(A <= insert(b,C)) = (A <= C | b:A & A-{b} <= C)";
by (fast_tac eq_cs 1);
-val subset_insert_iff = result();
+qed "subset_insert_iff";
(** Simple properties of Compl -- complement of a set **)
goal Set.thy "A Int Compl(A) = {}";
by (fast_tac eq_cs 1);
-val Compl_disjoint = result();
+qed "Compl_disjoint";
goal Set.thy "A Un Compl(A) = {x.True}";
by (fast_tac eq_cs 1);
-val Compl_partition = result();
+qed "Compl_partition";
goal Set.thy "Compl(Compl(A)) = A";
by (fast_tac eq_cs 1);
-val double_complement = result();
+qed "double_complement";
goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)";
by (fast_tac eq_cs 1);
-val Compl_Un = result();
+qed "Compl_Un";
goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)";
by (fast_tac eq_cs 1);
-val Compl_Int = result();
+qed "Compl_Int";
goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))";
by (fast_tac eq_cs 1);
-val Compl_UN = result();
+qed "Compl_UN";
goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))";
by (fast_tac eq_cs 1);
-val Compl_INT = result();
+qed "Compl_INT";
(*Halmos, Naive Set Theory, page 16.*)
goal Set.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_eq = result();
+qed "Un_Int_assoc_eq";
(** Big Union and Intersection **)
goal Set.thy "Union({}) = {}";
by (fast_tac eq_cs 1);
-val Union_empty = result();
+qed "Union_empty";
goal Set.thy "Union(insert(a,B)) = a Un Union(B)";
by (fast_tac eq_cs 1);
-val Union_insert = result();
+qed "Union_insert";
goal Set.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 Set.thy "Union(A Int B) <= Union(A) Int Union(B)";
by (fast_tac set_cs 1);
-val Union_Int_subset = result();
+qed "Union_Int_subset";
val prems = goal Set.thy
"(Union(C) Int A = {}) = (! B:C. B Int A = {})";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
-val Union_disjoint = result();
+qed "Union_disjoint";
goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)";
by (best_tac eq_cs 1);
-val Inter_Un_distrib = result();
+qed "Inter_Un_distrib";
(** Unions and Intersections of Families **)
@@ -183,141 +183,141 @@
goal Set.thy "Union(range(f)) = (UN x.f(x))";
by (fast_tac eq_cs 1);
-val Union_range_eq = result();
+qed "Union_range_eq";
goal Set.thy "Inter(range(f)) = (INT x.f(x))";
by (fast_tac eq_cs 1);
-val Inter_range_eq = result();
+qed "Inter_range_eq";
goal Set.thy "Union(B``A) = (UN x:A. B(x))";
by (fast_tac eq_cs 1);
-val Union_image_eq = result();
+qed "Union_image_eq";
goal Set.thy "Inter(B``A) = (INT x:A. B(x))";
by (fast_tac eq_cs 1);
-val Inter_image_eq = result();
+qed "Inter_image_eq";
goal Set.thy "!!A. a: A ==> (UN y:A. c) = c";
by (fast_tac eq_cs 1);
-val UN_constant = result();
+qed "UN_constant";
goal Set.thy "!!A. a: A ==> (INT y:A. c) = c";
by (fast_tac eq_cs 1);
-val INT_constant = result();
+qed "INT_constant";
goal Set.thy "(UN x.B) = B";
by (fast_tac eq_cs 1);
-val UN1_constant = result();
+qed "UN1_constant";
goal Set.thy "(INT x.B) = B";
by (fast_tac eq_cs 1);
-val INT1_constant = result();
+qed "INT1_constant";
goal Set.thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})";
by (fast_tac eq_cs 1);
-val UN_eq = result();
+qed "UN_eq";
(*Look: it has an EXISTENTIAL quantifier*)
goal Set.thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})";
by (fast_tac eq_cs 1);
-val INT_eq = result();
+qed "INT_eq";
(*Distributive laws...*)
goal Set.thy "A Int Union(B) = (UN C:B. A Int C)";
by (fast_tac eq_cs 1);
-val Int_Union = result();
+qed "Int_Union";
(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:
Union of a family of unions **)
goal Set.thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)";
by (fast_tac eq_cs 1);
-val Un_Union_image = result();
+qed "Un_Union_image";
(*Equivalent version*)
goal Set.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 Set.thy "A Un Inter(B) = (INT C:B. A Un C)";
by (fast_tac eq_cs 1);
-val Un_Inter = result();
+qed "Un_Inter";
goal Set.thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)";
by (best_tac eq_cs 1);
-val Int_Inter_image = result();
+qed "Int_Inter_image";
(*Equivalent version*)
goal Set.thy "(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";
(*Halmos, Naive Set Theory, page 35.*)
goal Set.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 Set.thy "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 Set.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 Set.thy
"(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";
(** Simple properties of Diff -- set difference **)
goal Set.thy "A-A = {}";
by (fast_tac eq_cs 1);
-val Diff_cancel = result();
+qed "Diff_cancel";
goal Set.thy "{}-A = {}";
by (fast_tac eq_cs 1);
-val empty_Diff = result();
+qed "empty_Diff";
goal Set.thy "A-{} = A";
by (fast_tac eq_cs 1);
-val Diff_empty = result();
+qed "Diff_empty";
(*NOT SUITABLE FOR REWRITING since {a} == insert(a,0)*)
goal Set.thy "A - insert(a,B) = A - B - {a}";
by (fast_tac eq_cs 1);
-val Diff_insert = result();
+qed "Diff_insert";
(*NOT SUITABLE FOR REWRITING since {a} == insert(a,0)*)
goal Set.thy "A - insert(a,B) = A - {a} - B";
by (fast_tac eq_cs 1);
-val Diff_insert2 = result();
+qed "Diff_insert2";
val prems = goal Set.thy "a:A ==> insert(a,A-{a}) = A";
by (fast_tac (eq_cs addSIs prems) 1);
-val insert_Diff = result();
+qed "insert_Diff";
goal Set.thy "A Int (B-A) = {}";
by (fast_tac eq_cs 1);
-val Diff_disjoint = result();
+qed "Diff_disjoint";
goal Set.thy "!!A. A<=B ==> A Un (B-A) = B";
by (fast_tac eq_cs 1);
-val Diff_partition = result();
+qed "Diff_partition";
goal Set.thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)";
by (fast_tac eq_cs 1);
-val double_diff = result();
+qed "double_diff";
goal Set.thy "A - (B Un C) = (A-B) Int (A-C)";
by (fast_tac eq_cs 1);
-val Diff_Un = result();
+qed "Diff_Un";
goal Set.thy "A - (B Int C) = (A-B) Un (A-C)";
by (fast_tac eq_cs 1);
-val Diff_Int = result();
+qed "Diff_Int";
val set_ss = set_ss addsimps
[in_empty,in_insert,insert_subset,
--- a/ex/Acc.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Acc.ML Mon Nov 21 17:50:34 1994 +0100
@@ -16,13 +16,13 @@
"[| !!b. <b,a>:r ==> b: acc(r) |] ==> a: acc(r)";
by (fast_tac (set_cs addIs (prems @
map (rewrite_rule [pred_def]) acc.intrs)) 1);
-val accI = result();
+qed "accI";
goal Acc.thy "!!a b r. [| b: acc(r); <a,b>: r |] ==> a: acc(r)";
by (etac acc.elim 1);
by (rewtac pred_def);
by (fast_tac set_cs 1);
-val acc_downward = result();
+qed "acc_downward";
val [major,indhyp] = goal Acc.thy
"[| a : acc(r); \
@@ -34,7 +34,7 @@
by (rewtac pred_def);
by (fast_tac set_cs 2);
be (Int_lower1 RS Pow_mono RS subsetD) 1;
-val acc_induct = result();
+qed "acc_induct";
val [major] = goal Acc.thy "r <= Sigma(acc(r), %u. acc(r)) ==> wf(r)";
@@ -42,22 +42,22 @@
by (etac acc.induct 1);
by (rewtac pred_def);
by (fast_tac set_cs 1);
-val acc_wfI = result();
+qed "acc_wfI";
val [major] = goal Acc.thy "wf(r) ==> ALL x. <x,y>: r | <y,x>:r --> y: acc(r)";
by (rtac (major RS wf_induct) 1);
br (impI RS allI) 1;
br accI 1;
by (fast_tac set_cs 1);
-val acc_wfD_lemma = result();
+qed "acc_wfD_lemma";
val [major] = goal Acc.thy "wf(r) ==> r <= Sigma(acc(r), %u. acc(r))";
by (rtac subsetI 1);
by (res_inst_tac [("p", "x")] PairE 1);
by (fast_tac (set_cs addSIs [SigmaI,
major RS acc_wfD_lemma RS spec RS mp]) 1);
-val acc_wfD = result();
+qed "acc_wfD";
goal Acc.thy "wf(r) = (r <= Sigma(acc(r), %u. acc(r)))";
by (EVERY1 [rtac iffI, etac acc_wfD, etac acc_wfI]);
-val wf_acc_iff = result();
+qed "wf_acc_iff";
--- a/ex/InSort.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/InSort.ML Mon Nov 21 17:50:34 1994 +0100
@@ -12,12 +12,12 @@
goalw InSort.thy [Sorting.total_def]
"!!f. [| total(f); ~f(x,y) |] ==> f(y,x)";
by(fast_tac HOL_cs 1);
-val totalD = result();
+qed "totalD";
goalw InSort.thy [Sorting.transf_def]
"!!f. [| transf(f); f(b,a) |] ==> !x. f(a,x) --> f(b,x)";
by(fast_tac HOL_cs 1);
-val transfD = result();
+qed "transfD";
goal InSort.thy "list_all(p,ins(f,x,xs)) = (list_all(p,xs) & p(x))";
by(list_ind_tac "xs" 1);
@@ -29,7 +29,7 @@
goal InSort.thy "(!x. p(x) --> q(x)) --> list_all(p,xs) --> list_all(q,xs)";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac (insort_ss setloop (split_tac [expand_if]))));
-val list_all_imp = result();
+qed "list_all_imp";
val prems = goal InSort.thy
"[| total(f); transf(f) |] ==> sorted(f,ins(f,x,xs)) = sorted(f,xs)";
--- a/ex/LList.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/LList.ML Mon Nov 21 17:50:34 1994 +0100
@@ -23,7 +23,7 @@
goalw LList.thy llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)";
by (rtac gfp_mono 1);
by (REPEAT (ares_tac basic_monos 1));
-val llist_mono = result();
+qed "llist_mono";
goal LList.thy "llist(A) = {Numb(0)} <+> (A <*> llist(A))";
@@ -31,7 +31,7 @@
by (fast_tac (univ_cs addSIs (equalityI :: map rew llist.intrs)
addEs [rew llist.elim]) 1)
end;
-val llist_unfold = result();
+qed "llist_unfold";
(*** Type checking by coinduction, using list_Fun
@@ -43,44 +43,44 @@
be llist.coinduct 1;
be (subsetD RS CollectD) 1;
ba 1;
-val llist_coinduct = result();
+qed "llist_coinduct";
goalw LList.thy [list_Fun_def, NIL_def] "NIL: list_Fun(A,X)";
by (fast_tac set_cs 1);
-val list_Fun_NIL_I = result();
+qed "list_Fun_NIL_I";
goalw LList.thy [list_Fun_def,CONS_def]
"!!M N. [| M: A; N: X |] ==> CONS(M,N) : list_Fun(A,X)";
by (fast_tac set_cs 1);
-val list_Fun_CONS_I = result();
+qed "list_Fun_CONS_I";
(*Utilise the "strong" part, i.e. gfp(f)*)
goalw LList.thy (llist.defs @ [list_Fun_def])
"!!M N. M: llist(A) ==> M : list_Fun(A, X Un llist(A))";
by (etac (llist.mono RS gfp_fun_UnI2) 1);
-val list_Fun_llist_I = result();
+qed "list_Fun_llist_I";
(*** LList_corec satisfies the desired recurion equation ***)
(*A continuity result?*)
goalw LList.thy [CONS_def] "CONS(M, UN x.f(x)) = (UN x. CONS(M, f(x)))";
by (simp_tac (univ_ss addsimps [In1_UN1, Scons_UN1_y]) 1);
-val CONS_UN1 = result();
+qed "CONS_UN1";
(*UNUSED; obsolete?
goal Prod.thy "split(p, %x y.UN z.f(x,y,z)) = (UN z. split(p, %x y.f(x,y,z)))";
by (simp_tac (prod_ss setloop (split_tac [expand_split])) 1);
-val split_UN1 = result();
+qed "split_UN1";
goal Sum.thy "sum_case(s,f,%y.UN z.g(y,z)) = (UN z.sum_case(s,f,%y. g(y,z)))";
by (simp_tac (sum_ss setloop (split_tac [expand_sum_case])) 1);
-val sum_case2_UN1 = result();
+qed "sum_case2_UN1";
*)
val prems = goalw LList.thy [CONS_def]
"[| M<=M'; N<=N' |] ==> CONS(M,N) <= CONS(M',N')";
by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1));
-val CONS_mono = result();
+qed "CONS_mono";
val corec_fun_simps = [LList_corec_fun_def RS def_nat_rec_0,
LList_corec_fun_def RS def_nat_rec_Suc];
@@ -95,7 +95,7 @@
by (res_inst_tac [("n","k")] natE 1);
by (ALLGOALS (asm_simp_tac corec_fun_ss));
by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1));
-val LList_corec_subset1 = result();
+qed "LList_corec_subset1";
goalw LList.thy [LList_corec_def]
"sum_case(%u.NIL, split(%z w. CONS(z, LList_corec(w,f))), f(a)) <= \
@@ -104,7 +104,7 @@
by (safe_tac set_cs);
by (ALLGOALS (res_inst_tac [("x","Suc(?k)")] UN1_I THEN'
asm_simp_tac corec_fun_ss));
-val LList_corec_subset2 = result();
+qed "LList_corec_subset2";
(*the recursion equation for LList_corec -- NOT SUITABLE FOR REWRITING!*)
goal LList.thy
@@ -112,7 +112,7 @@
\ split(%z w. CONS(z, LList_corec(w,f))), f(a))";
by (REPEAT (resolve_tac [equalityI, LList_corec_subset1,
LList_corec_subset2] 1));
-val LList_corec = result();
+qed "LList_corec";
(*definitional version of same*)
val [rew] = goal LList.thy
@@ -120,7 +120,7 @@
\ h(a) = sum_case(%u.NIL, split(%z w. CONS(z, h(w))), f(a))";
by (rewtac rew);
by (rtac LList_corec 1);
-val def_LList_corec = result();
+qed "def_LList_corec";
(*A typical use of co-induction to show membership in the gfp.
Bisimulation is range(%x. LList_corec(x,f)) *)
@@ -131,7 +131,7 @@
by (stac LList_corec 1);
by (simp_tac (llist_ss addsimps [list_Fun_NIL_I, list_Fun_CONS_I, CollectI]
|> add_eqI) 1);
-val LList_corec_type = result();
+qed "LList_corec_type";
(*Lemma for the proof of llist_corec*)
goal LList.thy
@@ -143,7 +143,7 @@
by (stac LList_corec 1);
by (asm_simp_tac (llist_ss addsimps [list_Fun_NIL_I]) 1);
by (fast_tac (set_cs addSIs [list_Fun_CONS_I]) 1);
-val LList_corec_type2 = result();
+qed "LList_corec_type2";
(**** llist equality as a gfp; the bisimulation principle ****)
@@ -154,7 +154,7 @@
by (fast_tac (univ_cs addSIs (equalityI :: map rew LListD.intrs)
addEs [rew LListD.elim]) 1)
end;
-val LListD_unfold = result();
+qed "LListD_unfold";
goal LList.thy "!M N. <M,N> : LListD(diag(A)) --> ntrunc(k,M) = ntrunc(k,N)";
by (res_inst_tac [("n", "k")] less_induct 1);
@@ -167,7 +167,7 @@
by (res_inst_tac [("n", "n'")] natE 1);
by (asm_simp_tac (univ_ss addsimps [CONS_def, ntrunc_one_In1]) 1);
by (asm_simp_tac (univ_ss addsimps [CONS_def, ntrunc_In1, ntrunc_Scons]) 1);
-val LListD_implies_ntrunc_equality = result();
+qed "LListD_implies_ntrunc_equality";
(*The domain of the LListD relation*)
goalw LList.thy (llist.defs @ [NIL_def, CONS_def])
@@ -177,7 +177,7 @@
by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1);
by (simp_tac fst_image_ss 1);
by (fast_tac univ_cs 1);
-val fst_image_LListD = result();
+qed "fst_image_LListD";
(*This inclusion justifies the use of coinduction to show M=N*)
goal LList.thy "LListD(diag(A)) <= diag(llist(A))";
@@ -189,7 +189,7 @@
ntrunc_equality) 1);
by (assume_tac 1);
by (etac (fst_imageI RS (fst_image_LListD RS subsetD)) 1);
-val LListD_subset_diag = result();
+qed "LListD_subset_diag";
(** Coinduction, using LListD_Fun
THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS!
@@ -200,25 +200,25 @@
be LListD.coinduct 1;
be (subsetD RS CollectD) 1;
ba 1;
-val LListD_coinduct = result();
+qed "LListD_coinduct";
goalw LList.thy [LListD_Fun_def,NIL_def] "<NIL,NIL> : LListD_Fun(r,s)";
by (fast_tac set_cs 1);
-val LListD_Fun_NIL_I = result();
+qed "LListD_Fun_NIL_I";
goalw LList.thy [LListD_Fun_def,CONS_def]
"!!x. [| x:A; <M,N>:s |] ==> <CONS(x,M), CONS(x,N)> : LListD_Fun(diag(A),s)";
by (fast_tac univ_cs 1);
-val LListD_Fun_CONS_I = result();
+qed "LListD_Fun_CONS_I";
(*Utilise the "strong" part, i.e. gfp(f)*)
goalw LList.thy (LListD.defs @ [LListD_Fun_def])
"!!M N. M: LListD(r) ==> M : LListD_Fun(r, X Un LListD(r))";
by (etac (LListD.mono RS gfp_fun_UnI2) 1);
-val LListD_Fun_LListD_I = result();
+qed "LListD_Fun_LListD_I";
-(*This converse inclusion helps to strengthen llist_equalityI*)
+(*This converse inclusion helps to strengthen LList_equalityI*)
goal LList.thy "diag(llist(A)) <= LListD(diag(A))";
by (rtac subsetI 1);
by (etac LListD_coinduct 1);
@@ -229,19 +229,19 @@
by (ALLGOALS
(asm_simp_tac (llist_ss addsimps [diagI, LListD_Fun_NIL_I,
LListD_Fun_CONS_I])));
-val diag_subset_LListD = result();
+qed "diag_subset_LListD";
goal LList.thy "LListD(diag(A)) = diag(llist(A))";
by (REPEAT (resolve_tac [equalityI, LListD_subset_diag,
diag_subset_LListD] 1));
-val LListD_eq_diag = result();
+qed "LListD_eq_diag";
goal LList.thy
"!!M N. M: llist(A) ==> <M,M> : LListD_Fun(diag(A), X Un diag(llist(A)))";
by (rtac (LListD_eq_diag RS subst) 1);
br LListD_Fun_LListD_I 1;
by (asm_simp_tac (HOL_ss addsimps [LListD_eq_diag, diagI]) 1);
-val LListD_Fun_diag_I = result();
+qed "LListD_Fun_diag_I";
(** To show two LLists are equal, exhibit a bisimulation!
@@ -254,7 +254,7 @@
by (etac LListD_coinduct 1);
by (asm_simp_tac (HOL_ss addsimps [LListD_eq_diag]) 1);
by (safe_tac prod_cs);
-val llist_equalityI = result();
+qed "LList_equalityI";
(*** Finality of llist(A): Uniqueness of functions defined by corecursion ***)
@@ -267,7 +267,7 @@
by (rtac ext 1);
(*next step avoids an unknown (and flexflex pair) in simplification*)
by (res_inst_tac [("A", "{u.True}"),
- ("r", "range(%u. <h1(u),h2(u)>)")] llist_equalityI 1);
+ ("r", "range(%u. <h1(u),h2(u)>)")] LList_equalityI 1);
by (rtac rangeI 1);
by (safe_tac set_cs);
by (stac prem1 1);
@@ -275,25 +275,25 @@
by (simp_tac (llist_ss addsimps [LListD_Fun_NIL_I,
CollectI RS LListD_Fun_CONS_I]
|> add_eqI) 1);
-val LList_corec_unique = result();
+qed "LList_corec_unique";
val [prem] = goal LList.thy
"[| !!x. h(x) = sum_case(%u.NIL, split(%z w. CONS(z,h(w))), f(x)) |] \
\ ==> h = (%x.LList_corec(x,f))";
by (rtac (LList_corec RS (prem RS LList_corec_unique)) 1);
-val equals_LList_corec = result();
+qed "equals_LList_corec";
(** Obsolete LList_corec_unique proof: complete induction, not coinduction **)
goalw LList.thy [CONS_def] "ntrunc(Suc(0), CONS(M,N)) = {}";
by (rtac ntrunc_one_In1 1);
-val ntrunc_one_CONS = result();
+qed "ntrunc_one_CONS";
goalw LList.thy [CONS_def]
"ntrunc(Suc(Suc(k)), CONS(M,N)) = CONS (ntrunc(k,M), ntrunc(k,N))";
by (simp_tac (HOL_ss addsimps [ntrunc_Scons,ntrunc_In1]) 1);
-val ntrunc_CONS = result();
+qed "ntrunc_CONS";
val [prem1,prem2] = goal LList.thy
"[| !!x. h1(x) = sum_case(%u.NIL, split(%z w. CONS(z,h1(w))), f(x)); \
@@ -311,14 +311,14 @@
by (res_inst_tac [("n", "xc")] natE 2);
by (ALLGOALS(asm_simp_tac(nat_ss addsimps
[ntrunc_0,ntrunc_one_CONS,ntrunc_CONS])));
-val LList_corec_unique = result();
+result();
(*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
goal LList.thy "mono(CONS(M))";
by (REPEAT (ares_tac [monoI, subset_refl, CONS_mono] 1));
-val Lconst_fun_mono = result();
+qed "Lconst_fun_mono";
(* Lconst(M) = CONS(M,Lconst(M)) *)
val Lconst = standard (Lconst_fun_mono RS (Lconst_def RS def_lfp_Tarski));
@@ -330,20 +330,20 @@
by (safe_tac set_cs);
by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1);
by (REPEAT (ares_tac [list_Fun_CONS_I, singletonI, UnI1] 1));
-val Lconst_type = result();
+qed "Lconst_type";
goal LList.thy "Lconst(M) = LList_corec(M, %x.Inr(<x,x>))";
by (rtac (equals_LList_corec RS fun_cong) 1);
by (simp_tac sum_ss 1);
by (rtac Lconst 1);
-val Lconst_eq_LList_corec = result();
+qed "Lconst_eq_LList_corec";
(*Thus we could have used gfp in the definition of Lconst*)
goal LList.thy "gfp(%N. CONS(M, N)) = LList_corec(M, %x.Inr(<x,x>))";
by (rtac (equals_LList_corec RS fun_cong) 1);
by (simp_tac sum_ss 1);
by (rtac (Lconst_fun_mono RS gfp_Tarski) 1);
-val gfp_Lconst_eq_LList_corec = result();
+qed "gfp_Lconst_eq_LList_corec";
(*** Isomorphisms ***)
@@ -351,19 +351,19 @@
goal LList.thy "inj(Rep_llist)";
by (rtac inj_inverseI 1);
by (rtac Rep_llist_inverse 1);
-val inj_Rep_llist = result();
+qed "inj_Rep_llist";
goal LList.thy "inj_onto(Abs_llist,llist(range(Leaf)))";
by (rtac inj_onto_inverseI 1);
by (etac Abs_llist_inverse 1);
-val inj_onto_Abs_llist = result();
+qed "inj_onto_Abs_llist";
(** Distinctness of constructors **)
goalw LList.thy [LNil_def,LCons_def] "~ LCons(x,xs) = LNil";
by (rtac (CONS_not_NIL RS (inj_onto_Abs_llist RS inj_onto_contraD)) 1);
by (REPEAT (resolve_tac (llist.intrs @ [rangeI, Rep_llist]) 1));
-val LCons_not_LNil = result();
+qed "LCons_not_LNil";
val LNil_not_LCons = standard (LCons_not_LNil RS not_sym);
@@ -375,19 +375,19 @@
goalw LList.thy [LNil_def]
"Rep_llist(LNil) = NIL";
by (rtac (llist.NIL_I RS Abs_llist_inverse) 1);
-val Rep_llist_LNil = result();
+qed "Rep_llist_LNil";
goalw LList.thy [LCons_def]
"Rep_llist(LCons(x,l)) = CONS(Leaf(x),Rep_llist(l))";
by (REPEAT (resolve_tac [llist.CONS_I RS Abs_llist_inverse,
rangeI, Rep_llist] 1));
-val Rep_llist_LCons = result();
+qed "Rep_llist_LCons";
(** Injectiveness of CONS and LCons **)
goalw LList.thy [CONS_def] "(CONS(M,N)=CONS(M',N')) = (M=M' & N=N')";
by (fast_tac (HOL_cs addSEs [Scons_inject, make_elim In1_inject]) 1);
-val CONS_CONS_eq = result();
+qed "CONS_CONS_eq";
val CONS_inject = standard (CONS_CONS_eq RS iffD1 RS conjE);
@@ -400,14 +400,14 @@
goalw LList.thy [LCons_def] "(LCons(x,xs)=LCons(y,ys)) = (x=y & xs=ys)";
by (fast_tac llist_cs 1);
-val LCons_LCons_eq = result();
+qed "LCons_LCons_eq";
val LCons_inject = standard (LCons_LCons_eq RS iffD1 RS conjE);
val [major] = goal LList.thy "CONS(M,N): llist(A) ==> M: A & N: llist(A)";
by (rtac (major RS llist.elim) 1);
by (etac CONS_neq_NIL 1);
by (fast_tac llist_cs 1);
-val CONS_D = result();
+qed "CONS_D";
(****** Reasoning about llist(A) ******)
@@ -424,7 +424,7 @@
\ LListD_Fun(diag(A), (%u.<f(u),g(u)>)``llist(A) Un \
\ diag(llist(A))) \
\ |] ==> f(M) = g(M)";
-by (rtac llist_equalityI 1);
+by (rtac LList_equalityI 1);
br (Mlist RS imageI) 1;
by (rtac subsetI 1);
by (etac imageE 1);
@@ -435,7 +435,7 @@
br (gMlist RS LListD_Fun_diag_I) 1;
by (etac ssubst 1);
by (REPEAT (ares_tac [CONScase] 1));
-val llist_fun_equalityI = result();
+qed "LList_fun_equalityI";
(*** The functional "Lmap" ***)
@@ -443,12 +443,12 @@
goal LList.thy "Lmap(f,NIL) = NIL";
by (rtac (Lmap_def RS def_LList_corec RS trans) 1);
by (simp_tac List_case_ss 1);
-val Lmap_NIL = result();
+qed "Lmap_NIL";
goal LList.thy "Lmap(f, CONS(M,N)) = CONS(f(M), Lmap(f,N))";
by (rtac (Lmap_def RS def_LList_corec RS trans) 1);
by (simp_tac List_case_ss 1);
-val Lmap_CONS = result();
+qed "Lmap_CONS";
(*Another type-checking proof by coinduction*)
val [major,minor] = goal LList.thy
@@ -459,34 +459,34 @@
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
by (REPEAT (ares_tac [list_Fun_NIL_I, list_Fun_CONS_I,
minor, imageI, UnI1] 1));
-val Lmap_type = result();
+qed "Lmap_type";
(*This type checking rule synthesises a sufficiently large set for f*)
val [major] = goal LList.thy "M: llist(A) ==> Lmap(f,M): llist(f``A)";
by (rtac (major RS Lmap_type) 1);
by (etac imageI 1);
-val Lmap_type2 = result();
+qed "Lmap_type2";
(** Two easy results about Lmap **)
val [prem] = goalw LList.thy [o_def]
"M: llist(A) ==> Lmap(f o g, M) = Lmap(f, Lmap(g, M))";
-by (rtac (prem RS imageI RS llist_equalityI) 1);
+by (rtac (prem RS imageI RS LList_equalityI) 1);
by (safe_tac set_cs);
by (etac llist.elim 1);
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, UnI1,
rangeI RS LListD_Fun_CONS_I] 1));
-val Lmap_compose = result();
+qed "Lmap_compose";
val [prem] = goal LList.thy "M: llist(A) ==> Lmap(%x.x, M) = M";
-by (rtac (prem RS imageI RS llist_equalityI) 1);
+by (rtac (prem RS imageI RS LList_equalityI) 1);
by (safe_tac set_cs);
by (etac llist.elim 1);
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps [Lmap_NIL,Lmap_CONS])));
by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI RS UnI1,
rangeI RS LListD_Fun_CONS_I] 1));
-val Lmap_ident = result();
+qed "Lmap_ident";
(*** Lappend -- its two arguments cause some complications! ***)
@@ -494,19 +494,19 @@
goalw LList.thy [Lappend_def] "Lappend(NIL,NIL) = NIL";
by (rtac (LList_corec RS trans) 1);
by (simp_tac List_case_ss 1);
-val Lappend_NIL_NIL = result();
+qed "Lappend_NIL_NIL";
goalw LList.thy [Lappend_def]
"Lappend(NIL,CONS(N,N')) = CONS(N, Lappend(NIL,N'))";
by (rtac (LList_corec RS trans) 1);
by (simp_tac List_case_ss 1);
-val Lappend_NIL_CONS = result();
+qed "Lappend_NIL_CONS";
goalw LList.thy [Lappend_def]
"Lappend(CONS(M,M'), N) = CONS(M, Lappend(M',N))";
by (rtac (LList_corec RS trans) 1);
by (simp_tac List_case_ss 1);
-val Lappend_CONS = result();
+qed "Lappend_CONS";
val Lappend_ss =
List_case_ss addsimps [llist.NIL_I, Lappend_NIL_NIL, Lappend_NIL_CONS,
@@ -514,14 +514,14 @@
|> add_eqI;
goal LList.thy "!!M. M: llist(A) ==> Lappend(NIL,M) = M";
-by (etac llist_fun_equalityI 1);
+by (etac LList_fun_equalityI 1);
by (ALLGOALS (asm_simp_tac Lappend_ss));
-val Lappend_NIL = result();
+qed "Lappend_NIL";
goal LList.thy "!!M. M: llist(A) ==> Lappend(M,NIL) = M";
-by (etac llist_fun_equalityI 1);
+by (etac LList_fun_equalityI 1);
by (ALLGOALS (asm_simp_tac Lappend_ss));
-val Lappend_NIL2 = result();
+qed "Lappend_NIL2";
(** Alternative type-checking proofs for Lappend **)
@@ -550,7 +550,7 @@
by (asm_simp_tac (Lappend_ss addsimps [Lappend_NIL, list_Fun_llist_I]) 1);
by (asm_simp_tac Lappend_ss 1);
by (fast_tac (set_cs addSIs [list_Fun_CONS_I]) 1);
-val Lappend_type = result();
+qed "Lappend_type";
(**** Lazy lists as the type 'a llist -- strongly typed versions of above ****)
@@ -565,12 +565,12 @@
goalw LList.thy [llist_case_def,LNil_def] "llist_case(c, d, LNil) = c";
by (simp_tac Rep_llist_ss 1);
-val llist_case_LNil = result();
+qed "llist_case_LNil";
goalw LList.thy [llist_case_def,LCons_def]
"llist_case(c, d, LCons(M,N)) = d(M,N)";
by (simp_tac Rep_llist_ss 1);
-val llist_case_LCons = result();
+qed "llist_case_LCons";
(*Elimination is case analysis, not induction.*)
val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def]
@@ -585,7 +585,7 @@
by (asm_simp_tac (HOL_ss addsimps [Rep_llist_LCons]) 1);
by (etac (Abs_llist_inverse RS ssubst) 1);
by (rtac refl 1);
-val llistE = result();
+qed "llistE";
(** llist_corec: corecursion for 'a llist **)
@@ -599,7 +599,7 @@
by (asm_simp_tac (llist_ss addsimps [LList_corec_type2,Abs_llist_inverse]) 1);
(*FIXME: correct case splits usd to be found automatically:
by (ASM_SIMP_TAC(llist_ss addsimps [LList_corec_type2,Abs_llist_inverse]) 1);*)
-val llist_corec = result();
+qed "llist_corec";
(*definitional version of same*)
val [rew] = goal LList.thy
@@ -607,7 +607,7 @@
\ h(a) = sum_case(%u.LNil, split(%z w. LCons(z, h(w))), f(a))";
by (rewtac rew);
by (rtac llist_corec 1);
-val def_llist_corec = result();
+qed "def_llist_corec";
(**** Proofs about type 'a llist functions ****)
@@ -619,13 +619,13 @@
by (stac llist_unfold 1);
by (simp_tac (HOL_ss addsimps [NIL_def, CONS_def]) 1);
by (fast_tac univ_cs 1);
-val LListD_Fun_subset_Sigma_llist = result();
+qed "LListD_Fun_subset_Sigma_llist";
goal LList.thy
"prod_fun(Rep_llist,Rep_llist) `` r <= \
\ Sigma(llist(range(Leaf)), %x.llist(range(Leaf)))";
by (fast_tac (prod_cs addIs [Rep_llist]) 1);
-val subset_Sigma_llist = result();
+qed "subset_Sigma_llist";
val [prem] = goal LList.thy
"r <= Sigma(llist(range(Leaf)), %x.llist(range(Leaf))) ==> \
@@ -634,7 +634,7 @@
by (rtac (prem RS subsetD RS SigmaE2) 1);
by (assume_tac 1);
by (asm_simp_tac (HOL_ss addsimps [o_def,prod_fun,Abs_llist_inverse]) 1);
-val prod_fun_lemma = result();
+qed "prod_fun_lemma";
goal LList.thy
"prod_fun(Rep_llist, Rep_llist) `` range(%x. <x, x>) = \
@@ -642,7 +642,7 @@
br equalityI 1;
by (fast_tac (univ_cs addIs [Rep_llist]) 1);
by (fast_tac (univ_cs addSEs [Abs_llist_inverse RS subst]) 1);
-val prod_fun_range_eq_diag = result();
+qed "prod_fun_range_eq_diag";
(** To show two llists are equal, exhibit a bisimulation!
[also admits true equality] **)
@@ -651,7 +651,7 @@
by (rtac (inj_Rep_llist RS injD) 1);
by (res_inst_tac [("r", "prod_fun(Rep_llist,Rep_llist)``r"),
("A", "range(Leaf)")]
- llist_equalityI 1);
+ LList_equalityI 1);
by (rtac (prem1 RS prod_fun_imageI) 1);
by (rtac (prem2 RS image_mono RS subset_trans) 1);
by (rtac (image_compose RS subst) 1);
@@ -661,18 +661,18 @@
by (rtac (LListD_Fun_subset_Sigma_llist RS prod_fun_lemma) 1);
by (rtac (subset_Sigma_llist RS Un_least) 1);
by (rtac diag_subset_Sigma 1);
-val llist_equalityI = result();
+qed "llist_equalityI";
(** Rules to prove the 2nd premise of llist_equalityI **)
goalw LList.thy [llistD_Fun_def,LNil_def] "<LNil,LNil> : llistD_Fun(r)";
by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1);
-val llistD_Fun_LNil_I = result();
+qed "llistD_Fun_LNil_I";
val [prem] = goalw LList.thy [llistD_Fun_def,LCons_def]
"<l1,l2>:r ==> <LCons(x,l1), LCons(x,l2)> : llistD_Fun(r)";
by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1);
by (rtac (prem RS prod_fun_imageI) 1);
-val llistD_Fun_LCons_I = result();
+qed "llistD_Fun_LCons_I";
(*Utilise the "strong" part, i.e. gfp(f)*)
goalw LList.thy [llistD_Fun_def]
@@ -682,7 +682,7 @@
by (rtac (image_Un RS ssubst) 1);
by (stac prod_fun_range_eq_diag 1);
br (Rep_llist RS LListD_Fun_diag_I) 1;
-val llistD_Fun_range_I = result();
+qed "llistD_Fun_range_I";
(*A special case of list_equality for functions over lazy lists*)
val [prem1,prem2] = goal LList.thy
@@ -701,7 +701,7 @@
by (rtac llistD_Fun_range_I 1);
by (etac ssubst 1);
by (rtac prem2 1);
-val llist_fun_equalityI = result();
+qed "llist_fun_equalityI";
(*simpset for llist bisimulations*)
val llistD_simps = [llist_case_LNil, llist_case_LCons,
@@ -715,12 +715,12 @@
goal LList.thy "lmap(f,LNil) = LNil";
by (rtac (lmap_def RS def_llist_corec RS trans) 1);
by (simp_tac llistD_ss 1);
-val lmap_LNil = result();
+qed "lmap_LNil";
goal LList.thy "lmap(f, LCons(M,N)) = LCons(f(M), lmap(f,N))";
by (rtac (lmap_def RS def_llist_corec RS trans) 1);
by (simp_tac llistD_ss 1);
-val lmap_LCons = result();
+qed "lmap_LCons";
(** Two easy results about lmap **)
@@ -728,12 +728,12 @@
goal LList.thy "lmap(f o g, l) = lmap(f, lmap(g, l))";
by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons])));
-val lmap_compose = result();
+qed "lmap_compose";
goal LList.thy "lmap(%x.x, l) = l";
by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
by (ALLGOALS (simp_tac (llistD_ss addsimps [lmap_LNil, lmap_LCons])));
-val lmap_ident = result();
+qed "lmap_ident";
(*** iterates -- llist_fun_equalityI cannot be used! ***)
@@ -741,7 +741,7 @@
goal LList.thy "iterates(f,x) = LCons(x, iterates(f,f(x)))";
by (rtac (iterates_def RS def_llist_corec RS trans) 1);
by (simp_tac sum_ss 1);
-val iterates = result();
+qed "iterates";
goal LList.thy "lmap(f, iterates(f,x)) = iterates(f,f(x))";
by (res_inst_tac [("r", "range(%u.<lmap(f,iterates(f,u)),iterates(f,f(u))>)")]
@@ -751,12 +751,12 @@
by (res_inst_tac [("x1", "f(u)")] (iterates RS ssubst) 1);
by (res_inst_tac [("x1", "u")] (iterates RS ssubst) 1);
by (simp_tac (llistD_ss addsimps [lmap_LCons]) 1);
-val lmap_iterates = result();
+qed "lmap_iterates";
goal LList.thy "iterates(f,x) = LCons(x, lmap(f, iterates(f,x)))";
br (lmap_iterates RS ssubst) 1;
br iterates 1;
-val iterates_lmap = result();
+qed "iterates_lmap";
(*** A rather complex proof about iterates -- cf Andy Pitts ***)
@@ -767,12 +767,12 @@
\ LCons(nat_rec(n, b, %m. f), nat_rec(n, l, %m. lmap(f)))";
by (nat_ind_tac "n" 1);
by (ALLGOALS (asm_simp_tac (nat_ss addsimps [lmap_LCons])));
-val fun_power_lmap = result();
+qed "fun_power_lmap";
goal Nat.thy "nat_rec(n, g(x), %m. g) = nat_rec(Suc(n), x, %m. g)";
by (nat_ind_tac "n" 1);
by (ALLGOALS (asm_simp_tac nat_ss));
-val fun_power_Suc = result();
+qed "fun_power_Suc";
val Pair_cong = read_instantiate_sg (sign_of Prod.thy)
[("f","Pair")] (standard(refl RS cong RS cong));
@@ -798,7 +798,7 @@
by (stac fun_power_Suc 1);
br (UN1_I RS UnI1) 1;
br rangeI 1;
-val iterates_equality = result();
+qed "iterates_equality";
(*** lappend -- its two arguments cause some complications! ***)
@@ -806,31 +806,31 @@
goalw LList.thy [lappend_def] "lappend(LNil,LNil) = LNil";
by (rtac (llist_corec RS trans) 1);
by (simp_tac llistD_ss 1);
-val lappend_LNil_LNil = result();
+qed "lappend_LNil_LNil";
goalw LList.thy [lappend_def]
"lappend(LNil,LCons(l,l')) = LCons(l, lappend(LNil,l'))";
by (rtac (llist_corec RS trans) 1);
by (simp_tac llistD_ss 1);
-val lappend_LNil_LCons = result();
+qed "lappend_LNil_LCons";
goalw LList.thy [lappend_def]
"lappend(LCons(l,l'), N) = LCons(l, lappend(l',N))";
by (rtac (llist_corec RS trans) 1);
by (simp_tac llistD_ss 1);
-val lappend_LCons = result();
+qed "lappend_LCons";
goal LList.thy "lappend(LNil,l) = l";
by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
by (ALLGOALS
(simp_tac (llistD_ss addsimps [lappend_LNil_LNil, lappend_LNil_LCons])));
-val lappend_LNil = result();
+qed "lappend_LNil";
goal LList.thy "lappend(l,LNil) = l";
by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
by (ALLGOALS
(simp_tac (llistD_ss addsimps [lappend_LNil_LNil, lappend_LCons])));
-val lappend_LNil2 = result();
+qed "lappend_LNil2";
(*The infinite first argument blocks the second*)
goal LList.thy "lappend(iterates(f,x), N) = iterates(f,x)";
@@ -840,7 +840,7 @@
by (safe_tac set_cs);
by (stac iterates 1);
by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1);
-val lappend_iterates = result();
+qed "lappend_iterates";
(** Two proofs that lmap distributes over lappend **)
@@ -870,11 +870,11 @@
by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
by (simp_tac (llistD_ss addsimps [lappend_LNil, lmap_LNil])1);
by (simp_tac (llistD_ss addsimps [lappend_LCons, lmap_LCons]) 1);
-val lmap_lappend_distrib = result();
+qed "lmap_lappend_distrib";
(*Without strong coinduction, three case analyses might be needed*)
goal LList.thy "lappend(lappend(l1,l2) ,l3) = lappend(l1, lappend(l2,l3))";
by (res_inst_tac [("l","l1")] llist_fun_equalityI 1);
by (simp_tac (llistD_ss addsimps [lappend_LNil])1);
by (simp_tac (llistD_ss addsimps [lappend_LCons]) 1);
-val lappend_assoc = result();
+qed "lappend_assoc";
--- a/ex/LexProd.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/LexProd.ML Mon Nov 21 17:50:34 1994 +0100
@@ -12,7 +12,7 @@
by (rtac allI 1);
by (rtac (surjective_pairing RS ssubst) 1);
by (fast_tac HOL_cs 1);
-val split_all_pair = result();
+qed "split_all_pair";
val [wfa,wfb] = goalw LexProd.thy [wf_def,LexProd.lex_prod_def]
"[| wf(ra); wf(rb) |] ==> wf(ra**rb)";
@@ -21,4 +21,4 @@
by (EVERY1 [rtac allI,rtac impI]);
by (rtac (wfb RS spec RS mp) 1);
by (fast_tac (set_cs addSEs [Pair_inject]) 1);
-val wf_lex_prod = result();
+qed "wf_lex_prod";
--- a/ex/MT.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/MT.ML Mon Nov 21 17:50:34 1994 +0100
@@ -22,7 +22,7 @@
by (hyp_subst_tac 1);
by (rtac singletonI 1);
by (assume_tac 1);
-val notsingletonI = result();
+qed "notsingletonI";
val prems = goalw MT.thy [Un_def]
"[| c : A Un B; c : A & ~c : B ==> P; c : B ==> P |] ==> P";
@@ -31,7 +31,7 @@
by (resolve_tac prems 1);br conjI 1;ba 1;ba 1;
by (eresolve_tac prems 1);
by (eresolve_tac prems 1);
-val UnSE = result();
+qed "UnSE";
(* ############################################################ *)
(* Inference systems *)
@@ -49,14 +49,14 @@
by (rtac (fst_conv RS ssubst) 1);
by (rtac (snd_conv RS ssubst) 1);
by (resolve_tac prems 1);
-val infsys_p1 = result();
+qed "infsys_p1";
val prems = goal MT.thy "P(fst(<a,b>),snd(<a,b>)) ==> P(a,b)";
by (cut_facts_tac prems 1);
by (dtac (fst_conv RS subst) 1);
by (dtac (snd_conv RS subst) 1);
by (assume_tac 1);
-val infsys_p2 = result();
+qed "infsys_p2";
val prems = goal MT.thy
"P(a,b,c) ==> P(fst(fst(<<a,b>,c>)),snd(fst(<<a,b>,c>)),snd(<<a,b>,c>))";
@@ -65,7 +65,7 @@
by (rtac (snd_conv RS ssubst) 1);
by (rtac (snd_conv RS ssubst) 1);
by (resolve_tac prems 1);
-val infsys_pp1 = result();
+qed "infsys_pp1";
val prems = goal MT.thy
"P(fst(fst(<<a,b>,c>)),snd(fst(<<a,b>,c>)),snd(<<a,b>,c>)) ==> P(a,b,c)";
@@ -75,7 +75,7 @@
by (dtac (snd_conv RS subst) 1);
by (dtac (snd_conv RS subst) 1);
by (assume_tac 1);
-val infsys_pp2 = result();
+qed "infsys_pp2";
(* ############################################################ *)
(* Fixpoints *)
@@ -87,7 +87,7 @@
by (rtac subsetD 1);
by (rtac lfp_lemma2 1);
by (resolve_tac prems 1);brs prems 1;
-val lfp_intro2 = result();
+qed "lfp_intro2";
val prems = goal MT.thy
" [| x:lfp(f); mono(f); !!y. y:f(lfp(f)) ==> P(y) |] ==> \
@@ -95,7 +95,7 @@
by (cut_facts_tac prems 1);
by (resolve_tac prems 1);br subsetD 1;
by (rtac lfp_lemma3 1);ba 1;ba 1;
-val lfp_elim2 = result();
+qed "lfp_elim2";
val prems = goal MT.thy
" [| x:lfp(f); mono(f); !!y. y:f(lfp(f) Int {x.P(x)}) ==> P(y) |] ==> \
@@ -103,7 +103,7 @@
by (cut_facts_tac prems 1);
by (etac induct 1);ba 1;
by (eresolve_tac prems 1);
-val lfp_ind2 = result();
+qed "lfp_ind2";
(* Greatest fixpoints *)
@@ -117,12 +117,12 @@
by (rtac (monoh RS monoD RS subsetD) 1);
by (rtac Un_upper2 1);
by (etac (monoh RS gfp_lemma2 RS subsetD) 1);
-val gfp_coind2 = result();
+qed "gfp_coind2";
val [gfph,monoh,caseh] = goal MT.thy
"[| x:gfp(f); mono(f); !! y. y:f(gfp(f)) ==> P(y) |] ==> P(x)";
by (rtac caseh 1);br subsetD 1;br gfp_lemma2 1;br monoh 1;br gfph 1;
-val gfp_elim2 =result();
+qed "gfp_elim2";
(* ############################################################ *)
(* Expressions *)
@@ -168,7 +168,7 @@
goalw MT.thy [mono_def, eval_fun_def] "mono(eval_fun)";
by infsys_mono_tac;
-val eval_fun_mono = result();
+qed "eval_fun_mono";
(* Introduction rules *)
@@ -178,7 +178,7 @@
by (rewtac eval_fun_def);
by (rtac CollectI 1);br disjI1 1;
by (fast_tac HOL_cs 1);
-val eval_const = result();
+qed "eval_const";
val prems = goalw MT.thy [eval_def, eval_rel_def]
"ev:ve_dom(ve) ==> ve |- e_var(ev) ---> ve_app(ve,ev)";
@@ -188,7 +188,7 @@
by (rewtac eval_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val eval_var = result();
+qed "eval_var";
val prems = goalw MT.thy [eval_def, eval_rel_def]
"ve |- fn ev => e ---> v_clos(<|ev,e,ve|>)";
@@ -198,7 +198,7 @@
by (rewtac eval_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val eval_fn = result();
+qed "eval_fn";
val prems = goalw MT.thy [eval_def, eval_rel_def]
" cl = <| ev1, e, ve + {ev2 |-> v_clos(cl)} |> ==> \
@@ -209,7 +209,7 @@
by (rewtac eval_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val eval_fix = result();
+qed "eval_fix";
val prems = goalw MT.thy [eval_def, eval_rel_def]
" [| ve |- e1 ---> v_const(c1); ve |- e2 ---> v_const(c2) |] ==> \
@@ -220,7 +220,7 @@
by (rewtac eval_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI2 1;br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val eval_app1 = result();
+qed "eval_app1";
val prems = goalw MT.thy [eval_def, eval_rel_def]
" [| ve |- e1 ---> v_clos(<|xm,em,vem|>); \
@@ -234,7 +234,7 @@
by (rewtac eval_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI2 1;br disjI2 1;br disjI2 1;
by (fast_tac HOL_cs 1);
-val eval_app2 = result();
+qed "eval_app2";
(* Strong elimination, induction on evaluations *)
@@ -264,7 +264,7 @@
by (safe_tac HOL_cs);
by (ALLGOALS (resolve_tac prems));
by (ALLGOALS (fast_tac set_cs));
-val eval_ind0 = result();
+qed "eval_ind0";
val prems = goal MT.thy
" [| ve |- e ---> v; \
@@ -288,7 +288,7 @@
by (ALLGOALS (rtac infsys_pp1));
by (ALLGOALS (resolve_tac prems));
by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_pp2 1)));
-val eval_ind = result();
+qed "eval_ind";
(* ############################################################ *)
(* Elaborations *)
@@ -296,7 +296,7 @@
goalw MT.thy [mono_def, elab_fun_def] "mono(elab_fun)";
by infsys_mono_tac;
-val elab_fun_mono = result();
+qed "elab_fun_mono";
(* Introduction rules *)
@@ -308,7 +308,7 @@
by (rewtac elab_fun_def);
by (rtac CollectI 1);br disjI1 1;
by (fast_tac HOL_cs 1);
-val elab_const = result();
+qed "elab_const";
val prems = goalw MT.thy [elab_def, elab_rel_def]
"x:te_dom(te) ==> te |- e_var(x) ===> te_app(te,x)";
@@ -318,7 +318,7 @@
by (rewtac elab_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val elab_var = result();
+qed "elab_var";
val prems = goalw MT.thy [elab_def, elab_rel_def]
"te + {x |=> ty1} |- e ===> ty2 ==> te |- fn x => e ===> ty1->ty2";
@@ -328,7 +328,7 @@
by (rewtac elab_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val elab_fn = result();
+qed "elab_fn";
val prems = goalw MT.thy [elab_def, elab_rel_def]
" te + {f |=> ty1->ty2} + {x |=> ty1} |- e ===> ty2 ==> \
@@ -339,7 +339,7 @@
by (rewtac elab_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI2 1;br disjI1 1;
by (fast_tac HOL_cs 1);
-val elab_fix = result();
+qed "elab_fix";
val prems = goalw MT.thy [elab_def, elab_rel_def]
" [| te |- e1 ===> ty1->ty2; te |- e2 ===> ty1 |] ==> \
@@ -350,7 +350,7 @@
by (rewtac elab_fun_def);
by (rtac CollectI 1);br disjI2 1;br disjI2 1;br disjI2 1;br disjI2 1;
by (fast_tac HOL_cs 1);
-val elab_app = result();
+qed "elab_app";
(* Strong elimination, induction on elaborations *)
@@ -380,7 +380,7 @@
by (safe_tac HOL_cs);
by (ALLGOALS (resolve_tac prems));
by (ALLGOALS (fast_tac set_cs));
-val elab_ind0 = result();
+qed "elab_ind0";
val prems = goal MT.thy
" [| te |- e ===> t; \
@@ -406,7 +406,7 @@
by (ALLGOALS (rtac infsys_pp1));
by (ALLGOALS (resolve_tac prems));
by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_pp2 1)));
-val elab_ind = result();
+qed "elab_ind";
(* Weak elimination, case analysis on elaborations *)
@@ -431,7 +431,7 @@
by (safe_tac HOL_cs);
by (ALLGOALS (resolve_tac prems));
by (ALLGOALS (fast_tac set_cs));
-val elab_elim0 = result();
+qed "elab_elim0";
val prems = goal MT.thy
" [| te |- e ===> t; \
@@ -452,7 +452,7 @@
by (ALLGOALS (rtac infsys_pp1));
by (ALLGOALS (resolve_tac prems));
by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_pp2 1)));
-val elab_elim = result();
+qed "elab_elim";
(* Elimination rules for each expression *)
@@ -464,25 +464,25 @@
val prems = goal MT.thy "te |- e ===> t ==> (e = e_const(c) --> c isof t)";
by (elab_e_elim_tac prems);
-val elab_const_elim_lem = result();
+qed "elab_const_elim_lem";
val prems = goal MT.thy "te |- e_const(c) ===> t ==> c isof t";
by (cut_facts_tac prems 1);
by (dtac elab_const_elim_lem 1);
by (fast_tac prop_cs 1);
-val elab_const_elim = result();
+qed "elab_const_elim";
val prems = goal MT.thy
"te |- e ===> t ==> (e = e_var(x) --> t=te_app(te,x) & x:te_dom(te))";
by (elab_e_elim_tac prems);
-val elab_var_elim_lem = result();
+qed "elab_var_elim_lem";
val prems = goal MT.thy
"te |- e_var(ev) ===> t ==> t=te_app(te,ev) & ev : te_dom(te)";
by (cut_facts_tac prems 1);
by (dtac elab_var_elim_lem 1);
by (fast_tac prop_cs 1);
-val elab_var_elim = result();
+qed "elab_var_elim";
val prems = goal MT.thy
" te |- e ===> t ==> \
@@ -490,7 +490,7 @@
\ (? t1 t2.t=t_fun(t1,t2) & te + {x1 |=> t1} |- e1 ===> t2) \
\ )";
by (elab_e_elim_tac prems);
-val elab_fn_elim_lem = result();
+qed "elab_fn_elim_lem";
val prems = goal MT.thy
" te |- fn x1 => e1 ===> t ==> \
@@ -498,14 +498,14 @@
by (cut_facts_tac prems 1);
by (dtac elab_fn_elim_lem 1);
by (fast_tac prop_cs 1);
-val elab_fn_elim = result();
+qed "elab_fn_elim";
val prems = goal MT.thy
" te |- e ===> t ==> \
\ (e = fix f(x) = e1 --> \
\ (? t1 t2. t=t1->t2 & te + {f |=> t1->t2} + {x |=> t1} |- e1 ===> t2))";
by (elab_e_elim_tac prems);
-val elab_fix_elim_lem = result();
+qed "elab_fix_elim_lem";
val prems = goal MT.thy
" te |- fix ev1(ev2) = e1 ===> t ==> \
@@ -513,20 +513,20 @@
by (cut_facts_tac prems 1);
by (dtac elab_fix_elim_lem 1);
by (fast_tac prop_cs 1);
-val elab_fix_elim = result();
+qed "elab_fix_elim";
val prems = goal MT.thy
" te |- e ===> t2 ==> \
\ (e = e1 @ e2 --> (? t1 . te |- e1 ===> t1->t2 & te |- e2 ===> t1))";
by (elab_e_elim_tac prems);
-val elab_app_elim_lem = result();
+qed "elab_app_elim_lem";
val prems = goal MT.thy
"te |- e1 @ e2 ===> t2 ==> (? t1 . te |- e1 ===> t1->t2 & te |- e2 ===> t1)";
by (cut_facts_tac prems 1);
by (dtac elab_app_elim_lem 1);
by (fast_tac prop_cs 1);
-val elab_app_elim = result();
+qed "elab_app_elim";
(* ############################################################ *)
(* The extended correspondence relation *)
@@ -536,7 +536,7 @@
goalw MT.thy [mono_def,MT.hasty_fun_def] "mono(hasty_fun)";
by infsys_mono_tac;
-val mono_hasty_fun = result();
+val mono_hasty_fun = store_thm("mono_hasty_fun", result());
(*
Because hasty_rel has been defined as the greatest fixpoint of hasty_fun it
@@ -552,7 +552,7 @@
by (rtac CollectI 1);br disjI1 1;
by (fast_tac HOL_cs 1);
by (rtac mono_hasty_fun 1);
-val hasty_rel_const_coind = result();
+qed "hasty_rel_const_coind";
(* Second strong introduction (co-induction) rule for hasty_rel *)
@@ -570,7 +570,7 @@
by (rtac CollectI 1);br disjI2 1;
by (fast_tac HOL_cs 1);
by (rtac mono_hasty_fun 1);
-val hasty_rel_clos_coind = result();
+qed "hasty_rel_clos_coind";
(* Elimination rule for hasty_rel *)
@@ -592,7 +592,7 @@
by (safe_tac HOL_cs);
by (ALLGOALS (resolve_tac prems));
by (ALLGOALS (fast_tac set_cs));
-val hasty_rel_elim0 = result();
+qed "hasty_rel_elim0";
val prems = goal MT.thy
" [| <v,t> : hasty_rel; \
@@ -608,20 +608,20 @@
by (ALLGOALS (rtac infsys_p1));
by (ALLGOALS (resolve_tac prems));
by (REPEAT ((assume_tac 1) ORELSE (dtac infsys_p2 1)));
-val hasty_rel_elim = result();
+qed "hasty_rel_elim";
(* Introduction rules for hasty *)
val prems = goalw MT.thy [hasty_def] "c isof t ==> v_const(c) hasty t";
by (resolve_tac (prems RL [hasty_rel_const_coind]) 1);
-val hasty_const = result();
+qed "hasty_const";
val prems = goalw MT.thy [hasty_def,hasty_env_def]
"te |- fn ev => e ===> t & ve hastyenv te ==> v_clos(<|ev,e,ve|>) hasty t";
by (cut_facts_tac prems 1);
by (rtac hasty_rel_clos_coind 1);
by (ALLGOALS (fast_tac set_cs));
-val hasty_clos = result();
+qed "hasty_clos";
(* Elimination on constants for hasty *)
@@ -630,12 +630,12 @@
by (cut_facts_tac prems 1);
by (rtac hasty_rel_elim 1);
by (ALLGOALS (fast_tac (v_ext_cs HOL_cs)));
-val hasty_elim_const_lem = result();
+qed "hasty_elim_const_lem";
val prems = goal MT.thy "v_const(c) hasty t ==> c isof t";
by (cut_facts_tac (prems RL [hasty_elim_const_lem]) 1);
by (fast_tac HOL_cs 1);
-val hasty_elim_const = result();
+qed "hasty_elim_const";
(* Elimination on closures for hasty *)
@@ -646,13 +646,13 @@
by (cut_facts_tac prems 1);
by (rtac hasty_rel_elim 1);
by (ALLGOALS (fast_tac (v_ext_cs HOL_cs)));
-val hasty_elim_clos_lem = result();
+qed "hasty_elim_clos_lem";
val prems = goal MT.thy
"v_clos(<|ev,e,ve|>) hasty t ==> ? te.te |- fn ev => e ===> t & ve hastyenv te ";
by (cut_facts_tac (prems RL [hasty_elim_clos_lem]) 1);
by (fast_tac HOL_cs 1);
-val hasty_elim_clos = result();
+qed "hasty_elim_clos";
(* ############################################################ *)
(* The pointwise extension of hasty to environments *)
@@ -679,7 +679,7 @@
by (rtac (ve_app_owr1 RS ssubst) 1);
by (rtac (te_app_owr1 RS ssubst) 1);
by (assume_tac 1);
-val hasty_env1 = result();
+qed "hasty_env1";
(* ############################################################ *)
(* The Consistency theorem *)
@@ -690,7 +690,7 @@
by (cut_facts_tac prems 1);
by (dtac elab_const_elim 1);
by (etac hasty_const 1);
-val consistency_const = result();
+qed "consistency_const";
val prems = goalw MT.thy [hasty_env_def]
" [| ev : ve_dom(ve); ve hastyenv te ; te |- e_var(ev) ===> t |] ==> \
@@ -698,7 +698,7 @@
by (cut_facts_tac prems 1);
by (dtac elab_var_elim 1);
by (fast_tac HOL_cs 1);
-val consistency_var = result();
+qed "consistency_var";
val prems = goal MT.thy
" [| ve hastyenv te ; te |- fn ev => e ===> t |] ==> \
@@ -706,7 +706,7 @@
by (cut_facts_tac prems 1);
by (rtac hasty_clos 1);
by (fast_tac prop_cs 1);
-val consistency_fn = result();
+qed "consistency_fn";
val prems = goalw MT.thy [hasty_env_def,hasty_def]
" [| cl = <| ev1, e, ve + { ev2 |-> v_clos(cl) } |>; \
@@ -742,7 +742,7 @@
by (rtac (te_app_owr1 RS ssubst) 1);
by (etac subst 1);
by (fast_tac set_cs 1);
-val consistency_fix = result();
+qed "consistency_fix";
val prems = goal MT.thy
" [| ! t te. ve hastyenv te --> te |- e1 ===> t --> v_const(c1) hasty t; \
@@ -759,7 +759,7 @@
by (fast_tac HOL_cs 1);
by (rtac hasty_elim_const 1);
by (fast_tac HOL_cs 1);
-val consistency_app1 = result();
+qed "consistency_app1";
val prems = goal MT.thy
" [| ! t te. \
@@ -784,7 +784,7 @@
by (dtac t_fun_inj 1);
by (safe_tac prop_cs);
by ((dtac hasty_env1 1) THEN (assume_tac 1) THEN (fast_tac HOL_cs 1));
-val consistency_app2 = result();
+qed "consistency_app2";
val prems = goal MT.thy
"ve |- e ---> v ==> (! t te. ve hastyenv te --> te |- e ===> t --> v hasty t)";
@@ -800,7 +800,7 @@
by (rtac consistency_fix 1);ba 1;ba 1;ba 1;
by (rtac consistency_app1 1);ba 1;ba 1;ba 1;ba 1;
by (rtac consistency_app2 1);ba 5;ba 4;ba 3;ba 2;ba 1;
-val consistency = result();
+qed "consistency";
(* ############################################################ *)
(* The Basic Consistency theorem *)
@@ -813,7 +813,7 @@
by (etac allE 1);be impE 1;ba 1;be exE 1;be conjE 1;
by (dtac hasty_const 1);
by ((rtac ssubst 1) THEN (assume_tac 1) THEN (assume_tac 1));
-val basic_consistency_lem = result();
+qed "basic_consistency_lem";
val prems = goal MT.thy
"[| ve isofenv te; ve |- e ---> v_const(c); te |- e ===> t |] ==> c isof t";
@@ -821,6 +821,6 @@
by (rtac hasty_elim_const 1);
by (dtac consistency 1);
by (fast_tac (HOL_cs addSIs [basic_consistency_lem]) 1);
-val basic_consistency = result();
+qed "basic_consistency";
--- a/ex/NatSum.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/NatSum.ML Mon Nov 21 17:50:34 1994 +0100
@@ -15,7 +15,7 @@
by (nat_ind_tac "n" 1);
by (simp_tac natsum_ss 1);
by (asm_simp_tac natsum_ss 1);
-val sum_of_naturals = result();
+qed "sum_of_naturals";
goal NatSum.thy
"Suc(Suc(Suc(Suc(Suc(Suc(0))))))*sum(%i.i*i,Suc(n)) = \
@@ -24,7 +24,7 @@
by (nat_ind_tac "n" 1);
by (simp_tac natsum_ss 1);
by (asm_simp_tac natsum_ss 1);
-val sum_of_squares = result();
+qed "sum_of_squares";
goal NatSum.thy
"Suc(Suc(Suc(Suc(0))))*sum(%i.i*i*i,Suc(n)) = n*n*Suc(n)*Suc(n)";
@@ -32,12 +32,12 @@
by (nat_ind_tac "n" 1);
by (simp_tac natsum_ss 1);
by (asm_simp_tac natsum_ss 1);
-val sum_of_cubes = result();
+qed "sum_of_cubes";
(*The sum of the first n odd numbers equals n squared.*)
goal NatSum.thy "sum(%i.Suc(i+i), n) = n*n";
by (nat_ind_tac "n" 1);
by (simp_tac natsum_ss 1);
by (asm_simp_tac natsum_ss 1);
-val sum_of_odds = result();
+qed "sum_of_odds";
--- a/ex/PropLog.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/PropLog.ML Mon Nov 21 17:50:34 1994 +0100
@@ -14,12 +14,12 @@
goalw PropLog.thy thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)";
by (rtac lfp_mono 1);
by (REPEAT (ares_tac basic_monos 1));
-val thms_mono = result();
+qed "thms_mono";
(*Rule is called I for Identity Combinator, not for Introduction*)
goal PropLog.thy "H |- p->p";
by(best_tac (HOL_cs addIs [thms.K,thms.S,thms.MP]) 1);
-val thms_I = result();
+qed "thms_I";
(** Weakening, left and right **)
@@ -36,7 +36,7 @@
goal PropLog.thy "!!H. H |- q ==> H |- p->q";
by(fast_tac (HOL_cs addIs [thms.K,thms.MP]) 1);
-val weaken_right = result();
+qed "weaken_right";
(*The deduction theorem*)
goal PropLog.thy "!!H. insert(p,H) |- q ==> H |- p->q";
@@ -46,7 +46,7 @@
by (fast_tac (set_cs addIs [thms.S RS weaken_right]) 1);
by (fast_tac (set_cs addIs [thms.DN RS weaken_right]) 1);
by (fast_tac (set_cs addIs [thms.S RS thms.MP RS thms.MP]) 1);
-val deduction = result();
+qed "deduction";
(* "[| insert(p,H) |- q; H |- p |] ==> H |- q"
@@ -68,15 +68,15 @@
goalw PropLog.thy [eval_def] "tt[false] = False";
by (simp_tac pl_ss 1);
-val eval_false = result();
+qed "eval_false";
goalw PropLog.thy [eval_def] "tt[#v] = (v:tt)";
by (simp_tac pl_ss 1);
-val eval_var = result();
+qed "eval_var";
goalw PropLog.thy [eval_def] "tt[p->q] = (tt[p]-->tt[q])";
by (simp_tac pl_ss 1);
-val eval_imp = result();
+qed "eval_imp";
val pl_ss = pl_ss addsimps [eval_false, eval_var, eval_imp];
@@ -85,7 +85,7 @@
by (etac thms.induct 1);
by (fast_tac (set_cs addSDs [eval_imp RS iffD1 RS mp]) 5);
by (ALLGOALS (asm_simp_tac pl_ss));
-val soundness = result();
+qed "soundness";
(*** Towards the completeness proof ***)
@@ -94,7 +94,7 @@
by (etac (weaken_left_insert RS thms_notE) 1);
by (rtac thms.H 1);
by (rtac insertI1 1);
-val false_imp = result();
+qed "false_imp";
val [premp,premq] = goal PropLog.thy
"[| H |- p; H |- q->false |] ==> H |- (p->q)->false";
@@ -103,7 +103,7 @@
by (rtac (thms.H RS thms.MP) 1);
by (rtac insertI1 1);
by (rtac (premp RS weaken_left_insert) 1);
-val imp_false = result();
+qed "imp_false";
(*This formulation is required for strong induction hypotheses*)
goal PropLog.thy "hyps(p,tt) |- if(tt[p], p, p->false)";
@@ -113,14 +113,14 @@
by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2,
weaken_right, imp_false]
addSEs [false_imp]) 1);
-val hyps_thms_if = result();
+qed "hyps_thms_if";
(*Key lemma for completeness; yields a set of assumptions satisfying p*)
val [sat] = goalw PropLog.thy [sat_def] "{} |= p ==> hyps(p,tt) |- p";
by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN
rtac hyps_thms_if 2);
by (fast_tac set_cs 1);
-val sat_thms_p = result();
+qed "sat_thms_p";
(*For proving certain theorems in our new propositional logic*)
val thms_cs =
@@ -131,14 +131,14 @@
by (rtac (deduction RS deduction) 1);
by (rtac (thms.DN RS thms.MP) 1);
by (ALLGOALS (best_tac (thms_cs addSIs prems)));
-val thms_excluded_middle = result();
+qed "thms_excluded_middle";
(*Hard to prove directly because it requires cuts*)
val prems = goal PropLog.thy
"[| insert(p,H) |- q; insert(p->false,H) |- q |] ==> H |- q";
by (rtac (thms_excluded_middle RS thms.MP RS thms.MP) 1);
by (REPEAT (resolve_tac (prems@[deduction]) 1));
-val thms_excluded_middle_rule = result();
+qed "thms_excluded_middle_rule";
(*** Completeness -- lemmas for reducing the set of assumptions ***)
@@ -150,7 +150,7 @@
by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
by (simp_tac pl_ss 1);
by (fast_tac set_cs 1);
-val hyps_Diff = result();
+qed "hyps_Diff";
(*For the case hyps(p,t)-insert(#v -> false,Y) |- p;
we also have hyps(p,t)-{#v->false} <= hyps(p, insert(v,t)) *)
@@ -160,17 +160,17 @@
by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
by (simp_tac pl_ss 1);
by (fast_tac set_cs 1);
-val hyps_insert = result();
+qed "hyps_insert";
(** Two lemmas for use with weaken_left **)
goal Set.thy "B-C <= insert(a, B-insert(a,C))";
by (fast_tac set_cs 1);
-val insert_Diff_same = result();
+qed "insert_Diff_same";
goal Set.thy "insert(a, B-{c}) - D <= insert(a, B-insert(c,D))";
by (fast_tac set_cs 1);
-val insert_Diff_subset2 = result();
+qed "insert_Diff_subset2";
(*The set hyps(p,t) is finite, and elements have the form #v or #v->false;
could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*)
@@ -178,7 +178,7 @@
by (PropLog.pl.induct_tac "p" 1);
by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN'
fast_tac (set_cs addSIs Fin.intrs@[Fin_UnI])));
-val hyps_finite = result();
+qed "hyps_finite";
val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;
@@ -203,19 +203,19 @@
by (rtac (insert_Diff_subset2 RS weaken_left) 1);
by (rtac (hyps_insert RS Diff_weaken_left) 1);
by (etac spec 1);
-val completeness_0_lemma = result();
+qed "completeness_0_lemma";
(*The base case for completeness*)
val [sat] = goal PropLog.thy "{} |= p ==> {} |- p";
by (rtac (Diff_cancel RS subst) 1);
by (rtac (sat RS (completeness_0_lemma RS spec)) 1);
-val completeness_0 = result();
+qed "completeness_0";
(*A semantic analogue of the Deduction Theorem*)
val [sat] = goalw PropLog.thy [sat_def] "insert(p,H) |= q ==> H |= p->q";
by (simp_tac pl_ss 1);
by (cfast_tac [sat] 1);
-val sat_imp = result();
+qed "sat_imp";
val [finite] = goal PropLog.thy "H: Fin({p.True}) ==> !p. H |= p --> H |- p";
by (rtac (finite RS Fin_induct) 1);
@@ -223,12 +223,12 @@
by (rtac (weaken_left_insert RS thms.MP) 1);
by (fast_tac (set_cs addSIs [sat_imp]) 1);
by (fast_tac thms_cs 1);
-val completeness_lemma = result();
+qed "completeness_lemma";
val completeness = completeness_lemma RS spec RS mp;
val [finite] = goal PropLog.thy "H: Fin({p.True}) ==> (H |- p) = (H |= p)";
by (fast_tac (set_cs addSEs [soundness, finite RS completeness]) 1);
-val thms_iff = result();
+qed "thms_iff";
writeln"Reached end of file.";
--- a/ex/Puzzle.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Puzzle.ML Mon Nov 21 17:50:34 1994 +0100
@@ -11,7 +11,7 @@
(*specialized form of induction needed below*)
val prems = goal Nat.thy "[| P(0); !!n. P(Suc(n)) |] ==> !n.P(n)";
by (EVERY1 [rtac (nat_induct RS allI), resolve_tac prems, resolve_tac prems]);
-val nat_exh = result();
+qed "nat_exh";
goal Puzzle.thy "! n. k=f(n) --> n <= f(n)";
by (res_inst_tac [("n","k")] less_induct 1);
@@ -23,33 +23,33 @@
by (subgoal_tac "f(na) <= f(f(na))" 1);
by (best_tac (HOL_cs addIs [lessD,Puzzle.f_ax,le_less_trans,le_trans]) 1);
by (fast_tac (HOL_cs addIs [Puzzle.f_ax]) 1);
-val lemma = result() RS spec RS mp;
+val lemma = store_thm("lemma", result() RS spec RS mp);
goal Puzzle.thy "n <= f(n)";
by (fast_tac (HOL_cs addIs [lemma]) 1);
-val lemma1 = result();
+qed "lemma1";
goal Puzzle.thy "f(n) < f(Suc(n))";
by (fast_tac (HOL_cs addIs [Puzzle.f_ax,le_less_trans,lemma1]) 1);
-val lemma2 = result();
+qed "lemma2";
val prems = goal Puzzle.thy "(!!n.f(n) <= f(Suc(n))) ==> m<n --> f(m) <= f(n)";
by (res_inst_tac[("n","n")]nat_induct 1);
by (simp_tac nat_ss 1);
by (simp_tac nat_ss 1);
by (fast_tac (HOL_cs addIs (le_trans::prems)) 1);
-val mono_lemma1 = result() RS mp;
+val mono_lemma1 = store_thm("mono_lemma1", result() RS mp);
val [p1,p2] = goal Puzzle.thy
"[| !! n. f(n)<=f(Suc(n)); m<=n |] ==> f(m) <= f(n)";
by (rtac (p2 RS le_imp_less_or_eq RS disjE) 1);
by (etac (p1 RS mono_lemma1) 1);
by (fast_tac (HOL_cs addIs [le_refl]) 1);
-val mono_lemma = result();
+qed "mono_lemma";
val prems = goal Puzzle.thy "m <= n ==> f(m) <= f(n)";
by (fast_tac (HOL_cs addIs ([mono_lemma,less_imp_le,lemma2]@prems)) 1);
-val f_mono = result();
+qed "f_mono";
goal Puzzle.thy "f(n) = n";
by (rtac le_anti_sym 1);
--- a/ex/Qsort.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Qsort.ML Mon Nov 21 17:50:34 1994 +0100
@@ -27,7 +27,7 @@
goal HOL.thy "((~P --> Q) & (P --> Q)) = Q";
by(fast_tac HOL_cs 1);
-val lemma = result();
+qed "lemma";
goal Qsort.thy "(Alls x:qsort(le,xs).P(x)) = (Alls x:xs.P(x))";
by(res_inst_tac[("xs","xs"),("p","le")]Qsort.qsort_ind 1);
--- a/ex/Rec.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Rec.ML Mon Nov 21 17:50:34 1994 +0100
@@ -2,4 +2,4 @@
goalw thy [mono_def,Domf_def] "mono(Domf(F))";
by (DEPTH_SOLVE (slow_step_tac set_cs 1));
-val mono_Domf = result();
+qed "mono_Domf";
--- a/ex/Simult.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Simult.ML Mon Nov 21 17:50:34 1994 +0100
@@ -19,19 +19,19 @@
\ <+> ({Numb(0)} <+> Part(Z,In0) <*> Part(Z,In1)))";
by (REPEAT (ares_tac [monoI, subset_refl, usum_mono, uprod_mono,
Part_mono] 1));
-val TF_fun_mono = result();
+qed "TF_fun_mono";
val TF_unfold = TF_fun_mono RS (TF_def RS def_lfp_Tarski);
goalw Simult.thy [TF_def] "!!A B. A<=B ==> TF(A) <= TF(B)";
by (REPEAT (ares_tac [lfp_mono, subset_refl, usum_mono, uprod_mono] 1));
-val TF_mono = result();
+qed "TF_mono";
goalw Simult.thy [TF_def] "TF(sexp) <= sexp";
by (rtac lfp_lowerbound 1);
by (fast_tac (univ_cs addIs sexp.intrs@[sexp_In0I, sexp_In1I]
addSEs [PartE]) 1);
-val TF_sexp = result();
+qed "TF_sexp";
(* A <= sexp ==> TF(A) <= sexp *)
val TF_subset_sexp = standard
@@ -52,14 +52,14 @@
by (rtac ([TF_def, TF_fun_mono, major] MRS def_induct) 1);
by (fast_tac (set_cs addIs (prems@[PartI])
addEs [usumE, uprodE, PartE]) 1);
-val TF_induct = result();
+qed "TF_induct";
(*This lemma replaces a use of subgoal_tac to prove tree_forest_induct*)
val prems = goalw Simult.thy [Part_def]
"! M: TF(A). (M: Part(TF(A),In0) --> P(M)) & (M: Part(TF(A),In1) --> Q(M)) \
\ ==> (! M: Part(TF(A),In0). P(M)) & (! M: Part(TF(A),In1). Q(M))";
by (cfast_tac prems 1);
-val TF_induct_lemma = result();
+qed "TF_induct_lemma";
val uplus_cs = set_cs addSIs [PartI]
addSDs [In0_inject, In1_inject]
@@ -79,7 +79,7 @@
by (rewrite_goals_tac TF_Rep_defs);
by (ALLGOALS (fast_tac (uplus_cs addIs prems)));
(*29 secs??*)
-val Tree_Forest_induct = result();
+qed "Tree_Forest_induct";
(*Induction for the abstract types 'a tree, 'a forest*)
val prems = goalw Simult.thy [Tcons_def,Fnil_def,Fcons_def]
@@ -98,7 +98,7 @@
by (ALLGOALS (fast_tac (set_cs addSEs [Abs_Tree_inverse RS subst,
Abs_Forest_inverse RS subst]
addSIs prems)));
-val tree_forest_induct = result();
+qed "tree_forest_induct";
@@ -107,22 +107,22 @@
goal Simult.thy "inj(Rep_Tree)";
by (rtac inj_inverseI 1);
by (rtac Rep_Tree_inverse 1);
-val inj_Rep_Tree = result();
+qed "inj_Rep_Tree";
goal Simult.thy "inj_onto(Abs_Tree,Part(TF(range(Leaf)),In0))";
by (rtac inj_onto_inverseI 1);
by (etac Abs_Tree_inverse 1);
-val inj_onto_Abs_Tree = result();
+qed "inj_onto_Abs_Tree";
goal Simult.thy "inj(Rep_Forest)";
by (rtac inj_inverseI 1);
by (rtac Rep_Forest_inverse 1);
-val inj_Rep_Forest = result();
+qed "inj_Rep_Forest";
goal Simult.thy "inj_onto(Abs_Forest,Part(TF(range(Leaf)),In1))";
by (rtac inj_onto_inverseI 1);
by (etac Abs_Forest_inverse 1);
-val inj_onto_Abs_Forest = result();
+qed "inj_onto_Abs_Forest";
(** Introduction rules for constructors **)
@@ -137,36 +137,36 @@
val prems = goalw Simult.thy TF_Rep_defs
"[| a: A; M: Part(TF(A),In1) |] ==> TCONS(a,M) : Part(TF(A),In0)";
by (fast_tac (TF_Rep_cs addIs prems) 1);
-val TCONS_I = result();
+qed "TCONS_I";
(* FNIL is a TF(A) -- this also justifies the type definition*)
goalw Simult.thy TF_Rep_defs "FNIL: Part(TF(A),In1)";
by (fast_tac TF_Rep_cs 1);
-val FNIL_I = result();
+qed "FNIL_I";
val prems = goalw Simult.thy TF_Rep_defs
"[| M: Part(TF(A),In0); N: Part(TF(A),In1) |] ==> \
\ FCONS(M,N) : Part(TF(A),In1)";
by (fast_tac (TF_Rep_cs addIs prems) 1);
-val FCONS_I = result();
+qed "FCONS_I";
(** Injectiveness of TCONS and FCONS **)
goalw Simult.thy TF_Rep_defs "(TCONS(K,M)=TCONS(L,N)) = (K=L & M=N)";
by (fast_tac TF_Rep_cs 1);
-val TCONS_TCONS_eq = result();
+qed "TCONS_TCONS_eq";
val TCONS_inject = standard (TCONS_TCONS_eq RS iffD1 RS conjE);
goalw Simult.thy TF_Rep_defs "(FCONS(K,M)=FCONS(L,N)) = (K=L & M=N)";
by (fast_tac TF_Rep_cs 1);
-val FCONS_FCONS_eq = result();
+qed "FCONS_FCONS_eq";
val FCONS_inject = standard (FCONS_FCONS_eq RS iffD1 RS conjE);
(** Distinctness of TCONS, FNIL and FCONS **)
goalw Simult.thy TF_Rep_defs "TCONS(M,N) ~= FNIL";
by (fast_tac TF_Rep_cs 1);
-val TCONS_not_FNIL = result();
+qed "TCONS_not_FNIL";
val FNIL_not_TCONS = standard (TCONS_not_FNIL RS not_sym);
val TCONS_neq_FNIL = standard (TCONS_not_FNIL RS notE);
@@ -174,7 +174,7 @@
goalw Simult.thy TF_Rep_defs "FCONS(M,N) ~= FNIL";
by (fast_tac TF_Rep_cs 1);
-val FCONS_not_FNIL = result();
+qed "FCONS_not_FNIL";
val FNIL_not_FCONS = standard (FCONS_not_FNIL RS not_sym);
val FCONS_neq_FNIL = standard (FCONS_not_FNIL RS notE);
@@ -182,7 +182,7 @@
goalw Simult.thy TF_Rep_defs "TCONS(M,N) ~= FCONS(K,L)";
by (fast_tac TF_Rep_cs 1);
-val TCONS_not_FCONS = result();
+qed "TCONS_not_FCONS";
val FCONS_not_TCONS = standard (TCONS_not_FCONS RS not_sym);
val TCONS_neq_FCONS = standard (TCONS_not_FCONS RS notE);
@@ -206,12 +206,12 @@
goalw Simult.thy [Tcons_def] "(Tcons(x,xs)=Tcons(y,ys)) = (x=y & xs=ys)";
by (fast_tac TF_cs 1);
-val Tcons_Tcons_eq = result();
+qed "Tcons_Tcons_eq";
val Tcons_inject = standard (Tcons_Tcons_eq RS iffD1 RS conjE);
goalw Simult.thy [Fcons_def,Fnil_def] "Fcons(x,xs) ~= Fnil";
by (fast_tac TF_cs 1);
-val Fcons_not_Fnil = result();
+qed "Fcons_not_Fnil";
val Fcons_neq_Fnil = standard (Fcons_not_Fnil RS notE);;
val Fnil_neq_Fcons = sym RS Fcons_neq_Fnil;
@@ -221,7 +221,7 @@
goalw Simult.thy [Fcons_def] "(Fcons(x,xs)=Fcons(y,ys)) = (x=y & xs=ys)";
by (fast_tac TF_cs 1);
-val Fcons_Fcons_eq = result();
+qed "Fcons_Fcons_eq";
val Fcons_inject = standard (Fcons_Fcons_eq RS iffD1 RS conjE);
@@ -238,12 +238,12 @@
by (rtac (TF_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [Case_In0, Split]) 1);
by (asm_simp_tac (pred_sexp_ss addsimps [In0_def]) 1);
-val TF_rec_TCONS = result();
+qed "TF_rec_TCONS";
goalw Simult.thy [FNIL_def] "TF_rec(FNIL,b,c,d) = c";
by (rtac (TF_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [Case_In1, List_case_NIL]) 1);
-val TF_rec_FNIL = result();
+qed "TF_rec_FNIL";
goalw Simult.thy [FCONS_def]
"!!M N. [| M: sexp; N: sexp |] ==> \
@@ -251,7 +251,7 @@
by (rtac (TF_rec_unfold RS trans) 1);
by (simp_tac (HOL_ss addsimps [Case_In1, List_case_CONS]) 1);
by (asm_simp_tac (pred_sexp_ss addsimps [CONS_def,In1_def]) 1);
-val TF_rec_FCONS = result();
+qed "TF_rec_FCONS";
(*** tree_rec, forest_rec -- by TF_rec ***)
@@ -274,14 +274,14 @@
goalw Simult.thy [tree_rec_def, forest_rec_def, Tcons_def]
"tree_rec(Tcons(a,tf),b,c,d) = b(a,tf,forest_rec(tf,b,c,d))";
by (simp_tac tf_rec_ss 1);
-val tree_rec_Tcons = result();
+qed "tree_rec_Tcons";
goalw Simult.thy [forest_rec_def, Fnil_def] "forest_rec(Fnil,b,c,d) = c";
by (simp_tac tf_rec_ss 1);
-val forest_rec_Fnil = result();
+qed "forest_rec_Fnil";
goalw Simult.thy [tree_rec_def, forest_rec_def, Fcons_def]
"forest_rec(Fcons(t,tf),b,c,d) = \
\ d(t,tf,tree_rec(t,b,c,d), forest_rec(tf,b,c,d))";
by (simp_tac tf_rec_ss 1);
-val forest_rec_Cons = result();
+qed "forest_rec_Cons";
--- a/ex/Sorting.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Sorting.ML Mon Nov 21 17:50:34 1994 +0100
@@ -14,13 +14,13 @@
goal Sorting.thy "!x.mset(xs@ys,x) = mset(xs,x)+mset(ys,x)";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac (sorting_ss setloop (split_tac [expand_if]))));
-val mset_app_distr = result();
+qed "mset_app_distr";
goal Sorting.thy "!x. mset([x:xs. ~p(x)], x) + mset([x:xs.p(x)], x) = \
\ mset(xs, x)";
by(list_ind_tac "xs" 1);
by(ALLGOALS(asm_simp_tac (sorting_ss setloop (split_tac [expand_if]))));
-val mset_compl_add = result();
+qed "mset_compl_add";
val sorting_ss = sorting_ss addsimps
[mset_app_distr, mset_compl_add];
--- a/ex/Term.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/Term.ML Mon Nov 21 17:50:34 1994 +0100
@@ -14,19 +14,19 @@
goal Term.thy "term(A) = A <*> list(term(A))";
by (fast_tac (univ_cs addSIs (equalityI :: term.intrs)
addEs [term.elim]) 1);
-val term_unfold = result();
+qed "term_unfold";
(*This justifies using term in other recursive type definitions*)
goalw Term.thy term.defs "!!A B. A<=B ==> term(A) <= term(B)";
by (REPEAT (ares_tac ([lfp_mono, list_mono] @ basic_monos) 1));
-val term_mono = result();
+qed "term_mono";
(** Type checking -- term creates well-founded sets **)
goalw Term.thy term.defs "term(sexp) <= sexp";
by (rtac lfp_lowerbound 1);
by (fast_tac (univ_cs addIs [sexp.SconsI, list_sexp RS subsetD]) 1);
-val term_sexp = result();
+qed "term_sexp";
(* A <= sexp ==> term(A) <= sexp *)
val term_subset_sexp = standard ([term_mono, term_sexp] MRS subset_trans);
@@ -56,7 +56,7 @@
by (rtac (major RS term_induct) 1);
by (etac list.induct 1);
by (REPEAT (ares_tac prems 1));
-val term_induct2 = result();
+qed "term_induct2";
(*** Structural Induction on the abstract type 'a term ***)
@@ -90,7 +90,7 @@
by (rtac list_subset_sexp 2);
by (fast_tac set_cs 2);
by (ALLGOALS (asm_simp_tac list_all_ss));
-val term_induct = result();
+qed "term_induct";
(*Induction for the abstract type 'a term*)
val prems = goal Term.thy
@@ -101,7 +101,7 @@
by (etac rev_mp 1);
by (rtac list_induct 1); (*types force good instantiation*)
by (ALLGOALS (asm_simp_tac (list_all_ss addsimps prems)));
-val term_induct2 = result();
+qed "term_induct2";
(*Perform induction on xs. *)
fun term_ind2_tac a i =
@@ -127,7 +127,7 @@
by (strip_tac 1);
by (etac (pred_sexp_CONS_D RS conjE) 1);
by (asm_simp_tac (list_all_ss addsimps [trancl_pred_sexpD1, cut_apply]) 1);
-val Abs_map_lemma = result();
+qed "Abs_map_lemma";
val [prem1,prem2,A_subset_sexp] = goal Term.thy
"[| M: sexp; N: list(term(A)); A<=sexp |] ==> \
@@ -138,7 +138,7 @@
prem2 RS Abs_map_lemma RS spec RS mp, pred_sexpI2 RS r_into_trancl,
prem1, prem2 RS rev_subsetD, list_subset_sexp,
term_subset_sexp, A_subset_sexp])1);
-val Term_rec = result();
+qed "Term_rec";
(*** term_rec -- by Term_rec ***)
--- a/ex/meson.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/meson.ML Mon Nov 21 17:50:34 1994 +0100
@@ -61,7 +61,7 @@
"! x. ? y. Q(x,y) ==> ? f. ! x. Q(x,f(x))";
by (cut_facts_tac [major] 1);
by (fast_tac (HOL_cs addEs [selectI]) 1);
-val choice = result();
+qed "choice";
(***** Generating clauses for the Meson Proof Procedure *****)
@@ -83,7 +83,7 @@
by (rtac notE 1);
by (etac minor 2);
by (ALLGOALS assume_tac);
-val make_neg_rule = result();
+qed "make_neg_rule";
(*For Plaisted's "Postive refinement" of the MESON procedure*)
val [major,minor] = goal HOL.thy "~P|Q ==> (P ==> Q)";
@@ -91,7 +91,7 @@
by (rtac notE 1);
by (rtac minor 2);
by (ALLGOALS assume_tac);
-val make_refined_neg_rule = result();
+qed "make_refined_neg_rule";
(*P should be a literal*)
val [major,minor] = goal HOL.thy "P|Q ==> ((P==>~P) ==> Q)";
@@ -99,7 +99,7 @@
by (rtac notE 1);
by (etac minor 1);
by (ALLGOALS assume_tac);
-val make_pos_rule = result();
+qed "make_pos_rule";
(*** Generation of a goal clause -- put away the final literal ***)
@@ -107,13 +107,13 @@
by (rtac notE 1);
by (rtac minor 2);
by (ALLGOALS (rtac major));
-val make_neg_goal = result();
+qed "make_neg_goal";
val [major,minor] = goal HOL.thy "P ==> ((P==>~P) ==> False)";
by (rtac notE 1);
by (rtac minor 1);
by (ALLGOALS (rtac major));
-val make_pos_goal = result();
+qed "make_pos_goal";
(**** Lemmas for forward proof (like congruence rules) ****)
@@ -125,28 +125,28 @@
by (rtac (major RS conjE) 1);
by (rtac conjI 1);
by (ALLGOALS (eresolve_tac prems));
-val conj_forward = result();
+qed "conj_forward";
val major::prems = goal HOL.thy
"[| P'|Q'; P' ==> P; Q' ==> Q |] ==> P|Q";
by (rtac (major RS disjE) 1);
by (ALLGOALS (dresolve_tac prems));
by (ALLGOALS (eresolve_tac [disjI1,disjI2]));
-val disj_forward = result();
+qed "disj_forward";
val major::prems = goal HOL.thy
"[| ! x. P'(x); !!x. P'(x) ==> P(x) |] ==> ! x. P(x)";
by (rtac allI 1);
by (resolve_tac prems 1);
by (rtac (major RS spec) 1);
-val all_forward = result();
+qed "all_forward";
val major::prems = goal HOL.thy
"[| ? x. P'(x); !!x. P'(x) ==> P(x) |] ==> ? x. P(x)";
by (rtac (major RS exE) 1);
by (rtac exI 1);
by (eresolve_tac prems 1);
-val ex_forward = result();
+qed "ex_forward";
(**** Operators for forward proof ****)
@@ -263,7 +263,7 @@
by (ALLGOALS (eresolve_tac prems));
by (etac notE 1);
by (assume_tac 1);
-val disj_forward2 = result();
+qed "disj_forward2";
(*Forward proof, passing extra assumptions as theorems to the tactic*)
fun forward_res2 nf hyps state =
--- a/ex/rel.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/rel.ML Mon Nov 21 17:50:34 1994 +0100
@@ -35,46 +35,46 @@
val [prem] = goalw Rel.thy [domain_def,Pair_def] "<a,b>: r ==> a : domain(r)";
by (fast_tac (set_cs addIs [prem]) 1);
-val domainI = result();
+qed "domainI";
val major::prems = goalw Rel.thy [domain_def]
"[| a : domain(r); !!y. <a,y>: r ==> P |] ==> P";
by (rtac (major RS CollectE) 1);
by (etac exE 1);
by (REPEAT (ares_tac prems 1));
-val domainE = result();
+qed "domainE";
(*** range2 ***)
val [prem] = goalw Rel.thy [range2_def,Pair_def] "<a,b>: r ==> b : range2(r)";
by (fast_tac (set_cs addIs [prem]) 1);
-val range2I = result();
+qed "range2I";
val major::prems = goalw Rel.thy [range2_def]
"[| b : range2(r); !!x. <x,b>: r ==> P |] ==> P";
by (rtac (major RS CollectE) 1);
by (etac exE 1);
by (REPEAT (ares_tac prems 1));
-val range2E = result();
+qed "range2E";
(*** field ***)
val [prem] = goalw Rel.thy [field_def] "<a,b>: r ==> a : field(r)";
by (rtac (prem RS domainI RS UnI1) 1);
-val fieldI1 = result();
+qed "fieldI1";
val [prem] = goalw Rel.thy [field_def] "<a,b>: r ==> b : field(r)";
by (rtac (prem RS range2I RS UnI2) 1);
-val fieldI2 = result();
+qed "fieldI2";
val [prem] = goalw Rel.thy [field_def]
"(~ <c,a>:r ==> <a,b>: r) ==> a : field(r)";
by (rtac (prem RS domainI RS UnCI) 1);
by (swap_res_tac [range2I] 1);
by (etac notnotD 1);
-val fieldCI = result();
+qed "fieldCI";
val major::prems = goalw Rel.thy [field_def]
"[| a : field(r); \
@@ -82,15 +82,15 @@
\ !!x. <x,a>: r ==> P |] ==> P";
by (rtac (major RS UnE) 1);
by (REPEAT (eresolve_tac (prems@[domainE,range2E]) 1));
-val fieldE = result();
+qed "fieldE";
goalw Rel.thy [field_def] "domain(r) <= field(r)";
by (rtac Un_upper1 1);
-val domain_in_field = result();
+qed "domain_in_field";
goalw Rel.thy [field_def] "range2(r) <= field(r)";
by (rtac Un_upper2 1);
-val range2_in_field = result();
+qed "range2_in_field";
????????????????????????????????????????????????????????????????;
@@ -105,5 +105,5 @@
by (fast_tac ZF_cs 3);
by (fast_tac ZF_cs 2);
by (fast_tac ZF_cs 1);
-val wfI2 = result();
+qed "wfI2";
--- a/ex/set.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/ex/set.ML Mon Nov 21 17:50:34 1994 +0100
@@ -21,7 +21,7 @@
goal Set.thy "~ (? f:: 'a=>'a set. ! S. ? x. f(x) = S)";
(*requires best-first search because it is undirectional*)
by (best_tac (set_cs addSEs [equalityCE]) 1);
-val cantor1 = result();
+qed "cantor1";
(*This form displays the diagonal term*)
goal Set.thy "! f:: 'a=>'a set. ! x. f(x) ~= ?S(f)";
@@ -48,15 +48,15 @@
by (rtac equalityI 1);
by (fast_tac (set_cs addEs [Inv_f_f RS ssubst]) 1);
by (fast_tac (set_cs addEs [Inv_f_f RS ssubst]) 1);
-val inv_image_comp = result();
+qed "inv_image_comp";
val prems = goal Set.thy "f(a) ~: (f``X) ==> a~:X";
by (cfast_tac prems 1);
-val contra_imageI = result();
+qed "contra_imageI";
goal Lfp.thy "(a ~: Compl(X)) = (a:X)";
by (fast_tac set_cs 1);
-val not_Compl = result();
+qed "not_Compl";
(*Lots of backtracking in this proof...*)
val [compl,fg,Xa] = goal Lfp.thy
@@ -64,7 +64,7 @@
by (EVERY1 [rtac (not_Compl RS subst), rtac contra_imageI,
rtac (compl RS subst), rtac (fg RS subst), stac not_Compl,
rtac imageI, rtac Xa]);
-val disj_lemma = result();
+qed "disj_lemma";
goal Lfp.thy "range(%z. if(z:X, f(z), g(z))) = f``X Un g``Compl(X)";
by (rtac equalityI 1);
@@ -77,21 +77,21 @@
by (rtac (excluded_middle RS disjE) 1);
by (EVERY' [rtac (if_P RS ssubst), atac, fast_tac set_cs] 2);
by (EVERY' [rtac (if_not_P RS ssubst), atac, fast_tac set_cs] 1);
-val range_if_then_else = result();
+qed "range_if_then_else";
goal Lfp.thy "a : X Un Compl(X)";
by (fast_tac set_cs 1);
-val X_Un_Compl = result();
+qed "X_Un_Compl";
goalw Lfp.thy [surj_def] "surj(f) = (!a. a : range(f))";
by (fast_tac (set_cs addEs [ssubst]) 1);
-val surj_iff_full_range = result();
+qed "surj_iff_full_range";
val [compl] = goal Lfp.thy
"Compl(f``X) = g``Compl(X) ==> surj(%z. if(z:X, f(z), g(z)))";
by (sstac [surj_iff_full_range, range_if_then_else, compl RS sym] 1);
by (rtac (X_Un_Compl RS allI) 1);
-val surj_if_then_else = result();
+qed "surj_if_then_else";
val [injf,injg,compl,bij] = goal Lfp.thy
"[| inj_onto(f,X); inj_onto(g,Compl(X)); Compl(f``X) = g``Compl(X); \
@@ -107,13 +107,13 @@
by (fast_tac (set_cs addEs [inj_ontoD, sym RS f_eq_gE]) 1);
by (stac expand_if 1);
by (fast_tac (set_cs addEs [inj_ontoD, f_eq_gE]) 1);
-val bij_if_then_else = result();
+qed "bij_if_then_else";
goal Lfp.thy "? X. X = Compl(g``Compl((f:: 'a=>'b)``X))";
by (rtac exI 1);
by (rtac lfp_Tarski 1);
by (REPEAT (ares_tac [monoI, image_mono, Compl_anti_mono] 1));
-val decomposition = result();
+qed "decomposition";
val [injf,injg] = goal Lfp.thy
"[| inj(f:: 'a=>'b); inj(g:: 'b=>'a) |] ==> \
@@ -127,6 +127,6 @@
etac imageE, etac ssubst, rtac rangeI]);
by (EVERY1 [etac ssubst, stac double_complement,
rtac (injg RS inv_image_comp RS sym)]);
-val schroeder_bernstein = result();
+qed "schroeder_bernstein";
writeln"Reached end of file.";
--- a/mono.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/mono.ML Mon Nov 21 17:50:34 1994 +0100
@@ -8,58 +8,58 @@
goal Set.thy "!!A B. A<=B ==> f``A <= f``B";
by (fast_tac set_cs 1);
-val image_mono = result();
+qed "image_mono";
goal Set.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)";
by (fast_tac set_cs 1);
-val Pow_mono = result();
+qed "Pow_mono";
goal Set.thy "!!A B. A<=B ==> Union(A) <= Union(B)";
by (fast_tac set_cs 1);
-val Union_mono = result();
+qed "Union_mono";
goal Set.thy "!!A B. B<=A ==> Inter(A) <= Inter(B)";
by (fast_tac set_cs 1);
-val Inter_anti_mono = result();
+qed "Inter_anti_mono";
val prems = goal Set.thy
"[| A<=B; !!x. x:A ==> f(x)<=g(x) |] ==> \
\ (UN x:A. f(x)) <= (UN x:B. g(x))";
by (fast_tac (set_cs addIs (prems RL [subsetD])) 1);
-val UN_mono = result();
+qed "UN_mono";
val [prem] = goal Set.thy
"[| !!x. f(x)<=g(x) |] ==> (UN x. f(x)) <= (UN x. g(x))";
by (fast_tac (set_cs addIs [prem RS subsetD]) 1);
-val UN1_mono = result();
+qed "UN1_mono";
val prems = goal Set.thy
"[| B<=A; !!x. x:A ==> f(x)<=g(x) |] ==> \
\ (INT x:A. f(x)) <= (INT x:A. g(x))";
by (fast_tac (set_cs addIs (prems RL [subsetD])) 1);
-val INT_anti_mono = result();
+qed "INT_anti_mono";
(*The inclusion is POSITIVE! *)
val [prem] = goal Set.thy
"[| !!x. f(x)<=g(x) |] ==> (INT x. f(x)) <= (INT x. g(x))";
by (fast_tac (set_cs addIs [prem RS subsetD]) 1);
-val INT1_mono = result();
+qed "INT1_mono";
goal Set.thy "!!A B. [| A<=C; B<=D |] ==> A Un B <= C Un D";
by (fast_tac set_cs 1);
-val Un_mono = result();
+qed "Un_mono";
goal Set.thy "!!A B. [| A<=C; B<=D |] ==> A Int B <= C Int D";
by (fast_tac set_cs 1);
-val Int_mono = result();
+qed "Int_mono";
goal Set.thy "!!A::'a set. [| A<=C; D<=B |] ==> A-B <= C-D";
by (fast_tac set_cs 1);
-val Diff_mono = result();
+qed "Diff_mono";
goal Set.thy "!!A B. A<=B ==> Compl(B) <= Compl(A)";
by (fast_tac set_cs 1);
-val Compl_anti_mono = result();
+qed "Compl_anti_mono";
val prems = goal Prod.thy
"[| A<=C; !!x. x:A ==> B<=D |] ==> Sigma(A,%x.B) <= Sigma(C,%x.D)";
@@ -67,7 +67,7 @@
by (fast_tac (set_cs addIs (prems RL [subsetD])
addSIs [SigmaI]
addSEs [SigmaE]) 1);
-val Sigma_mono = result();
+qed "Sigma_mono";
(** Monotonicity of implications. For inductive definitions **)
@@ -76,46 +76,46 @@
by (rtac impI 1);
by (etac subsetD 1);
by (assume_tac 1);
-val in_mono = result();
+qed "in_mono";
goal HOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)";
by (fast_tac HOL_cs 1);
-val conj_mono = result();
+qed "conj_mono";
goal HOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1|P2) --> (Q1|Q2)";
by (fast_tac HOL_cs 1);
-val disj_mono = result();
+qed "disj_mono";
goal HOL.thy "!!P1 P2 Q1 Q2.[| Q1-->P1; P2-->Q2 |] ==> (P1-->P2)-->(Q1-->Q2)";
by (fast_tac HOL_cs 1);
-val imp_mono = result();
+qed "imp_mono";
goal HOL.thy "P-->P";
by (rtac impI 1);
by (assume_tac 1);
-val imp_refl = result();
+qed "imp_refl";
val [PQimp] = goal HOL.thy
"[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))";
by (fast_tac (HOL_cs addIs [PQimp RS mp]) 1);
-val ex_mono = result();
+qed "ex_mono";
val [PQimp] = goal HOL.thy
"[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))";
by (fast_tac (HOL_cs addIs [PQimp RS mp]) 1);
-val all_mono = result();
+qed "all_mono";
val [PQimp] = goal Set.thy
"[| !!x. P(x) --> Q(x) |] ==> Collect(P) <= Collect(Q)";
by (fast_tac (set_cs addIs [PQimp RS mp]) 1);
-val Collect_mono = result();
+qed "Collect_mono";
(*Used in indrule.ML*)
val [subs,PQimp] = goal Set.thy
"[| A<=B; !!x. x:A ==> P(x) --> Q(x) \
\ |] ==> A Int Collect(P) <= B Int Collect(Q)";
by (fast_tac (set_cs addIs [subs RS subsetD, PQimp RS mp]) 1);
-val Int_Collect_mono = result();
+qed "Int_Collect_mono";
(*Used in intr_elim.ML and in individual datatype definitions*)
val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono,
--- a/subset.ML Fri Nov 11 10:35:03 1994 +0100
+++ b/subset.ML Mon Nov 21 17:50:34 1994 +0100
@@ -17,35 +17,35 @@
val prems = goal Set.thy
"B:A ==> B <= Union(A)";
by (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1));
-val Union_upper = result();
+qed "Union_upper";
val [prem] = goal Set.thy
"[| !!X. X:A ==> X<=C |] ==> Union(A) <= C";
br subsetI 1;
by (REPEAT (eresolve_tac [asm_rl, UnionE, prem RS subsetD] 1));
-val Union_least = result();
+qed "Union_least";
(** General union **)
val prems = goal Set.thy
"a:A ==> B(a) <= (UN x:A. B(x))";
by (REPEAT (ares_tac (prems@[UN_I RS subsetI]) 1));
-val UN_upper = result();
+qed "UN_upper";
val [prem] = goal Set.thy
"[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A. B(x)) <= C";
br subsetI 1;
by (REPEAT (eresolve_tac [asm_rl, UN_E, prem RS subsetD] 1));
-val UN_least = result();
+qed "UN_least";
goal Set.thy "B(a) <= (UN x. B(x))";
by (REPEAT (ares_tac [UN1_I RS subsetI] 1));
-val UN1_upper = result();
+qed "UN1_upper";
val [prem] = goal Set.thy "[| !!x. B(x)<=C |] ==> (UN x. B(x)) <= C";
br subsetI 1;
by (REPEAT (eresolve_tac [asm_rl, UN1_E, prem RS subsetD] 1));
-val UN1_least = result();
+qed "UN1_least";
(*** Big Intersection -- greatest lower bound of a set ***)
@@ -53,67 +53,67 @@
val prems = goal Set.thy "B:A ==> Inter(A) <= B";
br subsetI 1;
by (REPEAT (resolve_tac prems 1 ORELSE etac InterD 1));
-val Inter_lower = result();
+qed "Inter_lower";
val [prem] = goal Set.thy
"[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)";
br (InterI RS subsetI) 1;
by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
-val Inter_greatest = result();
+qed "Inter_greatest";
val prems = goal Set.thy "a:A ==> (INT x:A. B(x)) <= B(a)";
br subsetI 1;
by (REPEAT (resolve_tac prems 1 ORELSE etac INT_D 1));
-val INT_lower = result();
+qed "INT_lower";
val [prem] = goal Set.thy
"[| !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A. B(x))";
br (INT_I RS subsetI) 1;
by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
-val INT_greatest = result();
+qed "INT_greatest";
goal Set.thy "(INT x. B(x)) <= B(a)";
br subsetI 1;
by (REPEAT (resolve_tac prems 1 ORELSE etac INT1_D 1));
-val INT1_lower = result();
+qed "INT1_lower";
val [prem] = goal Set.thy
"[| !!x. C<=B(x) |] ==> C <= (INT x. B(x))";
br (INT1_I RS subsetI) 1;
by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
-val INT1_greatest = result();
+qed "INT1_greatest";
(*** Finite Union -- the least upper bound of 2 sets ***)
goal Set.thy "A <= A Un B";
by (REPEAT (ares_tac [subsetI,UnI1] 1));
-val Un_upper1 = result();
+qed "Un_upper1";
goal Set.thy "B <= A Un B";
by (REPEAT (ares_tac [subsetI,UnI2] 1));
-val Un_upper2 = result();
+qed "Un_upper2";
val prems = goal Set.thy "[| A<=C; B<=C |] ==> A Un B <= C";
by (cut_facts_tac prems 1);
by (DEPTH_SOLVE (ares_tac [subsetI] 1
ORELSE eresolve_tac [UnE,subsetD] 1));
-val Un_least = result();
+qed "Un_least";
(*** Finite Intersection -- the greatest lower bound of 2 sets *)
goal Set.thy "A Int B <= A";
by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1));
-val Int_lower1 = result();
+qed "Int_lower1";
goal Set.thy "A Int B <= B";
by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1));
-val Int_lower2 = result();
+qed "Int_lower2";
val prems = goal Set.thy "[| C<=A; C<=B |] ==> C <= A Int B";
by (cut_facts_tac prems 1);
by (REPEAT (ares_tac [subsetI,IntI] 1
ORELSE etac subsetD 1));
-val Int_greatest = result();
+qed "Int_greatest";
(*** Set difference ***)
@@ -126,10 +126,10 @@
by (rtac Un_least 1);
by (rtac (Un_upper1 RS (prem RS monoD)) 1);
by (rtac (Un_upper2 RS (prem RS monoD)) 1);
-val mono_Un = result();
+qed "mono_Un";
val [prem] = goal Set.thy "mono(f) ==> f(A Int B) <= f(A) Int f(B)";
by (rtac Int_greatest 1);
by (rtac (Int_lower1 RS (prem RS monoD)) 1);
by (rtac (Int_lower2 RS (prem RS monoD)) 1);
-val mono_Int = result();
+qed "mono_Int";