| author | haftmann | 
| Wed, 28 Nov 2007 09:01:42 +0100 | |
| changeset 25485 | 33840a854e63 | 
| parent 25245 | 1fcfcdcba53c | 
| child 26626 | c6231d64d264 | 
| permissions | -rw-r--r-- | 
| 11522 | 1 | (* Title: Pure/Proof/proof_syntax.ML | 
| 2 | ID: $Id$ | |
| 11539 | 3 | Author: Stefan Berghofer, TU Muenchen | 
| 11522 | 4 | |
| 5 | Function for parsing and printing proof terms. | |
| 6 | *) | |
| 7 | ||
| 8 | signature PROOF_SYNTAX = | |
| 9 | sig | |
| 17078 | 10 | val proofT: typ | 
| 11 | val add_proof_syntax: theory -> theory | |
| 12 | val disambiguate_names: theory -> Proofterm.proof -> | |
| 11522 | 13 | Proofterm.proof * Proofterm.proof Symtab.table | 
| 17078 | 14 | val proof_of_term: theory -> Proofterm.proof Symtab.table -> | 
| 11522 | 15 | bool -> term -> Proofterm.proof | 
| 17078 | 16 | val term_of_proof: Proofterm.proof -> term | 
| 17 | val cterm_of_proof: theory -> Proofterm.proof -> cterm * (cterm -> Proofterm.proof) | |
| 18 | val read_term: theory -> typ -> string -> term | |
| 19 | val read_proof: theory -> bool -> string -> Proofterm.proof | |
| 20 | val proof_syntax: Proofterm.proof -> theory -> theory | |
| 21 | val proof_of: bool -> thm -> Proofterm.proof | |
| 22 | val pretty_proof: theory -> Proofterm.proof -> Pretty.T | |
| 23 | val pretty_proof_of: bool -> thm -> Pretty.T | |
| 24 | val print_proof_of: bool -> thm -> unit | |
| 11522 | 25 | end; | 
| 26 | ||
| 27 | structure ProofSyntax : PROOF_SYNTAX = | |
| 28 | struct | |
| 29 | ||
| 30 | open Proofterm; | |
| 31 | ||
| 32 | (**** add special syntax for embedding proof terms ****) | |
| 33 | ||
| 34 | val proofT = Type ("proof", []);
 | |
| 11614 | 35 | val paramT = Type ("param", []);
 | 
| 36 | val paramsT = Type ("params", []);
 | |
| 11522 | 37 | val idtT = Type ("idt", []);
 | 
| 24848 | 38 | val aT = TFree (Name.aT, []); | 
| 11522 | 39 | |
| 40 | (** constants for theorems and axioms **) | |
| 41 | ||
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 42 | fun add_proof_atom_consts names thy = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 43 | thy | 
| 22796 | 44 | |> Sign.absolute_path | 
| 45 | |> Sign.add_consts_i (map (fn name => (name, proofT, NoSyn)) names); | |
| 11522 | 46 | |
| 47 | (** constants for application and abstraction **) | |
| 11614 | 48 | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 49 | fun add_proof_syntax thy = | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 50 | thy | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 51 | |> Theory.copy | 
| 22796 | 52 | |> Sign.root_path | 
| 53 | |> Sign.add_defsort_i [] | |
| 54 |   |> Sign.add_types [("proof", 0, NoSyn)]
 | |
| 55 | |> Sign.add_consts_i | |
| 11614 | 56 |       [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)),
 | 
| 57 |        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)),
 | |
| 11522 | 58 |        ("Abst", (aT --> proofT) --> proofT, NoSyn),
 | 
| 13199 | 59 |        ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn),
 | 
| 60 |        ("Hyp", propT --> proofT, NoSyn),
 | |
| 61 |        ("Oracle", propT --> proofT, NoSyn),
 | |
| 62 |        ("MinProof", proofT, Delimfix "?")]
 | |
| 22796 | 63 | |> Sign.add_nonterminals ["param", "params"] | 
| 64 | |> Sign.add_syntax_i | |
| 11640 | 65 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1Lam _./ _)", [0, 3], 3)),
 | 
| 11614 | 66 |        ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
 | 
| 67 |        ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
 | |
| 68 |        ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)),
 | |
| 69 |        ("", paramT --> paramT, Delimfix "'(_')"),
 | |
| 70 |        ("", idtT --> paramsT, Delimfix "_"),
 | |
| 71 |        ("", paramT --> paramsT, Delimfix "_")]
 | |
| 22796 | 72 |   |> Sign.add_modesyntax_i ("xsymbols", true)
 | 
