| author | wenzelm | 
| Sat, 01 Apr 2023 13:04:59 +0200 | |
| changeset 77768 | 65008644d394 | 
| parent 74411 | 20b0b27bc6c7 | 
| child 79134 | 5f0bbed1c606 | 
| permissions | -rw-r--r-- | 
| 45007 | 1 | (* Title: Pure/Proof/proof_checker.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 | ||
| 44057 | 13 | structure Proof_Checker : PROOF_CHECKER = | 
| 11522 | 14 | struct | 
| 15 | ||
| 16 | (***** construct a theorem out of a proof term *****) | |
| 17 | ||
| 18 | fun lookup_thm thy = | |
| 74232 | 19 | let val tab = Symtab.build (Global_Theory.all_thms_of thy true |> fold_rev Symtab.update) in | 
| 44057 | 20 | fn s => | 
| 21 | (case Symtab.lookup tab s of | |
| 22 |         NONE => error ("Unknown theorem " ^ quote s)
 | |
| 23 | | SOME thm => thm) | |
| 11522 | 24 | end; | 
| 25 | ||
| 13733 | 26 | val beta_eta_convert = | 
| 22910 | 27 | Conv.fconv_rule Drule.beta_eta_conversion; | 
| 11522 | 28 | |
| 37229 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 29 | (* equality modulo renaming of type variables *) | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 30 | fun is_equal t t' = | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 31 | let | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 32 | val atoms = fold_types (fold_atyps (insert (op =))) t []; | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 33 | val atoms' = fold_types (fold_atyps (insert (op =))) t' [] | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 34 | in | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 35 | length atoms = length atoms' andalso | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 36 | 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: 
36042diff
changeset | 37 | end; | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 38 | |
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 39 | fun pretty_prf thy vs Hs prf = | 
| 37310 | 40 | let val prf' = prf |> Proofterm.prf_subst_bounds (map Free vs) |> | 
| 59582 | 41 | Proofterm.prf_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: 
36042diff
changeset | 42 | in | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 43 | (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: 
70412diff
changeset | 44 | 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: 
36042diff
changeset | 45 | end; | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 46 | |
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 47 | fun pretty_term thy vs _ t = | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 48 | let val t' = subst_bounds (map Free vs, t) | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 49 | in | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 50 | (Syntax.pretty_term_global thy t', | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 51 | Syntax.pretty_typ_global thy (fastype_of t')) | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 52 | end; | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 53 | |
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 54 | 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: 
36042diff
changeset | 55 | let | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 56 | 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: 
36042diff
changeset | 57 | 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: 
36042diff
changeset | 58 | in | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 59 | error (cat_lines | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 60 | [s, | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 61 | "", | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 62 | Pretty.string_of (Pretty.block | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 63 | [Pretty.str "Operator:", Pretty.brk 2, pp_f, | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 64 | Pretty.str " ::", Pretty.brk 1, pp_fT]), | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 65 | Pretty.string_of (Pretty.block | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 66 | [Pretty.str "Operand:", Pretty.brk 3, pp_a, | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 67 | Pretty.str " ::", Pretty.brk 1, pp_aT]), | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 68 | ""]) | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 69 | end; | 
| 
47eb565796f4
- Equality check on propositions after lookup of theorem now takes type variable
 berghofe parents: 
36042diff
changeset | 70 | |
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 71 | fun thm_of_proof thy = | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 72 | let val lookup = lookup_thm thy in | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 73 | fn prf => | 
| 13670 | 74 | let | 
| 74411 | 75 | val prf_names = | 
| 76 | Name.build_context (prf |> Proofterm.fold_proof_terms Term.declare_term_frees); | |
| 13670 | 77 | |
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 78 | fun thm_of_atom thm Ts = | 
| 11522 | 79 | let | 
| 74235 | 80 | val tvars = build_rev (Term.add_tvars (Thm.full_prop_of thm)); | 
| 74266 | 81 | val (fmap, thm') = Thm.varifyT_global' TFrees.empty thm; | 
| 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: 
59621diff
changeset | 82 | val ctye = | 
| 
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: 
59621diff
changeset | 83 | (tvars @ map (fn ((_, S), ixn) => (ixn, S)) fmap ~~ map (Thm.global_ctyp_of thy) Ts); | 
| 11522 | 84 | in | 
| 74282 | 85 | Thm.instantiate (TVars.make ctye, Vars.empty) | 
| 86 | (Thm.forall_intr_vars (Thm.forall_intr_frees thm')) | |
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 87 | end; | 
| 11522 | 88 | |
| 70493 | 89 |         fun thm_of _ _ (PThm ({name, prop = prop', types = SOME Ts, ...}, _)) =
 | 
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 90 | let | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 91 | val thm = Thm.unconstrainT (Drule.implies_intr_hyps (lookup name)); | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 92 | val prop = Thm.prop_of thm; | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 93 | val _ = | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 94 | if is_equal prop prop' then () | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 95 | else | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 96 |                     error ("Duplicate use of theorem name " ^ quote name ^ "\n" ^
 | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 97 | Syntax.string_of_term_global thy prop ^ "\n\n" ^ | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 98 | Syntax.string_of_term_global thy prop'); | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 99 | in thm_of_atom thm Ts end | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 100 | |
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 101 | | thm_of _ _ (PAxm (name, _, SOME Ts)) = | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 102 | thm_of_atom (Thm.axiom thy name) Ts | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 103 | |
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 104 | | thm_of _ Hs (PBound i) = nth Hs i | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 105 | |
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 106 | | thm_of (vs, names) Hs (Abst (s, SOME T, prf)) = | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 107 | let | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 108 | val (x, names') = Name.variant s names; | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 109 | val thm = thm_of ((x, T) :: vs, names') Hs prf | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 110 | in | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 111 | 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: 
44058diff
changeset | 112 | end | 
| 11522 | 113 | |
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 114 | | thm_of (vs, names) Hs (prf % SOME t) = | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 115 | let | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 116 | val thm = thm_of (vs, names) Hs prf; | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 117 | 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: 
44058diff
changeset | 118 | in | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 119 | Thm.forall_elim ct thm | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 120 | 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: 
44058diff
changeset | 121 | end | 
| 11522 | 122 | |
| 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: 
59621diff
changeset | 123 | | thm_of (vs, names) Hs (AbsP (_, SOME t, prf)) = | 
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 124 | let | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 125 | 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: 
44058diff
changeset | 126 | val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf; | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 127 | in | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 128 | Thm.implies_intr ct thm | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 129 | end | 
| 11522 | 130 | |
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 131 | | thm_of vars Hs (prf %% prf') = | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 132 | let | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 133 | val thm = beta_eta_convert (thm_of vars Hs prf); | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 134 | val thm' = beta_eta_convert (thm_of vars Hs prf'); | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 135 | in | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 136 | Thm.implies_elim thm thm' | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 137 | 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: 
44058diff
changeset | 138 | end | 
| 11522 | 139 | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59582diff
changeset | 140 | | thm_of _ _ (Hyp t) = Thm.assume (Thm.global_cterm_of thy t) | 
| 11522 | 141 | |
| 71777 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 wenzelm parents: 
70843diff
changeset | 142 | | thm_of _ _ (PClass (T, c)) = | 
| 70838 | 143 | if Sign.of_sort thy (T, [c]) | 
| 144 | then Thm.of_class (Thm.global_ctyp_of thy T, c) | |
| 145 | else | |
| 71777 
3875815f5967
clarified signature: avoid clash with Isabelle/Scala Term.OFCLASS on case-insensible file-system;
 wenzelm parents: 
70843diff
changeset | 146 |                 error ("thm_of_proof: bad PClass proof " ^
 | 
| 70838 | 147 | Syntax.string_of_term_global thy (Logic.mk_of_class (T, c))) | 
| 148 | ||
| 44061 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 149 | | thm_of _ _ _ = error "thm_of_proof: partial proof term"; | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 150 | |
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 151 | in beta_eta_convert (thm_of ([], prf_names) [] prf) end | 
| 
9f17ede679e9
tuned thm_of_proof: build lookup table within closure;
 wenzelm parents: 
44058diff
changeset | 152 | end; | 
| 11522 | 153 | |
| 154 | end; |