author | nipkow |
Fri, 02 Dec 1994 11:43:20 +0100 | |
changeset 194 | b93cc55cb7ab |
parent 171 | 16c4ea954511 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: Substitutions/subst.ML |
2 |
Author: Martin Coen, Cambridge University Computer Laboratory |
|
3 |
Copyright 1993 University of Cambridge |
|
4 |
||
5 |
For subst.thy. |
|
6 |
*) |
|
7 |
||
8 |
open Subst; |
|
9 |
||
10 |
(***********) |
|
11 |
||
12 |
val subst_defs = [subst_def,comp_def,sdom_def]; |
|
13 |
||
14 |
val raw_subst_ss = utlemmas_ss addsimps al_rews; |
|
15 |
||
16 |
local fun mk_thm s = prove_goalw Subst.thy subst_defs s |
|
17 |
(fn _ => [simp_tac raw_subst_ss 1]) |
|
18 |
in val subst_rews = map mk_thm |
|
19 |
["Const(c) <| al = Const(c)", |
|
20 |
"Comb(t,u) <| al = Comb(t <| al, u <| al)", |
|
194
b93cc55cb7ab
small updates because datatype list is now used. In particular Nil -> []
nipkow
parents:
171
diff
changeset
|
21 |
"[] <> bl = bl", |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
22 |
"<a,b>#al <> bl = <a,b <| bl> # (al <> bl)", |
194
b93cc55cb7ab
small updates because datatype list is now used. In particular Nil -> []
nipkow
parents:
171
diff
changeset
|
23 |
"sdom([]) = {}", |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
24 |
"sdom(<a,b>#al) = if(Var(a)=b,sdom(al) Int Compl({a}),sdom(al) Un {a})" |
0 | 25 |
]; |
26 |
(* This rewrite isn't always desired *) |
|
27 |
val Var_subst = mk_thm "Var(x) <| al = assoc(x,Var(x),al)"; |
|
28 |
end; |
|
29 |
||
30 |
val subst_ss = raw_subst_ss addsimps subst_rews; |
|
31 |
||
32 |
(**** Substitutions ****) |
|
33 |
||
194
b93cc55cb7ab
small updates because datatype list is now used. In particular Nil -> []
nipkow
parents:
171
diff
changeset
|
34 |
goal Subst.thy "t <| [] = t"; |
0 | 35 |
by (uterm_ind_tac "t" 1); |
36 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); |
|
171 | 37 |
qed "subst_Nil"; |
0 | 38 |
|
39 |
goal Subst.thy "t <: u --> t <| s <: u <| s"; |
|
40 |
by (uterm_ind_tac "u" 1); |
|
41 |
by (ALLGOALS (asm_simp_tac subst_ss)); |
|
171 | 42 |
val subst_mono = store_thm("subst_mono", result() RS mp); |
0 | 43 |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
44 |
goal Subst.thy "~ (Var(v) <: t) --> t <| <v,t <| s>#s = t <| s"; |
0 | 45 |
by (imp_excluded_middle_tac "t = Var(v)" 1); |
46 |
by (res_inst_tac [("P", |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
47 |
"%x.~x=Var(v) --> ~(Var(v) <: x) --> x <| <v,t<|s>#s=x<|s")] |
0 | 48 |
uterm_induct 2); |
49 |
by (ALLGOALS (simp_tac (subst_ss addsimps [Var_subst]))); |
|
50 |
by (fast_tac HOL_cs 1); |
|
171 | 51 |
val Var_not_occs = store_thm("Var_not_occs", result() RS mp); |
0 | 52 |
|
53 |
goal Subst.thy |
|
54 |
"(t <|r = t <|s) = (! v.v : vars_of(t) --> Var(v) <|r = Var(v) <|s)"; |
|
55 |
by (uterm_ind_tac "t" 1); |
|
56 |
by (REPEAT (etac rev_mp 3)); |
|
57 |
by (ALLGOALS (asm_simp_tac subst_ss)); |
|
58 |
by (ALLGOALS (fast_tac HOL_cs)); |
|
171 | 59 |
qed "agreement"; |
0 | 60 |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
61 |
goal Subst.thy "~ v: vars_of(t) --> t <| <v,u>#s = t <| s"; |
0 | 62 |
by(simp_tac(subst_ss addsimps [agreement,Var_subst] |
63 |
setloop (split_tac [expand_if])) 1); |
|
171 | 64 |
val repl_invariance = store_thm("repl_invariance", result() RS mp); |
0 | 65 |
|
66 |
val asms = goal Subst.thy |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
67 |
"v : vars_of(t) --> w : vars_of(t <| <v,Var(w)>#s)"; |
0 | 68 |
by (uterm_ind_tac "t" 1); |
69 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); |
|
171 | 70 |
val Var_in_subst = store_thm("Var_in_subst", result() RS mp); |
0 | 71 |
|
72 |
(**** Equality between Substitutions ****) |
|
73 |
||
74 |
goalw Subst.thy [subst_eq_def] "r =s= s = (! t.t <| r = t <| s)"; |
|
75 |
by (simp_tac subst_ss 1); |
|
171 | 76 |
qed "subst_eq_iff"; |
0 | 77 |
|
78 |
local fun mk_thm s = prove_goal Subst.thy s |
|
79 |
(fn prems => [cut_facts_tac prems 1, |
|
80 |
REPEAT (etac rev_mp 1), |
|
81 |
simp_tac (subst_ss addsimps [subst_eq_iff]) 1]) |
|
82 |
in |
|
83 |
val subst_refl = mk_thm "r = s ==> r =s= s"; |
|
84 |
val subst_sym = mk_thm "r =s= s ==> s =s= r"; |
|
85 |
val subst_trans = mk_thm "[| q =s= r; r =s= s |] ==> q =s= s"; |
|
86 |
end; |
|
87 |
||
88 |
val eq::prems = goalw Subst.thy [subst_eq_def] |
|
89 |
"[| r =s= s; P(t <| r,u <| r) |] ==> P(t <| s,u <| s)"; |
|
90 |
by (resolve_tac [eq RS spec RS subst] 1); |
|
91 |
by (resolve_tac (prems RL [eq RS spec RS subst]) 1); |
|
171 | 92 |
qed "subst_subst2"; |
0 | 93 |
|
94 |
val ssubst_subst2 = subst_sym RS subst_subst2; |
|
95 |
||
96 |
(**** Composition of Substitutions ****) |
|
97 |
||
194
b93cc55cb7ab
small updates because datatype list is now used. In particular Nil -> []
nipkow
parents:
171
diff
changeset
|
98 |
goal Subst.thy "s <> [] = s"; |
0 | 99 |
by (alist_ind_tac "s" 1); |
100 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [subst_Nil]))); |
|
171 | 101 |
qed "comp_Nil"; |
0 | 102 |
|
103 |
goal Subst.thy "(t <| r <> s) = (t <| r <| s)"; |
|
104 |
by (uterm_ind_tac "t" 1); |
|
105 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst]))); |
|
106 |
by (alist_ind_tac "r" 1); |
|
107 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst,subst_Nil] |
|
108 |
setloop (split_tac [expand_if])))); |
|
171 | 109 |
qed "subst_comp"; |
0 | 110 |
|
111 |
goal Subst.thy "q <> r <> s =s= q <> (r <> s)"; |
|
112 |
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); |
|
171 | 113 |
qed "comp_assoc"; |
0 | 114 |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
115 |
goal Subst.thy "<w,Var(w) <| s>#s =s= s"; |
0 | 116 |
by (rtac (allI RS (subst_eq_iff RS iffD2)) 1); |
117 |
by (uterm_ind_tac "t" 1); |
|
118 |
by (REPEAT (etac rev_mp 3)); |
|
119 |
by (ALLGOALS (simp_tac (subst_ss addsimps[Var_subst] |
|
120 |
setloop (split_tac [expand_if])))); |
|
171 | 121 |
qed "Cons_trivial"; |
0 | 122 |
|
123 |
val [prem] = goal Subst.thy "q <> r =s= s ==> t <| q <| r = t <| s"; |
|
124 |
by (simp_tac (subst_ss addsimps [prem RS (subst_eq_iff RS iffD1), |
|
125 |
subst_comp RS sym]) 1); |
|
171 | 126 |
qed "comp_subst_subst"; |
0 | 127 |
|
128 |
(**** Domain and range of Substitutions ****) |
|
129 |
||
130 |
goal Subst.thy "(v : sdom(s)) = (~ Var(v) <| s = Var(v))"; |
|
131 |
by (alist_ind_tac "s" 1); |
|
132 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_subst] |
|
133 |
setloop (split_tac[expand_if])))); |
|
134 |
by (fast_tac HOL_cs 1); |
|
171 | 135 |
qed "sdom_iff"; |
0 | 136 |
|
137 |
goalw Subst.thy [srange_def] |
|
138 |
"v : srange(s) = (? w.w : sdom(s) & v : vars_of(Var(w) <| s))"; |
|
139 |
by (fast_tac set_cs 1); |
|
171 | 140 |
qed "srange_iff"; |
0 | 141 |
|
142 |
goal Subst.thy "(t <| s = t) = (sdom(s) Int vars_of(t) = {})"; |
|
143 |
by (uterm_ind_tac "t" 1); |
|
144 |
by (REPEAT (etac rev_mp 3)); |
|
145 |
by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,Var_subst]))); |
|
146 |
by (ALLGOALS (fast_tac set_cs)); |
|
171 | 147 |
qed "invariance"; |
0 | 148 |
|
149 |
goal Subst.thy "v : sdom(s) --> ~v : srange(s) --> ~v : vars_of(t <| s)"; |
|
150 |
by (uterm_ind_tac "t" 1); |
|
151 |
by (imp_excluded_middle_tac "x : sdom(s)" 1); |
|
152 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [sdom_iff,srange_iff]))); |
|
153 |
by (ALLGOALS (fast_tac set_cs)); |
|
171 | 154 |
val Var_elim = store_thm("Var_elim", result() RS mp RS mp); |
0 | 155 |
|
156 |
val asms = goal Subst.thy |
|
157 |
"[| v : sdom(s); v : vars_of(t <| s) |] ==> v : srange(s)"; |
|
158 |
by (REPEAT (ares_tac (asms @ [Var_elim RS swap RS classical]) 1)); |
|
171 | 159 |
qed "Var_elim2"; |
0 | 160 |
|
161 |
goal Subst.thy "v : vars_of(t <| s) --> v : srange(s) | v : vars_of(t)"; |
|
162 |
by (uterm_ind_tac "t" 1); |
|
163 |
by (REPEAT_SOME (etac rev_mp )); |
|
164 |
by (ALLGOALS (simp_tac (subst_ss addsimps [sdom_iff,srange_iff]))); |
|
165 |
by (REPEAT (step_tac (set_cs addIs [vars_var_iff RS iffD1 RS sym]) 1)); |
|
166 |
by (etac notE 1); |
|
167 |
by (etac subst 1); |
|
168 |
by (ALLGOALS (fast_tac set_cs)); |
|
171 | 169 |
val Var_intro = store_thm("Var_intro", result() RS mp); |
0 | 170 |
|
171 |
goal Subst.thy |
|
172 |
"v : srange(s) --> (? w.w : sdom(s) & v : vars_of(Var(w) <| s))"; |
|
173 |
by (simp_tac (subst_ss addsimps [srange_iff]) 1); |
|
171 | 174 |
val srangeE = store_thm("srangeE", make_elim (result() RS mp)); |
0 | 175 |
|
176 |
val asms = goal Subst.thy |
|
177 |
"sdom(s) Int srange(s) = {} = (! t.sdom(s) Int vars_of(t <| s) = {})"; |
|
178 |
by (simp_tac subst_ss 1); |
|
179 |
by (fast_tac (set_cs addIs [Var_elim2] addEs [srangeE]) 1); |
|
171 | 180 |
qed "dom_range_disjoint"; |
0 | 181 |
|
182 |
val asms = goal Subst.thy "~ u <| s = u --> (? x.x : sdom(s))"; |
|
183 |
by (simp_tac (subst_ss addsimps [invariance]) 1); |
|
184 |
by (fast_tac set_cs 1); |
|
171 | 185 |
val subst_not_empty = store_thm("subst_not_empty", result() RS mp); |