src/HOLCF/Tools/repdef.ML
changeset 39989 ad60d7311f43
parent 39986 38677db30cad
child 40038 9d061b3d8f46
--- a/src/HOLCF/Tools/repdef.ML	Mon Oct 11 07:09:42 2010 -0700
+++ b/src/HOLCF/Tools/repdef.ML	Mon Oct 11 08:32:09 2010 -0700
@@ -7,7 +7,7 @@
 signature REPDEF =
 sig
   type rep_info =
-    { emb_def: thm, prj_def: thm, sfp_def: thm, SFP: thm }
+    { emb_def: thm, prj_def: thm, defl_def: thm, DEFL: thm }
 
   val add_repdef: bool -> binding option -> binding * (string * sort) list * mixfix ->
     term -> (binding * binding) option -> theory ->
@@ -28,19 +28,19 @@
 (** type definitions **)
 
 type rep_info =
-  { emb_def: thm, prj_def: thm, sfp_def: thm, SFP: thm };
+  { emb_def: thm, prj_def: thm, defl_def: thm, DEFL: thm };
 
 (* building types and terms *)
 
 val udomT = @{typ udom};
-val sfpT = @{typ sfp};
+val deflT = @{typ defl};
 fun emb_const T = Const (@{const_name emb}, T ->> udomT);
 fun prj_const T = Const (@{const_name prj}, udomT ->> T);
-fun sfp_const T = Const (@{const_name sfp}, Term.itselfT T --> sfpT);
+fun defl_const T = Const (@{const_name defl}, Term.itselfT T --> deflT);
 
 fun mk_cast (t, x) =
   capply_const (udomT, udomT)
-  $ (capply_const (sfpT, udomT ->> udomT) $ @{const cast} $ t)
+  $ (capply_const (deflT, udomT ->> udomT) $ @{const cast} $ t)
   $ x;
 
 (* manipulating theorems *)
@@ -70,8 +70,8 @@
     val tmp_ctxt = tmp_ctxt |> Variable.declare_constraints defl;
 
     val deflT = Term.fastype_of defl;
-    val _ = if deflT = @{typ "sfp"} then ()
-            else error ("Not type sfp: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
+    val _ = if deflT = @{typ "defl"} then ()
+            else error ("Not type defl: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
 
     (*lhs*)
     val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt) raw_args;
@@ -84,12 +84,12 @@
       |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name);
 
     (*set*)
-    val in_defl = @{term "in_sfp :: udom => sfp => bool"};
+    val in_defl = @{term "in_defl :: udom => defl => bool"};
     val set = HOLogic.Collect_const udomT $ Abs ("x", udomT, in_defl $ Bound 0 $ defl);
 
     (*pcpodef*)
-    val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_sfp} 1;
-    val tac2 = rtac @{thm adm_mem_Collect_in_sfp} 1;
+    val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_defl} 1;
+    val tac2 = rtac @{thm adm_mem_Collect_in_defl} 1;
     val ((info, cpo_info, pcpo_info), thy) = thy
       |> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2);
 
@@ -99,12 +99,12 @@
     val emb_eqn = Logic.mk_equals (emb_const newT, cabs_const (newT, udomT) $ Rep_const);
     val prj_eqn = Logic.mk_equals (prj_const newT, cabs_const (udomT, newT) $
       Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0)));
-    val sfp_eqn = Logic.mk_equals (sfp_const newT,
+    val defl_eqn = Logic.mk_equals (defl_const newT,
       Abs ("x", Term.itselfT newT, defl));
     val name_def = Binding.suffix_name "_def" name;
     val emb_bind = (Binding.prefix_name "emb_" name_def, []);
     val prj_bind = (Binding.prefix_name "prj_" name_def, []);
-    val sfp_bind = (Binding.prefix_name "sfp_" name_def, []);
+    val defl_bind = (Binding.prefix_name "defl_" name_def, []);
 
     (*instantiate class rep*)
     val lthy = thy
@@ -113,34 +113,34 @@
         Specification.definition (NONE, (emb_bind, emb_eqn)) lthy;
     val ((_, (_, prj_ldef)), lthy) =
         Specification.definition (NONE, (prj_bind, prj_eqn)) lthy;
-    val ((_, (_, sfp_ldef)), lthy) =
-        Specification.definition (NONE, (sfp_bind, sfp_eqn)) lthy;
+    val ((_, (_, defl_ldef)), lthy) =
+        Specification.definition (NONE, (defl_bind, defl_eqn)) lthy;
     val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy);
     val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef;
     val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef;
-    val sfp_def = singleton (ProofContext.export lthy ctxt_thy) sfp_ldef;
+    val defl_def = singleton (ProofContext.export lthy ctxt_thy) defl_ldef;
     val type_definition_thm =
       MetaSimplifier.rewrite_rule
         (the_list (#set_def (#2 info)))
         (#type_definition (#2 info));
     val typedef_thms =
-      [type_definition_thm, #below_def cpo_info, emb_def, prj_def, sfp_def];
+      [type_definition_thm, #below_def cpo_info, emb_def, prj_def, defl_def];
     val thy = lthy
       |> Class.prove_instantiation_instance
           (K (Tactic.rtac (@{thm typedef_rep_class} OF typedef_thms) 1))
       |> Local_Theory.exit_global;
 
     (*other theorems*)
-    val sfp_thm' = Thm.transfer thy sfp_def;
-    val (SFP_thm, thy) = thy
+    val defl_thm' = Thm.transfer thy defl_def;
+    val (DEFL_thm, thy) = thy
       |> Sign.add_path (Binding.name_of name)
       |> Global_Theory.add_thm
-         ((Binding.prefix_name "SFP_" name,
-          Drule.zero_var_indexes (@{thm typedef_SFP} OF [sfp_thm'])), [])
+         ((Binding.prefix_name "DEFL_" name,
+          Drule.zero_var_indexes (@{thm typedef_DEFL} OF [defl_thm'])), [])
       ||> Sign.restore_naming thy;
 
     val rep_info =
-      { emb_def = emb_def, prj_def = prj_def, sfp_def = sfp_def, SFP = SFP_thm };
+      { emb_def = emb_def, prj_def = prj_def, defl_def = defl_def, DEFL = DEFL_thm };
   in
     ((info, cpo_info, pcpo_info, rep_info), thy)
   end