src/Pure/Proof/proofchecker.ML
author wenzelm
Sat, 17 Oct 2009 17:18:59 +0200
changeset 32971 55ba9b6648ef
parent 30146 a77fc0209723
child 35845 e5980f0ad025
permissions -rw-r--r--
less pervasive names;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     1
(*  Title:      Pure/Proof/proofchecker.ML
11539
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     3
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     4
Simple proof checker based only on the core inference rules
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     5
of Isabelle/Pure.
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     6
*)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     7
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     8
signature PROOF_CHECKER =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     9
sig
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    10
  val thm_of_proof : theory -> Proofterm.proof -> thm
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    11
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    12
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    13
structure ProofChecker : PROOF_CHECKER =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    14
struct
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    15
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    16
open Proofterm;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    17
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    18
(***** construct a theorem out of a proof term *****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    19
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    20
fun lookup_thm thy =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17223
diff changeset
    21
  let val tab = fold_rev Symtab.update (PureThy.all_thms_of thy) Symtab.empty
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    22
  in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17223
diff changeset
    23
    (fn s => case Symtab.lookup tab s of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15001
diff changeset
    24
       NONE => error ("Unknown theorem " ^ quote s)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15001
diff changeset
    25
     | SOME thm => thm)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    26
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    27
13733
8ea7388f66d4 - tuned beta_eta_convert
berghofe
parents: 13670
diff changeset
    28
val beta_eta_convert =
22910
54d231cbc19a moved conversions to structure Conv;
wenzelm
parents: 21646
diff changeset
    29
  Conv.fconv_rule Drule.beta_eta_conversion;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    30
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    31
fun thm_of_proof thy prf =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    32
  let
29277
29dd1c516337 Term.declare_term_frees;
wenzelm
parents: 29270
diff changeset
    33
    val prf_names = Proofterm.fold_proof_terms Term.declare_term_frees (K I) prf Name.context;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    34
    val lookup = lookup_thm thy;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    35
13670
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    36
    fun thm_of_atom thm Ts =
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    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
9f03d8a9a81b Thm.varifyT': natural argument order;
wenzelm
parents: 17412
diff changeset
    39
        val (fmap, thm') = Thm.varifyT' [] thm;
20151
b22c14181eb7 thm_of_proof: tuned Name operations;
wenzelm
parents: 20071
diff changeset
    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
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    42
      in
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    43
        Thm.instantiate (ctye, []) (forall_intr_vars (forall_intr_frees thm'))
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    44
      end;
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    45
28808
7925199a0226 adapted PThm;
wenzelm
parents: 28674
diff changeset
    46
    fun thm_of _ _ (PThm (_, ((name, prop', SOME Ts), _))) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    47
          let
16608
4f8d7b83c7e2 Drule.implies_intr_hyps;
wenzelm
parents: 16351
diff changeset
    48
            val thm = Drule.implies_intr_hyps (lookup name);
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    49
            val {prop, ...} = rep_thm thm;
13733
8ea7388f66d4 - tuned beta_eta_convert
berghofe
parents: 13670
diff changeset
    50
            val _ = if prop aconv prop' then () else
12238
09966ccbc84c Improved error message.
berghofe
parents: 11612
diff changeset
    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
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    54
          in thm_of_atom thm Ts end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    55
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15001
diff changeset
    56
      | thm_of _ _ (PAxm (name, _, SOME Ts)) =
28674
08a77c495dc1 renamed Thm.get_axiom_i to Thm.axiom;
wenzelm
parents: 26939
diff changeset
    57
          thm_of_atom (Thm.axiom thy name) Ts
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    58
30146
a77fc0209723 eliminated NJ's List.nth;
wenzelm
parents: 29277
diff changeset
    59
      | thm_of _ Hs (PBound i) = nth Hs i
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    60
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    61
      | thm_of (vs, names) Hs (Abst (s, SOME T, prf)) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    62
          let
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    63
            val ([x], names') = Name.variants [s] names;
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    64
            val thm = thm_of ((x, T) :: vs, names') Hs prf
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    65
          in
20151
b22c14181eb7 thm_of_proof: tuned Name operations;
wenzelm
parents: 20071
diff changeset
    66
            Thm.forall_intr (Thm.cterm_of thy (Free (x, T))) thm
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    67
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    68
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    69
      | thm_of (vs, names) Hs (prf % SOME t) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    70
          let
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    71
            val thm = thm_of (vs, names) Hs prf;
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    72
            val ct = Thm.cterm_of thy (Term.subst_bounds (map Free vs, t));
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    73
          in Thm.forall_elim ct thm end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    74
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    75
      | thm_of (vs, names) Hs (AbsP (s, SOME t, prf)) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    76
          let
20151
b22c14181eb7 thm_of_proof: tuned Name operations;
wenzelm
parents: 20071
diff changeset
    77
            val ct = Thm.cterm_of thy (Term.subst_bounds (map Free vs, t));
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    78
            val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    79
          in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    80
            Thm.implies_intr ct thm
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    81
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    82
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    83
      | thm_of vars Hs (prf %% prf') =
20151
b22c14181eb7 thm_of_proof: tuned Name operations;
wenzelm
parents: 20071
diff changeset
    84
          let
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    85
            val thm = beta_eta_convert (thm_of vars Hs prf);
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    86
            val thm' = beta_eta_convert (thm_of vars Hs prf');
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    87
          in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    88
            Thm.implies_elim thm thm'
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    89
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    90
20151
b22c14181eb7 thm_of_proof: tuned Name operations;
wenzelm
parents: 20071
diff changeset
    91
      | thm_of _ _ (Hyp t) = Thm.assume (Thm.cterm_of thy t)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    92
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    93
      | thm_of _ _ _ = error "thm_of_proof: partial proof term";
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    94
20164
928c8dc07216 thm_of_proof: improved generation of variables;
wenzelm
parents: 20151
diff changeset
    95
  in beta_eta_convert (thm_of ([], prf_names) [] prf) end;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    96
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    97
end;