| author | nipkow | 
| Mon, 14 May 2018 15:37:26 +0200 | |
| changeset 68175 | e0bd5089eabf | 
| parent 67649 | 1e1782c1aedf | 
| child 69575 | f77cc54f6d47 | 
| permissions | -rw-r--r-- | 
| 11522 | 1  | 
(* Title: Pure/Proof/reconstruct.ML  | 
| 11539 | 2  | 
Author: Stefan Berghofer, TU Muenchen  | 
| 11522 | 3  | 
|
4  | 
Reconstruction of partial proof terms.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature RECONSTRUCT =  | 
|
8  | 
sig  | 
|
| 62922 | 9  | 
val quiet_mode : bool Config.T  | 
10  | 
val reconstruct_proof : Proof.context -> term -> Proofterm.proof -> Proofterm.proof  | 
|
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
11  | 
val prop_of' : term list -> Proofterm.proof -> term  | 
| 13138 | 12  | 
val prop_of : Proofterm.proof -> term  | 
| 62922 | 13  | 
val proof_of : Proof.context -> thm -> Proofterm.proof  | 
14  | 
val expand_proof : Proof.context -> (string * term option) list ->  | 
|
| 
13342
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
15  | 
Proofterm.proof -> Proofterm.proof  | 
| 64986 | 16  | 
val clean_proof_of : Proof.context -> bool -> thm -> Proofterm.proof  | 
| 11522 | 17  | 
end;  | 
18  | 
||
19  | 
structure Reconstruct : RECONSTRUCT =  | 
|
20  | 
struct  | 
|
21  | 
||
| 62922 | 22  | 
val quiet_mode =  | 
| 64556 | 23  | 
  Config.bool (Config.declare ("Reconstruct.quiet_mode", \<^here>) (K (Config.Bool true)));
 | 
| 62922 | 24  | 
|
25  | 
fun message ctxt msg =  | 
|
26  | 
if Config.get ctxt quiet_mode then () else writeln (msg ());  | 
|
| 11522 | 27  | 
|
| 28813 | 28  | 
fun vars_of t = map Var (rev (Term.add_vars t []));  | 
29  | 
fun frees_of t = map Free (rev (Term.add_frees t []));  | 
|
| 11522 | 30  | 
|
| 27330 | 31  | 
fun forall_intr_vfs prop = fold_rev Logic.all  | 
| 28813 | 32  | 
(vars_of prop @ frees_of prop) prop;  | 
| 11522 | 33  | 
|
| 37310 | 34  | 
fun forall_intr_vfs_prf prop prf = fold_rev Proofterm.forall_intr_proof'  | 
| 28813 | 35  | 
(vars_of prop @ frees_of prop) prf;  | 
| 12001 | 36  | 
|
| 11522 | 37  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
38  | 
(**** generate constraints for proof term ****)  | 
| 11522 | 39  | 
|
| 23178 | 40  | 
fun mk_var env Ts T =  | 
| 11522 | 41  | 
let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T)  | 
| 33245 | 42  | 
in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end;  | 
| 11522 | 43  | 
|
| 33245 | 44  | 
fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) =
 | 
45  | 
  (TVar (("'t", maxidx + 1), S),
 | 
|
46  | 
    Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv});
 | 
|
| 11522 | 47  | 
|
| 19473 | 48  | 
val mk_abs = fold (fn T => fn u => Abs ("", T, u));
 | 
| 11522 | 49  | 
|
| 62922 | 50  | 
fun unifyT ctxt env T U =  | 
| 11522 | 51  | 
let  | 
| 32032 | 52  | 
    val Envir.Envir {maxidx, tenv, tyenv} = env;
 | 
| 62922 | 53  | 
val (tyenv', maxidx') = Sign.typ_unify (Proof_Context.theory_of ctxt) (T, U) (tyenv, maxidx);  | 
| 32032 | 54  | 
  in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end;
 | 
