berghofe@11522: (* Title: Pure/Proof/proof_syntax.ML berghofe@11522: ID: $Id$ wenzelm@11539: Author: Stefan Berghofer, TU Muenchen berghofe@11522: berghofe@11522: Function for parsing and printing proof terms. berghofe@11522: *) berghofe@11522: berghofe@11522: signature PROOF_SYNTAX = berghofe@11522: sig wenzelm@17078: val proofT: typ wenzelm@17078: val add_proof_syntax: theory -> theory wenzelm@17078: val disambiguate_names: theory -> Proofterm.proof -> berghofe@11522: Proofterm.proof * Proofterm.proof Symtab.table wenzelm@17078: val proof_of_term: theory -> Proofterm.proof Symtab.table -> berghofe@11522: bool -> term -> Proofterm.proof wenzelm@17078: val term_of_proof: Proofterm.proof -> term wenzelm@17078: val cterm_of_proof: theory -> Proofterm.proof -> cterm * (cterm -> Proofterm.proof) wenzelm@17078: val read_term: theory -> typ -> string -> term wenzelm@17078: val read_proof: theory -> bool -> string -> Proofterm.proof wenzelm@17078: val proof_syntax: Proofterm.proof -> theory -> theory wenzelm@17078: val proof_of: bool -> thm -> Proofterm.proof wenzelm@17078: val pretty_proof: theory -> Proofterm.proof -> Pretty.T wenzelm@17078: val pretty_proof_of: bool -> thm -> Pretty.T wenzelm@17078: val print_proof_of: bool -> thm -> unit berghofe@11522: end; berghofe@11522: berghofe@11522: structure ProofSyntax : PROOF_SYNTAX = berghofe@11522: struct berghofe@11522: berghofe@11522: open Proofterm; berghofe@11522: berghofe@11522: (**** add special syntax for embedding proof terms ****) berghofe@11522: berghofe@11522: val proofT = Type ("proof", []); berghofe@11614: val paramT = Type ("param", []); berghofe@11614: val paramsT = Type ("params", []); berghofe@11522: val idtT = Type ("idt", []); wenzelm@14854: val aT = TFree ("'a", []); berghofe@11522: berghofe@11522: (** constants for theorems and axioms **) berghofe@11522: wenzelm@16425: fun add_proof_atom_consts names thy = wenzelm@16425: thy wenzelm@22796: |> Sign.absolute_path wenzelm@22796: |> Sign.add_consts_i (map (fn name => (name, proofT, NoSyn)) names); berghofe@11522: berghofe@11522: (** constants for application and abstraction **) berghofe@11614: wenzelm@16425: fun add_proof_syntax thy = wenzelm@16425: thy wenzelm@16425: |> Theory.copy wenzelm@22796: |> Sign.root_path wenzelm@22796: |> Sign.add_defsort_i [] wenzelm@22796: |> Sign.add_types [("proof", 0, NoSyn)] wenzelm@22796: |> Sign.add_consts_i berghofe@11614: [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)), berghofe@11614: ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)), berghofe@11522: ("Abst", (aT --> proofT) --> proofT, NoSyn), berghofe@13199: ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn), berghofe@13199: ("Hyp", propT --> proofT, NoSyn), berghofe@13199: ("Oracle", propT --> proofT, NoSyn), berghofe@13199: ("MinProof", proofT, Delimfix "?")] wenzelm@22796: |> Sign.add_nonterminals ["param", "params"] wenzelm@22796: |> Sign.add_syntax_i berghofe@11640: [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1Lam _./ _)", [0, 3], 3)), berghofe@11614: ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)), berghofe@11614: ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)), berghofe@11614: ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)), berghofe@11614: ("", paramT --> paramT, Delimfix "'(_')"), berghofe@11614: ("", idtT --> paramsT, Delimfix "_"), berghofe@11614: ("", paramT --> paramsT, Delimfix "_")] wenzelm@22796: |> Sign.add_modesyntax_i ("xsymbols", true) berghofe@11640: [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\_./ _)", [0, 3], 3)), berghofe@11522: ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\/ _)", [4, 5], 4)), wenzelm@16425: ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\/ _)", [4, 5], 4))] wenzelm@22796: |> Sign.add_modesyntax_i ("latex", false) wenzelm@16425: [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(1\\<^bold>\\_./ _)", [0, 3], 3))] wenzelm@22796: |> Sign.add_trrules_i (map Syntax.ParsePrintRule berghofe@11522: [(Syntax.mk_appl (Constant "_Lam") berghofe@11614: [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"], berghofe@11614: Syntax.mk_appl (Constant "_Lam") berghofe@11614: [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]), berghofe@11614: (Syntax.mk_appl (Constant "_Lam") berghofe@11522: [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"], berghofe@11522: Syntax.mk_appl (Constant "AbsP") [Variable "A", berghofe@11522: (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]), berghofe@11614: (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"], berghofe@11522: Syntax.mk_appl (Constant "Abst") berghofe@11614: [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]); berghofe@11522: berghofe@11522: berghofe@11522: (**** create unambiguous theorem names ****) berghofe@11522: berghofe@11522: fun disambiguate_names thy prf = berghofe@11522: let berghofe@17019: val thms = thms_of_proof prf Symtab.empty; wenzelm@16866: val thms' = map (apsnd Thm.full_prop_of) (PureThy.all_thms_of thy); berghofe@11522: haftmann@21056: val tab = Symtab.fold (fn (key, ps) => fn tab => wenzelm@19473: let val prop = the_default (Bound 0) (AList.lookup (op =) thms' key) haftmann@21056: in fst (fold_rev (fn (prop', prf) => fn x as (tab, i) => berghofe@11522: if prop <> prop' then wenzelm@17412: (Symtab.update (key ^ "_" ^ string_of_int i, prf) tab, i+1) haftmann@21056: else x) ps (tab, 1)) haftmann@21056: end) thms Symtab.empty; berghofe@11522: berghofe@11522: fun rename (Abst (s, T, prf)) = Abst (s, T, rename prf) berghofe@11522: | rename (AbsP (s, t, prf)) = AbsP (s, t, rename prf) berghofe@11614: | rename (prf1 %% prf2) = rename prf1 %% rename prf2 berghofe@11614: | rename (prf % t) = rename prf % t wenzelm@21646: | rename (prf' as PThm (s, prf, prop, Ts)) = berghofe@11522: let wenzelm@19473: val prop' = the_default (Bound 0) (AList.lookup (op =) thms' s); wenzelm@19305: val ps = remove (op =) prop' (map fst (the (Symtab.lookup thms s))) berghofe@11522: in if prop = prop' then prf' else wenzelm@21646: PThm (s ^ "_" ^ string_of_int (length ps - find_index (fn p => p = prop) ps), berghofe@11522: prf, prop, Ts) berghofe@11522: end berghofe@11522: | rename prf = prf berghofe@11522: berghofe@11522: in (rename prf, tab) end; berghofe@11522: berghofe@11522: berghofe@11522: (**** translation between proof terms and pure terms ****) berghofe@11522: berghofe@11522: fun proof_of_term thy tab ty = berghofe@11522: let wenzelm@16350: val thms = PureThy.all_thms_of thy; wenzelm@16350: val axms = Theory.all_axioms_of thy; berghofe@11522: wenzelm@20548: fun mk_term t = (if ty then I else map_types (K dummyT)) berghofe@11614: (Term.no_dummy_patterns t); berghofe@11614: berghofe@11522: fun prf_of [] (Bound i) = PBound i berghofe@11522: | prf_of Ts (Const (s, Type ("proof", _))) = skalberg@15531: change_type (if ty then SOME Ts else NONE) wenzelm@21858: (case NameSpace.explode s of berghofe@11614: "axm" :: xs => berghofe@11522: let wenzelm@21858: val name = NameSpace.implode xs; wenzelm@17223: val prop = (case AList.lookup (op =) axms name of skalberg@15531: SOME prop => prop skalberg@15531: | NONE => error ("Unknown axiom " ^ quote name)) skalberg@15531: in PAxm (name, prop, NONE) end berghofe@11614: | "thm" :: xs => wenzelm@21858: let val name = NameSpace.implode xs; wenzelm@17223: in (case AList.lookup (op =) thms name of skalberg@15531: SOME thm => fst (strip_combt (Thm.proof_of thm)) wenzelm@17412: | NONE => (case Symtab.lookup tab name of skalberg@15531: SOME prf => prf skalberg@15531: | NONE => error ("Unknown theorem " ^ quote name))) berghofe@11522: end berghofe@11522: | _ => error ("Illegal proof constant name: " ^ quote s)) berghofe@13199: | prf_of Ts (Const ("Hyp", _) $ prop) = Hyp prop berghofe@11522: | prf_of Ts (v as Var ((_, Type ("proof", _)))) = Hyp v berghofe@11522: | prf_of [] (Const ("Abst", _) $ Abs (s, T, prf)) = skalberg@15531: Abst (s, if ty then SOME T else NONE, berghofe@11522: incr_pboundvars (~1) 0 (prf_of [] prf)) berghofe@11522: | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) = berghofe@11614: AbsP (s, case t of skalberg@15531: Const ("dummy_pattern", _) => NONE skalberg@15531: | _ $ Const ("dummy_pattern", _) => NONE skalberg@15531: | _ => SOME (mk_term t), berghofe@11522: incr_pboundvars 0 (~1) (prf_of [] prf)) berghofe@11522: | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) = berghofe@11614: prf_of [] prf1 %% prf_of [] prf2 berghofe@11522: | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) = berghofe@11522: prf_of (T::Ts) prf berghofe@11614: | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf % skalberg@15531: (case t of Const ("dummy_pattern", _) => NONE | _ => SOME (mk_term t)) berghofe@11522: | prf_of _ t = error ("Not a proof term:\n" ^ wenzelm@16425: Sign.string_of_term thy t) berghofe@11522: berghofe@11522: in prf_of [] end; berghofe@11522: berghofe@11522: berghofe@11522: val AbsPt = Const ("AbsP", [propT, proofT --> proofT] ---> proofT); berghofe@11522: val AppPt = Const ("AppP", [proofT, proofT] ---> proofT); berghofe@13199: val Hypt = Const ("Hyp", propT --> proofT); berghofe@13199: val Oraclet = Const ("Oracle", propT --> proofT); berghofe@13199: val MinProoft = Const ("MinProof", proofT); berghofe@11522: wenzelm@19473: val mk_tyapp = fold (fn T => fn prf => Const ("Appt", wenzelm@19391: [proofT, Term.itselfT T] ---> proofT) $ prf $ Logic.mk_type T); berghofe@11522: wenzelm@21646: fun term_of _ (PThm (name, _, _, NONE)) = wenzelm@16195: Const (NameSpace.append "thm" name, proofT) wenzelm@21646: | term_of _ (PThm (name, _, _, SOME Ts)) = wenzelm@19473: mk_tyapp Ts (Const (NameSpace.append "thm" name, proofT)) wenzelm@16195: | term_of _ (PAxm (name, _, NONE)) = Const (NameSpace.append "axm" name, proofT) skalberg@15531: | term_of _ (PAxm (name, _, SOME Ts)) = wenzelm@19473: mk_tyapp Ts (Const (NameSpace.append "axm" name, proofT)) berghofe@11522: | term_of _ (PBound i) = Bound i berghofe@11522: | term_of Ts (Abst (s, opT, prf)) = wenzelm@18939: let val T = the_default dummyT opT berghofe@11522: in Const ("Abst", (T --> proofT) --> proofT) $ berghofe@11522: Abs (s, T, term_of (T::Ts) (incr_pboundvars 1 0 prf)) berghofe@11522: end berghofe@11522: | term_of Ts (AbsP (s, t, prf)) = wenzelm@18939: AbsPt $ the_default (Term.dummy_pattern propT) t $ berghofe@11522: Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf)) berghofe@11614: | term_of Ts (prf1 %% prf2) = berghofe@11522: AppPt $ term_of Ts prf1 $ term_of Ts prf2 berghofe@11614: | term_of Ts (prf % opt) = wenzelm@18939: let val t = the_default (Term.dummy_pattern dummyT) opt berghofe@11522: in Const ("Appt", berghofe@11522: [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $ berghofe@11522: term_of Ts prf $ t berghofe@11522: end berghofe@11522: | term_of Ts (Hyp t) = Hypt $ t berghofe@11522: | term_of Ts (Oracle (_, t, _)) = Oraclet $ t berghofe@11522: | term_of Ts (MinProof _) = MinProoft; berghofe@11522: berghofe@11522: val term_of_proof = term_of []; berghofe@11522: berghofe@11522: fun cterm_of_proof thy prf = berghofe@11522: let berghofe@11522: val (prf', tab) = disambiguate_names thy prf; wenzelm@16350: val thm_names = filter_out (equal "") wenzelm@16350: (map fst (PureThy.all_thms_of thy) @ map fst (Symtab.dest tab)); wenzelm@16350: val axm_names = map fst (Theory.all_axioms_of thy); wenzelm@16425: val thy' = thy wenzelm@16425: |> add_proof_syntax wenzelm@16425: |> add_proof_atom_consts wenzelm@16195: (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names) berghofe@11522: in wenzelm@16425: (cterm_of thy' (term_of_proof prf'), berghofe@11522: proof_of_term thy tab true o Thm.term_of) berghofe@11522: end; berghofe@11522: berghofe@11522: fun read_term thy = berghofe@11522: let wenzelm@16350: val thm_names = filter_out (equal "") (map fst (PureThy.all_thms_of thy)); wenzelm@16350: val axm_names = map fst (Theory.all_axioms_of thy); wenzelm@16425: val thy' = thy wenzelm@16425: |> add_proof_syntax wenzelm@16425: |> add_proof_atom_consts wenzelm@16195: (map (NameSpace.append "axm") axm_names @ map (NameSpace.append "thm") thm_names) wenzelm@22675: in Sign.simple_read_term thy' end; berghofe@11522: berghofe@11522: fun read_proof thy = berghofe@11522: let val rd = read_term thy proofT berghofe@11522: in berghofe@11522: (fn ty => fn s => proof_of_term thy Symtab.empty ty (Logic.varify (rd s))) berghofe@11522: end; berghofe@11522: wenzelm@17078: fun proof_syntax prf = berghofe@11522: let wenzelm@19305: val thm_names = filter_out (equal "") wenzelm@19305: (map fst (Symtab.dest (thms_of_proof prf Symtab.empty))); berghofe@17019: val axm_names = map fst (Symtab.dest (axms_of_proof prf Symtab.empty)); berghofe@11522: in wenzelm@17078: add_proof_syntax #> wenzelm@17078: add_proof_atom_consts wenzelm@17078: (map (NameSpace.append "thm") thm_names @ map (NameSpace.append "axm") axm_names) berghofe@11522: end; berghofe@11522: wenzelm@17078: fun proof_of full thm = wenzelm@17078: let wenzelm@17078: val {thy, der = (_, prf), ...} = Thm.rep_thm thm; wenzelm@17078: val prop = Thm.full_prop_of thm; wenzelm@17078: val prf' = (case strip_combt (fst (strip_combP prf)) of wenzelm@17078: (PThm (_, prf', prop', _), _) => if prop = prop' then prf' else prf wenzelm@17078: | _ => prf) wenzelm@17078: in if full then Reconstruct.reconstruct_proof thy prop prf' else prf' end; wenzelm@17078: wenzelm@17078: fun pretty_proof thy prf = wenzelm@17078: Sign.pretty_term (proof_syntax prf thy) (term_of_proof prf); wenzelm@17078: berghofe@11522: fun pretty_proof_of full thm = wenzelm@17078: pretty_proof (Thm.theory_of_thm thm) (proof_of full thm); berghofe@11522: berghofe@11522: val print_proof_of = Pretty.writeln oo pretty_proof_of; berghofe@11522: berghofe@11522: end;