author | haftmann |
Sun, 26 Apr 2009 08:34:53 +0200 | |
changeset 30988 | b53800e3ee47 |
parent 30146 | a77fc0209723 |
child 35845 | e5980f0ad025 |
permissions | -rw-r--r-- |
11522 | 1 |
(* Title: Pure/Proof/proofchecker.ML |
11539 | 2 |
Author: Stefan Berghofer, TU Muenchen |
11522 | 3 |
|
4 |
Simple proof checker based only on the core inference rules |
|
5 |
of Isabelle/Pure. |
|
6 |
*) |
|
7 |
||
8 |
signature PROOF_CHECKER = |
|
9 |
sig |
|
10 |
val thm_of_proof : theory -> Proofterm.proof -> thm |
|
11 |
end; |
|
12 |
||
11612 | 13 |
structure ProofChecker : PROOF_CHECKER = |
11522 | 14 |
struct |
15 |
||
16 |
open Proofterm; |
|
17 |
||
18 |
(***** construct a theorem out of a proof term *****) |
|
19 |
||
20 |
fun lookup_thm thy = |
|
17412 | 21 |
let val tab = fold_rev Symtab.update (PureThy.all_thms_of thy) Symtab.empty |
11522 | 22 |
in |
17412 | 23 |
(fn s => case Symtab.lookup tab s of |
15531 | 24 |
NONE => error ("Unknown theorem " ^ quote s) |
25 |
| SOME thm => thm) |
|
11522 | 26 |
end; |
27 |
||
13733 | 28 |
val beta_eta_convert = |
22910 | 29 |
Conv.fconv_rule Drule.beta_eta_conversion; |
11522 | 30 |
|
31 |
fun thm_of_proof thy prf = |
|
32 |
let |
|
29277 | 33 |
val prf_names = Proofterm.fold_proof_terms Term.declare_term_frees (K I) prf Name.context; |
11522 | 34 |
val lookup = lookup_thm thy; |
35 |
||
13670 | 36 |
fun thm_of_atom thm Ts = |
37 |
let |
|
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
28808
diff
changeset
|
38 |
val tvars = OldTerm.term_tvars (Thm.full_prop_of thm); |
18127 | 39 |
val (fmap, thm') = Thm.varifyT' [] thm; |
20151 | 40 |
val ctye = map (pairself (Thm.ctyp_of thy)) |
15798
016f3be5a5ec
Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents:
15574
diff
changeset
|
41 |
(map TVar tvars @ map (fn ((_, S), ixn) => TVar (ixn, S)) fmap ~~ Ts) |
13670 | 42 |
in |
43 |
Thm.instantiate (ctye, []) (forall_intr_vars (forall_intr_frees thm')) |
|
44 |
end; |
|
45 |
||
28808 | 46 |
fun thm_of _ _ (PThm (_, ((name, prop', SOME Ts), _))) = |
11522 | 47 |
let |
16608 | 48 |
val thm = Drule.implies_intr_hyps (lookup name); |
11522 | 49 |
val {prop, ...} = rep_thm thm; |
13733 | 50 |
val _ = if prop aconv prop' then () else |
12238 | 51 |
error ("Duplicate use of theorem name " ^ quote name ^ "\n" ^ |
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
22910
diff
changeset
|
52 |
Syntax.string_of_term_global thy prop ^ "\n\n" ^ |
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
22910
diff
changeset
|
53 |
Syntax.string_of_term_global thy prop'); |
13670 | 54 |
in thm_of_atom thm Ts end |
11522 | 55 |
|
15531 | 56 |
| thm_of _ _ (PAxm (name, _, SOME Ts)) = |
28674 | 57 |
thm_of_atom (Thm.axiom thy name) Ts |
11522 | 58 |
|
30146 | 59 |
| thm_of _ Hs (PBound i) = nth Hs i |
11522 | 60 |
|
20164 | 61 |
| thm_of (vs, names) Hs (Abst (s, SOME T, prf)) = |
11522 | 62 |
let |
20164 | 63 |
val ([x], names') = Name.variants [s] names; |
64 |
val thm = thm_of ((x, T) :: vs, names') Hs prf |
|
11522 | 65 |
in |
20151 | 66 |
Thm.forall_intr (Thm.cterm_of thy (Free (x, T))) thm |
11522 | 67 |
end |
68 |
||
20164 | 69 |
| thm_of (vs, names) Hs (prf % SOME t) = |
11522 | 70 |
let |
20164 | 71 |
val thm = thm_of (vs, names) Hs prf; |
72 |
val ct = Thm.cterm_of thy (Term.subst_bounds (map Free vs, t)); |
|
11522 | 73 |
in Thm.forall_elim ct thm end |
74 |
||
20164 | 75 |
| thm_of (vs, names) Hs (AbsP (s, SOME t, prf)) = |
11522 | 76 |
let |
20151 | 77 |
val ct = Thm.cterm_of thy (Term.subst_bounds (map Free vs, t)); |
20164 | 78 |
val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf; |
11522 | 79 |
in |
80 |
Thm.implies_intr ct thm |
|
81 |
end |
|
82 |
||
20164 | 83 |
| thm_of vars Hs (prf %% prf') = |
20151 | 84 |
let |
20164 | 85 |
val thm = beta_eta_convert (thm_of vars Hs prf); |
86 |
val thm' = beta_eta_convert (thm_of vars Hs prf'); |
|
11522 | 87 |
in |
88 |
Thm.implies_elim thm thm' |
|
89 |
end |
|
90 |
||
20151 | 91 |
| thm_of _ _ (Hyp t) = Thm.assume (Thm.cterm_of thy t) |
11522 | 92 |
|
93 |
| thm_of _ _ _ = error "thm_of_proof: partial proof term"; |
|
94 |
||
20164 | 95 |
in beta_eta_convert (thm_of ([], prf_names) [] prf) end; |
11522 | 96 |
|
97 |
end; |