| 11522 | 55  | 
|
| 32032 | 56  | 
fun chaseT env (T as TVar v) =  | 
57  | 
(case Type.lookup (Envir.type_env env) v of  | 
|
58  | 
NONE => T  | 
|
59  | 
| SOME T' => chaseT env T')  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
60  | 
| chaseT _ T = T;  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
61  | 
|
| 62922 | 62  | 
fun infer_type ctxt (env as Envir.Envir {maxidx, tenv, tyenv}) Ts vTs
 | 
| 32032 | 63  | 
(t as Const (s, T)) = if T = dummyT then  | 
| 62922 | 64  | 
(case Sign.const_type (Proof_Context.theory_of ctxt) s of  | 
| 15531 | 65  | 
          NONE => error ("reconstruct_proof: No such constant: " ^ quote s)
 | 
| 23178 | 66  | 
| SOME T =>  | 
| 19419 | 67  | 
let val T' = Type.strip_sorts (Logic.incr_tvar (maxidx + 1) T)  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
68  | 
in (Const (s, T'), T', vTs,  | 
| 32032 | 69  | 
              Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv})
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
70  | 
end)  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
71  | 
else (t, T, vTs, env)  | 
| 62922 | 72  | 
| infer_type ctxt env Ts vTs (t as Free (s, T)) =  | 
| 17412 | 73  | 
if T = dummyT then (case Symtab.lookup vTs s of  | 
| 15531 | 74  | 
NONE =>  | 
| 33245 | 75  | 
let val (T, env') = mk_tvar [] env  | 
| 17412 | 76  | 
in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end  | 
| 15531 | 77  | 
| SOME T => (Free (s, T), T, vTs, env))  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
78  | 
else (t, T, vTs, env)  | 
| 62922 | 79  | 
| infer_type ctxt env Ts vTs (Var _) = error "reconstruct_proof: internal error"  | 
80  | 
| infer_type ctxt env Ts vTs (Abs (s, T, t)) =  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
81  | 
let  | 
| 33245 | 82  | 
val (T', env') = if T = dummyT then mk_tvar [] env else (T, env);  | 
| 62922 | 83  | 
val (t', U, vTs', env'') = infer_type ctxt env' (T' :: Ts) vTs t  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
84  | 
in (Abs (s, T', t'), T' --> U, vTs', env'') end  | 
| 62922 | 85  | 
| infer_type ctxt env Ts vTs (t $ u) =  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
86  | 
let  | 
| 62922 | 87  | 
val (t', T, vTs1, env1) = infer_type ctxt env Ts vTs t;  | 
88  | 
val (u', U, vTs2, env2) = infer_type ctxt env1 Ts vTs1 u;  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
89  | 
in (case chaseT env2 T of  | 
| 62922 | 90  | 
          Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT ctxt env2 U U')
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
91  | 
| _ =>  | 
| 33245 | 92  | 
let val (V, env3) = mk_tvar [] env2  | 
| 62922 | 93  | 
in (t' $ u', V, vTs2, unifyT ctxt env3 T (U --> V)) end)  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
94  | 
end  | 
| 62922 | 95  | 
| infer_type ctxt env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env)  | 
| 43278 | 96  | 
      handle General.Subscript => error ("infer_type: bad variable index " ^ string_of_int i));
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
97  | 
|
| 62922 | 98  | 
fun cantunify ctxt (t, u) =  | 
99  | 
  error ("Non-unifiable terms:\n" ^
 | 
|
100  | 
Syntax.string_of_term ctxt t ^ "\n\n" ^ Syntax.string_of_term ctxt u);  | 
|
| 11522 | 101  | 
|
| 62922 | 102  | 
fun decompose ctxt Ts (p as (t, u)) env =  | 
| 33245 | 103  | 
let  | 
104  | 
fun rigrig (a, T) (b, U) uT ts us =  | 
|
| 62922 | 105  | 
if a <> b then cantunify ctxt p  | 
106  | 
else apfst flat (fold_map (decompose ctxt Ts) (ts ~~ us) (uT env T U))  | 
|
| 33245 | 107  | 
in  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
108  | 
case apply2 (strip_comb o Envir.head_norm env) p of  | 
| 62922 | 109  | 
((Const c, ts), (Const d, us)) => rigrig c d (unifyT ctxt) ts us  | 
110  | 
| ((Free c, ts), (Free d, us)) => rigrig c d (unifyT ctxt) ts us  | 
|
| 13715 | 111  | 
| ((Bound i, ts), (Bound j, us)) =>  | 
112  | 
rigrig (i, dummyT) (j, dummyT) (K o K) ts us  | 
|
113  | 
| ((Abs (_, T, t), []), (Abs (_, U, u), [])) =>  | 
|
| 62922 | 114  | 
decompose ctxt (T::Ts) (t, u) (unifyT ctxt env T U)  | 
| 13715 | 115  | 
| ((Abs (_, T, t), []), _) =>  | 
| 62922 | 116  | 
decompose ctxt (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env  | 
| 13715 | 117  | 
| (_, (Abs (_, T, u), [])) =>  | 
| 62922 | 118  | 
decompose ctxt (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env  | 
| 33245 | 119  | 
| _ => ([(mk_abs Ts t, mk_abs Ts u)], env)  | 
| 13715 | 120  | 
end;  | 
| 11522 | 121  | 
|
| 62922 | 122  | 
fun make_constraints_cprf ctxt env cprf =  | 
| 11522 | 123  | 
let  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
124  | 
fun add_cnstrt Ts prop prf cs env vTs (t, u) =  | 
| 11522 | 125  | 
let  | 
126  | 
val t' = mk_abs Ts t;  | 
|
| 12236 | 127  | 
val u' = mk_abs Ts u  | 
| 11522 | 128  | 
in  | 
| 62922 | 129  | 
(prop, prf, cs, Pattern.unify (Context.Proof ctxt) (t', u') env, vTs)  | 
| 12236 | 130  | 
handle Pattern.Pattern =>  | 
| 62922 | 131  | 
let val (cs', env') = decompose ctxt [] (t', u') env  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
132  | 
in (prop, prf, cs @ cs', env', vTs) end  | 
| 12236 | 133  | 
| Pattern.Unif =>  | 
| 62922 | 134  | 
cantunify ctxt (Envir.norm_term env t', Envir.norm_term env u')  | 
| 11522 | 135  | 
end;  | 
136  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
137  | 
fun mk_cnstrts_atom env vTs prop opTs prf =  | 
| 11522 | 138  | 
let  | 
| 
36042
 
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
 
krauss 
parents: 
33245 
diff
changeset
 | 
139  | 
val tvars = Term.add_tvars prop [] |> rev;  | 
| 
 
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
 
krauss 
parents: 
33245 
diff
changeset
 | 
140  | 
val tfrees = Term.add_tfrees prop [] |> rev;  | 
| 33245 | 141  | 
val (Ts, env') =  | 
| 31903 | 142  | 
(case opTs of  | 
| 33245 | 143  | 
NONE => fold_map mk_tvar (map snd tvars @ map snd tfrees) env  | 
144  | 
| SOME Ts => (Ts, env));  | 
|
| 28813 | 145  | 
val prop' = subst_atomic_types (map TVar tvars @ map TFree tfrees ~~ Ts)  | 
| 
40722
 
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
 
wenzelm 
parents: 
37310 
diff
changeset
 | 
146  | 
(forall_intr_vfs prop) handle ListPair.UnequalLengths =>  | 
| 37310 | 147  | 
                error ("Wrong number of type arguments for " ^ quote (Proofterm.guess_name prf))
 | 
148  | 
in (prop', Proofterm.change_type (SOME Ts) prf, [], env', vTs) end;  | 
|
| 11522 | 149  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
150  | 
fun head_norm (prop, prf, cnstrts, env, vTs) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
151  | 
(Envir.head_norm env prop, prf, cnstrts, env, vTs);  | 
| 23178 | 152  | 
|
| 30146 | 153  | 
fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs)  | 
| 43278 | 154  | 
          handle General.Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i))
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
155  | 
| 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
 | 
