--- a/src/Pure/Isar/theory_target.ML Wed Aug 11 14:19:32 2010 +0200
+++ b/src/Pure/Isar/theory_target.ML Wed Aug 11 14:20:34 2010 +0200
@@ -7,18 +7,11 @@
signature THEORY_TARGET =
sig
- val peek: local_theory ->
- {target: string,
- is_locale: bool,
- is_class: bool,
- instantiation: string list * (string * sort) list * sort,
- overloading: (string * (string * typ) * bool) list}
+ val peek: local_theory -> {target: string, is_locale: bool, is_class: bool}
val init: string option -> theory -> local_theory
val context_cmd: xstring -> theory -> local_theory
val instantiation: string list * (string * sort) list * sort -> theory -> local_theory
val instantiation_cmd: xstring list * xstring list * xstring -> theory -> local_theory
- val overloading: (string * (string * typ) * bool) list -> theory -> local_theory
- val overloading_cmd: (string * string * bool) list -> theory -> local_theory
end;
structure Theory_Target: THEORY_TARGET =
@@ -26,15 +19,12 @@
(* context data *)
-datatype target = Target of {target: string, is_locale: bool,
- is_class: bool, instantiation: string list * (string * sort) list * sort,
- overloading: (string * (string * typ) * bool) list};
+datatype target = Target of {target: string, is_locale: bool, is_class: bool};
-fun make_target target is_locale is_class instantiation overloading =
- Target {target = target, is_locale = is_locale,
- is_class = is_class, instantiation = instantiation, overloading = overloading};
+fun make_target target is_locale is_class =
+ Target {target = target, is_locale = is_locale, is_class = is_class};
-val global_target = make_target "" false false ([], [], []) [];
+val global_target = make_target "" false false;
structure Data = Proof_Data
(
@@ -47,28 +37,18 @@
(* generic declarations *)
-fun theory_declaration decl lthy =
- let
- val global_decl = Morphism.form
- (Morphism.transform (Local_Theory.global_morphism lthy) decl);
- in
- lthy
- |> Local_Theory.theory (Context.theory_map global_decl)
- |> Local_Theory.target (Context.proof_map global_decl)
- end;
-
fun locale_declaration locale { syntax, pervasive } decl lthy =
let
val add = if syntax then Locale.add_syntax_declaration else Locale.add_declaration;
val locale_decl = Morphism.transform (Local_Theory.target_morphism lthy) decl;
in
lthy
- |> pervasive ? theory_declaration decl
+ |> pervasive ? Generic_Target.theory_declaration decl
|> Local_Theory.target (add locale locale_decl)
end;
fun target_declaration (Target {target, ...}) { syntax, pervasive } =
- if target = "" then theory_declaration
+ if target = "" then Generic_Target.theory_declaration
else locale_declaration target { syntax = syntax, pervasive = pervasive };
@@ -111,23 +91,14 @@
fun syntax_error c = error ("Illegal mixfix syntax for overloaded constant " ^ quote c);
-fun theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =
- let
- val (const, lthy2) = lthy |> Local_Theory.theory_result
- (Sign.declare_const ((b, U), mx));
- val lhs = list_comb (const, type_params @ term_params);
- val ((_, def), lthy3) = lthy2 |> Local_Theory.theory_result
- (Thm.add_def false false (b_def, Logic.mk_equals (lhs, rhs)));
- in ((lhs, def), lthy3) end;
-
fun locale_foundation ta (((b, U), mx), (b_def, rhs)) (type_params, term_params) =
- theory_foundation (((b, U), NoSyn), (b_def, rhs)) (type_params, term_params)
+ Generic_Target.theory_foundation (((b, U), NoSyn), (b_def, rhs)) (type_params, term_params)
#-> (fn (lhs, def) => locale_const_declaration ta Syntax.mode_default ((b, mx), lhs)
#> pair (lhs, def))
fun class_foundation (ta as Target {target, ...})
(((b, U), mx), (b_def, rhs)) (type_params, term_params) =
- theory_foundation (((b, U), NoSyn), (b_def, rhs)) (type_params, term_params)
+ Generic_Target.theory_foundation (((b, U), NoSyn), (b_def, rhs)) (type_params, term_params)
#-> (fn (lhs, def) => locale_const_declaration ta Syntax.mode_default ((b, NoSyn), lhs)
#> class_target ta (Class_Target.declare target ((b, mx), (type_params, lhs)))
#> pair (lhs, def))
@@ -139,16 +110,7 @@
##>> AxClass.define_overloaded b_def (c, rhs))
||> Class_Target.confirm_declaration b
| NONE => lthy |>
- theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params);
-
-fun overloading_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =
- case Overloading.operation lthy b
- of SOME (c, checked) => if mx <> NoSyn then syntax_error c
- else lthy |> Local_Theory.theory_result (Overloading.declare (c, U)
- ##>> Overloading.define checked b_def (c, rhs))
- ||> Overloading.confirm b
- | NONE => lthy |>
- theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params);
+ Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params);
fun fork_mixfix (Target {is_locale, is_class, ...}) mx =
if not is_locale then (NoSyn, NoSyn, mx)
@@ -192,16 +154,6 @@
(* notes *)
-fun theory_notes kind global_facts lthy =
- let
- val thy = ProofContext.theory_of lthy;
- val global_facts' = Attrib.map_facts (Attrib.attribute_i thy) global_facts;
- in
- lthy
- |> Local_Theory.theory (PureThy.note_thmss kind global_facts' #> snd)
- |> Local_Theory.target (ProofContext.note_thmss kind global_facts' #> snd)
- end;
-
fun locale_notes locale kind global_facts local_facts lthy =
let
val global_facts' = Attrib.map_facts (K I) global_facts;
@@ -215,15 +167,11 @@
fun target_notes (ta as Target {target, is_locale, ...}) =
if is_locale then locale_notes target
- else fn kind => fn global_facts => fn _ => theory_notes kind global_facts;
+ else fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts;
(* abbrev *)
-fun theory_abbrev prmode ((b, mx), t) = Local_Theory.theory
- (Sign.add_abbrev (#1 prmode) (b, t) #->
- (fn (lhs, _) => Sign.notation true prmode [(lhs, mx)]));
-
fun locale_abbrev ta prmode ((b, mx), t) xs = Local_Theory.theory_result
(Sign.add_abbrev Print_Mode.internal (b, t)) #->
(fn (lhs, _) => locale_const_declaration ta prmode
@@ -236,7 +184,7 @@
|> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), global_rhs) xs
|> is_class ? class_target ta (Class_Target.abbrev target prmode ((b, mx), t'))
else lthy
- |> theory_abbrev prmode ((b, mx), global_rhs);
+ |> Generic_Target.theory_abbrev prmode ((b, mx), global_rhs);
(* pretty *)
@@ -261,26 +209,22 @@
map (Pretty.chunks o Element.pretty_ctxt ctxt) elems))]
end;
-fun pretty (Target {target, is_class, instantiation, overloading, ...}) ctxt =
+fun pretty (Target {target, is_class, ...}) ctxt =
Pretty.block [Pretty.command "theory", Pretty.brk 1,
Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::
- (if not (null overloading) then [Overloading.pretty ctxt]
- else if not (null (#1 instantiation)) then [Class_Target.pretty_instantiation ctxt]
- else pretty_thy ctxt target is_class);
+ pretty_thy ctxt target is_class;
(* init various targets *)
local
-fun init_data (Target {target, is_locale, is_class, instantiation, overloading}) =
- if not (null (#1 instantiation)) then Class_Target.init_instantiation instantiation
- else if not (null overloading) then Overloading.init overloading
- else if not is_locale then ProofContext.init_global
+fun init_data (Target {target, is_locale, is_class}) =
+ if not is_locale then ProofContext.init_global
else if not is_class then Locale.init target
else Class_Target.init target;
-fun init_target (ta as Target {target, instantiation, overloading, ...}) thy =
+fun init_target (ta as Target {target, ...}) thy =
thy
|> init_data ta
|> Data.put ta
@@ -294,39 +238,14 @@
{ syntax = true, pervasive = pervasive },
pretty = pretty ta,
reinit = init_target ta o ProofContext.theory_of,
- exit = Local_Theory.target_of o
- (if not (null (#1 instantiation)) then Class_Target.conclude_instantiation
- else if not (null overloading) then Overloading.conclude
- else I)};
+ exit = Local_Theory.target_of};
fun named_target _ NONE = global_target
| named_target thy (SOME target) =
if Locale.defined thy target
- then make_target target true (Class_Target.is_class thy target) ([], [], []) []
+ then make_target target true (Class_Target.is_class thy target)
else error ("No such locale: " ^ quote target);
-fun gen_overloading prep_const raw_ops thy =
- let
- val ctxt = ProofContext.init_global thy;
- val ops = raw_ops |> map (fn (name, const, checked) =>
- (name, Term.dest_Const (prep_const ctxt const), checked));
- in
- thy
- |> Overloading.init ops
- |> Local_Theory.init NONE ""
- {define = Generic_Target.define overloading_foundation,
- notes = Generic_Target.notes
- (fn kind => fn global_facts => fn _ => theory_notes kind global_facts),
- abbrev = Generic_Target.abbrev
- (fn prmode => fn (b, mx) => fn (t, _) => fn _ =>
- theory_abbrev prmode ((b, mx), t)),
- declaration = fn pervasive => theory_declaration,
- syntax_declaration = fn pervasive => theory_declaration,
- pretty = single o Overloading.pretty,
- reinit = gen_overloading prep_const raw_ops o ProofContext.theory_of,
- exit = Local_Theory.target_of o Overloading.conclude}
- end;
-
in
fun init target thy = init_target (named_target thy target) thy;
@@ -340,11 +259,11 @@
|> Local_Theory.init NONE ""
{define = Generic_Target.define instantiation_foundation,
notes = Generic_Target.notes
- (fn kind => fn global_facts => fn _ => theory_notes kind global_facts),
+ (fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts),
abbrev = Generic_Target.abbrev
- (fn prmode => fn (b, mx) => fn (t, _) => fn _ => theory_abbrev prmode ((b, mx), t)),
- declaration = fn pervasive => theory_declaration,
- syntax_declaration = fn pervasive => theory_declaration,
+ (fn prmode => fn (b, mx) => fn (t, _) => fn _ => Generic_Target.theory_abbrev prmode ((b, mx), t)),
+ declaration = K Generic_Target.theory_declaration,
+ syntax_declaration = K Generic_Target.theory_declaration,
pretty = single o Class_Target.pretty_instantiation,
reinit = instantiation arities o ProofContext.theory_of,
exit = Local_Theory.target_of o Class_Target.conclude_instantiation};
@@ -352,9 +271,6 @@
fun instantiation_cmd arities thy =
instantiation (Class_Target.read_multi_arity thy arities) thy;
-val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);
-val overloading_cmd = gen_overloading Syntax.read_term;
-
end;
end;