src/HOLCF/Pcpodef.thy
changeset 31076 99fe356cbbc2
parent 29138 661a8db7e647
child 31738 7b9b9ba532ca
--- a/src/HOLCF/Pcpodef.thy	Fri May 08 19:28:11 2009 +0100
+++ b/src/HOLCF/Pcpodef.thy	Fri May 08 16:19:51 2009 -0700
@@ -16,22 +16,22 @@
   if the ordering is defined in the standard way.
 *}
 
-setup {* Sign.add_const_constraint (@{const_name Porder.sq_le}, NONE) *}
+setup {* Sign.add_const_constraint (@{const_name Porder.below}, NONE) *}
 
 theorem typedef_po:
   fixes Abs :: "'a::po \<Rightarrow> 'b::type"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
   shows "OFCLASS('b, po_class)"
- apply (intro_classes, unfold less)
-   apply (rule refl_less)
-  apply (erule (1) trans_less)
+ apply (intro_classes, unfold below)
+   apply (rule below_refl)
+  apply (erule (1) below_trans)
  apply (rule type_definition.Rep_inject [OF type, THEN iffD1])
- apply (erule (1) antisym_less)
+ apply (erule (1) below_antisym)
 done
 
-setup {* Sign.add_const_constraint (@{const_name Porder.sq_le},
-  SOME @{typ "'a::sq_ord \<Rightarrow> 'a::sq_ord \<Rightarrow> bool"}) *}
+setup {* Sign.add_const_constraint (@{const_name Porder.below},
+  SOME @{typ "'a::below \<Rightarrow> 'a::below \<Rightarrow> bool"}) *}
 
 subsection {* Proving a subtype is finite *}
 
@@ -58,9 +58,9 @@
 subsection {* Proving a subtype is chain-finite *}
 
 lemma monofun_Rep:
-  assumes less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+  assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
   shows "monofun Rep"
-by (rule monofunI, unfold less)
+by (rule monofunI, unfold below)
 
 lemmas ch2ch_Rep = ch2ch_monofun [OF monofun_Rep]
 lemmas ub2ub_Rep = ub2ub_monofun [OF monofun_Rep]
@@ -68,10 +68,10 @@
 theorem typedef_chfin:
   fixes Abs :: "'a::chfin \<Rightarrow> 'b::po"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
   shows "OFCLASS('b, chfin_class)"
  apply intro_classes
- apply (drule ch2ch_Rep [OF less])
+ apply (drule ch2ch_Rep [OF below])
  apply (drule chfin)
  apply (unfold max_in_chain_def)
  apply (simp add: type_definition.Rep_inject [OF type])
@@ -90,28 +90,28 @@
 lemma Abs_inverse_lub_Rep:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and adm:  "adm (\<lambda>x. x \<in> A)"
   shows "chain S \<Longrightarrow> Rep (Abs (\<Squnion>i. Rep (S i))) = (\<Squnion>i. Rep (S i))"
  apply (rule type_definition.Abs_inverse [OF type])
- apply (erule admD [OF adm ch2ch_Rep [OF less]])
+ apply (erule admD [OF adm ch2ch_Rep [OF below]])
  apply (rule type_definition.Rep [OF type])
 done
 
 theorem typedef_lub:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and adm: "adm (\<lambda>x. x \<in> A)"
   shows "chain S \<Longrightarrow> range S <<| Abs (\<Squnion>i. Rep (S i))"
- apply (frule ch2ch_Rep [OF less])
+ apply (frule ch2ch_Rep [OF below])
  apply (rule is_lubI)
   apply (rule ub_rangeI)
-  apply (simp only: less Abs_inverse_lub_Rep [OF type less adm])
+  apply (simp only: below Abs_inverse_lub_Rep [OF type below adm])
   apply (erule is_ub_thelub)
- apply (simp only: less Abs_inverse_lub_Rep [OF type less adm])
+ apply (simp only: below Abs_inverse_lub_Rep [OF type below adm])
  apply (erule is_lub_thelub)
- apply (erule ub2ub_Rep [OF less])
+ apply (erule ub2ub_Rep [OF below])
 done
 
 lemmas typedef_thelub = typedef_lub [THEN thelubI, standard]
@@ -119,13 +119,13 @@
 theorem typedef_cpo:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and adm: "adm (\<lambda>x. x \<in> A)"
   shows "OFCLASS('b, cpo_class)"
 proof
   fix S::"nat \<Rightarrow> 'b" assume "chain S"
   hence "range S <<| Abs (\<Squnion>i. Rep (S i))"
-    by (rule typedef_lub [OF type less adm])
+    by (rule typedef_lub [OF type below adm])
   thus "\<exists>x. range S <<| x" ..
 qed
 
@@ -136,14 +136,14 @@
 theorem typedef_cont_Rep:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and adm: "adm (\<lambda>x. x \<in> A)"
   shows "cont Rep"
  apply (rule contI)
- apply (simp only: typedef_thelub [OF type less adm])
- apply (simp only: Abs_inverse_lub_Rep [OF type less adm])
+ apply (simp only: typedef_thelub [OF type below adm])
+ apply (simp only: Abs_inverse_lub_Rep [OF type below adm])
  apply (rule cpo_lubI)
- apply (erule ch2ch_Rep [OF less])
+ apply (erule ch2ch_Rep [OF below])
 done
 
 text {*
@@ -153,28 +153,28 @@
 *}
 
 theorem typedef_is_lubI:
-  assumes less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+  assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
   shows "range (\<lambda>i. Rep (S i)) <<| Rep x \<Longrightarrow> range S <<| x"
  apply (rule is_lubI)
   apply (rule ub_rangeI)
