src/HOLCF/Tools/Domain/domain_axioms.ML
changeset 35529 089e438b925b
parent 35525 fa231b86cb1e
child 35556 730fdfbbd5f8
--- a/src/HOLCF/Tools/Domain/domain_axioms.ML	Tue Mar 02 19:45:37 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_axioms.ML	Tue Mar 02 20:04:17 2010 -0800
@@ -1,22 +1,22 @@
 (*  Title:      HOLCF/Tools/Domain/domain_axioms.ML
     Author:     David von Oheimb
+    Author:     Brian Huffman
 
 Syntax generator for domain command.
 *)
 
 signature DOMAIN_AXIOMS =
 sig
+  val axiomatize_isomorphism :
+      binding * (typ * typ) ->
+      theory -> Domain_Take_Proofs.iso_info * theory
+
   val copy_of_dtyp :
       string Symtab.table -> (int -> term) -> Datatype.dtyp -> term
 
-  val calc_axioms :
-      Domain_Library.eq list -> int -> Domain_Library.eq ->
-      string * (string * term) list
-
   val add_axioms :
-      ((string * typ list) *
-       (binding * (bool * binding option * typ) list * mixfix) list) list ->
-      Domain_Library.eq list -> theory -> theory
+      (binding * (typ * typ)) list ->
+      theory -> theory
 end;
 
 
@@ -46,24 +46,57 @@
       SOME f => list_ccomb (%%:f, map (copy_of_dtyp tab r) ds)
     | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID);
 
-fun calc_axioms
-    (eqs : eq list)
-    (n : int)
-    (eqn as ((dname,_),cons) : eq)
-    : string * (string * term) list =
+local open HOLCF_Library in
+
+fun axiomatize_isomorphism
+    (dbind : binding, (lhsT, rhsT))
+    (thy : theory)
+    : Domain_Take_Proofs.iso_info * theory =
   let
-    val dc_abs = %%:(dname^"_abs");
-    val dc_rep = %%:(dname^"_rep");
-    val x_name'= "x";
-    val x_name = idx_name eqs x_name' (n+1);
-    val dnam = Long_Name.base_name dname;
+    val dname = Long_Name.base_name (Binding.name_of dbind);
+
+    val abs_bind = Binding.suffix_name "_abs" dbind;
+    val rep_bind = Binding.suffix_name "_rep" dbind;
+
+    val (abs_const, thy) =
+        Sign.declare_const ((abs_bind, rhsT ->> lhsT), NoSyn) thy;
+    val (rep_const, thy) =
+        Sign.declare_const ((rep_bind, lhsT ->> rhsT), NoSyn) thy;
+
+    val x = Free ("x", lhsT);
+    val y = Free ("y", rhsT);
+
+    val abs_iso_eqn =
+        Logic.all y (mk_trp (mk_eq (rep_const ` (abs_const ` y), y)));
+    val rep_iso_eqn =
+        Logic.all x (mk_trp (mk_eq (abs_const ` (rep_const ` x), x)));
+
+    val thy = Sign.add_path dname thy;
 
-    val abs_iso_ax = ("abs_iso", mk_trp(dc_rep`(dc_abs`%x_name') === %:x_name'));
-    val rep_iso_ax = ("rep_iso", mk_trp(dc_abs`(dc_rep`%x_name') === %:x_name'));
+    val (abs_iso_thm, thy) =
+        yield_singleton PureThy.add_axioms
+        ((Binding.name "abs_iso", abs_iso_eqn), []) thy;
+
+    val (rep_iso_thm, thy) =
+        yield_singleton PureThy.add_axioms
+        ((Binding.name "rep_iso", rep_iso_eqn), []) thy;
+
+    val thy = Sign.parent_path thy;
+
+    val result =
+        {
+          absT = lhsT,
+          repT = rhsT,
+          abs_const = abs_const,
+          rep_const = rep_const,
+          abs_inverse = abs_iso_thm,
+          rep_inverse = rep_iso_thm
+        };
   in
-    (dnam, [abs_iso_ax, rep_iso_ax])
+    (result, thy)
   end;
 
+end;
 
 (* legacy type inference *)
 
@@ -78,52 +111,30 @@
 fun add_axioms_i x = snd o PureThy.add_axioms (map (Thm.no_attributes o apfst Binding.name) x);
 fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy;
 
-fun add_axioms eqs' (eqs : eq list) thy' =
+fun add_axioms
+    (dom_eqns : (binding * (typ * typ)) list)
+    (thy : theory) =
   let
-    val dnames = map (fst o fst) eqs;
-    val x_name = idx_name dnames "x"; 
+
+    (* declare and axiomatize abs/rep *)
+    val (iso_infos, thy) =
+        fold_map axiomatize_isomorphism dom_eqns thy;
 
     fun add_one (dnam, axs) =
         Sign.add_path dnam
           #> add_axioms_infer axs
           #> Sign.parent_path;
 
-    val axs = mapn (calc_axioms eqs) 0 eqs;
-    val thy = thy' |> fold add_one axs;
-
-    fun get_iso_info ((dname, tyvars), cons') =
-      let
-        fun opt_lazy (lazy,_,t) = if lazy then mk_uT t else t
-        fun prod     (_,args,_) =
-            case args of [] => oneT
-                       | _ => foldr1 mk_sprodT (map opt_lazy args);
-        val ax_abs_iso = PureThy.get_thm thy (dname ^ ".abs_iso");
-        val ax_rep_iso = PureThy.get_thm thy (dname ^ ".rep_iso");
-        val lhsT = Type(dname,tyvars);
-        val rhsT = foldr1 mk_ssumT (map prod cons');
-        val rep_const = Const(dname^"_rep", lhsT ->> rhsT);
-        val abs_const = Const(dname^"_abs", rhsT ->> lhsT);
-      in
-        {
-          absT = lhsT,
-          repT = rhsT,
-          abs_const = abs_const,
-          rep_const = rep_const,
-          abs_inverse = ax_abs_iso,
-          rep_inverse = ax_rep_iso
-        }
-      end;
-    val dom_binds = map (Binding.name o Long_Name.base_name) dnames;
+    (* define take function *)
     val (take_info, thy) =
         Domain_Take_Proofs.define_take_functions
-          (dom_binds ~~ map get_iso_info eqs') thy;
+          (map fst dom_eqns ~~ iso_infos) thy;
 
     (* declare lub_take axioms *)
     local
-      fun ax_lub_take dname =
+      fun ax_lub_take (dbind, take_const) =
         let
-          val dnam : string = Long_Name.base_name dname;
-          val take_const = %%:(dname^"_take");
+          val dnam = Long_Name.base_name (Binding.name_of dbind);
           val lub = %%: @{const_name lub};
           val image = %%: @{const_name image};
           val UNIV = @{term "UNIV :: nat set"};
@@ -132,11 +143,14 @@
         in
           add_one (dnam, [("lub_take", ax)])
         end
+      val dbinds = map fst dom_eqns;
+      val take_consts = #take_consts take_info;
     in
-      val thy = fold ax_lub_take dnames thy
+      val thy = fold ax_lub_take (dbinds ~~ take_consts) thy
     end;
+
   in
-    thy
-  end; (* let (add_axioms) *)
+    thy (* TODO: also return iso_infos, take_info, lub_take_thms *)
+  end;
 
 end; (* struct *)