author | wenzelm |
Tue, 30 Jul 2019 13:22:29 +0200 | |
changeset 70446 | 609b10dc1a7d |
parent 70445 | 5b3f8a64e0f4 |
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 |
|
70443 | 9 |
val reconstruct_proof: Proof.context -> term -> Proofterm.proof -> Proofterm.proof |
10 |
val prop_of': term list -> Proofterm.proof -> term |
|
11 |
val prop_of: Proofterm.proof -> term |
|
12 |
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
|
13 |
Proofterm.proof -> Proofterm.proof |
11522 | 14 |
end; |
15 |
||
16 |
structure Reconstruct : RECONSTRUCT = |
|
17 |
struct |
|
18 |
||
70397 | 19 |
(* generate constraints for proof term *) |
11522 | 20 |
|
23178 | 21 |
fun mk_var env Ts T = |
11522 | 22 |
let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) |
33245 | 23 |
in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end; |
11522 | 24 |
|
33245 | 25 |
fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) = |
26 |
(TVar (("'t", maxidx + 1), S), |
|
27 |
Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}); |
|
11522 | 28 |
|
19473 | 29 |
val mk_abs = fold (fn T => fn u => Abs ("", T, u)); |
11522 | 30 |
|
62922 | 31 |
fun unifyT ctxt env T U = |
11522 | 32 |
let |
32032 | 33 |
val Envir.Envir {maxidx, tenv, tyenv} = env; |
62922 | 34 |
val (tyenv', maxidx') = Sign.typ_unify (Proof_Context.theory_of ctxt) (T, U) (tyenv, maxidx); |
32032 | 35 |
in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end; |
11522 | 36 |
|
32032 | 37 |
fun chaseT env (T as TVar v) = |
38 |
(case Type.lookup (Envir.type_env env) v of |
|
39 |
NONE => T |
|
40 |
| SOME T' => chaseT env T') |
|
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
41 |
| chaseT _ T = T; |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
42 |
|
70446 | 43 |
fun infer_type ctxt (env as Envir.Envir {maxidx, tenv, tyenv}) _ vTs |
32032 | 44 |
(t as Const (s, T)) = if T = dummyT then |
62922 | 45 |
(case Sign.const_type (Proof_Context.theory_of ctxt) s of |
15531 | 46 |
NONE => error ("reconstruct_proof: No such constant: " ^ quote s) |
23178 | 47 |
| SOME T => |
19419 | 48 |
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
|
49 |
in (Const (s, T'), T', vTs, |
32032 | 50 |
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
|
51 |
end) |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
52 |
else (t, T, vTs, env) |
70446 | 53 |
| infer_type _ env _ vTs (t as Free (s, T)) = |
17412 | 54 |
if T = dummyT then (case Symtab.lookup vTs s of |
15531 | 55 |
NONE => |
33245 | 56 |
let val (T, env') = mk_tvar [] env |
17412 | 57 |
in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end |
15531 | 58 |
| SOME T => (Free (s, T), T, vTs, env)) |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
59 |
else (t, T, vTs, env) |
70446 | 60 |
| infer_type _ _ _ _ (Var _) = error "reconstruct_proof: internal error" |
62922 | 61 |
| 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
|
62 |
let |
33245 | 63 |
val (T', env') = if T = dummyT then mk_tvar [] env else (T, env); |
62922 | 64 |
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
|
65 |
in (Abs (s, T', t'), T' --> U, vTs', env'') end |
62922 | 66 |
| infer_type ctxt env Ts vTs (t $ u) = |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
67 |
let |
62922 | 68 |
val (t', T, vTs1, env1) = infer_type ctxt env Ts vTs t; |
69 |
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
|
70 |
in (case chaseT env2 T of |
62922 | 71 |
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
|
72 |
| _ => |
33245 | 73 |
let val (V, env3) = mk_tvar [] env2 |
62922 | 74 |
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
|
75 |
end |
70446 | 76 |
| infer_type _ env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env) |
43278 | 77 |
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
|
78 |
|
62922 | 79 |
fun cantunify ctxt (t, u) = |
80 |
error ("Non-unifiable terms:\n" ^ |
|
81 |
Syntax.string_of_term ctxt t ^ "\n\n" ^ Syntax.string_of_term ctxt u); |
|
11522 | 82 |
|
62922 | 83 |
fun decompose ctxt Ts (p as (t, u)) env = |
33245 | 84 |
let |
85 |
fun rigrig (a, T) (b, U) uT ts us = |
|
62922 | 86 |
if a <> b then cantunify ctxt p |
87 |
else apfst flat (fold_map (decompose ctxt Ts) (ts ~~ us) (uT env T U)) |
|
33245 | 88 |
in |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58950
diff
changeset
|
89 |
case apply2 (strip_comb o Envir.head_norm env) p of |
62922 | 90 |
((Const c, ts), (Const d, us)) => rigrig c d (unifyT ctxt) ts us |
91 |
| ((Free c, ts), (Free d, us)) => rigrig c d (unifyT ctxt) ts us |
|
13715 | 92 |
| ((Bound i, ts), (Bound j, us)) => |
93 |
rigrig (i, dummyT) (j, dummyT) (K o K) ts us |
|
94 |
| ((Abs (_, T, t), []), (Abs (_, U, u), [])) => |
|
62922 | 95 |
decompose ctxt (T::Ts) (t, u) (unifyT ctxt env T U) |
13715 | 96 |
| ((Abs (_, T, t), []), _) => |
62922 | 97 |
decompose ctxt (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env |
13715 | 98 |
| (_, (Abs (_, T, u), [])) => |
62922 | 99 |
decompose ctxt (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env |
33245 | 100 |
| _ => ([(mk_abs Ts t, mk_abs Ts u)], env) |
13715 | 101 |
end; |
11522 | 102 |
|
62922 | 103 |
fun make_constraints_cprf ctxt env cprf = |
11522 | 104 |
let |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
105 |
fun add_cnstrt Ts prop prf cs env vTs (t, u) = |
11522 | 106 |
let |
107 |
val t' = mk_abs Ts t; |
|
12236 | 108 |
val u' = mk_abs Ts u |
11522 | 109 |
in |
62922 | 110 |
(prop, prf, cs, Pattern.unify (Context.Proof ctxt) (t', u') env, vTs) |
12236 | 111 |
handle Pattern.Pattern => |
62922 | 112 |
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
|
113 |
in (prop, prf, cs @ cs', env', vTs) end |
12236 | 114 |
| Pattern.Unif => |
62922 | 115 |
cantunify ctxt (Envir.norm_term env t', Envir.norm_term env u') |
11522 | 116 |
end; |
117 |
||
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
118 |
fun mk_cnstrts_atom env vTs prop opTs prf = |
11522 | 119 |
let |
36042
85efdadee8ae
switched PThm/PAxm etc. to use canonical order of type variables (term variables unchanged)
krauss
parents:
33245
diff
changeset
|
120 |
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
|
121 |
val tfrees = Term.add_tfrees prop [] |> rev; |
33245 | 122 |
val (Ts, env') = |
31903 | 123 |
(case opTs of |
33245 | 124 |
NONE => fold_map mk_tvar (map snd tvars @ map snd tfrees) env |
125 |
| SOME Ts => (Ts, env)); |
|
28813 | 126 |
val prop' = subst_atomic_types (map TVar tvars @ map TFree tfrees ~~ Ts) |
70400
65bbef66e2ec
clarified treatment of unnamed PThm nodes (from close_derivation): retain full proof, publish when named;
wenzelm
parents:
70397
diff
changeset
|
127 |
(Proofterm.forall_intr_vfs prop) handle ListPair.UnequalLengths => |
37310 | 128 |
error ("Wrong number of type arguments for " ^ quote (Proofterm.guess_name prf)) |
70417 | 129 |
in (prop', Proofterm.change_types (SOME Ts) prf, [], env', vTs) end; |
11522 | 130 |
|
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
131 |
fun head_norm (prop, prf, cnstrts, env, vTs) = |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
132 |
(Envir.head_norm env prop, prf, cnstrts, env, vTs); |
23178 | 133 |
|
30146 | 134 |
fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs) |
43278 | 135 |
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
|
136 |
| 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
|
137 |
let |
33245 | 138 |
val (T, env') = |
139 |
(case opT of |
|
140 |
NONE => mk_tvar [] env |
|
141 |
| SOME T => (T, env)); |
|
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
142 |
val (t, prf, cnstrts, env'', vTs') = |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
143 |
mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf; |
62922 | 144 |
in |
145 |
(Const ("Pure.all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf), |
|
146 |
cnstrts, env'', vTs') |
|
11522 | 147 |
end |
15531 | 148 |
| mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) = |
11522 | 149 |
let |
62922 | 150 |
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
|
151 |
val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf; |
15531 | 152 |
in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'') |
11522 | 153 |
end |
15531 | 154 |
| mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) = |
11522 | 155 |
let |
33245 | 156 |
val (t, env') = mk_var env Ts propT; |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
157 |
val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf; |
15531 | 158 |
in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs') |
11522 | 159 |
end |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
160 |
| mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) = |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
161 |
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
|
162 |
in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of |
56245 | 163 |
(Const ("Pure.imp", _) $ u' $ t', prf1, cnstrts', env'', vTs'') => |
11613 | 164 |
add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts) |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
165 |
env'' vTs'' (u, u') |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
166 |
| (t, prf1, cnstrts', env'', vTs'') => |
33245 | 167 |
let val (v, env''') = mk_var env'' Ts propT |
11613 | 168 |
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
|
169 |
env''' vTs'' (t, Logic.mk_implies (u, v)) |
11522 | 170 |
end) |
171 |
end |
|
15531 | 172 |
| mk_cnstrts env Ts Hs vTs (cprf % SOME t) = |
62922 | 173 |
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
|
174 |
in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of |
56245 | 175 |
(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
|
176 |
prf, cnstrts, env2, vTs2) => |
62922 | 177 |
let val env3 = unifyT ctxt env2 T U |
15531 | 178 |
in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2) |
11522 | 179 |
end |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
180 |
| (u, prf, cnstrts, env2, vTs2) => |
33245 | 181 |
let val (v, env3) = mk_var env2 Ts (U --> propT); |
11522 | 182 |
in |
15531 | 183 |
add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2 |
56245 | 184 |
(u, Const ("Pure.all", (U --> propT) --> propT) $ v) |
11522 | 185 |
end) |
186 |
end |
|
15531 | 187 |
| mk_cnstrts env Ts Hs vTs (cprf % NONE) = |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
188 |
(case head_norm (mk_cnstrts env Ts Hs vTs cprf) of |
56245 | 189 |
(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
|
190 |
prf, cnstrts, env', vTs') => |
33245 | 191 |
let val (t, env'') = mk_var env' Ts T |
15531 | 192 |
in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs') |
11522 | 193 |
end |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
194 |
| (u, prf, cnstrts, env', vTs') => |
11522 | 195 |
let |
33245 | 196 |
val (T, env1) = mk_tvar [] env'; |
197 |
val (v, env2) = mk_var env1 Ts (T --> propT); |
|
198 |
val (t, env3) = mk_var env2 Ts T |
|
11522 | 199 |
in |
15531 | 200 |
add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs' |
56245 | 201 |
(u, Const ("Pure.all", (T --> propT) --> propT) $ v) |
11522 | 202 |
end) |
70412
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
203 |
| 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
|
204 |
mk_cnstrts_atom env vTs prop opTs prf |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
205 |
| mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) = |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
206 |
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
|
207 |
| 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
|
208 |
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
|
209 |
| mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) = |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
210 |
mk_cnstrts_atom env vTs prop opTs prf |
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
211 |
| mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs) |
11613 | 212 |
| mk_cnstrts _ _ _ _ _ = error "reconstruct_proof: minimal proof object" |
13669
a9f229eafba7
- reconstruct_proof no longer relies on TypeInfer.infer_types
berghofe
parents:
13610
diff
changeset
|
213 |
in mk_cnstrts env [] [] Symtab.empty cprf end; |
11522 | 214 |
|
215 |
||
70397 | 216 |
(* update list of free variables of constraints *) |
11522 | 217 |
|
218 |
fun upd_constrs env cs = |
|
219 |
let |
|
32032 | 220 |
val tenv = Envir.term_env env; |
221 |
val tyenv = Envir.type_env env; |
|
20675 | 222 |
val dom = [] |
32032 | 223 |
|> Vartab.fold (cons o #1) tenv |
224 |
|> Vartab.fold (cons o #1) tyenv; |
|
20675 | 225 |
val vran = [] |
32032 | 226 |
|> Vartab.fold (Term.add_var_names o #2 o #2) tenv |
227 |
|> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv; |
|
11522 | 228 |
fun check_cs [] = [] |
32032 | 229 |
| check_cs ((u, p, vs) :: ps) = |
230 |
let val vs' = subtract (op =) dom vs in |
|
231 |
if vs = vs' then (u, p, vs) :: check_cs ps |
|
232 |
else (true, p, fold (insert op =) vs' vran) :: check_cs ps |
|
233 |
end; |
|
11522 | 234 |
in check_cs cs end; |
235 |
||
32032 | 236 |
|
70397 | 237 |
(* solution of constraints *) |
11522 | 238 |
|
239 |
fun solve _ [] bigenv = bigenv |
|
62922 | 240 |
| solve ctxt cs bigenv = |
11522 | 241 |
let |
70446 | 242 |
fun search _ [] = error ("Unsolvable constraints:\n" ^ |
11660 | 243 |
Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => |
70443 | 244 |
Syntax.pretty_flexpair ctxt (apply2 (Envir.norm_term bigenv) p)) cs))) |
11522 | 245 |
| search env ((u, p as (t1, t2), vs)::ps) = |
246 |
if u then |
|
247 |
let |
|
248 |
val tn1 = Envir.norm_term bigenv t1; |
|
249 |
val tn2 = Envir.norm_term bigenv t2 |
|
250 |
in |
|
251 |
if Pattern.pattern tn1 andalso Pattern.pattern tn2 then |
|
62922 | 252 |
(Pattern.unify (Context.Proof ctxt) (tn1, tn2) env, ps) handle Pattern.Unif => |
253 |
cantunify ctxt (tn1, tn2) |
|
11522 | 254 |
else |
62922 | 255 |
let val (cs', env') = decompose ctxt [] (tn1, tn2) env |
11522 | 256 |
in if cs' = [(tn1, tn2)] then |
257 |
apsnd (cons (false, (tn1, tn2), vs)) (search env ps) |
|
258 |
else search env' (map (fn q => (true, q, vs)) cs' @ ps) |
|
259 |
end |
|
260 |
end |
|
261 |
else apsnd (cons (false, p, vs)) (search env ps); |
|
262 |
val Envir.Envir {maxidx, ...} = bigenv; |
|
263 |
val (env, cs') = search (Envir.empty maxidx) cs; |
|
264 |
in |
|
62922 | 265 |
solve ctxt (upd_constrs env cs') (Envir.merge (bigenv, env)) |
11522 | 266 |
end; |
267 |
||
268 |
||
70397 | 269 |
(* reconstruction of proofs *) |
11522 | 270 |
|
62922 | 271 |
fun reconstruct_proof ctxt prop cprf = |
11522 | 272 |
let |
37310 | 273 |
val (cprf' % SOME prop', thawf) = Proofterm.freeze_thaw_prf (cprf % SOME prop); |
62922 | 274 |
val (t, prf, cs, env, _) = make_constraints_cprf ctxt |
37310 | 275 |
(Envir.empty (Proofterm.maxidx_proof cprf ~1)) cprf'; |
44119
caddb5264048
avoid OldTerm operations -- with subtle changes of semantics;
wenzelm
parents:
44060
diff
changeset
|
276 |
val cs' = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58950
diff
changeset
|
277 |
map (apply2 (Envir.norm_term env)) ((t, prop') :: cs) |
44119
caddb5264048
avoid OldTerm operations -- with subtle changes of semantics;
wenzelm
parents:
44060
diff
changeset
|
278 |
|> map (fn p => (true, p, Term.add_var_names (#1 p) (Term.add_var_names (#2 p) []))); |
62922 | 279 |
val env' = solve ctxt cs' env |
11522 | 280 |
in |
37310 | 281 |
thawf (Proofterm.norm_proof env' prf) |
11522 | 282 |
end; |
283 |
||
28813 | 284 |
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
|
285 |
(map TVar (Term.add_tvars prop [] |> rev) @ map TFree (Term.add_tfrees prop [] |> rev) ~~ Ts) |
70400
65bbef66e2ec
clarified treatment of unnamed PThm nodes (from close_derivation): retain full proof, publish when named;
wenzelm
parents:
70397
diff
changeset
|
286 |
(Proofterm.forall_intr_vfs prop); |
13138 | 287 |
|
63616 | 288 |
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
|
289 |
|
30146 | 290 |
fun prop_of0 Hs (PBound i) = nth Hs i |
15531 | 291 |
| 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
|
292 |
Logic.all_const T $ (Abs (s, T, prop_of0 Hs prf)) |
70446 | 293 |
| prop_of0 Hs (AbsP (_, SOME t, prf)) = |
13734
50dcee1c509e
prop_of now returns proposition in beta-eta normal form.
berghofe
parents:
13715
diff
changeset
|
294 |
Logic.mk_implies (t, prop_of0 (t :: Hs) prf) |
15531 | 295 |
| prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of |
56245 | 296 |
Const ("Pure.all", _) $ f => f $ t |
13256
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
297 |
| _ => error "prop_of: all expected") |
70446 | 298 |
| prop_of0 Hs (prf1 %% _) = (case head_norm (prop_of0 Hs prf1) of |
299 |
Const ("Pure.imp", _) $ _ $ Q => Q |
|
13256
cf85c4f7dcf2
Added function prop_of' taking assumption context as an argument.
berghofe
parents:
13138
diff
changeset
|
300 |
| _ => error "prop_of: ==> expected") |
70446 | 301 |
| prop_of0 _ (Hyp t) = t |
302 |
| prop_of0 _ (PThm (_, ((_, prop, SOME Ts, _), _))) = prop_of_atom prop Ts |
|
303 |
| prop_of0 _ (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts |
|
304 |
| prop_of0 _ (OfClass (T, c)) = Logic.mk_of_class (T, c) |
|
305 |
| prop_of0 _ (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
|
306 |
| prop_of0 _ _ = error "prop_of: partial proof object"; |
13138 | 307 |
|
18929 | 308 |
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
|
309 |
val prop_of = prop_of' []; |
13138 | 310 |
|
44060 | 311 |
|
70397 | 312 |
(* expand and reconstruct subproofs *) |
11522 | 313 |
|
62922 | 314 |
fun expand_proof ctxt thms prf = |
11522 | 315 |
let |
23178 | 316 |
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
|
317 |
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
|
318 |
in (maxidx', prfs', AbsP (s, t, prf')) end |
23178 | 319 |
| 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
|
320 |
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
|
321 |
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
|
322 |
| expand maxidx prfs (prf1 %% prf2) = |
11522 | 323 |
let |
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
324 |
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
|
325 |
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
|
326 |
in (maxidx'', prfs'', prf1' %% prf2') end |
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
327 |
| expand maxidx prfs (prf % t) = |
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
328 |
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
|
329 |
in (maxidx', prfs', prf' % t) end |
70412
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
330 |
| expand maxidx prfs (prf as PThm (_, ((a, prop, SOME Ts, open_proof), body))) = |
13342
915d4d004643
expand_proof now also takes an optional term describing the proposition
berghofe
parents:
13256
diff
changeset
|
331 |
if not (exists |
15531 | 332 |
(fn (b, NONE) => a = b |
333 |
| (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
|
334 |
then (maxidx, prfs, prf) else |
11522 | 335 |
let |
16862 | 336 |
val (maxidx', prf, prfs') = |
17232 | 337 |
(case AList.lookup (op =) prfs (a, prop) of |
15531 | 338 |
NONE => |
11522 | 339 |
let |
70412
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
340 |
val prf' = |
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
341 |
Proofterm.join_proof body |
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
342 |
|> open_proof |
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
343 |
|> reconstruct_proof ctxt prop |
64ead6de6212
defer rew_proof on unnamed PThm node as open_proof operation: significant performance improvement;
wenzelm
parents:
70400
diff
changeset
|
344 |
|> Proofterm.forall_intr_vfs_prf prop; |
70419 | 345 |
val (maxidx', prfs', prf) = expand (Proofterm.maxidx_proof prf' ~1) prfs prf' |
346 |
in |
|
347 |
(maxidx' + maxidx + 1, Proofterm.incr_indexes (maxidx + 1) prf, |
|
348 |
((a, prop), (maxidx', prf)) :: prfs') |
|
13610
d4a2ac255447
Fixed bug in expand_proof which caused indexes to be incremented incorrectly.
berghofe
parents:
13342
diff
changeset
|
349 |
end |
70419 | 350 |
| SOME (maxidx', prf) => |
351 |
(maxidx' + maxidx + 1, Proofterm.incr_indexes (maxidx + 1) prf, prfs)); |
|
352 |
in (maxidx', prfs', Proofterm.app_types (maxidx + 1) prop Ts prf) end |
|
12870
3905bc0e9002
Indexes of variables in expanded proofs are now incremented to avoid clashes.
berghofe
parents:
12527
diff
changeset
|
353 |
| expand maxidx prfs prf = (maxidx, prfs, prf); |
11522 | 354 |
|
37310 | 355 |
in #3 (expand (Proofterm.maxidx_proof prf ~1) [] prf) end; |
11522 | 356 |
|
357 |
end; |