# HG changeset patch # User wenzelm # Date 1165510726 -3600 # Node ID f3c78a133fbbb6ab5326f263d6bca16f70488df1 # Parent 6c07cc87fe2b8e39a34a9d2c0c42527761f89dcd simplified add_abbrevs: no mixfix; diff -r 6c07cc87fe2b -r f3c78a133fbb src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Thu Dec 07 17:58:45 2006 +0100 +++ b/src/Pure/Isar/proof_context.ML Thu Dec 07 17:58:46 2006 +0100 @@ -142,8 +142,8 @@ val add_notation: Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context val expand_abbrevs: bool -> Proof.context -> Proof.context - val add_abbrevs: Syntax.mode -> ((bstring * mixfix) * term) list -> Proof.context -> - (term * term) list * Proof.context + val add_abbrevs: string -> (bstring * term) list -> Proof.context -> + ((string * typ) * term) list * Proof.context val verbose: bool ref val setmp_verbose: ('a -> 'b) -> 'a -> 'b val print_syntax: Proof.context -> unit @@ -877,21 +877,18 @@ val expand_abbrevs = map_consts o apsnd o Consts.expand_abbrevs; -fun add_abbrevs prmode = fold_map (fn ((raw_c, raw_mx), raw_t) => fn ctxt => +fun add_abbrevs mode = fold_map (fn (c, raw_t) => fn ctxt => let - val (c, mx) = Syntax.const_mixfix raw_c raw_mx; val t0 = cert_term (ctxt |> expand_abbrevs false) raw_t handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote c); val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0]; - val (((full_c, T), rhs), consts') = consts_of ctxt - |> Consts.abbreviate (pp ctxt) (tsig_of ctxt) (naming_of ctxt) (#1 prmode) ((c, t), true); + val (res, consts') = consts_of ctxt + |> Consts.abbreviate (pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode ((c, t), true); in ctxt |> Variable.declare_term t |> map_consts (apsnd (K consts')) - |> map_syntax (LocalSyntax.add_modesyntax (theory_of ctxt) - prmode [(false, (Syntax.constN ^ full_c, T, mx))]) - |> pair (Const (full_c, T), rhs) + |> pair res end); diff -r 6c07cc87fe2b -r f3c78a133fbb src/Pure/sign.ML --- a/src/Pure/sign.ML Thu Dec 07 17:58:45 2006 +0100 +++ b/src/Pure/sign.ML Thu Dec 07 17:58:46 2006 +0100 @@ -192,8 +192,8 @@ val read_prop: theory -> string -> term val add_consts_authentic: (bstring * typ * mixfix) list -> theory -> theory val add_notation: Syntax.mode -> (term * mixfix) list -> theory -> theory - val add_abbrevs: Syntax.mode -> ((bstring * mixfix) * term) list -> theory -> - (term * term) list * theory + val add_abbrevs: string -> (bstring * term) list -> theory -> + ((string * typ) * term) list * theory include SIGN_THEORY end @@ -751,22 +751,16 @@ end; -(* add abbreviations -- cf. ProofContext.add_abbrevs *) +(* add abbreviations *) -fun add_abbrevs prmode = fold_map (fn ((raw_c, raw_mx), raw_t) => fn thy => +fun add_abbrevs mode = fold_map (fn (c, raw_t) => fn thy => let - val (c, mx) = Syntax.const_mixfix raw_c raw_mx; val prep_tm = Compress.term thy o Term.no_dummy_patterns o cert_term_abbrev thy; val t = (prep_tm raw_t handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg) handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote c); - val (((full_c, T), rhs), consts') = consts_of thy - |> Consts.abbreviate (pp thy) (tsig_of thy) (naming_of thy) (#1 prmode) ((c, t), true); - in - thy - |> map_consts (K consts') - |> add_modesyntax_i prmode [(Syntax.constN ^ full_c, T, mx)] - |> pair (Const (full_c, T), rhs) - end); + val (res, consts') = consts_of thy + |> Consts.abbreviate (pp thy) (tsig_of thy) (naming_of thy) mode ((c, t), true); + in (res, thy |> map_consts (K consts')) end); (* add constraints *)