src/Pure/Proof/proof_checker.ML
author wenzelm
Sun, 09 Jun 2024 15:31:33 +0200
changeset 80306 c2537860ccf8
parent 80299 a397fd0c451a
child 80590 505f97165f52
permissions -rw-r--r--
more accurate thm "name_hint", using Thm_Name.T;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45007
cc86edb97c2c fixed headers;
wenzelm
parents: 44062
diff changeset
     1
(*  Title:      Pure/Proof/proof_checker.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
44057
fda143b5c2f5 modernized strcture Proof_Checker;
wenzelm
parents: 43326
diff changeset
    13
structure Proof_Checker : 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
(***** construct a theorem out of a proof term *****)
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
fun lookup_thm thy =
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    19
  let
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    20
    val tab =
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    21
      Thm_Name.Table.build (Global_Theory.all_thms_of thy true |> fold_rev Thm_Name.Table.update);
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    22
  in
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    23
    fn name =>
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    24
      (case Thm_Name.Table.lookup tab name of
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    25
        NONE => error ("Unknown theorem " ^ quote (Thm_Name.print name))
44057
fda143b5c2f5 modernized strcture Proof_Checker;
wenzelm
parents: 43326
diff changeset
    26
      | SOME thm => thm)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    27
  end;
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    28
13733
8ea7388f66d4 - tuned beta_eta_convert
berghofe
parents: 13670
diff changeset
    29
val beta_eta_convert =
22910
54d231cbc19a moved conversions to structure Conv;
wenzelm
parents: 21646
diff changeset
    30
  Conv.fconv_rule Drule.beta_eta_conversion;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    31
37229
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    32
(* equality modulo renaming of type variables *)
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    33
fun is_equal t t' =
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    34
  let
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    35
    val atoms = fold_types (fold_atyps (insert (op =))) t [];
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    36
    val atoms' = fold_types (fold_atyps (insert (op =))) t' []
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    37
  in
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    38
    length atoms = length atoms' andalso
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    39
    map_types (map_atyps (the o AList.lookup (op =) (atoms ~~ atoms'))) t aconv t'
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    40
  end;
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    41
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    42
fun pretty_prf thy vs Hs prf =
79167
4fb0723dc5fc clarified signature;
wenzelm
parents: 79134
diff changeset
    43
  let val prf' = prf |> Proofterm.subst_bounds (map Free vs) |>
4fb0723dc5fc clarified signature;
wenzelm
parents: 79134
diff changeset
    44
    Proofterm.subst_pbounds (map (Hyp o Thm.prop_of) Hs)
37229
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    45
  in
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    46
    (Proof_Syntax.pretty_proof (Syntax.init_pretty_global thy) prf',
70447
755d58b48cec clarified modules: provide reconstruct_proof / expand_proof at the bottom of proof term construction;
wenzelm
parents: 70412
diff changeset
    47
     Syntax.pretty_term_global thy (Proofterm.prop_of prf'))
37229
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    48
  end;
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    49
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    50
fun pretty_term thy vs _ t =
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    51
  let val t' = subst_bounds (map Free vs, t)
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    52
  in
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    53
    (Syntax.pretty_term_global thy t',
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    54
     Syntax.pretty_typ_global thy (fastype_of t'))
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    55
  end;
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    56
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    57
fun appl_error thy prt vs Hs s f a =
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    58
  let
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    59
    val (pp_f, pp_fT) = pretty_prf thy vs Hs f;
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    60
    val (pp_a, pp_aT) = prt thy vs Hs a
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    61
  in
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    62
    error (cat_lines
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    63
      [s,
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    64
       "",
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    65
       Pretty.string_of (Pretty.block
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    66
         [Pretty.str "Operator:", Pretty.brk 2, pp_f,
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    67
           Pretty.str " ::", Pretty.brk 1, pp_fT]),
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    68
       Pretty.string_of (Pretty.block
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    69
         [Pretty.str "Operand:", Pretty.brk 3, pp_a,
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    70
           Pretty.str " ::", Pretty.brk 1, pp_aT]),
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    71
       ""])
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    72
  end;
47eb565796f4 - Equality check on propositions after lookup of theorem now takes type variable
berghofe
parents: 36042
diff changeset
    73
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    74
fun thm_of_proof thy =
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    75
  let val lookup = lookup_thm thy in
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    76
    fn prf =>
13670
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    77
      let
74411
20b0b27bc6c7 clarified signature;
wenzelm
parents: 74282
diff changeset
    78
        val prf_names =
20b0b27bc6c7 clarified signature;
wenzelm
parents: 74282
diff changeset
    79
          Name.build_context (prf |> Proofterm.fold_proof_terms Term.declare_term_frees);
13670
c71b905a852a Fixed problem with theorems containing TFrees.
berghofe
parents: 12238
diff changeset
    80
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    81
        fun thm_of_atom thm Ts =
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    82
          let
74235
dbaed92fd8af tuned signature;
wenzelm
parents: 74233
diff changeset
    83
            val tvars = build_rev (Term.add_tvars (Thm.full_prop_of thm));
79134
5f0bbed1c606 misc tuning and clarification;
wenzelm
parents: 74411
diff changeset
    84
            val (names, thm') = Thm.varifyT_global' TFrees.empty thm;
5f0bbed1c606 misc tuning and clarification;
wenzelm
parents: 74411
diff changeset
    85
            val ctye = tvars @ map #2 names ~~ map (Thm.global_ctyp_of thy) Ts;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    86
          in
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74266
diff changeset
    87
            Thm.instantiate (TVars.make ctye, Vars.empty)
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74266
diff changeset
    88
              (Thm.forall_intr_vars (Thm.forall_intr_frees thm'))
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    89
          end;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
    90
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 79167
diff changeset
    91
        fun thm_of _ _ (PThm ({thm_name, prop = prop', types = SOME Ts, ...}, _)) =
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    92
              let
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80299
diff changeset
    93
                val thm = Thm.unconstrainT (Drule.implies_intr_hyps (lookup thm_name));
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    94
                val prop = Thm.prop_of thm;
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    95
                val _ =
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    96
                  if is_equal prop prop' then ()
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
    97
                  else
80299
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    98
                    error ("Duplicate use of theorem name " ^
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    99
                      quote (Global_Theory.print_thm_name thy thm_name) ^ "\n" ^
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   100
                      Syntax.string_of_term_global thy prop ^ "\n\n" ^
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   101
                      Syntax.string_of_term_global thy prop');
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   102
              in thm_of_atom thm Ts end
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   103
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   104
          | thm_of _ _ (PAxm (name, _, SOME Ts)) =
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   105
              thm_of_atom (Thm.axiom thy name) Ts
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   106
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   107
          | thm_of _ Hs (PBound i) = nth Hs i
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   108
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   109
          | thm_of (vs, names) Hs (Abst (s, SOME T, prf)) =
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   110
              let
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   111
                val (x, names') = Name.variant s names;
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   112
                val thm = thm_of ((x, T) :: vs, names') Hs prf
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   113
              in
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   114
                Thm.forall_intr (Thm.global_cterm_of thy (Free (x, T))) thm
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   115
              end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   116
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   117
          | thm_of (vs, names) Hs (prf % SOME t) =
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   118
              let
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   119
                val thm = thm_of (vs, names) Hs prf;
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   120
                val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t));
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   121
              in
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   122
                Thm.forall_elim ct thm
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   123
                handle THM (s, _, _) => appl_error thy pretty_term vs Hs s prf t
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   124
              end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   125
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 59621
diff changeset
   126
          | thm_of (vs, names) Hs (AbsP (_, SOME t, prf)) =
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   127
              let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   128
                val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t));
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   129
                val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf;
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   130
              in
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   131
                Thm.implies_intr ct thm
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   132
              end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   133
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   134
          | thm_of vars Hs (prf %% prf') =
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   135
              let
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   136
                val thm = beta_eta_convert (thm_of vars Hs prf);
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   137
                val thm' = beta_eta_convert (thm_of vars Hs prf');
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   138
              in
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   139
                Thm.implies_elim thm thm'
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   140
                handle THM (s, _, _) => appl_error thy pretty_prf (fst vars) Hs s prf prf'
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   141
              end
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   142
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   143
          | thm_of _ _ (Hyp t) = Thm.assume (Thm.global_cterm_of thy t)
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   144
71777
3875815f5967 clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents: 70843
diff changeset
   145
          | thm_of _ _ (PClass (T, c)) =
70838
e381e2624077 some treatment of OfClass proofs;
wenzelm
parents: 70493
diff changeset
   146
              if Sign.of_sort thy (T, [c])
e381e2624077 some treatment of OfClass proofs;
wenzelm
parents: 70493
diff changeset
   147
              then Thm.of_class (Thm.global_ctyp_of thy T, c)
e381e2624077 some treatment of OfClass proofs;
wenzelm
parents: 70493
diff changeset
   148
              else
71777
3875815f5967 clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
wenzelm
parents: 70843
diff changeset
   149
                error ("thm_of_proof: bad PClass proof " ^
70838
e381e2624077 some treatment of OfClass proofs;
wenzelm
parents: 70493
diff changeset
   150
                  Syntax.string_of_term_global thy (Logic.mk_of_class (T, c)))
e381e2624077 some treatment of OfClass proofs;
wenzelm
parents: 70493
diff changeset
   151
44061
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   152
          | thm_of _ _ _ = error "thm_of_proof: partial proof term";
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   153
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   154
      in beta_eta_convert (thm_of ([], prf_names) [] prf) end
9f17ede679e9 tuned thm_of_proof: build lookup table within closure;
wenzelm
parents: 44058
diff changeset
   155
  end;
11522
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   156
42fbb6abed5a Initial revision of tools for proof terms.
berghofe
parents:
diff changeset
   157
end;