-  apply (subst less)
+  apply (subst below)
   apply (erule is_ub_lub)
- apply (subst less)
+ apply (subst below)
  apply (erule is_lub_lub)
- apply (erule ub2ub_Rep [OF less])
+ apply (erule ub2ub_Rep [OF below])
 done
 
 theorem typedef_cont_Abs:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
   fixes f :: "'c::cpo \<Rightarrow> 'a::cpo"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and adm: "adm (\<lambda>x. x \<in> A)" (* not used *)
     and f_in_A: "\<And>x. f x \<in> A"
     and cont_f: "cont f"
   shows "cont (\<lambda>x. Abs (f x))"
  apply (rule contI)
- apply (rule typedef_is_lubI [OF less])
+ apply (rule typedef_is_lubI [OF below])
  apply (simp only: type_definition.Abs_inverse [OF type f_in_A])
  apply (erule cont_f [THEN contE])
 done
@@ -184,15 +184,15 @@
 theorem typedef_compact:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and adm: "adm (\<lambda>x. x \<in> A)"
   shows "compact (Rep k) \<Longrightarrow> compact k"
 proof (unfold compact_def)
   have cont_Rep: "cont Rep"
-    by (rule typedef_cont_Rep [OF type less adm])
+    by (rule typedef_cont_Rep [OF type below adm])
   assume "adm (\<lambda>x. \<not> Rep k \<sqsubseteq> x)"
   with cont_Rep have "adm (\<lambda>x. \<not> Rep k \<sqsubseteq> Rep x)" by (rule adm_subst)
-  thus "adm (\<lambda>x. \<not> k \<sqsubseteq> x)" by (unfold less)
+  thus "adm (\<lambda>x. \<not> k \<sqsubseteq> x)" by (unfold below)
 qed
 
 subsection {* Proving a subtype is pointed *}
@@ -205,13 +205,13 @@
 theorem typedef_pcpo_generic:
   fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and z_in_A: "z \<in> A"
     and z_least: "\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x"
   shows "OFCLASS('b, pcpo_class)"
  apply (intro_classes)
  apply (rule_tac x="Abs z" in exI, rule allI)
- apply (unfold less)
+ apply (unfold below)
  apply (subst type_definition.Abs_inverse [OF type z_in_A])
  apply (rule z_least [OF type_definition.Rep [OF type]])
 done
@@ -224,10 +224,10 @@
 theorem typedef_pcpo:
   fixes Abs :: "'a::pcpo \<Rightarrow> 'b::cpo"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "OFCLASS('b, pcpo_class)"
-by (rule typedef_pcpo_generic [OF type less UU_in_A], rule minimal)
+by (rule typedef_pcpo_generic [OF type below UU_in_A], rule minimal)
 
 subsubsection {* Strictness of @{term Rep} and @{term Abs} *}
 
@@ -238,66 +238,66 @@
 
 theorem typedef_Abs_strict:
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "Abs \<bottom> = \<bottom>"
- apply (rule UU_I, unfold less)
+ apply (rule UU_I, unfold below)
  apply (simp add: type_definition.Abs_inverse [OF type UU_in_A])
 done
 
 theorem typedef_Rep_strict:
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "Rep \<bottom> = \<bottom>"
- apply (rule typedef_Abs_strict [OF type less UU_in_A, THEN subst])
+ apply (rule typedef_Abs_strict [OF type below UU_in_A, THEN subst])
  apply (rule type_definition.Abs_inverse [OF type UU_in_A])
 done
 
 theorem typedef_Abs_strict_iff:
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "x \<in> A \<Longrightarrow> (Abs x = \<bottom>) = (x = \<bottom>)"
- apply (rule typedef_Abs_strict [OF type less UU_in_A, THEN subst])
+ apply (rule typedef_Abs_strict [OF type below UU_in_A, THEN subst])
  apply (simp add: type_definition.Abs_inject [OF type] UU_in_A)
 done
 
 theorem typedef_Rep_strict_iff:
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "(Rep x = \<bottom>) = (x = \<bottom>)"
- apply (rule typedef_Rep_strict [OF type less UU_in_A, THEN subst])
+ apply (rule typedef_Rep_strict [OF type below UU_in_A, THEN subst])
  apply (simp add: type_definition.Rep_inject [OF type])
 done
 
 theorem typedef_Abs_defined:
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "\<lbrakk>x \<noteq> \<bottom>; x \<in> A\<rbrakk> \<Longrightarrow> Abs x \<noteq> \<bottom>"
-by (simp add: typedef_Abs_strict_iff [OF type less UU_in_A])
+by (simp add: typedef_Abs_strict_iff [OF type below UU_in_A])
 
 theorem typedef_Rep_defined:
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "x \<noteq> \<bottom> \<Longrightarrow> Rep x \<noteq> \<bottom>"
-by (simp add: typedef_Rep_strict_iff [OF type less UU_in_A])
+by (simp add: typedef_Rep_strict_iff [OF type below UU_in_A])
 
 subsection {* Proving a subtype is flat *}
 
 theorem typedef_flat:
   fixes Abs :: "'a::flat \<Rightarrow> 'b::pcpo"
   assumes type: "type_definition Rep Abs A"
-    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
+    and below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
     and UU_in_A: "\<bottom> \<in> A"
   shows "OFCLASS('b, flat_class)"
  apply (intro_classes)
- apply (unfold less)
+ apply (unfold below)
  apply (simp add: type_definition.Rep_inject [OF type, symmetric])
- apply (simp add: typedef_Rep_strict [OF type less UU_in_A])
+ apply (simp add: typedef_Rep_strict [OF type below UU_in_A])
  apply (simp add: ax_flat)
 done