nipkow@4
|
1 |
(* Title: Provers/splitter
|
nipkow@4
|
2 |
ID: $Id$
|
nipkow@4
|
3 |
Author: Tobias Nipkow
|
nipkow@1030
|
4 |
Copyright 1995 TU Munich
|
nipkow@4
|
5 |
|
nipkow@4
|
6 |
Generic case-splitter, suitable for most logics.
|
nipkow@4
|
7 |
|
clasohm@0
|
8 |
Use:
|
clasohm@0
|
9 |
|
clasohm@0
|
10 |
val split_tac = mk_case_split_tac iffD;
|
clasohm@0
|
11 |
|
clasohm@0
|
12 |
by(case_split_tac splits i);
|
clasohm@0
|
13 |
|
clasohm@0
|
14 |
where splits = [P(elim(...)) == rhs, ...]
|
clasohm@0
|
15 |
iffD = [| P <-> Q; Q |] ==> P (* is called iffD2 in HOL *)
|
clasohm@0
|
16 |
|
clasohm@0
|
17 |
*)
|
clasohm@0
|
18 |
|
berghofe@1721
|
19 |
local
|
berghofe@1721
|
20 |
|
berghofe@1721
|
21 |
fun mk_case_split_tac_2 iffD order =
|
clasohm@0
|
22 |
let
|
clasohm@0
|
23 |
|
berghofe@1686
|
24 |
|
berghofe@1686
|
25 |
(************************************************************
|
berghofe@1686
|
26 |
Create lift-theorem "trlift" :
|
berghofe@1686
|
27 |
|
berghofe@1686
|
28 |
[| !! x. Q(x)==R(x) ; P(R) == C |] ==> P(Q)==C
|
berghofe@1686
|
29 |
|
berghofe@1686
|
30 |
*************************************************************)
|
berghofe@1686
|
31 |
|
nipkow@943
|
32 |
val lift =
|
nipkow@943
|
33 |
let val ct = read_cterm (#sign(rep_thm iffD))
|
nipkow@943
|
34 |
("[| !!x::'b::logic. Q(x) == R(x) |] ==> \
|
wenzelm@3835
|
35 |
\P(%x. Q(x)) == P(%x. R(x))::'a::logic",propT)
|
nipkow@943
|
36 |
in prove_goalw_cterm [] ct
|
nipkow@943
|
37 |
(fn [prem] => [rewtac prem, rtac reflexive_thm 1])
|
nipkow@943
|
38 |
end;
|
nipkow@4
|
39 |
|
clasohm@0
|
40 |
val trlift = lift RS transitive_thm;
|
clasohm@0
|
41 |
val _ $ (Var(P,PT)$_) $ _ = concl_of trlift;
|
clasohm@0
|
42 |
|
clasohm@0
|
43 |
|
berghofe@1686
|
44 |
(************************************************************************
|
berghofe@1686
|
45 |
Set up term for instantiation of P in the lift-theorem
|
berghofe@1686
|
46 |
|
berghofe@1686
|
47 |
Ts : types of parameters (i.e. variables bound by meta-quantifiers)
|
berghofe@1686
|
48 |
t : lefthand side of meta-equality in subgoal
|
berghofe@1686
|
49 |
the lift theorem is applied to (see select)
|
berghofe@1686
|
50 |
pos : "path" leading to abstraction, coded as a list
|
berghofe@1686
|
51 |
T : type of body of P(...)
|
berghofe@1686
|
52 |
maxi : maximum index of Vars
|
berghofe@1686
|
53 |
*************************************************************************)
|
berghofe@1686
|
54 |
|
nipkow@1030
|
55 |
fun mk_cntxt Ts t pos T maxi =
|
nipkow@1030
|
56 |
let fun var (t,i) = Var(("X",i),type_of1(Ts,t));
|
nipkow@1030
|
57 |
fun down [] t i = Bound 0
|
nipkow@1030
|
58 |
| down (p::ps) t i =
|
nipkow@1030
|
59 |
let val (h,ts) = strip_comb t
|
paulson@2266
|
60 |
val v1 = ListPair.map var (take(p,ts), i upto (i+p-1))
|
nipkow@1030
|
61 |
val u::us = drop(p,ts)
|
paulson@2266
|
62 |
val v2 = ListPair.map var (us, (i+p) upto (i+length(ts)-2))
|
nipkow@1030
|
63 |
in list_comb(h,v1@[down ps u (i+length ts)]@v2) end;
|
nipkow@1030
|
64 |
in Abs("", T, down (rev pos) t maxi) end;
|
nipkow@1030
|
65 |
|
berghofe@1686
|
66 |
|
berghofe@1686
|
67 |
(************************************************************************
|
berghofe@1686
|
68 |
Set up term for instantiation of P in the split-theorem
|
berghofe@1686
|
69 |
P(...) == rhs
|
berghofe@1686
|
70 |
|
berghofe@1686
|
71 |
Ts : types of parameters (i.e. variables bound by meta-quantifiers)
|
berghofe@1686
|
72 |
t : lefthand side of meta-equality in subgoal
|
berghofe@1686
|
73 |
the split theorem is applied to (see select)
|
berghofe@1686
|
74 |
T : type of body of P(...)
|
berghofe@1721
|
75 |
tt : the term Const(..,..) $ ...
|
berghofe@1686
|
76 |
maxi : maximum index of Vars
|
berghofe@1686
|
77 |
|
berghofe@1721
|
78 |
lev : abstraction level
|
berghofe@1686
|
79 |
*************************************************************************)
|
berghofe@1686
|
80 |
|
berghofe@1721
|
81 |
fun mk_cntxt_splitthm Ts t tt T maxi =
|
berghofe@1721
|
82 |
let fun down lev (Abs(v,T2,t)) = Abs(v,T2,down (lev+1) t)
|
berghofe@1721
|
83 |
| down lev (Bound i) = if i >= lev
|
berghofe@1721
|
84 |
then Var(("X",maxi+i-lev),nth_elem(i-lev,Ts))
|
berghofe@1721
|
85 |
else Bound i
|
berghofe@1721
|
86 |
| down lev t =
|
berghofe@1721
|
87 |
let val (h,ts) = strip_comb t
|
berghofe@1721
|
88 |
val h2 = (case h of Bound _ => down lev h | _ => h)
|
berghofe@1721
|
89 |
in if incr_bv(lev,0,tt)=t
|
berghofe@1721
|
90 |
then
|
berghofe@1721
|
91 |
Bound (lev)
|
berghofe@1721
|
92 |
else
|
berghofe@1721
|
93 |
list_comb(h2,map (down lev) ts)
|
berghofe@1721
|
94 |
end;
|
berghofe@1721
|
95 |
in Abs("",T,down 0 t) end;
|
berghofe@1686
|
96 |
|
berghofe@1686
|
97 |
|
berghofe@1686
|
98 |
(* add all loose bound variables in t to list is *)
|
nipkow@1030
|
99 |
fun add_lbnos(is,t) = add_loose_bnos(t,0,is);
|
nipkow@1030
|
100 |
|
nipkow@1064
|
101 |
(* check if the innermost quantifier that needs to be removed
|
nipkow@1064
|
102 |
has a body of type T; otherwise the expansion thm will fail later on
|
nipkow@1064
|
103 |
*)
|
nipkow@1064
|
104 |
fun type_test(T,lbnos,apsns) =
|
paulson@2143
|
105 |
let val (_,U,_) = nth_elem(foldl Int.min (hd lbnos, tl lbnos), apsns)
|
nipkow@1064
|
106 |
in T=U end;
|
clasohm@0
|
107 |
|
berghofe@1686
|
108 |
(*************************************************************************
|
berghofe@1686
|
109 |
Create a "split_pack".
|
berghofe@1686
|
110 |
|
berghofe@1686
|
111 |
thm : the relevant split-theorem, i.e. P(...) == rhs , where P(...)
|
berghofe@1686
|
112 |
is of the form
|
berghofe@1686
|
113 |
P( Const(key,...) $ t_1 $ ... $ t_n ) (e.g. key = "if")
|
berghofe@1686
|
114 |
T : type of P(...)
|
berghofe@1686
|
115 |
n : number of arguments expected by Const(key,...)
|
berghofe@1686
|
116 |
ts : list of arguments actually found
|
berghofe@1686
|
117 |
apsns : list of tuples of the form (T,U,pos), one tuple for each
|
berghofe@1686
|
118 |
abstraction that is encountered on the way to the position where
|
berghofe@1686
|
119 |
Const(key, ...) $ ... occurs, where
|
berghofe@1686
|
120 |
T : type of the variable bound by the abstraction
|
berghofe@1686
|
121 |
U : type of the abstraction's body
|
berghofe@1686
|
122 |
pos : "path" leading to the body of the abstraction
|
berghofe@1686
|
123 |
pos : "path" leading to the position where Const(key, ...) $ ... occurs.
|
berghofe@1686
|
124 |
TB : type of Const(key,...) $ t_1 $ ... $ t_n
|
berghofe@1721
|
125 |
t : the term Const(key,...) $ t_1 $ ... $ t_n
|
berghofe@1686
|
126 |
|
berghofe@1686
|
127 |
A split pack is a tuple of the form
|
berghofe@1686
|
128 |
(thm, apsns, pos, TB)
|
berghofe@1686
|
129 |
Note : apsns is reversed, so that the outermost quantifier's position
|
berghofe@1686
|
130 |
comes first ! If the terms in ts don't contain variables bound
|
berghofe@1686
|
131 |
by other than meta-quantifiers, apsns is empty, because no further
|
berghofe@1686
|
132 |
lifting is required before applying the split-theorem.
|
berghofe@1686
|
133 |
******************************************************************************)
|
berghofe@1686
|
134 |
|
berghofe@1721
|
135 |
fun mk_split_pack(thm,T,n,ts,apsns,pos,TB,t) =
|
nipkow@1064
|
136 |
if n > length ts then []
|
nipkow@1064
|
137 |
else let val lev = length apsns
|
nipkow@1030
|
138 |
val lbnos = foldl add_lbnos ([],take(n,ts))
|
nipkow@1030
|
139 |
val flbnos = filter (fn i => i < lev) lbnos
|
berghofe@1721
|
140 |
val tt = incr_bv(~lev,0,t)
|
berghofe@1721
|
141 |
in if null flbnos then [(thm,[],pos,TB,tt)]
|
paulson@2143
|
142 |
else if type_test(T,flbnos,apsns) then [(thm, rev apsns,pos,TB,tt)]
|
paulson@2143
|
143 |
else []
|
nipkow@1064
|
144 |
end;
|
clasohm@0
|
145 |
|
berghofe@1686
|
146 |
|
berghofe@1686
|
147 |
(****************************************************************************
|
berghofe@1686
|
148 |
Recursively scans term for occurences of Const(key,...) $ ...
|
berghofe@1686
|
149 |
Returns a list of "split-packs" (one for each occurence of Const(key,...) )
|
berghofe@1686
|
150 |
|
berghofe@1686
|
151 |
cmap : association list of split-theorems that should be tried.
|
berghofe@1686
|
152 |
The elements have the format (key,(thm,T,n)) , where
|
berghofe@1686
|
153 |
key : the theorem's key constant ( Const(key,...) $ ... )
|
berghofe@1686
|
154 |
thm : the theorem itself
|
berghofe@1686
|
155 |
T : type of P( Const(key,...) $ ... )
|
berghofe@1686
|
156 |
n : number of arguments expected by Const(key,...)
|
berghofe@1686
|
157 |
Ts : types of parameters
|
berghofe@1686
|
158 |
t : the term to be scanned
|
berghofe@1686
|
159 |
******************************************************************************)
|
berghofe@1686
|
160 |
|
nipkow@1030
|
161 |
fun split_posns cmap Ts t =
|
nipkow@1030
|
162 |
let fun posns Ts pos apsns (Abs(_,T,t)) =
|
nipkow@1030
|
163 |
let val U = fastype_of1(T::Ts,t)
|
nipkow@1030
|
164 |
in posns (T::Ts) (0::pos) ((T,U,pos)::apsns) t end
|
nipkow@1030
|
165 |
| posns Ts pos apsns t =
|
nipkow@1030
|
166 |
let val (h,ts) = strip_comb t
|
nipkow@1030
|
167 |
fun iter((i,a),t) = (i+1, (posns Ts (i::pos) apsns t) @ a);
|
nipkow@1030
|
168 |
val a = case h of
|
nipkow@1030
|
169 |
Const(c,_) =>
|
nipkow@1030
|
170 |
(case assoc(cmap,c) of
|
berghofe@1721
|
171 |
Some(thm,T,n) => mk_split_pack(thm,T,n,ts,apsns,pos,type_of1(Ts,t),t)
|
nipkow@1030
|
172 |
| None => [])
|
nipkow@1030
|
173 |
| _ => []
|
nipkow@1030
|
174 |
in snd(foldl iter ((0,a),ts)) end
|
nipkow@1030
|
175 |
in posns Ts [] [] t end;
|
clasohm@0
|
176 |
|
berghofe@1686
|
177 |
|
clasohm@0
|
178 |
fun nth_subgoal i thm = nth_elem(i-1,prems_of thm);
|
clasohm@0
|
179 |
|
berghofe@1721
|
180 |
fun shorter((_,ps,pos,_,_),(_,qs,qos,_,_)) =
|
berghofe@1686
|
181 |
let val ms = length ps and ns = length qs
|
berghofe@1721
|
182 |
in ms < ns orelse (ms = ns andalso order(length pos,length qos)) end;
|
berghofe@1686
|
183 |
|
berghofe@1686
|
184 |
|
berghofe@1686
|
185 |
(************************************************************
|
berghofe@1686
|
186 |
call split_posns with appropriate parameters
|
berghofe@1686
|
187 |
*************************************************************)
|
clasohm@0
|
188 |
|
nipkow@1030
|
189 |
fun select cmap state i =
|
nipkow@1030
|
190 |
let val goali = nth_subgoal i state
|
nipkow@1030
|
191 |
val Ts = rev(map #2 (Logic.strip_params goali))
|
nipkow@1030
|
192 |
val _ $ t $ _ = Logic.strip_assums_concl goali;
|
nipkow@1030
|
193 |
in (Ts,t,sort shorter (split_posns cmap Ts t)) end;
|
nipkow@1030
|
194 |
|
berghofe@1686
|
195 |
|
berghofe@1686
|
196 |
(*************************************************************
|
berghofe@1686
|
197 |
instantiate lift theorem
|
berghofe@1686
|
198 |
|
berghofe@1686
|
199 |
if t is of the form
|
berghofe@1686
|
200 |
... ( Const(...,...) $ Abs( .... ) ) ...
|
berghofe@1686
|
201 |
then
|
berghofe@1686
|
202 |
P = %a. ... ( Const(...,...) $ a ) ...
|
berghofe@1686
|
203 |
where a has type T --> U
|
berghofe@1686
|
204 |
|
berghofe@1686
|
205 |
Ts : types of parameters
|
berghofe@1686
|
206 |
t : lefthand side of meta-equality in subgoal
|
berghofe@1686
|
207 |
the split theorem is applied to (see cmap)
|
berghofe@1686
|
208 |
T,U,pos : see mk_split_pack
|
berghofe@1686
|
209 |
state : current proof state
|
berghofe@1686
|
210 |
lift : the lift theorem
|
berghofe@1686
|
211 |
i : no. of subgoal
|
berghofe@1686
|
212 |
**************************************************************)
|
berghofe@1686
|
213 |
|
nipkow@1030
|
214 |
fun inst_lift Ts t (T,U,pos) state lift i =
|
clasohm@0
|
215 |
let val sg = #sign(rep_thm state)
|
clasohm@0
|
216 |
val tsig = #tsig(Sign.rep_sg sg)
|
nipkow@1030
|
217 |
val cntxt = mk_cntxt Ts t pos (T-->U) (#maxidx(rep_thm lift))
|
lcp@231
|
218 |
val cu = cterm_of sg cntxt
|
lcp@231
|
219 |
val uT = #T(rep_cterm cu)
|
lcp@231
|
220 |
val cP' = cterm_of sg (Var(P,uT))
|
clasohm@0
|
221 |
val ixnTs = Type.typ_match tsig ([],(PT,uT));
|
lcp@231
|
222 |
val ixncTs = map (fn (x,y) => (x,ctyp_of sg y)) ixnTs;
|
clasohm@0
|
223 |
in instantiate (ixncTs, [(cP',cu)]) lift end;
|
clasohm@0
|
224 |
|
clasohm@0
|
225 |
|
berghofe@1686
|
226 |
(*************************************************************
|
berghofe@1686
|
227 |
instantiate split theorem
|
berghofe@1686
|
228 |
|
berghofe@1686
|
229 |
Ts : types of parameters
|
berghofe@1686
|
230 |
t : lefthand side of meta-equality in subgoal
|
berghofe@1686
|
231 |
the split theorem is applied to (see cmap)
|
berghofe@1686
|
232 |
pos : "path" to the body of P(...)
|
berghofe@1686
|
233 |
thm : the split theorem
|
berghofe@1686
|
234 |
TB : type of body of P(...)
|
berghofe@1686
|
235 |
state : current proof state
|
berghofe@1686
|
236 |
**************************************************************)
|
berghofe@1686
|
237 |
|
berghofe@1721
|
238 |
fun inst_split Ts t tt thm TB state =
|
berghofe@1686
|
239 |
let val _$((Var(P2,PT2))$_)$_ = concl_of thm
|
berghofe@1686
|
240 |
val sg = #sign(rep_thm state)
|
berghofe@1686
|
241 |
val tsig = #tsig(Sign.rep_sg sg)
|
berghofe@1721
|
242 |
val cntxt = mk_cntxt_splitthm Ts t tt TB (#maxidx(rep_thm thm))
|
berghofe@1686
|
243 |
val cu = cterm_of sg cntxt
|
berghofe@1686
|
244 |
val uT = #T(rep_cterm cu)
|
berghofe@1686
|
245 |
val cP' = cterm_of sg (Var(P2,uT))
|
berghofe@1686
|
246 |
val ixnTs = Type.typ_match tsig ([],(PT2,uT));
|
berghofe@1686
|
247 |
val ixncTs = map (fn (x,y) => (x,ctyp_of sg y)) ixnTs;
|
berghofe@1686
|
248 |
in instantiate (ixncTs, [(cP',cu)]) thm end;
|
berghofe@1686
|
249 |
|
berghofe@1686
|
250 |
|
berghofe@1686
|
251 |
(*****************************************************************************
|
berghofe@1686
|
252 |
The split-tactic
|
berghofe@1686
|
253 |
|
berghofe@1686
|
254 |
splits : list of split-theorems to be tried
|
berghofe@1686
|
255 |
i : number of subgoal the tactic should be applied to
|
berghofe@1686
|
256 |
*****************************************************************************)
|
berghofe@1686
|
257 |
|
clasohm@0
|
258 |
fun split_tac [] i = no_tac
|
clasohm@0
|
259 |
| split_tac splits i =
|
nipkow@1030
|
260 |
let fun const(thm) = let val _$(t as _$lhs)$_ = concl_of thm
|
clasohm@0
|
261 |
val (Const(a,_),args) = strip_comb lhs
|
nipkow@1030
|
262 |
in (a,(thm,fastype_of t,length args)) end
|
clasohm@0
|
263 |
val cmap = map const splits;
|
paulson@3537
|
264 |
fun lift_tac Ts t p st = (rtac (inst_lift Ts t p st trlift i) i) st
|
paulson@3537
|
265 |
fun lift_split_tac st = st |>
|
paulson@3537
|
266 |
let val (Ts,t,splits) = select cmap st i
|
nipkow@1030
|
267 |
in case splits of
|
nipkow@1030
|
268 |
[] => no_tac
|
berghofe@1721
|
269 |
| (thm,apsns,pos,TB,tt)::_ =>
|
nipkow@1030
|
270 |
(case apsns of
|
paulson@3537
|
271 |
[] => (fn state => state |>
|
paulson@3537
|
272 |
rtac (inst_split Ts t tt thm TB state) i)
|
paulson@3537
|
273 |
| p::_ => EVERY[lift_tac Ts t p,
|
nipkow@1030
|
274 |
rtac reflexive_thm (i+1),
|
paulson@3537
|
275 |
lift_split_tac])
|
nipkow@1030
|
276 |
end
|
paulson@3537
|
277 |
in COND (has_fewer_prems i) no_tac
|
paulson@3537
|
278 |
(rtac iffD i THEN lift_split_tac)
|
clasohm@0
|
279 |
end;
|
clasohm@0
|
280 |
|
clasohm@0
|
281 |
in split_tac end;
|
berghofe@1721
|
282 |
|
berghofe@1721
|
283 |
in
|
berghofe@1721
|
284 |
|
berghofe@1721
|
285 |
fun mk_case_split_tac iffD = mk_case_split_tac_2 iffD (op <=) ;
|
berghofe@1721
|
286 |
|
berghofe@1721
|
287 |
fun mk_case_split_inside_tac iffD = mk_case_split_tac_2 iffD (op >=) ;
|
berghofe@1721
|
288 |
|
berghofe@1721
|
289 |
end;
|