src/Pure/Proof/proof_syntax.ML
changeset 11614 3131fa12d425
parent 11539 0f17da240450
child 11640 be1bc3b88480
equal deleted inserted replaced
11613:c92a99a2b5b1 11614:3131fa12d425
    29 open Proofterm;
    29 open Proofterm;
    30 
    30 
    31 (**** add special syntax for embedding proof terms ****)
    31 (**** add special syntax for embedding proof terms ****)
    32 
    32 
    33 val proofT = Type ("proof", []);
    33 val proofT = Type ("proof", []);
    34 val lamT = Type ("lam_syn", []);
    34 val paramT = Type ("param", []);
       
    35 val paramsT = Type ("params", []);
    35 val idtT = Type ("idt", []);
    36 val idtT = Type ("idt", []);
    36 val aT = TFree ("'a", ["logic"]);
    37 val aT = TFree ("'a", ["logic"]);
    37 
    38 
    38 (** constants for theorems and axioms **)
    39 (** constants for theorems and axioms **)
    39 
    40 
    41 
    42 
    42 fun add_proof_atom_consts names sg = Sign.add_consts_i
    43 fun add_proof_atom_consts names sg = Sign.add_consts_i
    43   (map (fn name => (name, proofT, NoSyn)) names) (Sign.add_path "//" sg);
    44   (map (fn name => (name, proofT, NoSyn)) names) (Sign.add_path "//" sg);
    44 
    45 
    45 (** constants for application and abstraction **)
    46 (** constants for application and abstraction **)
    46   
    47 
    47 fun add_proof_syntax sg =
    48 fun add_proof_syntax sg =
    48   sg
    49   sg
    49   |> Sign.copy
    50   |> Sign.copy
    50   |> Sign.add_path "/"
    51   |> Sign.add_path "/"
    51   |> Sign.add_defsort_i ["logic"]
    52   |> Sign.add_defsort_i ["logic"]
    52   |> Sign.add_types [("proof", 0, NoSyn)]
    53   |> Sign.add_types [("proof", 0, NoSyn)]
    53   |> Sign.add_arities [("proof", [], "logic")]
    54   |> Sign.add_arities [("proof", [], "logic")]
    54   |> Sign.add_consts_i
    55   |> Sign.add_consts_i
    55       [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)),
    56       [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)),
    56        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)),
    57        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)),
    57        ("Abst", (aT --> proofT) --> proofT, NoSyn),
    58        ("Abst", (aT --> proofT) --> proofT, NoSyn),
    58        ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn)]
    59        ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn)]
    59   |> Sign.add_nonterminals ["lam_syn"]
    60   |> Sign.add_nonterminals ["param", "params"]
    60   |> Sign.add_syntax_i
    61   |> Sign.add_syntax_i
    61       [("_Lam", [lamT, proofT] ---> proofT, Mixfix ("(3Lam _./ _)", [0,0], 1)),
    62       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(3Lam _./ _)", [0, 3], 3)),
    62        ("_Lam0", [lamT, lamT] ---> lamT, Mixfix ("_,/ _", [1, 0], 0)),
    63        ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
    63        ("_Lam1", [idtT, propT] ---> lamT, Mixfix ("_ : _", [0, 0], 1)),
    64        ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)),
    64        ("_Lam2", idtT --> lamT, Mixfix ("_", [0], 1))]
    65        ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)),
       
    66        ("", paramT --> paramT, Delimfix "'(_')"),
       
    67        ("", idtT --> paramsT, Delimfix "_"),
       
    68        ("", paramT --> paramsT, Delimfix "_")]
    65   |> Sign.add_modesyntax_i (("xsymbols", true),
    69   |> Sign.add_modesyntax_i (("xsymbols", true),
    66       [("_Lam", [lamT, proofT] ---> proofT, Mixfix ("(3\\<Lambda>_./ _)", [0,0], 1)),
    70       [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(3\\<Lambda>_./ _)", [0, 3], 3)),
    67        ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
    71        ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\<cdot>/ _)", [4, 5], 4)),
    68        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))])
    72        ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\<bullet>/ _)", [4, 5], 4))])
    69   |> Sign.add_trrules_i (map Syntax.ParsePrintRule
    73   |> Sign.add_trrules_i (map Syntax.ParsePrintRule
    70       [(Syntax.mk_appl (Constant "_Lam")
    74       [(Syntax.mk_appl (Constant "_Lam")
       
    75           [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"],
       
    76         Syntax.mk_appl (Constant "_Lam")
       
    77           [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]),
       
    78        (Syntax.mk_appl (Constant "_Lam")
    71           [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"],
    79           [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"],
    72         Syntax.mk_appl (Constant "AbsP") [Variable "A",
    80         Syntax.mk_appl (Constant "AbsP") [Variable "A",
    73           (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]),
    81           (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]),
    74        (Syntax.mk_appl (Constant "_Lam")
    82        (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"],
    75           [Syntax.mk_appl (Constant "_Lam2") [Variable "x"], Variable "A"],
       
    76         Syntax.mk_appl (Constant "Abst")
    83         Syntax.mk_appl (Constant "Abst")
    77           [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])]),
    84           [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]);
    78        (Syntax.mk_appl (Constant "_Lam")
       
    79           [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"],
       
    80         Syntax.mk_appl (Constant "_Lam")
       
    81           [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]])]);
       
    82 
    85 
    83 
    86 
    84 (**** create unambiguous theorem names ****)
    87 (**** create unambiguous theorem names ****)
    85 
    88 
    86 fun disambiguate_names thy prf =
    89 fun disambiguate_names thy prf =
    97         else x) (ps, (tab, 1)))
   100         else x) (ps, (tab, 1)))
    98       end) (Symtab.empty, thms);
   101       end) (Symtab.empty, thms);
    99 
   102 
   100     fun rename (Abst (s, T, prf)) = Abst (s, T, rename prf)
   103     fun rename (Abst (s, T, prf)) = Abst (s, T, rename prf)
   101       | rename (AbsP (s, t, prf)) = AbsP (s, t, rename prf)
   104       | rename (AbsP (s, t, prf)) = AbsP (s, t, rename prf)
   102       | rename (prf1 % prf2) = rename prf1 % rename prf2
   105       | rename (prf1 %% prf2) = rename prf1 %% rename prf2
   103       | rename (prf %% t) = rename prf %% t
   106       | rename (prf % t) = rename prf % t
   104       | rename (prf' as PThm ((s, tags), prf, prop, Ts)) =
   107       | rename (prf' as PThm ((s, tags), prf, prop, Ts)) =
   105           let
   108           let
   106             val prop' = if_none (assoc (thms', s)) (Bound 0);
   109             val prop' = if_none (assoc (thms', s)) (Bound 0);
   107             val ps = map fst (the (Symtab.lookup (thms, s))) \ prop'
   110             val ps = map fst (the (Symtab.lookup (thms, s))) \ prop'
   108           in if prop = prop' then prf' else
   111           in if prop = prop' then prf' else
   124   let
   127   let
   125     val thys = thy :: Theory.ancestors_of thy;
   128     val thys = thy :: Theory.ancestors_of thy;
   126     val thms = flat (map thms_of thys);
   129     val thms = flat (map thms_of thys);
   127     val axms = flat (map (Symtab.dest o #axioms o rep_theory) thys);
   130     val axms = flat (map (Symtab.dest o #axioms o rep_theory) thys);
   128 
   131 
       
   132     fun mk_term t = (if ty then I else map_term_types (K dummyT))
       
   133       (Term.no_dummy_patterns t);
       
   134 
   129     fun prf_of [] (Bound i) = PBound i
   135     fun prf_of [] (Bound i) = PBound i
   130       | prf_of Ts (Const (s, Type ("proof", _))) =
   136       | prf_of Ts (Const (s, Type ("proof", _))) =
   131           change_type (if ty then Some Ts else None)
   137           change_type (if ty then Some Ts else None)
   132             (case NameSpace.unpack s of
   138             (case NameSpace.unpack s of
   133                "Axm" :: xs =>
   139                "axm" :: xs =>
   134                  let
   140                  let
   135                    val name = NameSpace.pack xs;
   141                    val name = NameSpace.pack xs;
   136                    val prop = (case assoc (axms, name) of
   142                    val prop = (case assoc (axms, name) of
   137                        Some prop => prop
   143                        Some prop => prop
   138                      | None => error ("Unknown axiom " ^ quote name))
   144                      | None => error ("Unknown axiom " ^ quote name))
   139                  in PAxm (name, prop, None) end
   145                  in PAxm (name, prop, None) end
   140              | "Thm" :: xs =>
   146              | "thm" :: xs =>
   141                  let val name = NameSpace.pack xs;
   147                  let val name = NameSpace.pack xs;
   142                  in (case assoc (thms, name) of
   148                  in (case assoc (thms, name) of
   143                      Some thm => fst (strip_combt (#2 (#der (rep_thm thm))))
   149                      Some thm => fst (strip_combt (#2 (#der (rep_thm thm))))
   144                    | None => (case Symtab.lookup (tab, name) of
   150                    | None => (case Symtab.lookup (tab, name) of
   145                          Some prf => prf
   151                          Some prf => prf
   149       | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
   155       | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v
   150       | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
   156       | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) =
   151           Abst (s, if ty then Some T else None,
   157           Abst (s, if ty then Some T else None,
   152             incr_pboundvars (~1) 0 (prf_of [] prf))
   158             incr_pboundvars (~1) 0 (prf_of [] prf))
   153       | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
   159       | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) =
   154           AbsP (s, case t of Const ("dummy_pattern", _) => None | _ => Some t,
   160           AbsP (s, case t of
       
   161                 Const ("dummy_pattern", _) => None
       
   162               | _ $ Const ("dummy_pattern", _) => None
       
   163               | _ => Some (mk_term t),
   155             incr_pboundvars 0 (~1) (prf_of [] prf))
   164             incr_pboundvars 0 (~1) (prf_of [] prf))
   156       | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
   165       | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) =
   157           prf_of [] prf1 % prf_of [] prf2
   166           prf_of [] prf1 %% prf_of [] prf2
   158       | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
   167       | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) =
   159           prf_of (T::Ts) prf
   168           prf_of (T::Ts) prf
   160       | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %%
   169       | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %
   161           (case t of Const ("dummy_pattern", _) => None | _ => Some t)
   170           (case t of Const ("dummy_pattern", _) => None | _ => Some (mk_term t))
   162       | prf_of _ t = error ("Not a proof term:\n" ^
   171       | prf_of _ t = error ("Not a proof term:\n" ^
   163           Sign.string_of_term (sign_of thy) t)
   172           Sign.string_of_term (sign_of thy) t)
   164 
   173 
   165   in prf_of [] end;
   174   in prf_of [] end;
   166 
   175 
   173 
   182 
   174 val mk_tyapp = foldl (fn (prf, T) => Const ("Appt",
   183 val mk_tyapp = foldl (fn (prf, T) => Const ("Appt",
   175   [proofT, itselfT T] ---> proofT) $ prf $ Logic.mk_type T);
   184   [proofT, itselfT T] ---> proofT) $ prf $ Logic.mk_type T);
   176 
   185 
   177 fun term_of _ (PThm ((name, _), _, _, None)) =
   186 fun term_of _ (PThm ((name, _), _, _, None)) =
   178       Const (add_prefix "Thm" name, proofT)
   187       Const (add_prefix "thm" name, proofT)
   179   | term_of _ (PThm ((name, _), _, _, Some Ts)) =
   188   | term_of _ (PThm ((name, _), _, _, Some Ts)) =
   180       mk_tyapp (Const (add_prefix "Thm" name, proofT), Ts)
   189       mk_tyapp (Const (add_prefix "thm" name, proofT), Ts)
   181   | term_of _ (PAxm (name, _, None)) = Const (add_prefix "Axm" name, proofT)
   190   | term_of _ (PAxm (name, _, None)) = Const (add_prefix "axm" name, proofT)
   182   | term_of _ (PAxm (name, _, Some Ts)) =
   191   | term_of _ (PAxm (name, _, Some Ts)) =
   183       mk_tyapp (Const (add_prefix "Axm" name, proofT), Ts)
   192       mk_tyapp (Const (add_prefix "axm" name, proofT), Ts)
   184   | term_of _ (PBound i) = Bound i
   193   | term_of _ (PBound i) = Bound i
   185   | term_of Ts (Abst (s, opT, prf)) = 
   194   | term_of Ts (Abst (s, opT, prf)) = 
   186       let val T = if_none opT dummyT
   195       let val T = if_none opT dummyT
   187       in Const ("Abst", (T --> proofT) --> proofT) $
   196       in Const ("Abst", (T --> proofT) --> proofT) $
   188         Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf))
   197         Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf))
   189       end
   198       end
   190   | term_of Ts (AbsP (s, t, prf)) =
   199   | term_of Ts (AbsP (s, t, prf)) =
   191       AbsPt $ if_none t (Const ("dummy_pattern", propT)) $
   200       AbsPt $ if_none t (Const ("dummy_pattern", propT)) $
   192         Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf))
   201         Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf))
   193   | term_of Ts (prf1 % prf2) =
   202   | term_of Ts (prf1 %% prf2) =
   194       AppPt $ term_of Ts prf1 $ term_of Ts prf2
   203       AppPt $ term_of Ts prf1 $ term_of Ts prf2
   195   | term_of Ts (prf %% opt) = 
   204   | term_of Ts (prf % opt) = 
   196       let val t = if_none opt (Const ("dummy_pattern", dummyT))
   205       let val t = if_none opt (Const ("dummy_pattern", dummyT))
   197       in Const ("Appt",
   206       in Const ("Appt",
   198         [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $
   207         [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $
   199           term_of Ts prf $ t
   208           term_of Ts prf $ t
   200       end
   209       end
   212       map fst (Symtab.dest tab);
   221       map fst (Symtab.dest tab);
   213     val axm_names = map fst (flat (map (Symtab.dest o #axioms o rep_theory) thys));
   222     val axm_names = map fst (flat (map (Symtab.dest o #axioms o rep_theory) thys));
   214     val sg = sign_of thy |>
   223     val sg = sign_of thy |>
   215       add_proof_syntax |>
   224       add_proof_syntax |>
   216       add_proof_atom_consts
   225       add_proof_atom_consts
   217         (map (add_prefix "Thm") thm_names @ map (add_prefix "Axm") axm_names)
   226         (map (add_prefix "thm") thm_names @ map (add_prefix "axm") axm_names)
   218   in
   227   in
   219     (cterm_of sg (term_of_proof prf'),
   228     (cterm_of sg (term_of_proof prf'),
   220      proof_of_term thy tab true o Thm.term_of)
   229      proof_of_term thy tab true o Thm.term_of)
   221   end;
   230   end;
   222 
   231 
   226     val thm_names = filter_out (equal "") (map fst (flat (map thms_of thys)));
   235     val thm_names = filter_out (equal "") (map fst (flat (map thms_of thys)));
   227     val axm_names = map fst (flat (map (Symtab.dest o #axioms o rep_theory) thys));
   236     val axm_names = map fst (flat (map (Symtab.dest o #axioms o rep_theory) thys));
   228     val sg = sign_of thy |>
   237     val sg = sign_of thy |>
   229       add_proof_syntax |>
   238       add_proof_syntax |>
   230       add_proof_atom_consts
   239       add_proof_atom_consts
   231         (map (add_prefix "Thm") thm_names @ map (add_prefix "Axm") axm_names)
   240         (map (add_prefix "thm") thm_names @ map (add_prefix "axm") axm_names)
   232   in
   241   in
   233     (fn T => fn s => Thm.term_of (read_cterm sg (s, T)))
   242     (fn T => fn s => Thm.term_of (read_cterm sg (s, T)))
   234   end;
   243   end;
   235 
   244 
   236 fun read_proof thy =
   245 fun read_proof thy =
   244     val thm_names = map fst (Symtab.dest (thms_of_proof Symtab.empty prf)) \ "";
   253     val thm_names = map fst (Symtab.dest (thms_of_proof Symtab.empty prf)) \ "";
   245     val axm_names = map fst (Symtab.dest (axms_of_proof Symtab.empty prf));
   254     val axm_names = map fst (Symtab.dest (axms_of_proof Symtab.empty prf));
   246     val sg' = sg |>
   255     val sg' = sg |>
   247       add_proof_syntax |>
   256       add_proof_syntax |>
   248       add_proof_atom_consts
   257       add_proof_atom_consts
   249         (map (add_prefix "Thm") thm_names @ map (add_prefix "Axm") axm_names)
   258         (map (add_prefix "thm") thm_names @ map (add_prefix "axm") axm_names)
   250   in
   259   in
   251     Sign.pretty_term sg' (term_of_proof prf)
   260     Sign.pretty_term sg' (term_of_proof prf)
   252   end;
   261   end;
   253 
   262 
   254 fun pretty_proof_of full thm =
   263 fun pretty_proof_of full thm =
   257     val prf' = (case strip_combt (fst (strip_combP prf)) of
   266     val prf' = (case strip_combt (fst (strip_combP prf)) of
   258         (PThm (_, prf', prop', _), _) => if prop=prop' then prf' else prf
   267         (PThm (_, prf', prop', _), _) => if prop=prop' then prf' else prf
   259       | _ => prf)
   268       | _ => prf)
   260   in
   269   in
   261     pretty_proof sign
   270     pretty_proof sign
   262       (if full then Reconstruct.reconstruct_prf sign prop prf' else prf')
   271       (if full then Reconstruct.reconstruct_proof sign prop prf' else prf')
   263   end;
   272   end;
   264 
   273 
   265 val print_proof_of = Pretty.writeln oo pretty_proof_of;
   274 val print_proof_of = Pretty.writeln oo pretty_proof_of;
   266 
   275 
   267 end;
   276 end;