src/Pure/Proof/proof_syntax.ML
author berghofe
Tue, 01 Jun 2010 11:30:57 +0200
changeset 37236 739d8b9c59da
parent 37153 01aa36932739
parent 37227 bdd8dd217b1f
child 37310 96e2b9a6f074
permissions -rw-r--r--
merged
     1 (*  Title:      Pure/Proof/proof_syntax.ML
     2     Author:     Stefan Berghofer, TU Muenchen
     3 
     4 Function for parsing and printing proof terms.
     5 *)
     6 
     7 signature PROOF_SYNTAX =
     8 sig
     9   val proofT: typ
    10   val add_proof_syntax: theory -> theory
    11   val proof_of_term: theory -> bool -> term -> Proofterm.proof
    12   val term_of_proof: Proofterm.proof -> term
    13   val cterm_of_proof: theory -> Proofterm.proof -> cterm * (cterm -> Proofterm.proof)
    14   val strip_sorts_consttypes: Proof.context -> Proof.context
    15   val read_term: theory -> bool -> typ -> string -> term
    16   val read_proof: theory -> bool -> bool -> string -> Proofterm.proof
    17   val proof_syntax: Proofterm.proof -> theory -> theory
    18   val proof_of: bool -> thm -> Proofterm.proof
    19   val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T
    20   val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T
    21 end;
    22 
    23 structure Proof_Syntax : PROOF_SYNTAX =
    24 struct
    25 
    26 open Proofterm;
    27 
    28 (**** add special syntax for embedding proof terms ****)
    29 
    30 val proofT = Type ("proof", []);
    31 val paramT = Type ("param", []);
    32 val paramsT = Type ("params", []);
    33 val idtT = Type ("idt", []);
    34 val aT = TFree (Name.aT, []);
    35 
    36 (** constants for theorems and axioms **)
    37 
    38 fun add_proof_atom_consts names thy =
    39   thy
    40   |> Sign.root_path
    41   |> Sign.add_consts_i (map (fn name => (Binding.qualified_name name, proofT, NoSyn)) names);
    42 
    43 (** constants for application and abstraction **)
    44 
    45 fun add_proof_syntax thy =
    46   thy
    47   |> Theory.copy
    48   |> Sign.root_path
    49   |> Sign.set_defsort []
    50   |> Sign.add_types [(Binding.name "proof", 0, NoSyn)]
    51   |> fold (snd oo Sign.declare_const)
    52       [((Binding.name "Appt", [proofT, aT] ---> proofT), Mixfix ("(1_ %/ _)", [4, 5], 4)),
    53        ((Binding.name "AppP", [proofT, proofT] ---> proofT), Mixfix ("(1_ %%/ _)", [4, 5], 4)),
    54        ((Binding.name "Abst", (aT --> proofT) --> proofT), NoSyn),
    55        ((Binding.name "AbsP", [propT, proofT --> proofT] ---> proofT), NoSyn),
    56        ((Binding.name "Hyp", propT --> proofT), NoSyn),
    57        ((Binding.name "Oracle", propT --> proofT), NoSyn),
    58        ((Binding.name "OfClass", (Term.a_itselfT --> propT) --> proofT), NoSyn),
    59        ((Binding.name "MinProof", proofT), Delimfix "?")]
    60   |> Sign.add_nonterminals [Binding.name "param", Binding.name "params"]
    61   |> Sign.add_syntax_i
    62       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1Lam _./ _)", [0, 3], 3)),
    63        ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
    64        ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
    65        ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)),
    66        ("", paramT --> paramT, Delimfix "'(_')"),
    67        ("", idtT --> paramsT, Delimfix "_"),
    68        ("", paramT --> paramsT, Delimfix "_")]
    69   |> Sign.add_modesyntax_i (Symbol.xsymbolsN, true)
    70       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<Lambda>_./ _)", [0, 3], 3)),
    71        (Syntax.mark_const "Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
    72        (Syntax.mark_const "AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))]
    73   |> Sign.add_modesyntax_i ("latex", false)
    74       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\<lambda>_./ _)", [0, 3], 3))]
    75   |> Sign.add_trrules_i (map Syntax.ParsePrintRule
    76       [(Syntax.mk_appl (Constant "_Lam")
    77           [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"],
    78         Syntax.mk_appl (Constant "_Lam")
    79           [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]),
    80        (Syntax.mk_appl (Constant "_Lam")
    81           [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"],
    82         Syntax.mk_appl (Constant (Syntax.mark_const "AbsP")) [Variable "A",
    83           (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]),
    84        (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"],
    85         Syntax.mk_appl (Constant (Syntax.mark_const "Abst"))
    86           [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]);
    87 
    88 
    89 (**** translation between proof terms and pure terms ****)
    90 
    91 fun proof_of_term thy ty =
    92   let
    93     val thms = PureThy.all_thms_of thy;
    94     val axms = Theory.all_axioms_of thy;
    95 
    96     fun mk_term t = (if ty then I else map_types (K dummyT))
    97       (Term.no_dummy_patterns t);
    98 
    99     fun prf_of [] (Bound i) = PBound i
   100       | prf_of Ts (Const (s, Type ("proof", _))) =
   101           change_type (if ty then SOME Ts else NONE)
   102             (case Long_Name.explode s of
   103                "axm" :: xs =>
   104                  let
   105                    val name = Long_Name.implode xs;
   106                    val prop = (case AList.lookup (op =) axms name of
   107                        SOME prop => prop
   108                      | NONE => error ("Unknown axiom " ^ quote name))
   109                  in PAxm (name, prop, NONE) end
   110              | "thm" :: xs =>
   111                  let val name = Long_Name.implode xs;
   112                  in (case AList.lookup (op =) thms name of
   113                      SOME thm => fst (strip_combt (fst (strip_combP (Thm.proof_of thm))))
   114                    | NONE => error ("Unknown theorem " ^ quote name))
   115                  end
   116              | _ => error ("Illegal proof constant name: " ^ quote s))
   117       | prf_of Ts (Const ("OfClass", _) $ Const (c_class, _)) =
   118           (case try Logic.class_of_const c_class of
   119             SOME c =>
   120               change_type (if ty then SOME Ts else NONE)
   121                 (OfClass (TVar ((Name.aT, 0), []), c))
   122           | NONE => error ("Bad class constant: " ^ quote c_class))
   123       | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop
   124       | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
   125       | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
   126           if T = proofT then
   127             error ("Term variable abstraction may not bind proof variable " ^ quote s)
   128           else Abst (s, if ty then SOME T else NONE,
   129             incr_pboundvars (~1) 0 (prf_of [] prf))
   130       | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
   131           AbsP (s, case t of
   132                 Const ("dummy_pattern", _) => NONE
   133               | _ $ Const ("dummy_pattern", _) => NONE
   134               | _ => SOME (mk_term t),
   135             incr_pboundvars 0 (~1) (prf_of [] prf))
   136       | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
   137           prf_of [] prf1 %% prf_of [] prf2
   138       | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
   139           prf_of (T::Ts) prf
   140       | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
   141           (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t))
   142       | prf_of _ t = error ("Not a proof term:\n" ^
   143           Syntax.string_of_term_global thy t)
   144 
   145   in prf_of [] end;
   146 
   147 
   148 val AbsPt = Const ("AbsP", [propT, proofT --> proofT] ---> proofT);
   149 val AppPt = Const ("AppP", [proofT, proofT] ---> proofT);
   150 val Hypt = Const ("Hyp", propT --> proofT);
   151 val Oraclet = Const ("Oracle", propT --> proofT);
   152 val OfClasst = Const ("OfClass", (Term.itselfT dummyT --> propT) --> proofT);
   153 val MinProoft = Const ("MinProof", proofT);
   154 
   155 val mk_tyapp = fold (fn T => fn prf => Const ("Appt",
   156   [proofT, Term.itselfT T] ---> proofT) $ prf $ Logic.mk_type T);
   157 
   158 fun term_of _ (PThm (_, ((name, _, NONE), _))) =
   159       Const (Long_Name.append "thm" name, proofT)
   160   | term_of _ (PThm (_, ((name, _, SOME Ts), _))) =
   161       mk_tyapp Ts (Const (Long_Name.append "thm" name, proofT))
   162   | term_of _ (PAxm (name, _, NONE)) = Const (Long_Name.append "axm" name, proofT)
   163   | term_of _ (PAxm (name, _, SOME Ts)) =
   164       mk_tyapp Ts (Const (Long_Name.append "axm" name, proofT))
   165   | term_of _ (OfClass (T, c)) =
   166       mk_tyapp [T] (OfClasst $ Const (Logic.const_of_class c, Term.itselfT dummyT --> propT))
   167   | term_of _ (PBound i) = Bound i
   168   | term_of Ts (Abst (s, opT, prf)) =
   169       let val T = the_default dummyT opT
   170       in Const ("Abst", (T --> proofT) --> proofT) $
   171         Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf))
   172       end
   173   | term_of Ts (AbsP (s, t, prf)) =
   174       AbsPt $ the_default (Term.dummy_pattern propT) t $
   175         Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf))
   176   | term_of Ts (prf1 %% prf2) =
   177       AppPt $ term_of Ts prf1 $ term_of Ts prf2
   178   | term_of Ts (prf % opt) =
   179       let val t = the_default (Term.dummy_pattern dummyT) opt
   180       in Const ("Appt",
   181         [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $
   182           term_of Ts prf $ t
   183       end
   184   | term_of Ts (Hyp t) = Hypt $ t
   185   | term_of Ts (Oracle (_, t, _)) = Oraclet $ t
   186   | term_of Ts MinProof = MinProoft;
   187 
   188 val term_of_proof = term_of [];
   189 
   190 fun cterm_of_proof thy prf =
   191   let
   192     val thm_names = map fst (PureThy.all_thms_of thy);
   193     val axm_names = map fst (Theory.all_axioms_of thy);
   194     val thy' = thy
   195       |> add_proof_syntax
   196       |> add_proof_atom_consts
   197         (map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names);
   198   in
   199     (cterm_of thy' (term_of_proof prf), proof_of_term thy true o Thm.term_of)
   200   end;
   201 
   202 fun strip_sorts_consttypes ctxt =
   203   let val {constants = (_, tab), ...} = Consts.dest (ProofContext.consts_of ctxt)
   204   in Symtab.fold (fn (s, (T, _)) =>
   205       ProofContext.add_const_constraint (s, SOME (Type.strip_sorts T)))
   206     tab ctxt
   207   end;
   208 
   209 fun read_term thy topsort =
   210   let
   211     val thm_names = filter_out (fn s => s = "") (map fst (PureThy.all_thms_of thy));
   212     val axm_names = map fst (Theory.all_axioms_of thy);
   213     val ctxt = thy
   214       |> add_proof_syntax
   215       |> add_proof_atom_consts
   216         (map (Long_Name.append "axm") axm_names @ map (Long_Name.append "thm") thm_names)
   217       |> ProofContext.init_global
   218       |> ProofContext.allow_dummies
   219       |> ProofContext.set_mode ProofContext.mode_schematic
   220       |> (if topsort then
   221             strip_sorts_consttypes #>
   222             ProofContext.set_defsort []
   223           else I);
   224   in
   225     fn ty => fn s =>
   226       (if ty = propT then Syntax.parse_prop else Syntax.parse_term) ctxt s
   227       |> Type_Infer.constrain ty |> Syntax.check_term ctxt
   228   end;
   229 
   230 fun read_proof thy topsort =
   231   let val rd = read_term thy topsort proofT
   232   in fn ty => fn s => proof_of_term thy ty (Logic.varify_global (rd s)) end;
   233 
   234 fun proof_syntax prf =
   235   let
   236     val thm_names = Symtab.keys (fold_proof_atoms true
   237       (fn PThm (_, ((name, _, _), _)) => if name <> "" then Symtab.update (name, ()) else I
   238         | _ => I) [prf] Symtab.empty);
   239     val axm_names = Symtab.keys (fold_proof_atoms true
   240       (fn PAxm (name, _, _) => Symtab.update (name, ()) | _ => I) [prf] Symtab.empty);
   241   in
   242     add_proof_syntax #>
   243     add_proof_atom_consts
   244       (map (Long_Name.append "thm") thm_names @ map (Long_Name.append "axm") axm_names)
   245   end;
   246 
   247 fun proof_of full thm =
   248   let
   249     val thy = Thm.theory_of_thm thm;
   250     val prop = Thm.full_prop_of thm;
   251     val prf = Thm.proof_of thm;
   252     val prf' = (case strip_combt (fst (strip_combP prf)) of
   253         (PThm (_, ((_, prop', _), body)), _) => if prop = prop' then join_proof body else prf
   254       | _ => prf)
   255   in if full then Reconstruct.reconstruct_proof thy prop prf' else prf' end;
   256 
   257 fun pretty_proof ctxt prf =
   258   ProofContext.pretty_term_abbrev
   259     (ProofContext.transfer_syntax (proof_syntax prf (ProofContext.theory_of ctxt)) ctxt)
   260     (term_of_proof prf);
   261 
   262 fun pretty_proof_of ctxt full th =
   263   pretty_proof ctxt (proof_of full th);
   264 
   265 end;