--- 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";