| author | berghofe |
| Tue, 09 Jul 2002 18:03:26 +0200 | |
| changeset 13325 | 5b5e12f0aee0 |
| parent 13256 | cf85c4f7dcf2 |
| child 13342 | 915d4d004643 |
| permissions | -rw-r--r-- |
| 11522 | 1 |
(* Title: Pure/Proof/reconstruct.ML |
2 |
ID: $Id$ |
|
| 11539 | 3 |
Author: Stefan Berghofer, TU Muenchen |
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
| 11522 | 5 |
|
6 |
Reconstruction of partial proof terms. |
|
7 |
*) |
|
8 |
||
9 |
signature RECONSTRUCT = |
|
10 |
sig |
|
11 |
val quiet_mode : bool ref |
|
| 11613 | 12 |
val reconstruct_proof : Sign.sg -> term -> Proofterm.proof -> Proofterm.proof |
|
13256
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
13 |
val prop_of' : term list -> Proofterm.proof -> term |
| 13138 | 14 |
val prop_of : Proofterm.proof -> term |
| 11522 | 15 |
val expand_proof : Sign.sg -> string list -> Proofterm.proof -> Proofterm.proof |
16 |
end; |
|
17 |
||
18 |
structure Reconstruct : RECONSTRUCT = |
|
19 |
struct |
|
20 |
||
21 |
open Proofterm; |
|
22 |
||
23 |
val quiet_mode = ref true; |
|
24 |
fun message s = if !quiet_mode then () else writeln s; |
|
25 |
||
26 |
fun vars_of t = rev (foldl_aterms |
|
27 |
(fn (vs, v as Var _) => v ins vs | (vs, _) => vs) ([], t)); |
|
28 |
||
29 |
fun forall_intr (t, prop) = |
|
30 |
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) |
|
31 |
in all T $ Abs (a, T, abstract_over (t, prop)) end; |
|
32 |
||
33 |
fun forall_intr_vfs prop = foldr forall_intr |
|
34 |
(vars_of prop @ sort (make_ord atless) (term_frees prop), prop); |
|
35 |
||
| 12001 | 36 |
fun forall_intr_prf (t, prf) = |
37 |
let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p) |
|
38 |
in Abst (a, Some T, prf_abstract_over t prf) end; |
|
39 |
||
40 |
fun forall_intr_vfs_prf prop prf = foldr forall_intr_prf |
|
41 |
(vars_of prop @ sort (make_ord atless) (term_frees prop), prf); |
|
42 |
||
| 11522 | 43 |
fun merge_envs (Envir.Envir {asol=asol1, iTs=iTs1, maxidx=maxidx1})
|
44 |
(Envir.Envir {asol=asol2, iTs=iTs2, maxidx=maxidx2}) =
|
|
45 |
Envir.Envir {asol=Vartab.merge (op aconv) (asol1, asol2),
|
|
46 |
iTs=Vartab.merge (op =) (iTs1, iTs2), |
|
47 |
maxidx=Int.max (maxidx1, maxidx2)}; |
|
48 |
||
49 |
fun strip_abs (_::Ts) (Abs (_, _, t)) = strip_abs Ts t |
|
50 |
| strip_abs _ t = t; |
|
51 |
||
52 |
||
53 |
(******************************************************************************** |
|
54 |
generate constraints for proof term |
|
55 |
*********************************************************************************) |
|
56 |
||
57 |
fun mk_var env Ts T = |
|
58 |
let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) |
|
59 |
in (env', list_comb (v, map Bound (length Ts - 1 downto 0))) end; |
|
60 |
||
61 |
fun mk_tvar (Envir.Envir {iTs, asol, maxidx}, s) =
|
|
62 |
(Envir.Envir {iTs = iTs, asol = asol, maxidx = maxidx+1},
|
|
63 |
TVar (("'t", maxidx+1), s));
|
|
64 |
||
65 |
fun mk_abs Ts t = foldl (fn (u, T) => Abs ("", T, u)) (t, Ts);
|
|
66 |
||
67 |
fun make_Tconstraints_cprf maxidx cprf = |
|
68 |
let |
|
69 |
fun mk_Tcnstrts maxidx Ts (Abst (s, Some T, cprf)) = |
|
70 |
let val (cs, cprf', maxidx') = mk_Tcnstrts maxidx (T::Ts) cprf; |
|
71 |
in (cs, Abst (s, Some T, cprf'), maxidx') end |
|
72 |
| mk_Tcnstrts maxidx Ts (Abst (s, None, cprf)) = |
|
73 |
let |
|
74 |
val T' = TVar (("'t", maxidx+1), ["logic"]);
|
|
75 |
val (cs, cprf', maxidx') = mk_Tcnstrts (maxidx+1) (T'::Ts) cprf; |
|
76 |
in (cs, Abst (s, Some T', cprf'), maxidx') end |
|
77 |
| mk_Tcnstrts maxidx Ts (AbsP (s, Some t, cprf)) = |
|
78 |
let val (cs, cprf', maxidx') = mk_Tcnstrts maxidx Ts cprf; |
|
79 |
in ((mk_abs Ts t, rev Ts ---> propT)::cs, AbsP (s, Some t, cprf'), maxidx') end |
|
80 |
| mk_Tcnstrts maxidx Ts (AbsP (s, None, cprf)) = |
|
81 |
let val (cs, cprf', maxidx') = mk_Tcnstrts maxidx Ts cprf; |
|
82 |
in (cs, AbsP (s, None, cprf'), maxidx') end |
|
| 11613 | 83 |
| mk_Tcnstrts maxidx Ts (cprf1 %% cprf2) = |
| 11522 | 84 |
let |
85 |
val (cs, cprf1', maxidx') = mk_Tcnstrts maxidx Ts cprf1; |
|
86 |
val (cs', cprf2', maxidx'') = mk_Tcnstrts maxidx' Ts cprf2; |
|
| 11613 | 87 |
in (cs' @ cs, cprf1' %% cprf2', maxidx'') end |
88 |
| mk_Tcnstrts maxidx Ts (cprf % Some t) = |
|
| 11522 | 89 |
let val (cs, cprf', maxidx') = mk_Tcnstrts maxidx Ts cprf; |
90 |
in ((mk_abs Ts t, rev Ts ---> TypeInfer.logicT)::cs, |
|
| 11613 | 91 |
cprf' % Some t, maxidx') |
| 11522 | 92 |
end |
| 11613 | 93 |
| mk_Tcnstrts maxidx Ts (cprf % None) = |
| 11522 | 94 |
let val (cs, cprf', maxidx') = mk_Tcnstrts maxidx Ts cprf; |
| 11613 | 95 |
in (cs, cprf % None, maxidx') end |
| 11522 | 96 |
| mk_Tcnstrts maxidx _ cprf = ([], cprf, maxidx); |
97 |
in mk_Tcnstrts maxidx [] cprf end; |
|
98 |
||
99 |
fun unifyT sg env T U = |
|
100 |
let |
|
101 |
val Envir.Envir {asol, iTs, maxidx} = env;
|
|
| 12527 | 102 |
val (iTs', maxidx') = Type.unify (Sign.tsig_of sg) (iTs, maxidx) (T, U) |
| 11660 | 103 |
in Envir.Envir {asol=asol, iTs=iTs', maxidx=maxidx'} end
|
104 |
handle Type.TUNIFY => error ("Non-unifiable types:\n" ^
|
|
105 |
Sign.string_of_typ sg T ^ "\n\n" ^ Sign.string_of_typ sg U); |
|
| 11522 | 106 |
|
| 12236 | 107 |
fun decompose sg env Ts t u = case (Envir.head_norm env t, Envir.head_norm env u) of |
108 |
(Const ("all", _) $ t, Const ("all", _) $ u) => decompose sg env Ts t u
|
|
109 |
| (Const ("==>", _) $ t1 $ t2, Const ("==>", _) $ u1 $ u2) =>
|
|
| 11660 | 110 |
let val (env', cs) = decompose sg env Ts t1 u1 |
111 |
in apsnd (curry op @ cs) (decompose sg env' Ts t2 u2) end |
|
| 12236 | 112 |
| (Abs (_, T, t), Abs (_, U, u)) => |
| 11660 | 113 |
decompose sg (unifyT sg env T U) (T::Ts) t u |
| 12236 | 114 |
| (t, u) => (env, [(mk_abs Ts t, mk_abs Ts u)]); |
| 11522 | 115 |
|
| 11660 | 116 |
fun cantunify sg t u = error ("Non-unifiable terms:\n" ^
|
| 11522 | 117 |
Sign.string_of_term sg t ^ "\n\n" ^ Sign.string_of_term sg u); |
118 |
||
119 |
fun make_constraints_cprf sg env ts cprf = |
|
120 |
let |
|
121 |
fun add_cnstrt Ts prop prf cs env ts (t, u) = |
|
122 |
let |
|
123 |
val t' = mk_abs Ts t; |
|
| 12236 | 124 |
val u' = mk_abs Ts u |
| 11522 | 125 |
in |
| 12236 | 126 |
(prop, prf, cs, Pattern.unify (sg, env, [(t', u')]), ts) |
127 |
handle Pattern.Pattern => |
|
128 |
let val (env', cs') = decompose sg env [] t' u' |
|
129 |
in (prop, prf, cs @ cs', env', ts) end |
|
130 |
| Pattern.Unif => |
|
131 |
cantunify sg (Envir.norm_term env t') (Envir.norm_term env u') |
|
| 11522 | 132 |
end; |
133 |
||
| 13138 | 134 |
fun mk_cnstrts_atom env ts prop opTs prf = |
| 11522 | 135 |
let |
136 |
val tvars = term_tvars prop; |
|
137 |
val (env', Ts) = if_none (apsome (pair env) opTs) |
|
138 |
(foldl_map (mk_tvar o apsnd snd) (env, tvars)); |
|
| 13138 | 139 |
val prop' = subst_TVars (map fst tvars ~~ Ts) (forall_intr_vfs prop) |
140 |
handle LIST _ => error ("Wrong number of type arguments for " ^
|
|
141 |
quote (fst (get_name_tags [] prop prf))) |
|
142 |
in (prop', change_type (Some Ts) prf, [], env', ts) end; |
|
| 11522 | 143 |
|
144 |
fun mk_cnstrts env _ Hs ts (PBound i) = (nth_elem (i, Hs), PBound i, [], env, ts) |
|
145 |
| mk_cnstrts env Ts Hs ts (Abst (s, Some T, cprf)) = |
|
146 |
let val (t, prf, cnstrts, env', ts') = |
|
147 |
mk_cnstrts env (T::Ts) (map (incr_boundvars 1) Hs) ts cprf; |
|
148 |
in (Const ("all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, Some T, prf),
|
|
149 |
cnstrts, env', ts') |
|
150 |
end |
|
151 |
| mk_cnstrts env Ts Hs (t::ts) (AbsP (s, Some _, cprf)) = |
|
152 |
let |
|
153 |
val (u, prf, cnstrts, env', ts') = mk_cnstrts env Ts (t::Hs) ts cprf; |
|
154 |
val t' = strip_abs Ts t; |
|
155 |
in (Logic.mk_implies (t', u), AbsP (s, Some t', prf), cnstrts, env', ts') |
|
156 |
end |
|
157 |
| mk_cnstrts env Ts Hs ts (AbsP (s, None, cprf)) = |
|
158 |
let |
|
159 |
val (env', t) = mk_var env Ts propT; |
|
160 |
val (u, prf, cnstrts, env'', ts') = mk_cnstrts env' Ts (t::Hs) ts cprf; |
|
161 |
in (Logic.mk_implies (t, u), AbsP (s, Some t, prf), cnstrts, env'', ts') |
|
162 |
end |
|
| 11613 | 163 |
| mk_cnstrts env Ts Hs ts (cprf1 %% cprf2) = |
| 11522 | 164 |
let val (u, prf2, cnstrts, env', ts') = mk_cnstrts env Ts Hs ts cprf2 |
165 |
in (case mk_cnstrts env' Ts Hs ts' cprf1 of |
|
166 |
(Const ("==>", _) $ u' $ t', prf1, cnstrts', env'', ts'') =>
|
|
| 11613 | 167 |
add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts) |
| 11522 | 168 |
env'' ts'' (u, u') |
169 |
| (t, prf1, cnstrts', env'', ts'') => |
|
170 |
let val (env''', v) = mk_var env'' Ts propT |
|
| 11613 | 171 |
in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts) |
| 11522 | 172 |
env''' ts'' (t, Logic.mk_implies (u, v)) |
173 |
end) |
|
174 |
end |
|
| 11613 | 175 |
| mk_cnstrts env Ts Hs (t::ts) (cprf % Some _) = |
| 11522 | 176 |
let val t' = strip_abs Ts t |
177 |
in (case mk_cnstrts env Ts Hs ts cprf of |
|
178 |
(Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
|
|
179 |
prf, cnstrts, env', ts') => |
|
| 12236 | 180 |
let val env'' = unifyT sg env' T (Envir.fastype env' Ts t') |
| 11613 | 181 |
in (betapply (f, t'), prf % Some t', cnstrts, env'', ts') |
| 11522 | 182 |
end |
183 |
| (u, prf, cnstrts, env', ts') => |
|
184 |
let |
|
| 12236 | 185 |
val T = Envir.fastype env' Ts t'; |
| 11522 | 186 |
val (env'', v) = mk_var env' Ts (T --> propT); |
187 |
in |
|
| 11613 | 188 |
add_cnstrt Ts (v $ t') (prf % Some t') cnstrts env'' ts' |
| 11522 | 189 |
(u, Const ("all", (T --> propT) --> propT) $ v)
|
190 |
end) |
|
191 |
end |
|
| 11613 | 192 |
| mk_cnstrts env Ts Hs ts (cprf % None) = |
| 11522 | 193 |
(case mk_cnstrts env Ts Hs ts cprf of |
194 |
(Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
|
|
195 |
prf, cnstrts, env', ts') => |
|
196 |
let val (env'', t) = mk_var env' Ts T |
|
| 11613 | 197 |
in (betapply (f, t), prf % Some t, cnstrts, env'', ts') |
| 11522 | 198 |
end |
199 |
| (u, prf, cnstrts, env', ts') => |
|
200 |
let |
|
201 |
val (env1, T) = mk_tvar (env', ["logic"]); |
|
202 |
val (env2, v) = mk_var env1 Ts (T --> propT); |
|
203 |
val (env3, t) = mk_var env2 Ts T |
|
204 |
in |
|
| 11613 | 205 |
add_cnstrt Ts (v $ t) (prf % Some t) cnstrts env3 ts' |
| 11522 | 206 |
(u, Const ("all", (T --> propT) --> propT) $ v)
|
207 |
end) |
|
| 13138 | 208 |
| mk_cnstrts env _ _ ts (prf as PThm (_, _, prop, opTs)) = |
209 |
mk_cnstrts_atom env ts prop opTs prf |
|
210 |
| mk_cnstrts env _ _ ts (prf as PAxm (_, prop, opTs)) = |
|
211 |
mk_cnstrts_atom env ts prop opTs prf |
|
212 |
| mk_cnstrts env _ _ ts (prf as Oracle (_, prop, opTs)) = |
|
213 |
mk_cnstrts_atom env ts prop opTs prf |
|
| 11522 | 214 |
| mk_cnstrts env _ _ ts (Hyp t) = (t, Hyp t, [], env, ts) |
| 11613 | 215 |
| mk_cnstrts _ _ _ _ _ = error "reconstruct_proof: minimal proof object" |
| 11522 | 216 |
in mk_cnstrts env [] [] ts cprf end; |
217 |
||
218 |
fun add_term_ixns (is, Var (i, T)) = add_typ_ixns (i ins is, T) |
|
219 |
| add_term_ixns (is, Free (_, T)) = add_typ_ixns (is, T) |
|
220 |
| add_term_ixns (is, Const (_, T)) = add_typ_ixns (is, T) |
|
221 |
| add_term_ixns (is, t1 $ t2) = add_term_ixns (add_term_ixns (is, t1), t2) |
|
222 |
| add_term_ixns (is, Abs (_, T, t)) = add_term_ixns (add_typ_ixns (is, T), t) |
|
223 |
| add_term_ixns (is, _) = is; |
|
224 |
||
225 |
||
226 |
(******************************************************************************** |
|
227 |
update list of free variables of constraints |
|
228 |
*********************************************************************************) |
|
229 |
||
230 |
fun upd_constrs env cs = |
|
231 |
let |
|
232 |
val Envir.Envir {asol, iTs, ...} = env;
|
|
233 |
val dom = Vartab.foldl (uncurry (cons o fst) o Library.swap) |
|
234 |
(Vartab.foldl (uncurry (cons o fst) o Library.swap) ([], asol), iTs); |
|
235 |
val vran = Vartab.foldl (add_typ_ixns o apsnd snd) |
|
236 |
(Vartab.foldl (add_term_ixns o apsnd snd) ([], asol), iTs); |
|
237 |
fun check_cs [] = [] |
|
238 |
| check_cs ((u, p, vs)::ps) = |
|
239 |
let val vs' = vs \\ dom; |
|
240 |
in if vs = vs' then (u, p, vs)::check_cs ps |
|
241 |
else (true, p, vs' union vran)::check_cs ps |
|
242 |
end |
|
243 |
in check_cs cs end; |
|
244 |
||
245 |
(******************************************************************************** |
|
246 |
solution of constraints |
|
247 |
*********************************************************************************) |
|
248 |
||
249 |
fun solve _ [] bigenv = bigenv |
|
250 |
| solve sg cs bigenv = |
|
251 |
let |
|
| 11660 | 252 |
fun search env [] = error ("Unsolvable constraints:\n" ^
|
253 |
Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => |
|
254 |
Sign.pretty_term sg (Logic.mk_flexpair (pairself |
|
255 |
(Envir.norm_term bigenv) p))) cs))) |
|
| 11522 | 256 |
| search env ((u, p as (t1, t2), vs)::ps) = |
257 |
if u then |
|
258 |
let |
|
259 |
val tn1 = Envir.norm_term bigenv t1; |
|
260 |
val tn2 = Envir.norm_term bigenv t2 |
|
261 |
in |
|
262 |
if Pattern.pattern tn1 andalso Pattern.pattern tn2 then |
|
263 |
((Pattern.unify (sg, env, [(tn1, tn2)]), ps) handle Pattern.Unif => |
|
264 |
cantunify sg tn1 tn2) |
|
265 |
else |
|
266 |
let val (env', cs') = decompose sg env [] tn1 tn2 |
|
267 |
in if cs' = [(tn1, tn2)] then |
|
268 |
apsnd (cons (false, (tn1, tn2), vs)) (search env ps) |
|
269 |
else search env' (map (fn q => (true, q, vs)) cs' @ ps) |
|
270 |
end |
|
271 |
end |
|
272 |
else apsnd (cons (false, p, vs)) (search env ps); |
|
273 |
val Envir.Envir {maxidx, ...} = bigenv;
|
|
274 |
val (env, cs') = search (Envir.empty maxidx) cs; |
|
275 |
in |
|
276 |
solve sg (upd_constrs env cs') (merge_envs bigenv env) |
|
277 |
end; |
|
278 |
||
279 |
||
280 |
(******************************************************************************** |
|
281 |
reconstruction of proofs |
|
282 |
*********************************************************************************) |
|
283 |
||
| 11613 | 284 |
fun reconstruct_proof sg prop cprf = |
| 11522 | 285 |
let |
| 11613 | 286 |
val (cprf' % Some prop', thawf) = freeze_thaw_prf (cprf % Some prop); |
| 11522 | 287 |
val _ = message "Collecting type constraints..."; |
288 |
val (Tcs, cprf'', maxidx) = make_Tconstraints_cprf 0 cprf'; |
|
289 |
val (ts, Ts) = ListPair.unzip Tcs; |
|
290 |
val tsig = Sign.tsig_of sg; |
|
291 |
val {classrel, arities, ...} = Type.rep_tsig tsig;
|
|
292 |
val _ = message "Solving type constraints..."; |
|
293 |
val (ts', _, unifier) = TypeInfer.infer_types (Sign.pretty_term sg) (Sign.pretty_typ sg) |
|
294 |
(Sign.const_type sg) classrel arities [] false (K true) ts Ts; |
|
295 |
val env = Envir.Envir {asol = Vartab.empty, iTs = Vartab.make unifier, maxidx = maxidx};
|
|
296 |
val _ = message "Collecting term constraints..."; |
|
297 |
val (t, prf, cs, env, _) = make_constraints_cprf sg env ts' cprf''; |
|
298 |
val cs' = map (fn p => (true, p, op union |
|
299 |
(pairself (map (fst o dest_Var) o term_vars) p))) (map (pairself (Envir.norm_term env)) ((t, prop')::cs)); |
|
300 |
val _ = message ("Solving remaining constraints (" ^ string_of_int (length cs') ^ ") ...");
|
|
301 |
val env' = solve sg cs' env |
|
302 |
in |
|
303 |
thawf (norm_proof env' prf) |
|
304 |
end; |
|
305 |
||
|
13256
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
306 |
fun prop_of_atom prop Ts = |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
307 |
subst_TVars (map fst (term_tvars prop) ~~ Ts) (forall_intr_vfs prop); |
| 13138 | 308 |
|
|
13256
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
309 |
fun prop_of' Hs (PBound i) = nth_elem (i, Hs) |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
310 |
| prop_of' Hs (Abst (s, Some T, prf)) = |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
311 |
all T $ (Abs (s, T, prop_of' Hs prf)) |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
312 |
| prop_of' Hs (AbsP (s, Some t, prf)) = |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
313 |
Logic.mk_implies (t, prop_of' (t :: Hs) prf) |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
314 |
| prop_of' Hs (prf % Some t) = (case prop_of' Hs prf of |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
315 |
Const ("all", _) $ f => betapply (f, t)
|
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
316 |
| _ => error "prop_of: all expected") |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
317 |
| prop_of' Hs (prf1 %% prf2) = (case prop_of' Hs prf1 of |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
318 |
Const ("==>", _) $ P $ Q => Q
|
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
319 |
| _ => error "prop_of: ==> expected") |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
320 |
| prop_of' Hs (Hyp t) = t |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
321 |
| prop_of' Hs (PThm (_, _, prop, Some Ts)) = prop_of_atom prop Ts |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
322 |
| prop_of' Hs (PAxm (_, prop, Some Ts)) = prop_of_atom prop Ts |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
323 |
| prop_of' Hs (Oracle (_, prop, Some Ts)) = prop_of_atom prop Ts |
|
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
324 |
| prop_of' _ _ = error "prop_of: partial proof object"; |
| 13138 | 325 |
|
|
13256
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
326 |
val prop_of = prop_of' []; |
| 13138 | 327 |
|
| 11522 | 328 |
|
329 |
(******************************************************************************** |
|
330 |
expand and reconstruct subproofs |
|
331 |
*********************************************************************************) |
|
332 |
||
333 |
fun expand_proof sg names prf = |
|
334 |
let |
|
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
335 |
fun expand maxidx prfs (AbsP (s, t, prf)) = |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
336 |
let val (maxidx', prfs', prf') = expand maxidx prfs prf |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
337 |
in (maxidx', prfs', AbsP (s, t, prf')) end |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
338 |
| expand maxidx prfs (Abst (s, T, prf)) = |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
339 |
let val (maxidx', prfs', prf') = expand maxidx prfs prf |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
340 |
in (maxidx', prfs', Abst (s, T, prf')) end |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
341 |
| expand maxidx prfs (prf1 %% prf2) = |
| 11522 | 342 |
let |
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
343 |
val (maxidx', prfs', prf1') = expand maxidx prfs prf1; |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
344 |
val (maxidx'', prfs'', prf2') = expand maxidx' prfs' prf2; |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
345 |
in (maxidx'', prfs'', prf1' %% prf2') end |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
346 |
| expand maxidx prfs (prf % t) = |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
347 |
let val (maxidx', prfs', prf') = expand maxidx prfs prf |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
348 |
in (maxidx', prfs', prf' % t) end |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
349 |
| expand maxidx prfs (prf as PThm ((a, _), cprf, prop, Some Ts)) = |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
350 |
if not (a mem names) then (maxidx, prfs, prf) else |
| 11522 | 351 |
let |
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
352 |
val (maxidx', (i, prf), prfs') = (case assoc (prfs, (a, prop)) of |
| 11522 | 353 |
None => |
354 |
let |
|
355 |
val _ = message ("Reconstructing proof of " ^ a);
|
|
356 |
val _ = message (Sign.string_of_term sg prop); |
|
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
357 |
val i = maxidx + 1; |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
358 |
val prf' = map_proof_terms (Logic.incr_indexes ([], i)) |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
359 |
(incr_tvar i) (forall_intr_vfs_prf prop |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
360 |
(reconstruct_proof sg prop cprf)); |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
361 |
val (maxidx', prfs', prf) = expand |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
362 |
(maxidx_of_proof prf') prfs prf' |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
363 |
in (maxidx', (i, prf), ((a, prop), (i, prf)) :: prfs') end |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
364 |
| Some p => (maxidx, p, prfs)); |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
365 |
val tye = map (fn ((s, j), _) => (s, i + j)) (term_tvars prop) ~~ Ts |
| 11522 | 366 |
in |
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
367 |
(maxidx', prfs', |
|
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
368 |
map_proof_terms (subst_TVars tye) (typ_subst_TVars tye) prf) |
| 11522 | 369 |
end |
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
370 |
| expand maxidx prfs prf = (maxidx, prfs, prf); |
| 11522 | 371 |
|
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
372 |
in #3 (expand (maxidx_of_proof prf) [] prf) end; |
| 11522 | 373 |
|
374 |
end; |