| author | huffman | 
| Tue, 12 Jul 2005 18:26:44 +0200 | |
| changeset 16777 | 555c8951f05c | 
| parent 15798 | 016f3be5a5ec | 
| child 16787 | b6b6e2faaa41 | 
| permissions | -rw-r--r-- | 
| 11522 | 1  | 
(* Title: Pure/Proof/reconstruct.ML  | 
2  | 
ID: $Id$  | 
|
| 11539 | 3  | 
Author: Stefan Berghofer, TU Muenchen  | 
| 11522 | 4  | 
|
5  | 
Reconstruction of partial proof terms.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature RECONSTRUCT =  | 
|
9  | 
sig  | 
|
10  | 
val quiet_mode : bool ref  | 
|
| 11613 | 11  | 
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
 | 
12  | 
val prop_of' : term list -> Proofterm.proof -> term  | 
| 13138 | 13  | 
val prop_of : Proofterm.proof -> term  | 
| 
13342
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
14  | 
val expand_proof : Sign.sg -> (string * term option) list ->  | 
| 
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
15  | 
Proofterm.proof -> Proofterm.proof  | 
| 11522 | 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  | 
||
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
33  | 
fun forall_intr_vfs prop = foldr forall_intr prop  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
34  | 
(vars_of prop @ sort (make_ord atless) (term_frees prop));  | 
| 11522 | 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)  | 
|
| 15531 | 38  | 
in Abst (a, SOME T, prf_abstract_over t prf) end;  | 
| 12001 | 39  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
40  | 
fun forall_intr_vfs_prf prop prf = foldr forall_intr_prf prf  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
41  | 
(vars_of prop @ sort (make_ord atless) (term_frees prop));  | 
| 12001 | 42  | 
|
| 11522 | 43  | 
fun merge_envs (Envir.Envir {asol=asol1, iTs=iTs1, maxidx=maxidx1})
 | 
44  | 
  (Envir.Envir {asol=asol2, iTs=iTs2, maxidx=maxidx2}) =
 | 
