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