| 11640 | 73 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<Lambda>_./ _)", [0, 3], 3)),
 | 
| 11522 | 74 |        ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 75 |        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
 | 
| 22796 | 76 |   |> Sign.add_modesyntax_i ("latex", false)
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 77 |       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\<lambda>_./ _)", [0, 3], 3))]
 | 
| 22796 | 78 | |> Sign.add_trrules_i (map Syntax.ParsePrintRule | 
| 11522 | 79 | [(Syntax.mk_appl (Constant "_Lam") | 
| 11614 | 80 | [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"], | 
| 81 | Syntax.mk_appl (Constant "_Lam") | |
| 82 | [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]), | |
| 83 | (Syntax.mk_appl (Constant "_Lam") | |
| 11522 | 84 | [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"], | 
| 85 | Syntax.mk_appl (Constant "AbsP") [Variable "A", | |
| 86 | (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]), | |
| 11614 | 87 | (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"], | 
| 11522 | 88 | Syntax.mk_appl (Constant "Abst") | 
| 11614 | 89 | [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]); | 
| 11522 | 90 | |
| 91 | ||
| 92 | (**** create unambiguous theorem names ****) | |
| 93 | ||
| 94 | fun disambiguate_names thy prf = | |
| 95 | let | |
| 17019 
f68598628d08
Adapted to new interface og thms_of_proof / axms_of_proof.
 berghofe parents: 
16866diff
changeset | 96 | val thms = thms_of_proof prf Symtab.empty; | 
| 16866 | 97 | val thms' = map (apsnd Thm.full_prop_of) (PureThy.all_thms_of thy); | 
| 11522 | 98 | |
| 21056 | 99 | val tab = Symtab.fold (fn (key, ps) => fn tab => | 
| 19473 | 100 | let val prop = the_default (Bound 0) (AList.lookup (op =) thms' key) | 
| 21056 | 101 | in fst (fold_rev (fn (prop', prf) => fn x as (tab, i) => | 
| 11522 | 102 | if prop <> prop' then | 
| 17412 | 103 | (Symtab.update (key ^ "_" ^ string_of_int i, prf) tab, i+1) | 
| 21056 | 104 | else x) ps (tab, 1)) | 
| 105 | end) thms Symtab.empty; | |
| 11522 | 106 | |
| 107 | fun rename (Abst (s, T, prf)) = Abst (s, T, rename prf) | |
| 108 | | rename (AbsP (s, t, prf)) = AbsP (s, t, rename prf) | |
| 11614 | 109 | | rename (prf1 %% prf2) = rename prf1 %% rename prf2 | 
| 110 | | rename (prf % t) = rename prf % t | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21056diff
changeset | 111 | | rename (prf' as PThm (s, prf, prop, Ts)) = | 
| 11522 | 112 | let | 
| 19473 | 113 | val prop' = the_default (Bound 0) (AList.lookup (op =) thms' s); | 
| 19305 | 114 | val ps = remove (op =) prop' (map fst (the (Symtab.lookup thms s))) | 
| 11522 | 115 | in if prop = prop' then prf' else | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21056diff
changeset | 116 | PThm (s ^ "_" ^ string_of_int (length ps - find_index (fn p => p = prop) ps), | 
| 11522 | 117 | prf, prop, Ts) | 
| 118 | end | |
| 119 | | rename prf = prf | |
| 120 | ||
| 121 | in (rename prf, tab) end; | |
| 122 | ||
| 123 | ||
| 124 | (**** translation between proof terms and pure terms ****) | |
| 125 | ||
| 126 | fun proof_of_term thy tab ty = | |
| 127 | let | |
| 16350 | 128 | val thms = PureThy.all_thms_of thy; | 
| 129 | val axms = Theory.all_axioms_of thy; | |
| 11522 | 130 | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
19618diff
changeset | 131 | fun mk_term t = (if ty then I else map_types (K dummyT)) | 
| 11614 | 132 | (Term.no_dummy_patterns t); | 
| 133 | ||
| 11522 | 134 | fun prf_of [] (Bound i) = PBound i | 
| 135 |       | prf_of Ts (Const (s, Type ("proof", _))) =
 | |
| 15531 | 136 | change_type (if ty then SOME Ts else NONE) | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 137 | (case NameSpace.explode s of | 
| 11614 | 138 | "axm" :: xs => | 
| 11522 | 139 | let | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 140 | val name = NameSpace.implode xs; | 
| 17223 | 141 | val prop = (case AList.lookup (op =) axms name of | 
| 15531 | 142 | SOME prop => prop | 
| 143 |                      | NONE => error ("Unknown axiom " ^ quote name))
 | |
| 144 | in PAxm (name, prop, NONE) end | |
| 11614 | 145 | | "thm" :: xs => | 
| 21858 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 wenzelm parents: 
21646diff
changeset | 146 | let val name = NameSpace.implode xs; | 
| 17223 | 147 | in (case AList.lookup (op =) thms name of | 
| 15531 | 148 | SOME thm => fst (strip_combt (Thm.proof_of thm)) | 
| 17412 | 149 | | NONE => (case Symtab.lookup tab name of | 
| 15531 | 150 | SOME prf => prf | 
| 151 |                        | NONE => error ("Unknown theorem " ^ quote name)))
 | |
| 11522 | 152 | end | 
| 153 |              | _ => error ("Illegal proof constant name: " ^ quote s))
 | |
| 13199 | 154 |       | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
 | 
| 11522 | 155 |       | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
 | 
| 156 |       | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
 | |
| 25245 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 berghofe parents: 
24848diff
changeset | 157 | if T = proofT then | 
| 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 berghofe parents: 
24848diff
changeset | 158 |             error ("Term variable abstraction may not bind proof variable " ^ quote s)
 | 
| 
1fcfcdcba53c
Added well-formedness check to Abst case in function prf_of.
 berghofe parents: 
24848diff
changeset | 159 | else Abst (s, if ty then SOME T else NONE, | 
| 11522 | 160 | incr_pboundvars (~1) 0 (prf_of [] prf)) | 
| 161 |       | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
 | |
| 11614 | 162 | AbsP (s, case t of | 
| 15531 | 163 |                 Const ("dummy_pattern", _) => NONE
 | 
| 164 |               | _ $ Const ("dummy_pattern", _) => NONE
 | |
| 165 | | _ => SOME (mk_term t), | |
| 11522 | 166 | incr_pboundvars 0 (~1) (prf_of [] prf)) | 
| 167 |       | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
 | |
| 11614 | 168 | prf_of [] prf1 %% prf_of [] prf2 | 
| 11522 | 169 |       | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
 | 
| 170 | prf_of (T::Ts) prf | |
| 11614 | 171 |       | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
 | 
| 15531 | 172 |           (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t))
 | 
| 11522 | 173 |       | prf_of _ t = error ("Not a proof term:\n" ^
 | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 174 | Sign.string_of_term thy t) | 
| 11522 | 175 | |
| 176 | in prf_of [] end; | |
| 177 | ||
| 178 | ||
| 179 | val AbsPt = Const ("AbsP", [propT, proofT --> proofT] ---> proofT);
 | |
| 180 | val AppPt = Const ("AppP", [proofT, proofT] ---> proofT);
 | |
| 13199 | 181 | val Hypt = Const ("Hyp", propT --> proofT);
 | 
| 182 | val Oraclet = Const ("Oracle", propT --> proofT);
 | |
| 183 | val MinProoft = Const ("MinProof", proofT);
 | |
| 11522 | 184 | |
| 19473 | 185 | val mk_tyapp = fold (fn T => fn prf => Const ("Appt",
 | 
| 19391 | 186 | [proofT, Term.itselfT T] ---> proofT) $ prf $ Logic.mk_type T); | 
| 11522 | 187 | |
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21056diff
changeset | 188 | fun term_of _ (PThm (name, _, _, NONE)) = | 
| 16195 | 189 | Const (NameSpace.append "thm" name, proofT) | 
| 21646 
c07b5b0e8492
thm/prf: separate official name vs. additional tags;
 wenzelm parents: 
21056diff
changeset | 190 | | term_of _ (PThm (name, _, _, SOME Ts)) = | 
| 19473 | 191 | mk_tyapp Ts (Const (NameSpace.append "thm" name, proofT)) | 
| 16195 | 192 | | term_of _ (PAxm (name, _, NONE)) = Const (NameSpace.append "axm" name, proofT) | 
| 15531 | 193 | | term_of _ (PAxm (name, _, SOME Ts)) = | 
| 19473 | 194 | mk_tyapp Ts (Const (NameSpace.append "axm" name, proofT)) | 
| 11522 | 195 | | term_of _ (PBound i) = Bound i | 
| 196 | | term_of Ts (Abst (s, opT, prf)) = | |
| 18939 | 197 | let val T = the_default dummyT opT | 
| 11522 | 198 |       in Const ("Abst", (T --> proofT) --> proofT) $
 | 
| 199 | Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf)) | |
| 200 | end | |
| 201 | | term_of Ts (AbsP (s, t, prf)) = | |
| 18939 | 202 | AbsPt $ the_default (Term.dummy_pattern propT) t $ | 
| 11522 | 203 | Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf)) | 
| 11614 | 204 | | term_of Ts (prf1 %% prf2) = | 
| 11522 | 205 | AppPt $ term_of Ts prf1 $ term_of Ts prf2 | 
| 11614 | 206 | | term_of Ts (prf % opt) = | 
| 18939 | 207 | let val t = the_default (Term.dummy_pattern dummyT) opt | 
| 11522 | 208 |       in Const ("Appt",
 | 
| 209 | [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $ | |
| 210 | term_of Ts prf $ t | |
| 211 | end | |
| 212 | | term_of Ts (Hyp t) = Hypt $ t | |
| 213 | | term_of Ts (Oracle (_, t, _)) = Oraclet $ t | |
| 214 | | term_of Ts (MinProof _) = MinProoft; | |
| 215 | ||
| 216 | val term_of_proof = term_of []; | |
| 217 | ||
| 218 | fun cterm_of_proof thy prf = | |
| 219 | let | |
| 220 | val (prf', tab) = disambiguate_names thy prf; | |
| 16350 | 221 | val thm_names = filter_out (equal "") | 
| 222 | (map fst (PureThy.all_thms_of thy) @ map fst (Symtab.dest tab)); | |
| 223 | val axm_names = map fst (Theory.all_axioms_of thy); | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 224 | val thy' = thy | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 225 | |> add_proof_syntax | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 226 | |> add_proof_atom_consts | 
| 16195 | 227 | (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names) | 
| 11522 | 228 | in | 
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 229 | (cterm_of thy' (term_of_proof prf'), | 
| 11522 | 230 | proof_of_term thy tab true o Thm.term_of) | 
| 231 | end; | |
| 232 | ||
| 233 | fun read_term thy = | |
| 234 | let | |
| 16350 | 235 | val thm_names = filter_out (equal "") (map fst (PureThy.all_thms_of thy)); | 
| 236 | val axm_names = map fst (Theory.all_axioms_of thy); | |
| 16425 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 237 | val thy' = thy | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 238 | |> add_proof_syntax | 
| 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 wenzelm parents: 
16350diff
changeset | 239 | |> add_proof_atom_consts | 
| 16195 | 240 | (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names) | 
| 22675 
acf10be7dcca
cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
 wenzelm parents: 
21858diff
changeset | 241 | in Sign.simple_read_term thy' end; | 
| 11522 | 242 | |
| 243 | fun read_proof thy = | |
| 244 | let val rd = read_term thy proofT | |
| 245 | in | |
| 246 | (fn ty => fn s => proof_of_term thy Symtab.empty ty (Logic.varify (rd s))) | |
| 247 | end; | |
| 248 | ||
| 17078 | 249 | fun proof_syntax prf = | 
| 11522 | 250 | let | 
| 19305 | 251 | val thm_names = filter_out (equal "") | 
| 252 | (map fst (Symtab.dest (thms_of_proof prf Symtab.empty))); | |
| 17019 
f68598628d08
Adapted to new interface og thms_of_proof / axms_of_proof.
 berghofe parents: 
16866diff
changeset | 253 | val axm_names = map fst (Symtab.dest (axms_of_proof prf Symtab.empty)); | 
| 11522 | 254 | in | 
| 17078 | 255 | add_proof_syntax #> | 
| 256 | add_proof_atom_consts | |
| 257 | (map (NameSpace.append "thm") thm_names @ map (NameSpace.append "axm") axm_names) | |
| 11522 | 258 | end; | 
| 259 | ||
| 17078 | 260 | fun proof_of full thm = | 
| 261 | let | |
| 262 |     val {thy, der = (_, prf), ...} = Thm.rep_thm thm;
 | |
| 263 | val prop = Thm.full_prop_of thm; | |
| 264 | val prf' = (case strip_combt (fst (strip_combP prf)) of | |
| 265 | (PThm (_, prf', prop', _), _) => if prop = prop' then prf' else prf | |
| 266 | | _ => prf) | |
| 267 | in if full then Reconstruct.reconstruct_proof thy prop prf' else prf' end; | |
| 268 | ||
| 269 | fun pretty_proof thy prf = | |
| 270 | Sign.pretty_term (proof_syntax prf thy) (term_of_proof prf); | |
| 271 | ||
| 11522 | 272 | fun pretty_proof_of full thm = | 
| 17078 | 273 | pretty_proof (Thm.theory_of_thm thm) (proof_of full thm); | 
| 11522 | 274 | |
| 275 | val print_proof_of = Pretty.writeln oo pretty_proof_of; | |
| 276 | ||
| 277 | end; |