author | paulson |
Thu, 10 Oct 1996 11:59:01 +0200 | |
changeset 2088 | e814e03bbbb2 |
parent 1673 | d22110ddd0af |
child 3192 | a75558a4ed37 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/Subst/unifier.ML |
1266 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Martin Coen, Cambridge University Computer Laboratory |
968 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
For unifier.thy. |
|
7 |
Properties of unifiers. |
|
8 |
Cases for partial correctness of algorithm and conditions for termination. |
|
9 |
*) |
|
10 |
||
11 |
open Unifier; |
|
12 |
||
13 |
val unify_defs = |
|
14 |
[Idem_def,Unifier_def,MoreGeneral_def,MGUnifier_def,MGIUnifier_def]; |
|
15 |
||
16 |
(**** Unifiers ****) |
|
17 |
||
18 |
goalw Unifier.thy [Unifier_def] "Unifier s t u = (t <| s = u <| s)"; |
|
19 |
by (rtac refl 1); |
|
20 |
qed "Unifier_iff"; |
|
21 |
||
22 |
goal Unifier.thy |
|
23 |
"Unifier s (Comb t u) (Comb v w) --> Unifier s t v & Unifier s u w"; |
|
24 |
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); |
|
25 |
val Unifier_Comb = store_thm("Unifier_Comb", result() RS mp RS conjE); |
|
26 |
||
27 |
goal Unifier.thy |
|
28 |
"~v : vars_of(t) --> ~v : vars_of(u) -->Unifier s t u --> \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
29 |
\ Unifier ((v,r)#s) t u"; |
968 | 30 |
by (simp_tac (subst_ss addsimps [Unifier_iff,repl_invariance]) 1); |
31 |
val Cons_Unifier = store_thm("Cons_Unifier", result() RS mp RS mp RS mp); |
|
32 |
||
33 |
(**** Most General Unifiers ****) |
|
34 |
||
35 |
goalw Unifier.thy [MoreGeneral_def] "r >> s = (EX q. s =s= r <> q)"; |
|
36 |
by (rtac refl 1); |
|
37 |
qed "MoreGen_iff"; |
|
38 |
||
39 |
goal Unifier.thy "[] >> s"; |
|
40 |
by (simp_tac (subst_ss addsimps [MoreGen_iff]) 1); |
|
41 |
by (fast_tac (set_cs addIs [refl RS subst_refl]) 1); |
|
42 |
qed "MoreGen_Nil"; |
|
43 |
||
44 |
goalw Unifier.thy unify_defs |
|
45 |
"MGUnifier s t u = (ALL r.Unifier r t u = s >> r)"; |
|
46 |
by (REPEAT (ares_tac [iffI,allI] 1 ORELSE |
|
47 |
eresolve_tac [conjE,allE,mp,exE,ssubst_subst2] 1)); |
|
48 |
by (asm_simp_tac (subst_ss addsimps [subst_comp]) 1); |
|
49 |
by (fast_tac (set_cs addIs [comp_Nil RS sym RS subst_refl]) 1); |
|
50 |
qed "MGU_iff"; |
|
51 |
||
52 |
val [prem] = goal Unifier.thy |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
53 |
"~ Var(v) <: t ==> MGUnifier [(v,t)] (Var v) t"; |
968 | 54 |
by (simp_tac (subst_ss addsimps [MGU_iff,MoreGen_iff,Unifier_iff]) 1); |
55 |
by (REPEAT_SOME (step_tac set_cs)); |
|
56 |
by (etac subst 1); |
|
57 |
by (etac ssubst_subst2 2); |
|
58 |
by (rtac (Cons_trivial RS subst_sym) 1); |
|
59 |
by (simp_tac (subst_ss addsimps [prem RS Var_not_occs,Var_subst]) 1); |
|
60 |
qed "MGUnifier_Var"; |
|
61 |
||
62 |
(**** Most General Idempotent Unifiers ****) |
|
63 |
||
64 |
goal Unifier.thy "r <> r =s= r --> s =s= r <> q --> r <> s =s= s"; |
|
65 |
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); |
|
66 |
val MGIU_iff_lemma = store_thm("MGIU_iff_lemma", result() RS mp RS mp); |
|
67 |
||
68 |
goalw Unifier.thy unify_defs |
|
69 |
"MGIUnifier s t u = \ |
|
70 |
\ (Idem(s) & Unifier s t u & (ALL r.Unifier r t u --> s<>r=s=r))"; |
|
71 |
by (fast_tac (set_cs addEs [subst_sym,MGIU_iff_lemma]) 1); |
|
72 |
qed "MGIU_iff"; |
|
73 |
||
74 |
(**** Idempotence ****) |
|
75 |
||
76 |
goalw Unifier.thy unify_defs "Idem(s) = (s <> s =s= s)"; |
|
77 |
by (rtac refl 1); |
|
78 |
qed "raw_Idem_iff"; |
|
79 |
||
80 |
goal Unifier.thy "Idem(s) = (sdom(s) Int srange(s) = {})"; |
|
81 |
by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp, |
|
1673
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1465
diff
changeset
|
82 |
invariance,dom_range_disjoint] |
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
oheimb
parents:
1465
diff
changeset
|
83 |
delsimps (empty_iff::mem_simps)) 1); |
968 | 84 |
qed "Idem_iff"; |
85 |
||
86 |
goal Unifier.thy "Idem([])"; |
|
87 |
by (simp_tac (subst_ss addsimps [raw_Idem_iff,refl RS subst_refl]) 1); |
|
88 |
qed "Idem_Nil"; |
|
89 |
||
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
90 |
goal Unifier.thy "~ (Var(v) <: t) --> Idem([(v,t)])"; |
968 | 91 |
by (simp_tac (subst_ss addsimps [Var_subst,vars_iff_occseq,Idem_iff,srange_iff] |
92 |
setloop (split_tac [expand_if])) 1); |
|
93 |
by (fast_tac set_cs 1); |
|
94 |
val Var_Idem = store_thm("Var_Idem", result() RS mp); |
|
95 |
||
96 |
val [prem] = goalw Unifier.thy [Idem_def] |
|
97 |
"Idem(r) ==> Unifier s (t <| r) (u <| r) --> Unifier (r <> s) (t <| r) (u <| r)"; |
|
98 |
by (simp_tac (subst_ss addsimps |
|
1465 | 99 |
[Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1); |
968 | 100 |
val Unifier_Idem_subst = store_thm("Unifier_Idem_subst", result() RS mp); |
101 |
||
102 |
val [prem] = goal Unifier.thy |
|
103 |
"r <> s =s= s ==> Unifier s t u --> Unifier s (t <| r) (u <| r)"; |
|
104 |
by (simp_tac (subst_ss addsimps |
|
1465 | 105 |
[Unifier_iff,subst_comp,prem RS comp_subst_subst]) 1); |
968 | 106 |
val Unifier_comp_subst = store_thm("Unifier_comp_subst", result() RS mp); |
107 |
||
108 |
(*** The domain of a MGIU is a subset of the variables in the terms ***) |
|
109 |
(*** NB this and one for range are only needed for termination ***) |
|
110 |
||
111 |
val [prem] = goal Unifier.thy |
|
112 |
"~ vars_of(Var(x) <| r) = vars_of(Var(x) <| s) ==> ~r =s= s"; |
|
113 |
by (rtac (prem RS contrapos) 1); |
|
114 |
by (fast_tac (set_cs addEs [subst_subst2]) 1); |
|
115 |
qed "lemma_lemma"; |
|
116 |
||
117 |
val prems = goal Unifier.thy |
|
118 |
"x : sdom(s) --> ~x : srange(s) --> \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
119 |
\ ~vars_of(Var(x) <| s<> (x,Var(x))#s) = \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
120 |
\ vars_of(Var(x) <| (x,Var(x))#s)"; |
968 | 121 |
by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); |
122 |
by (REPEAT (resolve_tac [impI,disjI2] 1)); |
|
123 |
by(res_inst_tac [("x","x")] exI 1); |
|
1465 | 124 |
by (rtac conjI 1); |
968 | 125 |
by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1); |
126 |
by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1); |
|
127 |
val MGIU_sdom_lemma = store_thm("MGIU_sdom_lemma", result() RS mp RS mp RS lemma_lemma RS notE); |
|
128 |
||
129 |
goal Unifier.thy "MGIUnifier s t u --> sdom(s) <= vars_of(t) Un vars_of(u)"; |
|
130 |
by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1); |
|
131 |
by (asm_simp_tac (subst_ss addsimps [MGIU_iff,Idem_iff,subset_iff]) 1); |
|
132 |
by (safe_tac set_cs); |
|
133 |
by (eresolve_tac ([spec] RL [impE]) 1); |
|
134 |
by (rtac Cons_Unifier 1); |
|
135 |
by (ALLGOALS (fast_tac (set_cs addIs [Cons_Unifier,MGIU_sdom_lemma]))); |
|
136 |
val MGIU_sdom = store_thm("MGIU_sdom", result() RS mp); |
|
137 |
||
2088 | 138 |
|
139 |
(** COULD REPLACE THE TWO THEOREMS ABOVE BY THE FOLLOWING, IN THE PRESENCE |
|
140 |
OF DEMORGAN LAWS. DON'T KNOW WHAT TO DO WITH THE SIMILAR PROOF BELOW. |
|
141 |
val prems = goal Unifier.thy |
|
142 |
"x : sdom(s) --> ~x : srange(s) --> \ |
|
143 |
\ ~vars_of(Var(x) <| s<> (x,Var(x))#s) = \ |
|
144 |
\ vars_of(Var(x) <| (x,Var(x))#s)"; |
|
145 |
by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); |
|
146 |
by (REPEAT (resolve_tac [impI,disjI2] 1)); |
|
147 |
by(res_inst_tac [("x","x")] exI 1); |
|
148 |
by (rtac conjI 1); |
|
149 |
by (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp,repl_invariance]) 1); |
|
150 |
by (asm_simp_tac (subst_ss addsimps [Var_subst]) 1); |
|
151 |
val MGIU_sdom_lemma = result() RS mp RS mp RS lemma_lemma RSN (2, rev_notE); |
|
152 |
||
153 |
goal Unifier.thy "MGIUnifier s t u --> sdom(s) <= vars_of(t) Un vars_of(u)"; |
|
154 |
by (asm_simp_tac (subst_ss addsimps [MGIU_iff,Idem_iff,subset_iff]) 1); |
|
155 |
by (deepen_tac (set_cs addIs [Cons_Unifier] addEs [MGIU_sdom_lemma]) 3 1); |
|
156 |
val MGIU_sdom = store_thm("MGIU_sdom", result() RS mp); |
|
157 |
**) |
|
158 |
||
159 |
||
968 | 160 |
(*** The range of a MGIU is a subset of the variables in the terms ***) |
161 |
||
162 |
val prems = goal HOL.thy "P = Q ==> (~P) = (~Q)"; |
|
1266 | 163 |
by (simp_tac (subst_ss addsimps prems) 1); |
968 | 164 |
qed "not_cong"; |
165 |
||
166 |
val prems = goal Unifier.thy |
|
167 |
"~w=x --> x : vars_of(Var(w) <| s) --> w : sdom(s) --> ~w : srange(s) --> \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
168 |
\ ~vars_of(Var(w) <| s<> (x,Var(w))#s) = \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
169 |
\ vars_of(Var(w) <| (x,Var(w))#s)"; |
968 | 170 |
by (simp_tac (subst_ss addsimps [not_equal_iff]) 1); |
171 |
by (REPEAT (resolve_tac [impI,disjI1] 1)); |
|
172 |
by(res_inst_tac [("x","w")] exI 1); |
|
173 |
by (ALLGOALS (asm_simp_tac (subst_ss addsimps [Var_elim,subst_comp, |
|
174 |
vars_var_iff RS not_cong RS iffD2 RS repl_invariance]) )); |
|
175 |
by (fast_tac (set_cs addIs [Var_in_subst]) 1); |
|
176 |
val MGIU_srange_lemma = store_thm("MGIU_srange_lemma", result() RS mp RS mp RS mp RS mp RS lemma_lemma RS notE); |
|
177 |
||
178 |
goal Unifier.thy "MGIUnifier s t u --> srange(s) <= vars_of(t) Un vars_of(u)"; |
|
179 |
by (subgoal_tac "! P Q.(P | Q) = (~( ~P & ~Q))" 1); |
|
180 |
by (asm_simp_tac (subst_ss addsimps [MGIU_iff,srange_iff,subset_iff]) 1); |
|
181 |
by (simp_tac (subst_ss addsimps [Idem_iff]) 1); |
|
182 |
by (safe_tac set_cs); |
|
183 |
by (eresolve_tac ([spec] RL [impE]) 1); |
|
184 |
by (rtac Cons_Unifier 1); |
|
185 |
by (imp_excluded_middle_tac "w=ta" 4); |
|
186 |
by (fast_tac (set_cs addEs [MGIU_srange_lemma]) 5); |
|
187 |
by (ALLGOALS (fast_tac (set_cs addIs [Var_elim2]))); |
|
188 |
val MGIU_srange = store_thm("MGIU_srange", result() RS mp); |
|
189 |
||
190 |
(*************** Correctness of a simple unification algorithm ***************) |
|
191 |
(* *) |
|
192 |
(* fun unify Const(m) Const(n) = if m=n then Nil else Fail *) |
|
193 |
(* | unify Const(m) _ = Fail *) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
194 |
(* | unify Var(v) t = if Var(v)<:t then Fail else (v,t)#Nil *) |
968 | 195 |
(* | unify Comb(t,u) Const(n) = Fail *) |
196 |
(* | unify Comb(t,u) Var(v) = if Var(v) <: Comb(t,u) then Fail *) |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
197 |
(* else (v,Comb(t,u)#Nil *) |
968 | 198 |
(* | unify Comb(t,u) Comb(v,w) = let s = unify t v *) |
199 |
(* in if s=Fail then Fail *) |
|
200 |
(* else unify (u<|s) (w<|s); *) |
|
201 |
||
202 |
(**** Cases for the partial correctness of the algorithm ****) |
|
203 |
||
204 |
goalw Unifier.thy unify_defs "MGIUnifier s t u = MGIUnifier s u t"; |
|
205 |
by (safe_tac (HOL_cs addSEs ([sym]@([spec] RL [mp])))); |
|
206 |
qed "Unify_comm"; |
|
207 |
||
208 |
goal Unifier.thy "MGIUnifier [] (Const n) (Const n)"; |
|
209 |
by (simp_tac (subst_ss addsimps |
|
1465 | 210 |
[MGIU_iff,MGU_iff,Unifier_iff,subst_eq_iff,Idem_Nil]) 1); |
968 | 211 |
qed "Unify1"; |
212 |
||
213 |
goal Unifier.thy "~m=n --> (ALL l.~Unifier l (Const m) (Const n))"; |
|
214 |
by (simp_tac (subst_ss addsimps[Unifier_iff]) 1); |
|
215 |
val Unify2 = store_thm("Unify2", result() RS mp); |
|
216 |
||
217 |
val [prem] = goalw Unifier.thy [MGIUnifier_def] |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
968
diff
changeset
|
218 |
"~Var(v) <: t ==> MGIUnifier [(v,t)] (Var v) t"; |
968 | 219 |
by (fast_tac (HOL_cs addSIs [prem RS MGUnifier_Var,prem RS Var_Idem]) 1); |
220 |
qed "Unify3"; |
|
221 |
||
222 |
val [prem] = goal Unifier.thy "Var(v) <: t ==> (ALL l.~Unifier l (Var v) t)"; |
|
223 |
by (simp_tac (subst_ss addsimps |
|
1465 | 224 |
[Unifier_iff,prem RS subst_mono RS occs_irrefl2]) 1); |
968 | 225 |
qed "Unify4"; |
226 |
||
227 |
goal Unifier.thy "ALL l.~Unifier l (Const m) (Comb t u)"; |
|
228 |
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); |
|
229 |
qed "Unify5"; |
|
230 |
||
231 |
goal Unifier.thy |
|
232 |
"(ALL l.~Unifier l t v) --> (ALL l.~Unifier l (Comb t u) (Comb v w))"; |
|
233 |
by (simp_tac (subst_ss addsimps [Unifier_iff]) 1); |
|
234 |
val Unify6 = store_thm("Unify6", result() RS mp); |
|
235 |
||
236 |
goal Unifier.thy "MGIUnifier s t v --> (ALL l.~Unifier l (u <| s) (w <| s)) \ |
|
237 |
\ --> (ALL l.~Unifier l (Comb t u) (Comb v w))"; |
|
238 |
by (simp_tac (subst_ss addsimps [MGIU_iff]) 1); |
|
239 |
by (fast_tac (set_cs addIs [Unifier_comp_subst] addSEs [Unifier_Comb]) 1); |
|
240 |
val Unify7 = store_thm("Unify7", result() RS mp RS mp); |
|
241 |
||
242 |
val [p1,p2,p3] = goal Unifier.thy |
|
243 |
"[| Idem(r); Unifier s (t <| r) (u <| r); \ |
|
244 |
\ (! q.Unifier q (t <| r) (u <| r) --> s <> q =s= q) |] ==> \ |
|
245 |
\ Idem(r <> s)"; |
|
246 |
by (cut_facts_tac [p1, |
|
1465 | 247 |
p2 RS (p1 RS Unifier_Idem_subst RS (p3 RS spec RS mp))] 1); |
968 | 248 |
by (REPEAT_SOME (etac rev_mp)); |
249 |
by (simp_tac (subst_ss addsimps [raw_Idem_iff,subst_eq_iff,subst_comp]) 1); |
|
250 |
qed "Unify8_lemma1"; |
|
251 |
||
252 |
val [p1,p2,p3,p4] = goal Unifier.thy |
|
253 |
"[| Unifier q t v; Unifier q u w; (! q.Unifier q t v --> r <> q =s= q); \ |
|
254 |
\ (! q.Unifier q (u <| r) (w <| r) --> s <> q =s= q) |] ==> \ |
|
255 |
\ r <> s <> q =s= q"; |
|
256 |
val pp = p1 RS (p3 RS spec RS mp); |
|
257 |
by (cut_facts_tac [pp, |
|
1465 | 258 |
p2 RS (pp RS Unifier_comp_subst) RS (p4 RS spec RS mp)] 1); |
968 | 259 |
by (REPEAT_SOME (etac rev_mp)); |
260 |
by (simp_tac (subst_ss addsimps [subst_eq_iff,subst_comp]) 1); |
|
261 |
qed "Unify8_lemma2"; |
|
262 |
||
263 |
goal Unifier.thy "MGIUnifier r t v --> MGIUnifier s (u <| r) (w <| r) --> \ |
|
264 |
\ MGIUnifier (r <> s) (Comb t u) (Comb v w)"; |
|
265 |
by (simp_tac (subst_ss addsimps [MGIU_iff,subst_comp,comp_assoc]) 1); |
|
266 |
by (safe_tac HOL_cs); |
|
267 |
by (REPEAT (etac rev_mp 2)); |
|
268 |
by (simp_tac (subst_ss addsimps |
|
1465 | 269 |
[Unifier_iff,MGIU_iff,subst_comp,comp_assoc]) 2); |
968 | 270 |
by (ALLGOALS (fast_tac (set_cs addEs |
1465 | 271 |
[Unifier_Comb,Unify8_lemma1,Unify8_lemma2]))); |
968 | 272 |
qed "Unify8"; |
273 |
||
274 |
||
275 |
(********************** Termination of the algorithm *************************) |
|
276 |
(* *) |
|
277 |
(*UWFD is a well-founded relation that orders the 2 recursive calls in unify *) |
|
278 |
(* NB well-foundedness of UWFD isn't proved *) |
|
279 |
||
280 |
||
281 |
goalw Unifier.thy [UWFD_def] "UWFD t t' (Comb t u) (Comb t' u')"; |
|
282 |
by (simp_tac subst_ss 1); |
|
283 |
by (fast_tac set_cs 1); |
|
284 |
qed "UnifyWFD1"; |
|
285 |
||
286 |
val [prem] = goal Unifier.thy |
|
287 |
"MGIUnifier s t t' ==> vars_of(u <| s) Un vars_of(u' <| s) <= \ |
|
288 |
\ vars_of (Comb t u) Un vars_of (Comb t' u')"; |
|
289 |
by (subgoal_tac "vars_of(u <| s) Un vars_of(u' <| s) <= \ |
|
290 |
\ srange(s) Un vars_of(u) Un srange(s) Un vars_of(u')" 1); |
|
291 |
by (etac subset_trans 1); |
|
292 |
by (ALLGOALS (simp_tac (subst_ss addsimps [Var_intro,subset_iff]))); |
|
293 |
by (ALLGOALS (fast_tac (set_cs addDs |
|
1465 | 294 |
[Var_intro,prem RS MGIU_srange RS subsetD]))); |
968 | 295 |
qed "UWFD2_lemma1"; |
296 |
||
297 |
val [major,minor] = goal Unifier.thy |
|
298 |
"[| MGIUnifier s t t'; ~ u <| s = u |] ==> \ |
|
299 |
\ ~ vars_of(u <| s) Un vars_of(u' <| s) = \ |
|
300 |
\ (vars_of(t) Un vars_of(u)) Un (vars_of(t') Un vars_of(u'))"; |
|
301 |
by (cut_facts_tac |
|
302 |
[major RS (MGIU_iff RS iffD1) RS conjunct1 RS (Idem_iff RS iffD1)] 1); |
|
303 |
by (rtac (minor RS subst_not_empty RS exE) 1); |
|
304 |
by (rtac (make_elim ((major RS MGIU_sdom) RS subsetD)) 1 THEN assume_tac 1); |
|
305 |
by (rtac (disjI2 RS (not_equal_iff RS iffD2)) 1); |
|
306 |
by (REPEAT (etac rev_mp 1)); |
|
307 |
by (asm_simp_tac subst_ss 1); |
|
308 |
by (fast_tac (set_cs addIs [Var_elim2]) 1); |
|
309 |
qed "UWFD2_lemma2"; |
|
310 |
||
311 |
val [prem] = goalw Unifier.thy [UWFD_def] |
|
312 |
"MGIUnifier s t t' ==> UWFD (u <| s) (u' <| s) (Comb t u) (Comb t' u')"; |
|
313 |
by (cut_facts_tac |
|
314 |
[prem RS UWFD2_lemma1 RS (subseteq_iff_subset_eq RS iffD1)] 1); |
|
315 |
by (imp_excluded_middle_tac "u <| s = u" 1); |
|
1266 | 316 |
by (simp_tac (subst_ss delsimps (ssubset_iff :: utlemmas_rews) |
317 |
addsimps [occs_Comb2]) 1); |
|
968 | 318 |
by (rtac impI 1 THEN etac subst 1 THEN assume_tac 1); |
319 |
by (rtac impI 1); |
|
320 |
by (rtac (conjI RS (ssubset_iff RS iffD2) RS disjI1) 1); |
|
1266 | 321 |
by (asm_simp_tac (subst_ss delsimps (ssubset_iff :: utlemmas_rews) addsimps [subseteq_iff_subset_eq]) 1); |
968 | 322 |
by (asm_simp_tac subst_ss 1); |
323 |
by (fast_tac (set_cs addDs [prem RS UWFD2_lemma2]) 1); |
|
324 |
qed "UnifyWFD2"; |