Add mixin to base morphism, required by class package; cf ab324ffd6f3d.
authorballarin
Thu, 13 May 2010 13:30:16 +0200
changeset 36905 b47fd7148b57
parent 36904 3e200347a22e
child 36906 9eff24f4e5db
Add mixin to base morphism, required by class package; cf ab324ffd6f3d.
src/FOL/ex/LocaleTest.thy
src/Pure/Isar/locale.ML
--- a/src/FOL/ex/LocaleTest.thy	Thu May 13 13:29:43 2010 +0200
+++ b/src/FOL/ex/LocaleTest.thy	Thu May 13 13:30:16 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
--- a/src/Pure/Isar/locale.ML	Thu May 13 13:29:43 2010 +0200
+++ b/src/Pure/Isar/locale.ML	Thu May 13 13:30:16 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;