# HG changeset patch # User berghofe # Date 1001668060 -7200 # Node ID 3131fa12d42599181ecad84afe7a323cde401182 # Parent c92a99a2b5b1a0a0fc5475bb4ce6d3cb8ad4d5c6 - Tuned syntax - proof_of_term: fixed problems with dummy patterns and typing information diff -r c92a99a2b5b1 -r 3131fa12d425 src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML Fri Sep 28 11:05:37 2001 +0200 +++ b/src/Pure/Proof/proof_syntax.ML Fri Sep 28 11:07:40 2001 +0200 @@ -31,7 +31,8 @@ (**** add special syntax for embedding proof terms ****) val proofT = Type ("proof", []); -val lamT = Type ("lam_syn", []); +val paramT = Type ("param", []); +val paramsT = Type ("params", []); val idtT = Type ("idt", []); val aT = TFree ("'a", ["logic"]); @@ -43,7 +44,7 @@ (map (fn name => (name, proofT, NoSyn)) names) (Sign.add_path "//" sg); (** constants for application and abstraction **) - + fun add_proof_syntax sg = sg |> Sign.copy @@ -52,33 +53,35 @@ |> Sign.add_types [("proof", 0, NoSyn)] |> Sign.add_arities [("proof", [], "logic")] |> Sign.add_consts_i - [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)), - ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)), + [("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ %/ _)", [4, 5], 4)), + ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ %%/ _)", [4, 5], 4)), ("Abst", (aT --> proofT) --> proofT, NoSyn), ("AbsP", [propT, proofT --> proofT] ---> proofT, NoSyn)] - |> Sign.add_nonterminals ["lam_syn"] + |> Sign.add_nonterminals ["param", "params"] |> Sign.add_syntax_i - [("_Lam", [lamT, proofT] ---> proofT, Mixfix ("(3Lam _./ _)", [0,0], 1)), - ("_Lam0", [lamT, lamT] ---> lamT, Mixfix ("_,/ _", [1, 0], 0)), - ("_Lam1", [idtT, propT] ---> lamT, Mixfix ("_ : _", [0, 0], 1)), - ("_Lam2", idtT --> lamT, Mixfix ("_", [0], 1))] + [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(3Lam _./ _)", [0, 3], 3)), + ("_Lam0", [paramT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)), + ("_Lam0", [idtT, paramsT] ---> paramsT, Mixfix ("_/ _", [1, 0], 0)), + ("_Lam1", [idtT, propT] ---> paramT, Mixfix ("_: _", [0, 0], 0)), + ("", paramT --> paramT, Delimfix "'(_')"), + ("", idtT --> paramsT, Delimfix "_"), + ("", paramT --> paramsT, Delimfix "_")] |> Sign.add_modesyntax_i (("xsymbols", true), - [("_Lam", [lamT, proofT] ---> proofT, Mixfix ("(3\\_./ _)", [0,0], 1)), + [("_Lam", [paramsT, proofT] ---> proofT, Mixfix ("(3\\_./ _)", [0, 3], 3)), ("Appt", [proofT, aT] ---> proofT, Mixfix ("(1_ \\/ _)", [4, 5], 4)), ("AppP", [proofT, proofT] ---> proofT, Mixfix ("(1_ \\/ _)", [4, 5], 4))]) |> Sign.add_trrules_i (map Syntax.ParsePrintRule [(Syntax.mk_appl (Constant "_Lam") + [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"], + Syntax.mk_appl (Constant "_Lam") + [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]]), + (Syntax.mk_appl (Constant "_Lam") [Syntax.mk_appl (Constant "_Lam1") [Variable "x", Variable "A"], Variable "B"], Syntax.mk_appl (Constant "AbsP") [Variable "A", (Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "B"])]), - (Syntax.mk_appl (Constant "_Lam") - [Syntax.mk_appl (Constant "_Lam2") [Variable "x"], Variable "A"], + (Syntax.mk_appl (Constant "_Lam") [Variable "x", Variable "A"], Syntax.mk_appl (Constant "Abst") - [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])]), - (Syntax.mk_appl (Constant "_Lam") - [Syntax.mk_appl (Constant "_Lam0") [Variable "l", Variable "m"], Variable "A"], - Syntax.mk_appl (Constant "_Lam") - [Variable "l", Syntax.mk_appl (Constant "_Lam") [Variable "m", Variable "A"]])]); + [(Syntax.mk_appl (Constant "_abs") [Variable "x", Variable "A"])])]); (**** create unambiguous theorem names ****) @@ -99,8 +102,8 @@ fun rename (Abst (s, T, prf)) = Abst (s, T, rename prf) | rename (AbsP (s, t, prf)) = AbsP (s, t, rename prf) - | rename (prf1 % prf2) = rename prf1 % rename prf2 - | rename (prf %% t) = rename prf %% t + | rename (prf1 %% prf2) = rename prf1 %% rename prf2 + | rename (prf % t) = rename prf % t | rename (prf' as PThm ((s, tags), prf, prop, Ts)) = let val prop' = if_none (assoc (thms', s)) (Bound 0); @@ -126,18 +129,21 @@ val thms = flat (map thms_of thys); val axms = flat (map (Symtab.dest o #axioms o rep_theory) thys); + fun mk_term t = (if ty then I else map_term_types (K dummyT)) + (Term.no_dummy_patterns t); + fun prf_of [] (Bound i) = PBound i | prf_of Ts (Const (s, Type ("proof", _))) = change_type (if ty then Some Ts else None) (case NameSpace.unpack s of - "Axm" :: xs => + "axm" :: xs => let val name = NameSpace.pack xs; val prop = (case assoc (axms, name) of Some prop => prop | None => error ("Unknown axiom " ^ quote name)) in PAxm (name, prop, None) end - | "Thm" :: xs => + | "thm" :: xs => let val name = NameSpace.pack xs; in (case assoc (thms, name) of Some thm => fst (strip_combt (#2 (#der (rep_thm thm)))) @@ -151,14 +157,17 @@ Abst (s, if ty then Some T else None, incr_pboundvars (~1) 0 (prf_of [] prf)) | prf_of [] (Const ("AbsP", _) $ t $ Abs (s, _, prf)) = - AbsP (s, case t of Const ("dummy_pattern", _) => None | _ => Some t, + AbsP (s, case t of + Const ("dummy_pattern", _) => None + | _ $ Const ("dummy_pattern", _) => None + | _ => Some (mk_term t), incr_pboundvars 0 (~1) (prf_of [] prf)) | prf_of [] (Const ("AppP", _) $ prf1 $ prf2) = - prf_of [] prf1 % prf_of [] prf2 + prf_of [] prf1 %% prf_of [] prf2 | prf_of Ts (Const ("Appt", _) $ prf $ Const ("TYPE", Type (_, [T]))) = prf_of (T::Ts) prf - | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf %% - (case t of Const ("dummy_pattern", _) => None | _ => Some t) + | prf_of [] (Const ("Appt", _) $ prf $ t) = prf_of [] prf % + (case t of Const ("dummy_pattern", _) => None | _ => Some (mk_term t)) | prf_of _ t = error ("Not a proof term:\n" ^ Sign.string_of_term (sign_of thy) t) @@ -175,12 +184,12 @@ [proofT, itselfT T] ---> proofT) $ prf $ Logic.mk_type T); fun term_of _ (PThm ((name, _), _, _, None)) = - Const (add_prefix "Thm" name, proofT) + Const (add_prefix "thm" name, proofT) | term_of _ (PThm ((name, _), _, _, Some Ts)) = - mk_tyapp (Const (add_prefix "Thm" name, proofT), Ts) - | term_of _ (PAxm (name, _, None)) = Const (add_prefix "Axm" name, proofT) + mk_tyapp (Const (add_prefix "thm" name, proofT), Ts) + | term_of _ (PAxm (name, _, None)) = Const (add_prefix "axm" name, proofT) | term_of _ (PAxm (name, _, Some Ts)) = - mk_tyapp (Const (add_prefix "Axm" name, proofT), Ts) + mk_tyapp (Const (add_prefix "axm" name, proofT), Ts) | term_of _ (PBound i) = Bound i | term_of Ts (Abst (s, opT, prf)) = let val T = if_none opT dummyT @@ -190,9 +199,9 @@ | term_of Ts (AbsP (s, t, prf)) = AbsPt $ if_none t (Const ("dummy_pattern", propT)) $ Abs (s, proofT, term_of (proofT::Ts) (incr_pboundvars 0 1 prf)) - | term_of Ts (prf1 % prf2) = + | term_of Ts (prf1 %% prf2) = AppPt $ term_of Ts prf1 $ term_of Ts prf2 - | term_of Ts (prf %% opt) = + | term_of Ts (prf % opt) = let val t = if_none opt (Const ("dummy_pattern", dummyT)) in Const ("Appt", [proofT, fastype_of1 (Ts, t) handle TERM _ => dummyT] ---> proofT) $ @@ -214,7 +223,7 @@ val sg = sign_of thy |> add_proof_syntax |> add_proof_atom_consts - (map (add_prefix "Thm") thm_names @ map (add_prefix "Axm") axm_names) + (map (add_prefix "thm") thm_names @ map (add_prefix "axm") axm_names) in (cterm_of sg (term_of_proof prf'), proof_of_term thy tab true o Thm.term_of) @@ -228,7 +237,7 @@ val sg = sign_of thy |> add_proof_syntax |> add_proof_atom_consts - (map (add_prefix "Thm") thm_names @ map (add_prefix "Axm") axm_names) + (map (add_prefix "thm") thm_names @ map (add_prefix "axm") axm_names) in (fn T => fn s => Thm.term_of (read_cterm sg (s, T))) end; @@ -246,7 +255,7 @@ val sg' = sg |> add_proof_syntax |> add_proof_atom_consts - (map (add_prefix "Thm") thm_names @ map (add_prefix "Axm") axm_names) + (map (add_prefix "thm") thm_names @ map (add_prefix "axm") axm_names) in Sign.pretty_term sg' (term_of_proof prf) end; @@ -259,7 +268,7 @@ | _ => prf) in pretty_proof sign - (if full then Reconstruct.reconstruct_prf sign prop prf' else prf') + (if full then Reconstruct.reconstruct_proof sign prop prf' else prf') end; val print_proof_of = Pretty.writeln oo pretty_proof_of;