diff -r 69d815b0e1eb -r 21291189b51e Subst/subst.ML
--- a/Subst/subst.ML Thu Feb 24 14:45:57 1994 +0100
+++ b/Subst/subst.ML Wed Mar 02 12:26:55 1994 +0100
@@ -19,9 +19,9 @@
["Const(c) <| al = Const(c)",
"Comb(t,u) <| al = Comb(t <| al, u <| al)",
"Nil <> bl = bl",
- "Cons(,al) <> bl = Cons(, al <> bl)",
+ "#al <> bl = # (al <> bl)",
"sdom(Nil) = {}",
- "sdom(Cons(,al)) = if(Var(a)=b,sdom(al) Int Compl({a}),sdom(al) Un {a})"
+ "sdom(#al) = if(Var(a)=b,sdom(al) Int Compl({a}),sdom(al) Un {a})"
];
(* This rewrite isn't always desired *)
val Var_subst = mk_thm "Var(x) <| al = assoc(x,Var(x),al)";
@@ -41,10 +41,10 @@
by (ALLGOALS (asm_simp_tac subst_ss));
val subst_mono = result() RS mp;
-goal Subst.thy "~ (Var(v) <: t) --> t <| Cons(,s) = t <| s";
+goal Subst.thy "~ (Var(v) <: t) --> t <| #s = t <| s";
by (imp_excluded_middle_tac "t = Var(v)" 1);
by (res_inst_tac [("P",
- "%x.~x=Var(v) --> ~(Var(v) <: x) --> x <| Cons(,s)=x<|s")]
+ "%x.~x=Var(v) --> ~(Var(v) <: x) --> x <| #s=x<|s")]
uterm_induct 2);
by (ALLGOALS (simp_tac (subst_ss addsimps [Var_subst])));
by (fast_tac HOL_cs 1);
@@ -58,13 +58,13 @@
by (ALLGOALS (fast_tac HOL_cs));
val agreement = result();
-goal Subst.thy "~ v: vars_of(t) --> t <| Cons(,s) = t <| s";
+goal Subst.thy "~ v: vars_of(t) --> t <| #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 asms = goal Subst.thy
- "v : vars_of(t) --> w : vars_of(t <| Cons(,s))";
+ "v : vars_of(t) --> w : vars_of(t <| #s)";
by (uterm_ind_tac "t" 1);
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst])));
val Var_in_subst = result() RS mp;
@@ -112,7 +112,7 @@
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1);
val comp_assoc = result();
-goal Subst.thy "Cons(,s) =s= s";
+goal Subst.thy "#s =s= s";
by (rtac (allI RS (subst_eq_iff RS iffD2)) 1);
by (uterm_ind_tac "t" 1);
by (REPEAT (etac rev_mp 3));