rename constant 'one_when' to 'one_case'
authorhuffman
Sun, 24 Oct 2010 15:19:17 -0700
changeset 40212 20df78048db5
parent 40099 0fb7891f0c7c
child 40213 b63e966564da
rename constant 'one_when' to 'one_case'
src/HOLCF/One.thy
src/HOLCF/Tools/Domain/domain_constructors.ML
src/HOLCF/Tools/holcf_library.ML
--- a/src/HOLCF/One.thy	Sun Oct 24 15:11:24 2010 -0700
+++ b/src/HOLCF/One.thy	Sun Oct 24 15:19:17 2010 -0700
@@ -53,20 +53,20 @@
 text {* Case analysis function for type @{typ one} *}
 
 definition
-  one_when :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a" where
-  "one_when = (\<Lambda> a x. strict\<cdot>x\<cdot>a)"
+  one_case :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a" where
+  "one_case = (\<Lambda> a x. strict\<cdot>x\<cdot>a)"
 
 translations
-  "case x of XCONST ONE \<Rightarrow> t" == "CONST one_when\<cdot>t\<cdot>x"
-  "\<Lambda> (XCONST ONE). t" == "CONST one_when\<cdot>t"
+  "case x of XCONST ONE \<Rightarrow> t" == "CONST one_case\<cdot>t\<cdot>x"
+  "\<Lambda> (XCONST ONE). t" == "CONST one_case\<cdot>t"
 
-lemma one_when1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
-by (simp add: one_when_def)
+lemma one_case1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
+by (simp add: one_case_def)
 
-lemma one_when2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
-by (simp add: one_when_def)
+lemma one_case2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
+by (simp add: one_case_def)
 
-lemma one_when3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
+lemma one_case3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
 by (induct x rule: one_induct) simp_all
 
 end
--- a/src/HOLCF/Tools/Domain/domain_constructors.ML	Sun Oct 24 15:11:24 2010 -0700
+++ b/src/HOLCF/Tools/Domain/domain_constructors.ML	Sun Oct 24 15:19:17 2010 -0700
@@ -414,7 +414,7 @@
       val case_bind = Binding.suffix_name "_when" dbind;
       fun lambda_arg (lazy, v) t =
           (if lazy then mk_fup else I) (big_lambda v t);
-      fun lambda_args []      t = mk_one_when t
+      fun lambda_args []      t = mk_one_case t
         | lambda_args (x::[]) t = lambda_arg x t
         | lambda_args (x::xs) t = mk_ssplit (lambda_arg x (lambda_args xs t));
       fun one_con f (_, args) =
@@ -490,7 +490,7 @@
       let
         val defs = case_beta :: map mk_meta_eq [rep_strict, @{thm cfcomp2}];
         val goal = mk_trp (mk_strict case_app);
-        val rules = @{thms sscase1 ssplit1 strictify1 one_when1};
+        val rules = @{thms sscase1 ssplit1 strictify1 one_case1};
         val tacs = [resolve_tac rules 1];
       in prove thy defs goal (K tacs) end;
         
@@ -507,7 +507,7 @@
           val defs = case_beta :: con_betas;
           val rules1 = @{thms strictify2 sscase2 sscase3 ssplit2 fup2 ID1};
           val rules2 = @{thms con_defined_iff_rules};
-          val rules3 = @{thms cfcomp2 one_when2};
+          val rules3 = @{thms cfcomp2 one_case2};
           val rules = abs_inverse :: rules1 @ rules2 @ rules3;
           val tacs = [asm_simp_tac (beta_ss addsimps rules) 1];
         in prove thy defs goal (K tacs) end;
--- a/src/HOLCF/Tools/holcf_library.ML	Sun Oct 24 15:11:24 2010 -0700
+++ b/src/HOLCF/Tools/holcf_library.ML	Sun Oct 24 15:19:17 2010 -0700
@@ -174,8 +174,8 @@
 
 val oneT = @{typ "one"};
 
-fun one_when_const T = Const (@{const_name one_when}, T ->> oneT ->> T);
-fun mk_one_when t = one_when_const (fastype_of t) ` t;
+fun one_case_const T = Const (@{const_name one_case}, T ->> oneT ->> T);
+fun mk_one_case t = one_case_const (fastype_of t) ` t;
 
 
 (*** Strict product type ***)