# HG changeset patch # User ballarin # Date 1274690912 -7200 # Node ID 7099a9ed3be2797e5156e852d0939150ea286b69 # Parent c11cdb5e7e9764e8ee48498d7482770aa4996a76 Reapply mixin patch: base for performance improvements. diff -r c11cdb5e7e97 -r 7099a9ed3be2 src/FOL/ex/LocaleTest.thy --- a/src/FOL/ex/LocaleTest.thy Sun May 23 19:30:29 2010 -0700 +++ b/src/FOL/ex/LocaleTest.thy Mon May 24 10:48:32 2010 +0200 @@ -533,7 +533,7 @@ grefl: "gle(x, x)" gless_def: "gless(x, y) <-> gle(x, y) & x ~= y" grefl': "gle'(x, x)" gless'_def: "gless'(x, y) <-> gle'(x, y) & x ~= y" -text {* Mixin not applied to locales further up the hierachy. *} +text {* Setup *} locale mixin = reflexive begin @@ -548,16 +548,6 @@ by (simp add: reflexive.less_def[OF reflexive] gless_def) qed -thm le.less_def (* mixin not applied *) -lemma "reflexive.less(gle, x, y) <-> gle(x, y) & x ~= y" by (rule le.less_def) -thm le.less_thm (* mixin applied *) -lemma "gless(x, y) <-> gle(x, y) & x ~= y" by (rule le.less_thm) - -lemma "reflexive.less(gle, x, y) <-> gle(x, y) & x ~= y" - by (rule le.less_def) -lemma "gless(x, y) <-> gle(x, y) & x ~= y" - by (rule le.less_thm) - text {* Mixin propagated along the locale hierarchy *} locale mixin2 = mixin diff -r c11cdb5e7e97 -r 7099a9ed3be2 src/Pure/Isar/class.ML --- a/src/Pure/Isar/class.ML Sun May 23 19:30:29 2010 -0700 +++ b/src/Pure/Isar/class.ML Mon May 24 10:48:32 2010 +0200 @@ -293,8 +293,7 @@ |-> (fn (param_map, params, assm_axiom) => `(fn thy => calculate thy class sups base_sort param_map assm_axiom) #-> (fn (base_morph, eq_morph, export_morph, axiom, assm_intro, of_class) => - Locale.add_registration (class, case eq_morph of SOME eq_morph => base_morph $> eq_morph | NONE => base_morph) - (*FIXME should avoid modification of base morphism, but then problem with sublocale linorder < distrib_lattice*) + Locale.add_registration (class, base_morph) (Option.map (rpair true) eq_morph) export_morph #> register class sups params base_sort base_morph export_morph axiom assm_intro of_class)) |> Theory_Target.init (SOME class) diff -r c11cdb5e7e97 -r 7099a9ed3be2 src/Pure/Isar/locale.ML --- a/src/Pure/Isar/locale.ML Sun May 23 19:30:29 2010 -0700 +++ b/src/Pure/Isar/locale.ML Mon May 24 10:48:32 2010 +0200 @@ -384,9 +384,10 @@ |-> (fn regs => fn _ => map (fn ((name, (base, export)) ,_) => (name, base $> (collect_mixins thy (name, base $> export) |> compose_mixins) $> export)) regs); -fun all_registrations thy = Registrations.get thy - |-> (fn regs => fn mixins => map (reg_morph mixins) regs); - (* without inherited mixins *) +fun all_registrations thy = Registrations.get thy (* FIXME clone *) + (* with inherited mixins *) + |-> (fn regs => fn _ => map (fn ((name, (base, export)) ,_) => + (name, base $> (collect_mixins thy (name, base $> export) |> compose_mixins) $> export)) regs); fun activate_notes' activ_elem transfer thy export (name, morph) input = let @@ -462,19 +463,21 @@ fun add_registration (name, base_morph) mixin export thy = let val base = instance_of thy name base_morph; + val mix = case mixin of NONE => Morphism.identity + | SOME (mix, _) => mix; in if member (ident_eq thy) (get_idents (Context.Theory thy)) (name, base) then thy else (get_idents (Context.Theory thy), thy) (* add new registrations with inherited mixins *) - |> roundup thy (add_reg export) export (name, base_morph) + |> roundup thy (add_reg export) export (name, base_morph $> mix) (* FIXME *) |> snd (* add mixin *) |> (case mixin of NONE => I - | SOME mixin => amend_registration (name, base_morph) mixin export) + | SOME mixin => amend_registration (name, base_morph $> mix) mixin export) (* activate import hierarchy as far as not already active *) - |> Context.theory_map (activate_facts' export (name, base_morph)) + |> Context.theory_map (activate_facts' export (name, base_morph $> mix)) end;