|
| 
15798
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
45  | 
    Envir.Envir {asol=Vartab.merge (op =) (asol1, asol2),
 | 
| 11522 | 46  | 
iTs=Vartab.merge (op =) (iTs1, iTs2),  | 
47  | 
maxidx=Int.max (maxidx1, maxidx2)};  | 
|
48  | 
||
49  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
50  | 
(**** generate constraints for proof term ****)  | 
| 11522 | 51  | 
|
52  | 
fun mk_var env Ts T =  | 
|
53  | 
let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T)  | 
|
54  | 
in (env', list_comb (v, map Bound (length Ts - 1 downto 0))) end;  | 
|
55  | 
||
56  | 
fun mk_tvar (Envir.Envir {iTs, asol, maxidx}, s) =
 | 
|
57  | 
  (Envir.Envir {iTs = iTs, asol = asol, maxidx = maxidx+1},
 | 
|
58  | 
   TVar (("'t", maxidx+1), s));
 | 
|
59  | 
||
| 15570 | 60  | 
fun mk_abs Ts t = Library.foldl (fn (u, T) => Abs ("", T, u)) (t, Ts);
 | 
| 11522 | 61  | 
|
62  | 
fun unifyT sg env T U =  | 
|
63  | 
let  | 
|
64  | 
    val Envir.Envir {asol, iTs, maxidx} = env;
 | 
|
| 12527 | 65  | 
val (iTs', maxidx') = Type.unify (Sign.tsig_of sg) (iTs, maxidx) (T, U)  | 
| 11660 | 66  | 
  in Envir.Envir {asol=asol, iTs=iTs', maxidx=maxidx'} end
 | 
67  | 
  handle Type.TUNIFY => error ("Non-unifiable types:\n" ^
 | 
|
68  | 
Sign.string_of_typ sg T ^ "\n\n" ^ Sign.string_of_typ sg U);  | 
|
| 11522 | 69  | 
|
| 
15798
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
70  | 
fun chaseT (env as Envir.Envir {iTs, ...}) (T as TVar ixnS) =
 | 
| 
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
71  | 
(case Type.lookup (iTs, ixnS) of NONE => T | SOME T' => chaseT env T')  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
72  | 
| chaseT _ T = T;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
73  | 
|
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
74  | 
fun infer_type sg (env as Envir.Envir {maxidx, asol, iTs}) Ts vTs
 | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
75  | 
(t as Const (s, T)) = if T = dummyT then (case Sign.const_type sg s of  | 
| 15531 | 76  | 
          NONE => error ("reconstruct_proof: No such constant: " ^ quote s)
 | 
77  | 
| SOME T =>  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
78  | 
let val T' = incr_tvar (maxidx + 1) T  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
79  | 
in (Const (s, T'), T', vTs,  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
80  | 
              Envir.Envir {maxidx = maxidx + 1, asol = asol, iTs = iTs})
 | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
81  | 
end)  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
82  | 
else (t, T, vTs, env)  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
83  | 
| infer_type sg env Ts vTs (t as Free (s, T)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
84  | 
if T = dummyT then (case Symtab.lookup (vTs, s) of  | 
| 15531 | 85  | 
NONE =>  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
86  | 
let val (env', T) = mk_tvar (env, [])  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
87  | 
in (Free (s, T), T, Symtab.update_new ((s, T), vTs), env') end  | 
| 15531 | 88  | 
| SOME T => (Free (s, T), T, vTs, env))  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
89  | 
else (t, T, vTs, env)  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
90  | 
| infer_type sg env Ts vTs (Var _) = error "reconstruct_proof: internal error"  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
91  | 
| infer_type sg env Ts vTs (Abs (s, T, t)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
92  | 
let  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
93  | 
val (env', T') = if T = dummyT then mk_tvar (env, []) else (env, T);  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
94  | 
val (t', U, vTs', env'') = infer_type sg env' (T' :: Ts) vTs t  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
95  | 
in (Abs (s, T', t'), T' --> U, vTs', env'') end  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
96  | 
| infer_type sg env Ts vTs (t $ u) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
97  | 
let  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
98  | 
val (t', T, vTs1, env1) = infer_type sg env Ts vTs t;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
99  | 
val (u', U, vTs2, env2) = infer_type sg env1 Ts vTs1 u;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
100  | 
in (case chaseT env2 T of  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
101  | 
          Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT sg env2 U U')
 | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
102  | 
| _ =>  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
103  | 
let val (env3, V) = mk_tvar (env2, [])  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
104  | 
in (t' $ u', V, vTs2, unifyT sg env3 T (U --> V)) end)  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
105  | 
end  | 
| 15570 | 106  | 
| infer_type sg env Ts vTs (t as Bound i) = (t, List.nth (Ts, i), vTs, env);  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
107  | 
|
| 13715 | 108  | 
fun cantunify sg (t, u) = error ("Non-unifiable terms:\n" ^
 | 
109  | 
Sign.string_of_term sg t ^ "\n\n" ^ Sign.string_of_term sg u);  | 
|
| 11522 | 110  | 
|
| 13715 | 111  | 
fun decompose sg Ts (env, p as (t, u)) =  | 
112  | 
let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify sg p  | 
|
| 15570 | 113  | 
else apsnd List.concat (foldl_map (decompose sg Ts) (uT env T U, ts ~~ us))  | 
| 13715 | 114  | 
in case pairself (strip_comb o Envir.head_norm env) p of  | 
115  | 
((Const c, ts), (Const d, us)) => rigrig c d (unifyT sg) ts us  | 
|
116  | 
| ((Free c, ts), (Free d, us)) => rigrig c d (unifyT sg) ts us  | 
|
117  | 
| ((Bound i, ts), (Bound j, us)) =>  | 
|
118  | 
rigrig (i, dummyT) (j, dummyT) (K o K) ts us  | 
|
119  | 
| ((Abs (_, T, t), []), (Abs (_, U, u), [])) =>  | 
|
120  | 
decompose sg (T::Ts) (unifyT sg env T U, (t, u))  | 
|
121  | 
| ((Abs (_, T, t), []), _) =>  | 
|
122  | 
decompose sg (T::Ts) (env, (t, incr_boundvars 1 u $ Bound 0))  | 
|
123  | 
| (_, (Abs (_, T, u), [])) =>  | 
|
124  | 
decompose sg (T::Ts) (env, (incr_boundvars 1 t $ Bound 0, u))  | 
|
125  | 
| _ => (env, [(mk_abs Ts t, mk_abs Ts u)])  | 
|
126  | 
end;  | 
|
| 11522 | 127  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
128  | 
fun make_constraints_cprf sg env cprf =  | 
| 11522 | 129  | 
let  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
130  | 
fun add_cnstrt Ts prop prf cs env vTs (t, u) =  | 
| 11522 | 131  | 
let  | 
132  | 
val t' = mk_abs Ts t;  | 
|
| 12236 | 133  | 
val u' = mk_abs Ts u  | 
| 11522 | 134  | 
in  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
135  | 
(prop, prf, cs, Pattern.unify (sg, env, [(t', u')]), vTs)  | 
| 12236 | 136  | 
handle Pattern.Pattern =>  | 
| 13715 | 137  | 
let val (env', cs') = decompose sg [] (env, (t', u'))  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
138  | 
in (prop, prf, cs @ cs', env', vTs) end  | 
| 12236 | 139  | 
| Pattern.Unif =>  | 
| 13715 | 140  | 
cantunify sg (Envir.norm_term env t', Envir.norm_term env u')  | 
| 11522 | 141  | 
end;  | 
142  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
143  | 
fun mk_cnstrts_atom env vTs prop opTs prf =  | 
| 11522 | 144  | 
let  | 
145  | 
val tvars = term_tvars prop;  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
146  | 
val tfrees = term_tfrees prop;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
147  | 
val (prop', fmap) = Type.varify (prop, []);  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
148  | 
val (env', Ts) = (case opTs of  | 
| 15531 | 149  | 
NONE => foldl_map mk_tvar (env, map snd tvars @ map snd tfrees)  | 
150  | 
| SOME Ts => (env, Ts));  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
151  | 
val prop'' = subst_TVars (map fst tvars @ map snd fmap ~~ Ts)  | 
| 15570 | 152  | 
(forall_intr_vfs prop') handle UnequalLengths =>  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
153  | 
                error ("Wrong number of type arguments for " ^
 | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
154  | 
quote (fst (get_name_tags [] prop prf)))  | 
| 15531 | 155  | 
in (prop'', change_type (SOME Ts) prf, [], env', vTs) end;  | 
| 11522 | 156  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
157  | 
fun head_norm (prop, prf, cnstrts, env, vTs) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
158  | 
(Envir.head_norm env prop, prf, cnstrts, env, vTs);  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
159  | 
|
| 15570 | 160  | 
fun mk_cnstrts env _ Hs vTs (PBound i) = (List.nth (Hs, i), PBound i, [], env, vTs)  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
161  | 
| mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
162  | 
let  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
163  | 
val (env', T) = (case opT of  | 
| 15531 | 164  | 
NONE => mk_tvar (env, []) | SOME T => (env, T));  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
165  | 
val (t, prf, cnstrts, env'', vTs') =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
166  | 
mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf;  | 
| 15531 | 167  | 
          in (Const ("all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf),
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
168  | 
cnstrts, env'', vTs')  | 
| 11522 | 169  | 
end  | 
| 15531 | 170  | 
| mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) =  | 
| 11522 | 171  | 
let  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
172  | 
val (t', _, vTs', env') = infer_type sg env Ts vTs t;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
173  | 
val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf;  | 
| 15531 | 174  | 
in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'')  | 
| 11522 | 175  | 
end  | 
| 15531 | 176  | 
| mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) =  | 
| 11522 | 177  | 
let  | 
178  | 
val (env', t) = mk_var env Ts propT;  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
179  | 
val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf;  | 
| 15531 | 180  | 
in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs')  | 
| 11522 | 181  | 
end  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
182  | 
| mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
183  | 
let val (u, prf2, cnstrts, env', vTs') = mk_cnstrts env Ts Hs vTs cprf2  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
184  | 
in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
185  | 
              (Const ("==>", _) $ u' $ t', prf1, cnstrts', env'', vTs'') =>
 | 
| 11613 | 186  | 
add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts)  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
187  | 
env'' vTs'' (u, u')  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
188  | 
| (t, prf1, cnstrts', env'', vTs'') =>  | 
| 11522 | 189  | 
let val (env''', v) = mk_var env'' Ts propT  | 
| 11613 | 190  | 
in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts)  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
191  | 
env''' vTs'' (t, Logic.mk_implies (u, v))  | 
| 11522 | 192  | 
end)  | 
193  | 
end  | 
|
| 15531 | 194  | 
| mk_cnstrts env Ts Hs vTs (cprf % SOME t) =  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
195  | 
let val (t', U, vTs1, env1) = infer_type sg env Ts vTs t  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
196  | 
in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of  | 
| 11522 | 197  | 
             (Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
198  | 
prf, cnstrts, env2, vTs2) =>  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
199  | 
let val env3 = unifyT sg env2 T U  | 
| 15531 | 200  | 
in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2)  | 
| 11522 | 201  | 
end  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
202  | 
| (u, prf, cnstrts, env2, vTs2) =>  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
203  | 
let val (env3, v) = mk_var env2 Ts (U --> propT);  | 
| 11522 | 204  | 
in  | 
| 15531 | 205  | 
add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
206  | 
                   (u, Const ("all", (U --> propT) --> propT) $ v)
 | 
| 11522 | 207  | 
end)  | 
208  | 
end  | 
|
| 15531 | 209  | 
| mk_cnstrts env Ts Hs vTs (cprf % NONE) =  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
210  | 
(case head_norm (mk_cnstrts env Ts Hs vTs cprf) of  | 
| 11522 | 211  | 
             (Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
212  | 
prf, cnstrts, env', vTs') =>  | 
| 11522 | 213  | 
let val (env'', t) = mk_var env' Ts T  | 
| 15531 | 214  | 
in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs')  | 
| 11522 | 215  | 
end  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
216  | 
| (u, prf, cnstrts, env', vTs') =>  | 
| 11522 | 217  | 
let  | 
| 14854 | 218  | 
val (env1, T) = mk_tvar (env', []);  | 
| 11522 | 219  | 
val (env2, v) = mk_var env1 Ts (T --> propT);  | 
220  | 
val (env3, t) = mk_var env2 Ts T  | 
|
221  | 
in  | 
|
| 15531 | 222  | 
add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs'  | 
| 11522 | 223  | 
                   (u, Const ("all", (T --> propT) --> propT) $ v)
 | 
224  | 
end)  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
225  | 
| mk_cnstrts env _ _ vTs (prf as PThm (_, _, prop, opTs)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
226  | 
mk_cnstrts_atom env vTs prop opTs prf  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
227  | 
| mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
228  | 
mk_cnstrts_atom env vTs prop opTs prf  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
229  | 
| mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
230  | 
mk_cnstrts_atom env vTs prop opTs prf  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
231  | 
| mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs)  | 
| 11613 | 232  | 
| mk_cnstrts _ _ _ _ _ = error "reconstruct_proof: minimal proof object"  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
233  | 
in mk_cnstrts env [] [] Symtab.empty cprf end;  | 
| 11522 | 234  | 
|
235  | 
fun add_term_ixns (is, Var (i, T)) = add_typ_ixns (i ins is, T)  | 
|
236  | 
| add_term_ixns (is, Free (_, T)) = add_typ_ixns (is, T)  | 
|
237  | 
| add_term_ixns (is, Const (_, T)) = add_typ_ixns (is, T)  | 
|
238  | 
| add_term_ixns (is, t1 $ t2) = add_term_ixns (add_term_ixns (is, t1), t2)  | 
|
239  | 
| add_term_ixns (is, Abs (_, T, t)) = add_term_ixns (add_typ_ixns (is, T), t)  | 
|
240  | 
| add_term_ixns (is, _) = is;  | 
|
241  | 
||
242  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
243  | 
(**** update list of free variables of constraints ****)  | 
| 11522 | 244  | 
|
245  | 
fun upd_constrs env cs =  | 
|
246  | 
let  | 
|
247  | 
    val Envir.Envir {asol, iTs, ...} = env;
 | 
|
248  | 
val dom = Vartab.foldl (uncurry (cons o fst) o Library.swap)  | 
|
249  | 
(Vartab.foldl (uncurry (cons o fst) o Library.swap) ([], asol), iTs);  | 
|
| 
15798
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
250  | 
val vran = Vartab.foldl (add_typ_ixns o apsnd (snd o snd))  | 
| 
 
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
 
berghofe 
parents: 
15574 
diff
changeset
 | 
251  | 
(Vartab.foldl (add_term_ixns o apsnd (snd o snd)) ([], asol), iTs);  | 
| 11522 | 252  | 
fun check_cs [] = []  | 
253  | 
| check_cs ((u, p, vs)::ps) =  | 
|
254  | 
let val vs' = vs \\ dom;  | 
|
255  | 
in if vs = vs' then (u, p, vs)::check_cs ps  | 
|
256  | 
else (true, p, vs' union vran)::check_cs ps  | 
|
257  | 
end  | 
|
258  | 
in check_cs cs end;  | 
|
259  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
260  | 
(**** solution of constraints ****)  | 
| 11522 | 261  | 
|
262  | 
fun solve _ [] bigenv = bigenv  | 
|
263  | 
| solve sg cs bigenv =  | 
|
264  | 
let  | 
|
| 11660 | 265  | 
        fun search env [] = error ("Unsolvable constraints:\n" ^
 | 
266  | 
Pretty.string_of (Pretty.chunks (map (fn (_, p, _) =>  | 
|
| 14876 | 267  | 
Display.pretty_flexpair (Sign.pp sg) (pairself  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
268  | 
(Envir.norm_term bigenv) p)) cs)))  | 
| 11522 | 269  | 
| search env ((u, p as (t1, t2), vs)::ps) =  | 
270  | 
if u then  | 
|
271  | 
let  | 
|
272  | 
val tn1 = Envir.norm_term bigenv t1;  | 
|
273  | 
val tn2 = Envir.norm_term bigenv t2  | 
|
274  | 
in  | 
|
275  | 
if Pattern.pattern tn1 andalso Pattern.pattern tn2 then  | 
|
276  | 
((Pattern.unify (sg, env, [(tn1, tn2)]), ps) handle Pattern.Unif =>  | 
|
| 13715 | 277  | 
cantunify sg (tn1, tn2))  | 
| 11522 | 278  | 
else  | 
| 13715 | 279  | 
let val (env', cs') = decompose sg [] (env, (tn1, tn2))  | 
| 11522 | 280  | 
in if cs' = [(tn1, tn2)] then  | 
281  | 
apsnd (cons (false, (tn1, tn2), vs)) (search env ps)  | 
|
282  | 
else search env' (map (fn q => (true, q, vs)) cs' @ ps)  | 
|
283  | 
end  | 
|
284  | 
end  | 
|
285  | 
else apsnd (cons (false, p, vs)) (search env ps);  | 
|
286  | 
        val Envir.Envir {maxidx, ...} = bigenv;
 | 
|
287  | 
val (env, cs') = search (Envir.empty maxidx) cs;  | 
|
288  | 
in  | 
|
289  | 
solve sg (upd_constrs env cs') (merge_envs bigenv env)  | 
|
290  | 
end;  | 
|
291  | 
||
292  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
293  | 
(**** reconstruction of proofs ****)  | 
| 11522 | 294  | 
|
| 11613 | 295  | 
fun reconstruct_proof sg prop cprf =  | 
| 11522 | 296  | 
let  | 
| 15531 | 297  | 
val (cprf' % SOME prop', thawf) = freeze_thaw_prf (cprf % SOME prop);  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
298  | 
val _ = message "Collecting constraints...";  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
299  | 
val (t, prf, cs, env, _) = make_constraints_cprf sg  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
300  | 
(Envir.empty (maxidx_of_proof cprf)) cprf';  | 
| 11522 | 301  | 
val cs' = map (fn p => (true, p, op union  | 
302  | 
(pairself (map (fst o dest_Var) o term_vars) p))) (map (pairself (Envir.norm_term env)) ((t, prop')::cs));  | 
|
303  | 
    val _ = message ("Solving remaining constraints (" ^ string_of_int (length cs') ^ ") ...");
 | 
|
304  | 
val env' = solve sg cs' env  | 
|
305  | 
in  | 
|
306  | 
thawf (norm_proof env' prf)  | 
|
307  | 
end;  | 
|
308  | 
||
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
309  | 
fun prop_of_atom prop Ts =  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
310  | 
let val (prop', fmap) = Type.varify (prop, []);  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
311  | 
in subst_TVars (map fst (term_tvars prop) @ map snd fmap ~~ Ts)  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
312  | 
(forall_intr_vfs prop')  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
313  | 
end;  | 
| 13138 | 314  | 
|
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
315  | 
val head_norm = Envir.head_norm (Envir.empty 0);  | 
| 
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
316  | 
|
| 15570 | 317  | 
fun prop_of0 Hs (PBound i) = List.nth (Hs, i)  | 
| 15531 | 318  | 
| prop_of0 Hs (Abst (s, SOME T, prf)) =  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
319  | 
all T $ (Abs (s, T, prop_of0 Hs prf))  | 
| 15531 | 320  | 
| prop_of0 Hs (AbsP (s, SOME t, prf)) =  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
321  | 
Logic.mk_implies (t, prop_of0 (t :: Hs) prf)  | 
| 15531 | 322  | 
| prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
323  | 
      Const ("all", _) $ f => f $ t
 | 
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
324  | 
| _ => error "prop_of: all expected")  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
325  | 
| prop_of0 Hs (prf1 %% prf2) = (case head_norm (prop_of0 Hs prf1) of  | 
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
326  | 
      Const ("==>", _) $ P $ Q => Q
 | 
| 
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
327  | 
| _ => error "prop_of: ==> expected")  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
328  | 
| prop_of0 Hs (Hyp t) = t  | 
| 15531 | 329  | 
| prop_of0 Hs (PThm (_, _, prop, SOME Ts)) = prop_of_atom prop Ts  | 
330  | 
| prop_of0 Hs (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts  | 
|
331  | 
| prop_of0 Hs (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts  | 
|
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
332  | 
| prop_of0 _ _ = error "prop_of: partial proof object";  | 
| 13138 | 333  | 
|
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
334  | 
val prop_of' = Pattern.eta_contract oo (Envir.beta_norm oo prop_of0);  | 
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
335  | 
val prop_of = prop_of' [];  | 
| 13138 | 336  | 
|
| 11522 | 337  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
338  | 
(**** expand and reconstruct subproofs ****)  | 
| 11522 | 339  | 
|
| 
13342
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
340  | 
fun expand_proof sg thms prf =  | 
| 11522 | 341  | 
let  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
342  | 
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
 | 
343  | 
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
 | 
344  | 
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
 | 
345  | 
| 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
 | 
346  | 
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
 | 
347  | 
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
 | 
348  | 
| expand maxidx prfs (prf1 %% prf2) =  | 
| 11522 | 349  | 
let  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
350  | 
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
 | 
351  | 
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
 | 
352  | 
in (maxidx'', prfs'', prf1' %% prf2') end  | 
| 
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
353  | 
| expand maxidx prfs (prf % t) =  | 
| 
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
354  | 
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
 | 
355  | 
in (maxidx', prfs', prf' % t) end  | 
| 15531 | 356  | 
| expand maxidx prfs (prf as PThm ((a, _), cprf, prop, SOME Ts)) =  | 
| 
13342
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
357  | 
if not (exists  | 
| 15531 | 358  | 
(fn (b, NONE) => a = b  | 
359  | 
| (b, SOME prop') => a = b andalso prop = prop') thms)  | 
|
| 
13342
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
360  | 
then (maxidx, prfs, prf) else  | 
| 11522 | 361  | 
let  | 
| 
13610
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
362  | 
fun inc i =  | 
| 
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
363  | 
map_proof_terms (Logic.incr_indexes ([], i)) (incr_tvar i);  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
364  | 
val (maxidx', prf, prfs') = (case assoc (prfs, (a, prop)) of  | 
| 15531 | 365  | 
NONE =>  | 
| 11522 | 366  | 
let  | 
367  | 
                    val _ = message ("Reconstructing proof of " ^ a);
 | 
|
368  | 
val _ = message (Sign.string_of_term sg prop);  | 
|
| 
13610
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
369  | 
val prf' = forall_intr_vfs_prf prop  | 
| 
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
370  | 
(reconstruct_proof sg prop cprf);  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
371  | 
val (maxidx', prfs', prf) = expand  | 
| 
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
372  | 
(maxidx_of_proof prf') prfs prf'  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
373  | 
in (maxidx' + maxidx + 1, inc (maxidx + 1) prf,  | 
| 
13610
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
374  | 
((a, prop), (maxidx', prf)) :: prfs')  | 
| 
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
375  | 
end  | 
| 15531 | 376  | 
| SOME (maxidx', prf) => (maxidx' + maxidx + 1,  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
377  | 
inc (maxidx + 1) prf, prfs));  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
378  | 
val tfrees = term_tfrees prop;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
379  | 
val tye = map (fn ((s, j), _) => (s, maxidx + 1 + j))  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
380  | 
(term_tvars prop) @ map (rpair ~1 o fst) tfrees ~~ Ts;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
381  | 
val varify = map_type_tfree (fn p as (a, S) =>  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
382  | 
if p mem tfrees then TVar ((a, ~1), S) else TFree p)  | 
| 11522 | 383  | 
in  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
384  | 
(maxidx', prfs', map_proof_terms (subst_TVars tye o  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
385  | 
map_term_types varify) (typ_subst_TVars tye o varify) prf)  | 
| 11522 | 386  | 
end  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
387  | 
| expand maxidx prfs prf = (maxidx, prfs, prf);  | 
| 11522 | 388  | 
|
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
389  | 
in #3 (expand (maxidx_of_proof prf) [] prf) end;  | 
| 11522 | 390  | 
|
391  | 
end;  |