156  | 
let  | 
| 33245 | 157  | 
val (T, env') =  | 
158  | 
(case opT of  | 
|
159  | 
NONE => mk_tvar [] env  | 
|
160  | 
| SOME T => (T, env));  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
161  | 
val (t, prf, cnstrts, env'', vTs') =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
162  | 
mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf;  | 
| 62922 | 163  | 
in  | 
164  | 
            (Const ("Pure.all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf),
 | 
|
165  | 
cnstrts, env'', vTs')  | 
|
| 11522 | 166  | 
end  | 
| 15531 | 167  | 
| mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) =  | 
| 11522 | 168  | 
let  | 
| 62922 | 169  | 
val (t', _, vTs', env') = infer_type ctxt env Ts vTs t;  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
170  | 
val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf;  | 
| 15531 | 171  | 
in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'')  | 
| 11522 | 172  | 
end  | 
| 15531 | 173  | 
| mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) =  | 
| 11522 | 174  | 
let  | 
| 33245 | 175  | 
val (t, env') = mk_var env Ts propT;  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
176  | 
val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf;  | 
| 15531 | 177  | 
in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs')  | 
| 11522 | 178  | 
end  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
179  | 
| mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
180  | 
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
 | 
181  | 
in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of  | 
| 56245 | 182  | 
              (Const ("Pure.imp", _) $ u' $ t', prf1, cnstrts', env'', vTs'') =>
 | 
| 11613 | 183  | 
add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts)  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
184  | 
env'' vTs'' (u, u')  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
185  | 
| (t, prf1, cnstrts', env'', vTs'') =>  | 
| 33245 | 186  | 
let val (v, env''') = mk_var env'' Ts propT  | 
| 11613 | 187  | 
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
 | 
188  | 
env''' vTs'' (t, Logic.mk_implies (u, v))  | 
| 11522 | 189  | 
end)  | 
190  | 
end  | 
|
| 15531 | 191  | 
| mk_cnstrts env Ts Hs vTs (cprf % SOME t) =  | 
| 62922 | 192  | 
let val (t', U, vTs1, env1) = infer_type ctxt env Ts vTs t  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
193  | 
in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of  | 
| 56245 | 194  | 
             (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
195  | 
prf, cnstrts, env2, vTs2) =>  | 
| 62922 | 196  | 
let val env3 = unifyT ctxt env2 T U  | 
| 15531 | 197  | 
in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2)  | 
| 11522 | 198  | 
end  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
199  | 
| (u, prf, cnstrts, env2, vTs2) =>  | 
| 33245 | 200  | 
let val (v, env3) = mk_var env2 Ts (U --> propT);  | 
| 11522 | 201  | 
in  | 
| 15531 | 202  | 
add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2  | 
| 56245 | 203  | 
                   (u, Const ("Pure.all", (U --> propT) --> propT) $ v)
 | 
| 11522 | 204  | 
end)  | 
205  | 
end  | 
|
| 15531 | 206  | 
| mk_cnstrts env Ts Hs vTs (cprf % NONE) =  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
207  | 
(case head_norm (mk_cnstrts env Ts Hs vTs cprf) of  | 
| 56245 | 208  | 
             (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f,
 | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
209  | 
prf, cnstrts, env', vTs') =>  | 
| 33245 | 210  | 
let val (t, env'') = mk_var env' Ts T  | 
| 15531 | 211  | 
in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs')  | 
| 11522 | 212  | 
end  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
213  | 
| (u, prf, cnstrts, env', vTs') =>  | 
| 11522 | 214  | 
let  | 
| 33245 | 215  | 
val (T, env1) = mk_tvar [] env';  | 
216  | 
val (v, env2) = mk_var env1 Ts (T --> propT);  | 
|
217  | 
val (t, env3) = mk_var env2 Ts T  | 
|
| 11522 | 218  | 
in  | 
| 15531 | 219  | 
add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs'  | 
| 56245 | 220  | 
                   (u, Const ("Pure.all", (T --> propT) --> propT) $ v)
 | 
| 11522 | 221  | 
end)  | 
| 28808 | 222  | 
| mk_cnstrts env _ _ vTs (prf as PThm (_, ((_, prop, opTs), _))) =  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
223  | 
mk_cnstrts_atom env vTs prop opTs prf  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
224  | 
| mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
225  | 
mk_cnstrts_atom env vTs prop opTs prf  | 
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31903 
diff
changeset
 | 
226  | 
| mk_cnstrts env _ _ vTs (prf as OfClass (T, c)) =  | 
| 
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31903 
diff
changeset
 | 
227  | 
mk_cnstrts_atom env vTs (Logic.mk_of_class (T, c)) NONE prf  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
228  | 
| mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) =  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
229  | 
mk_cnstrts_atom env vTs prop opTs prf  | 
| 
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
230  | 
| mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs)  | 
| 11613 | 231  | 
| mk_cnstrts _ _ _ _ _ = error "reconstruct_proof: minimal proof object"  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
232  | 
in mk_cnstrts env [] [] Symtab.empty cprf end;  | 
| 11522 | 233  | 
|
234  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
235  | 
(**** update list of free variables of constraints ****)  | 
| 11522 | 236  | 
|
237  | 
fun upd_constrs env cs =  | 
|
238  | 
let  | 
|
| 32032 | 239  | 
val tenv = Envir.term_env env;  | 
240  | 
val tyenv = Envir.type_env env;  | 
|
| 20675 | 241  | 
val dom = []  | 
| 32032 | 242  | 
|> Vartab.fold (cons o #1) tenv  | 
243  | 
|> Vartab.fold (cons o #1) tyenv;  | 
|
| 20675 | 244  | 
val vran = []  | 
| 32032 | 245  | 
|> Vartab.fold (Term.add_var_names o #2 o #2) tenv  | 
246  | 
|> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv;  | 
|
| 11522 | 247  | 
fun check_cs [] = []  | 
| 32032 | 248  | 
| check_cs ((u, p, vs) :: ps) =  | 
249  | 
let val vs' = subtract (op =) dom vs in  | 
|
250  | 
if vs = vs' then (u, p, vs) :: check_cs ps  | 
|
251  | 
else (true, p, fold (insert op =) vs' vran) :: check_cs ps  | 
|
252  | 
end;  | 
|
| 11522 | 253  | 
in check_cs cs end;  | 
254  | 
||
| 32032 | 255  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
256  | 
(**** solution of constraints ****)  | 
| 11522 | 257  | 
|
258  | 
fun solve _ [] bigenv = bigenv  | 
|
| 62922 | 259  | 
| solve ctxt cs bigenv =  | 
| 11522 | 260  | 
let  | 
| 11660 | 261  | 
        fun search env [] = error ("Unsolvable constraints:\n" ^
 | 
262  | 
Pretty.string_of (Pretty.chunks (map (fn (_, p, _) =>  | 
|
| 62922 | 263  | 
Thm.pretty_flexpair ctxt (apply2 (Envir.norm_term bigenv) p)) cs)))  | 
| 11522 | 264  | 
| search env ((u, p as (t1, t2), vs)::ps) =  | 
265  | 
if u then  | 
|
266  | 
let  | 
|
267  | 
val tn1 = Envir.norm_term bigenv t1;  | 
|
268  | 
val tn2 = Envir.norm_term bigenv t2  | 
|
269  | 
in  | 
|
270  | 
if Pattern.pattern tn1 andalso Pattern.pattern tn2 then  | 
|
| 62922 | 271  | 
(Pattern.unify (Context.Proof ctxt) (tn1, tn2) env, ps) handle Pattern.Unif =>  | 
272  | 
cantunify ctxt (tn1, tn2)  | 
|
| 11522 | 273  | 
else  | 
| 62922 | 274  | 
let val (cs', env') = decompose ctxt [] (tn1, tn2) env  | 
| 11522 | 275  | 
in if cs' = [(tn1, tn2)] then  | 
276  | 
apsnd (cons (false, (tn1, tn2), vs)) (search env ps)  | 
|
277  | 
else search env' (map (fn q => (true, q, vs)) cs' @ ps)  | 
|
278  | 
end  | 
|
279  | 
end  | 
|
280  | 
else apsnd (cons (false, p, vs)) (search env ps);  | 
|
281  | 
        val Envir.Envir {maxidx, ...} = bigenv;
 | 
|
282  | 
val (env, cs') = search (Envir.empty maxidx) cs;  | 
|
283  | 
in  | 
|
| 62922 | 284  | 
solve ctxt (upd_constrs env cs') (Envir.merge (bigenv, env))  | 
| 11522 | 285  | 
end;  | 
286  | 
||
287  | 
||
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
288  | 
(**** reconstruction of proofs ****)  | 
| 11522 | 289  | 
|
| 62922 | 290  | 
fun reconstruct_proof ctxt prop cprf =  | 
| 11522 | 291  | 
let  | 
| 37310 | 292  | 
val (cprf' % SOME prop', thawf) = Proofterm.freeze_thaw_prf (cprf % SOME prop);  | 
| 62922 | 293  | 
val _ = message ctxt (fn _ => "Collecting constraints ...");  | 
294  | 
val (t, prf, cs, env, _) = make_constraints_cprf ctxt  | 
|
| 37310 | 295  | 
(Envir.empty (Proofterm.maxidx_proof cprf ~1)) cprf';  | 
| 
44119
 
caddb5264048
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
44060 
diff
changeset
 | 
296  | 
val cs' =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58950 
diff
changeset
 | 
297  | 
map (apply2 (Envir.norm_term env)) ((t, prop') :: cs)  | 
| 
44119
 
caddb5264048
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
44060 
diff
changeset
 | 
298  | 
|> map (fn p => (true, p, Term.add_var_names (#1 p) (Term.add_var_names (#2 p) [])));  | 
| 62922 | 299  | 
val _ =  | 
300  | 
message ctxt  | 
|
301  | 
        (fn () => "Solving remaining constraints (" ^ string_of_int (length cs') ^ ") ...");
 | 
|
302  | 
val env' = solve ctxt cs' env  | 
|
| 11522 | 303  | 
in  | 
| 37310 | 304  | 
thawf (Proofterm.norm_proof env' prf)  | 
| 11522 | 305  | 
end;  | 
306  | 
||
| 28813 | 307  | 
fun prop_of_atom prop Ts = subst_atomic_types  | 
| 
36042
 
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
 
krauss 
parents: 
33245 
diff
changeset
 | 
308  | 
(map TVar (Term.add_tvars prop [] |> rev) @ map TFree (Term.add_tfrees prop [] |> rev) ~~ Ts)  | 
| 28813 | 309  | 
(forall_intr_vfs prop);  | 
| 13138 | 310  | 
|
| 63616 | 311  | 
val head_norm = Envir.head_norm Envir.init;  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
312  | 
|
| 30146 | 313  | 
fun prop_of0 Hs (PBound i) = nth Hs i  | 
| 15531 | 314  | 
| prop_of0 Hs (Abst (s, SOME T, prf)) =  | 
| 
46217
 
7b19666f0e3d
renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
 
wenzelm 
parents: 
44119 
diff
changeset
 | 
315  | 
Logic.all_const T $ (Abs (s, T, prop_of0 Hs prf))  | 
| 15531 | 316  | 
| 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
 | 
317  | 
Logic.mk_implies (t, prop_of0 (t :: Hs) prf)  | 
| 15531 | 318  | 
| prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of  | 
| 56245 | 319  | 
      Const ("Pure.all", _) $ f => f $ t
 | 
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
320  | 
| _ => error "prop_of: all expected")  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
321  | 
| prop_of0 Hs (prf1 %% prf2) = (case head_norm (prop_of0 Hs prf1) of  | 
| 56245 | 322  | 
      Const ("Pure.imp", _) $ P $ Q => Q
 | 
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
323  | 
| _ => error "prop_of: ==> expected")  | 
| 
13734
 
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
 
berghofe 
parents: 
13715 
diff
changeset
 | 
324  | 
| prop_of0 Hs (Hyp t) = t  | 
| 28808 | 325  | 
| prop_of0 Hs (PThm (_, ((_, prop, SOME Ts), _))) = prop_of_atom prop Ts  | 
| 15531 | 326  | 
| prop_of0 Hs (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts  | 
| 
31943
 
5e960a0780a2
renamed inclass/Inclass to of_class/OfClass, in accordance to of_sort;
 
wenzelm 
parents: 
31903 
diff
changeset
 | 
327  | 
| prop_of0 Hs (OfClass (T, c)) = Logic.mk_of_class (T, c)  | 
| 15531 | 328  | 
| 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
 | 
329  | 
| prop_of0 _ _ = error "prop_of: partial proof object";  | 
| 13138 | 330  | 
|
| 18929 | 331  | 
val prop_of' = Envir.beta_eta_contract oo prop_of0;  | 
| 
13256
 
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
 
berghofe 
parents: 
13138 
diff
changeset
 | 
332  | 
val prop_of = prop_of' [];  | 
| 13138 | 333  | 
|
| 62922 | 334  | 
fun proof_of ctxt raw_thm =  | 
| 67649 | 335  | 
let val thm = Thm.transfer' ctxt raw_thm  | 
| 62922 | 336  | 
in reconstruct_proof ctxt (Thm.prop_of thm) (Thm.proof_of thm) end;  | 
| 44060 | 337  | 
|
338  | 
||
| 11522 | 339  | 
|
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
340  | 
(**** expand and reconstruct subproofs ****)  | 
| 11522 | 341  | 
|
| 62922 | 342  | 
fun expand_proof ctxt thms prf =  | 
| 11522 | 343  | 
let  | 
| 23178 | 344  | 
fun expand maxidx prfs (AbsP (s, t, prf)) =  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
345  | 
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
 | 
346  | 
in (maxidx', prfs', AbsP (s, t, prf')) end  | 
| 23178 | 347  | 
| expand maxidx prfs (Abst (s, T, prf)) =  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
348  | 
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
 | 
349  | 
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
 | 
350  | 
| expand maxidx prfs (prf1 %% prf2) =  | 
| 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', prfs', prf1') = expand maxidx prfs prf1;  | 
| 
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
353  | 
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
 | 
354  | 
in (maxidx'', prfs'', prf1' %% prf2') end  | 
| 
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
355  | 
| expand maxidx prfs (prf % t) =  | 
| 
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
356  | 
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
 | 
357  | 
in (maxidx', prfs', prf' % t) end  | 
| 28808 | 358  | 
| expand maxidx prfs (prf as PThm (_, ((a, prop, SOME Ts), body))) =  | 
| 
13342
 
915d4d004643
expand_proof now also takes an optional term describing the proposition
 
berghofe 
parents: 
13256 
diff
changeset
 | 
359  | 
if not (exists  | 
| 15531 | 360  | 
(fn (b, NONE) => a = b  | 
361  | 
| (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
 | 
362  | 
then (maxidx, prfs, prf) else  | 
| 11522 | 363  | 
let  | 
| 16862 | 364  | 
val (maxidx', prf, prfs') =  | 
| 17232 | 365  | 
(case AList.lookup (op =) prfs (a, prop) of  | 
| 15531 | 366  | 
NONE =>  | 
| 11522 | 367  | 
let  | 
| 44059 | 368  | 
val _ =  | 
| 62922 | 369  | 
message ctxt (fn () =>  | 
370  | 
"Reconstructing proof of " ^ a ^ "\n" ^ Syntax.string_of_term ctxt prop);  | 
|
| 
13610
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
371  | 
val prf' = forall_intr_vfs_prf prop  | 
| 62922 | 372  | 
(reconstruct_proof ctxt prop (Proofterm.join_proof body));  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
373  | 
val (maxidx', prfs', prf) = expand  | 
| 37310 | 374  | 
(Proofterm.maxidx_proof prf' ~1) prfs prf'  | 
375  | 
in (maxidx' + maxidx + 1, Proofterm.incr_indexes (maxidx + 1) prf,  | 
|
| 
13610
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
376  | 
((a, prop), (maxidx', prf)) :: prfs')  | 
| 
 
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
 
berghofe 
parents: 
13342 
diff
changeset
 | 
377  | 
end  | 
| 15531 | 378  | 
| SOME (maxidx', prf) => (maxidx' + maxidx + 1,  | 
| 37310 | 379  | 
Proofterm.incr_indexes (maxidx + 1) prf, prfs));  | 
| 
36042
 
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
 
krauss 
parents: 
33245 
diff
changeset
 | 
380  | 
val tfrees = Term.add_tfrees prop [] |> rev;  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
381  | 
val tye = map (fn ((s, j), _) => (s, maxidx + 1 + j))  | 
| 
36042
 
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
 
krauss 
parents: 
33245 
diff
changeset
 | 
382  | 
(Term.add_tvars prop [] |> rev) @ map (rpair ~1 o fst) tfrees ~~ Ts;  | 
| 
13669
 
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
 
berghofe 
parents: 
13610 
diff
changeset
 | 
383  | 
val varify = map_type_tfree (fn p as (a, S) =>  | 
| 20675 | 384  | 
if member (op =) tfrees p then TVar ((a, ~1), S) else TFree p)  | 
| 11522 | 385  | 
in  | 
| 37310 | 386  | 
(maxidx', prfs', Proofterm.map_proof_types (typ_subst_TVars tye o varify) prf)  | 
| 11522 | 387  | 
end  | 
| 
12870
 
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
 
berghofe 
parents: 
12527 
diff
changeset
 | 
388  | 
| expand maxidx prfs prf = (maxidx, prfs, prf);  | 
| 11522 | 389  | 
|
| 37310 | 390  | 
in #3 (expand (Proofterm.maxidx_proof prf ~1) [] prf) end;  | 
| 11522 | 391  | 
|
| 64986 | 392  | 
|
393  | 
(* cleanup for output etc. *)  | 
|
394  | 
||
395  | 
fun clean_proof_of ctxt full thm =  | 
|
396  | 
let  | 
|
397  | 
val (_, prop) =  | 
|
398  | 
Logic.unconstrainT (Thm.shyps_of thm)  | 
|
399  | 
(Logic.list_implies (Thm.hyps_of thm, Thm.prop_of thm));  | 
|
400  | 
in  | 
|
401  | 
Proofterm.proof_of (Proofterm.strip_thm (Thm.proof_body_of thm))  | 
|
402  | 
|> reconstruct_proof ctxt prop  | 
|
403  | 
    |> expand_proof ctxt [("", NONE)]
 | 
|
404  | 
|> Proofterm.rew_proof (Proof_Context.theory_of ctxt)  | 
|
405  | 
|> Proofterm.no_thm_proofs  | 
|
406  | 
|> not full ? Proofterm.shrink_proof  | 
|
407  | 
end;  | 
|
408  | 
||
| 11522 | 409  | 
end;  |