src/Pure/Proof/proofchecker.ML
author berghofe
Fri, 31 May 2002 18:52:23 +0200
changeset 13199 18402c1f76bf
parent 12238 09966ccbc84c
child 13670 c71b905a852a
permissions -rw-r--r--
Added constants for Hyp, Oracle and MinProof.
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
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
11539
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
0f17da240450 tuned headers;
wenzelm
parents: 11522
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     5
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     6
Simple proof checker based only on the core inference rules
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     7
of Isabelle/Pure.
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     8
*)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
     9
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    10
signature PROOF_CHECKER =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    11
sig
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    12
  val thm_of_proof : theory -> Proofterm.proof -> thm
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    13
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    14
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    15
structure ProofChecker : PROOF_CHECKER =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    16
struct
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
open Proofterm;
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
(***** construct a theorem out of a proof term *****)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    21
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    22
fun lookup_thm thy =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    23
  let val tab = foldr Symtab.update
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    24
    (flat (map thms_of (thy :: Theory.ancestors_of thy)), Symtab.empty)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    25
  in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    26
    (fn s => case Symtab.lookup (tab, s) of
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    27
       None => error ("Unknown theorem " ^ quote s)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    28
     | Some thm => thm)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    29
  end;
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 beta_eta_convert thm =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    32
  let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    33
    val beta_thm = beta_conversion true (cprop_of thm);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    34
    val (_, rhs) = Drule.dest_equals (cprop_of beta_thm);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    35
  in Thm.equal_elim (Thm.transitive beta_thm (eta_conversion rhs)) thm end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    36
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    37
fun thm_of_proof thy prf =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    38
  let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    39
    val names = add_prf_names ([], prf);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    40
    val sg = sign_of thy;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    41
    val lookup = lookup_thm thy;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    42
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    43
    fun thm_of _ _ (PThm ((name, _), _, prop', Some Ts)) =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    44
          let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    45
            val thm = lookup name;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    46
            val {prop, ...} = rep_thm thm;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    47
            val _ = if prop=prop' then () else
12238
09966ccbc84c Improved error message.
berghofe
parents: 11612
diff changeset
    48
              error ("Duplicate use of theorem name " ^ quote name ^ "\n" ^
09966ccbc84c Improved error message.
berghofe
parents: 11612
diff changeset
    49
                Sign.string_of_term sg prop ^ "\n\n" ^
09966ccbc84c Improved error message.
berghofe
parents: 11612
diff changeset
    50
                Sign.string_of_term sg prop');
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    51
            val tvars = term_tvars prop;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    52
            val ctye = map fst tvars ~~ map (Thm.ctyp_of sg) Ts
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    53
          in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    54
            Thm.instantiate (ctye, []) (forall_intr_vars thm)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    55
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    56
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    57
      | thm_of _ _ (PAxm (name, _, Some Ts)) =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    58
          let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    59
            val thm = get_axiom thy name;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    60
            val {prop, ...} = rep_thm thm;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    61
            val tvars = term_tvars prop;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    62
            val ctye = map fst tvars ~~ map (Thm.ctyp_of sg) Ts
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    63
          in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    64
            Thm.instantiate (ctye, []) (forall_intr_vars thm)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    65
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    66
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    67
      | thm_of _ Hs (PBound i) = nth_elem (i, Hs)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    68
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    69
      | thm_of vs Hs (Abst (s, Some T, prf)) =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    70
          let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    71
            val x = variant (names @ map fst vs) s;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    72
            val thm = thm_of ((x, T) :: vs) Hs prf
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    73
          in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    74
            Thm.forall_intr (Thm.cterm_of sg (Free (x, T))) thm
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    75
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    76
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    77
      | thm_of vs Hs (prf % Some t) =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    78
          let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    79
            val thm = thm_of vs Hs prf
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    80
            val ct = Thm.cterm_of sg (Term.subst_bounds (map Free vs, t))
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    81
          in Thm.forall_elim ct thm end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    82
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    83
      | thm_of vs Hs (AbsP (s, Some t, prf)) =
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    84
          let
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    85
            val ct = Thm.cterm_of sg (Term.subst_bounds (map Free vs, t));
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    86
            val thm = thm_of vs (Thm.assume ct :: Hs) prf
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_intr ct 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
11612
ae8450657bf0 Exchanged % and %%.
berghofe
parents: 11539
diff changeset
    91
      | thm_of vs Hs (prf %% prf') =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    92
          let 
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    93
            val thm = beta_eta_convert (thm_of vs Hs prf);
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    94
            val thm' = beta_eta_convert (thm_of vs Hs prf')
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    95
          in
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    96
            Thm.implies_elim thm thm'
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    97
          end
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    98
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    99
      | thm_of _ _ (Hyp t) = Thm.assume (Thm.cterm_of sg t)
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   100
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   101
      | thm_of _ _ _ = error "thm_of_proof: partial proof term";
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   102
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   103
  in thm_of [] [] prf end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   104
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   105
end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   106