src/HOLCF/Sprod.thy
author huffman
Sat, 27 Nov 2010 12:26:18 -0800
changeset 41015 a3e505b236e7
parent 40737 8e92772bc0e8
permissions -rw-r--r--
rename function 'strict' to 'seq', which is its name in Haskell
     1 (*  Title:      HOLCF/Sprod.thy
     2     Author:     Franz Regensburger
     3     Author:     Brian Huffman
     4 *)
     5 
     6 header {* The type of strict products *}
     7 
     8 theory Sprod
     9 imports Cfun
    10 begin
    11 
    12 default_sort pcpo
    13 
    14 subsection {* Definition of strict product type *}
    15 
    16 pcpodef ('a, 'b) sprod (infixr "**" 20) =
    17         "{p::'a \<times> 'b. p = \<bottom> \<or> (fst p \<noteq> \<bottom> \<and> snd p \<noteq> \<bottom>)}"
    18 by simp_all
    19 
    20 instance sprod :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
    21 by (rule typedef_chfin [OF type_definition_sprod below_sprod_def])
    22 
    23 type_notation (xsymbols)
    24   sprod  ("(_ \<otimes>/ _)" [21,20] 20)
    25 type_notation (HTML output)
    26   sprod  ("(_ \<otimes>/ _)" [21,20] 20)
    27 
    28 subsection {* Definitions of constants *}
    29 
    30 definition
    31   sfst :: "('a ** 'b) \<rightarrow> 'a" where
    32   "sfst = (\<Lambda> p. fst (Rep_sprod p))"
    33 
    34 definition
    35   ssnd :: "('a ** 'b) \<rightarrow> 'b" where
    36   "ssnd = (\<Lambda> p. snd (Rep_sprod p))"
    37 
    38 definition
    39   spair :: "'a \<rightarrow> 'b \<rightarrow> ('a ** 'b)" where
    40   "spair = (\<Lambda> a b. Abs_sprod (seq\<cdot>b\<cdot>a, seq\<cdot>a\<cdot>b))"
    41 
    42 definition
    43   ssplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a ** 'b) \<rightarrow> 'c" where
    44   "ssplit = (\<Lambda> f p. seq\<cdot>p\<cdot>(f\<cdot>(sfst\<cdot>p)\<cdot>(ssnd\<cdot>p)))"
    45 
    46 syntax
    47   "_stuple" :: "['a, args] => 'a ** 'b"  ("(1'(:_,/ _:'))")
    48 translations
    49   "(:x, y, z:)" == "(:x, (:y, z:):)"
    50   "(:x, y:)"    == "CONST spair\<cdot>x\<cdot>y"
    51 
    52 translations
    53   "\<Lambda>(CONST spair\<cdot>x\<cdot>y). t" == "CONST ssplit\<cdot>(\<Lambda> x y. t)"
    54 
    55 subsection {* Case analysis *}
    56 
    57 lemma spair_sprod: "(seq\<cdot>b\<cdot>a, seq\<cdot>a\<cdot>b) \<in> sprod"
    58 by (simp add: sprod_def seq_conv_if)
    59 
    60 lemma Rep_sprod_spair: "Rep_sprod (:a, b:) = (seq\<cdot>b\<cdot>a, seq\<cdot>a\<cdot>b)"
    61 by (simp add: spair_def cont_Abs_sprod Abs_sprod_inverse spair_sprod)
    62 
    63 lemmas Rep_sprod_simps =
    64   Rep_sprod_inject [symmetric] below_sprod_def
    65   Pair_fst_snd_eq below_prod_def
    66   Rep_sprod_strict Rep_sprod_spair
    67 
    68 lemma sprodE [case_names bottom spair, cases type: sprod]:
    69   obtains "p = \<bottom>" | x y where "p = (:x, y:)" and "x \<noteq> \<bottom>" and "y \<noteq> \<bottom>"
    70 using Rep_sprod [of p] by (auto simp add: sprod_def Rep_sprod_simps)
    71 
    72 lemma sprod_induct [case_names bottom spair, induct type: sprod]:
    73   "\<lbrakk>P \<bottom>; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> P (:x, y:)\<rbrakk> \<Longrightarrow> P x"
    74 by (cases x, simp_all)
    75 
    76 subsection {* Properties of \emph{spair} *}
    77 
    78 lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>"
    79 by (simp add: Rep_sprod_simps)
    80 
    81 lemma spair_strict2 [simp]: "(:x, \<bottom>:) = \<bottom>"
    82 by (simp add: Rep_sprod_simps)
    83 
    84 lemma spair_bottom_iff [simp]: "((:x, y:) = \<bottom>) = (x = \<bottom> \<or> y = \<bottom>)"
    85 by (simp add: Rep_sprod_simps seq_conv_if)
    86 
    87 lemma spair_below_iff:
    88   "((:a, b:) \<sqsubseteq> (:c, d:)) = (a = \<bottom> \<or> b = \<bottom> \<or> (a \<sqsubseteq> c \<and> b \<sqsubseteq> d))"
    89 by (simp add: Rep_sprod_simps seq_conv_if)
    90 
    91 lemma spair_eq_iff:
    92   "((:a, b:) = (:c, d:)) =
    93     (a = c \<and> b = d \<or> (a = \<bottom> \<or> b = \<bottom>) \<and> (c = \<bottom> \<or> d = \<bottom>))"
    94 by (simp add: Rep_sprod_simps seq_conv_if)
    95 
    96 lemma spair_strict: "x = \<bottom> \<or> y = \<bottom> \<Longrightarrow> (:x, y:) = \<bottom>"
    97 by simp
    98 
    99 lemma spair_strict_rev: "(:x, y:) \<noteq> \<bottom> \<Longrightarrow> x \<noteq> \<bottom> \<and> y \<noteq> \<bottom>"
   100 by simp
   101 
   102 lemma spair_defined: "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<noteq> \<bottom>"
   103 by simp
   104 
   105 lemma spair_defined_rev: "(:x, y:) = \<bottom> \<Longrightarrow> x = \<bottom> \<or> y = \<bottom>"
   106 by simp
   107 
   108 lemma spair_below:
   109   "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<sqsubseteq> (:a, b:) = (x \<sqsubseteq> a \<and> y \<sqsubseteq> b)"
   110 by (simp add: spair_below_iff)
   111 
   112 lemma spair_eq:
   113   "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ((:x, y:) = (:a, b:)) = (x = a \<and> y = b)"
   114 by (simp add: spair_eq_iff)
   115 
   116 lemma spair_inject:
   117   "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; (:x, y:) = (:a, b:)\<rbrakk> \<Longrightarrow> x = a \<and> y = b"
   118 by (rule spair_eq [THEN iffD1])
   119 
   120 lemma inst_sprod_pcpo2: "UU = (:UU,UU:)"
   121 by simp
   122 
   123 lemma sprodE2: "(\<And>x y. p = (:x, y:) \<Longrightarrow> Q) \<Longrightarrow> Q"
   124 by (cases p, simp only: inst_sprod_pcpo2, simp)
   125 
   126 subsection {* Properties of \emph{sfst} and \emph{ssnd} *}
   127 
   128 lemma sfst_strict [simp]: "sfst\<cdot>\<bottom> = \<bottom>"
   129 by (simp add: sfst_def cont_Rep_sprod Rep_sprod_strict)
   130 
   131 lemma ssnd_strict [simp]: "ssnd\<cdot>\<bottom> = \<bottom>"
   132 by (simp add: ssnd_def cont_Rep_sprod Rep_sprod_strict)
   133 
   134 lemma sfst_spair [simp]: "y \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>(:x, y:) = x"
   135 by (simp add: sfst_def cont_Rep_sprod Rep_sprod_spair)
   136 
   137 lemma ssnd_spair [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>(:x, y:) = y"
   138 by (simp add: ssnd_def cont_Rep_sprod Rep_sprod_spair)
   139 
   140 lemma sfst_bottom_iff [simp]: "(sfst\<cdot>p = \<bottom>) = (p = \<bottom>)"
   141 by (cases p, simp_all)
   142 
   143 lemma ssnd_bottom_iff [simp]: "(ssnd\<cdot>p = \<bottom>) = (p = \<bottom>)"
   144 by (cases p, simp_all)
   145 
   146 lemma sfst_defined: "p \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>p \<noteq> \<bottom>"
   147 by simp
   148 
   149 lemma ssnd_defined: "p \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>p \<noteq> \<bottom>"
   150 by simp
   151 
   152 lemma spair_sfst_ssnd: "(:sfst\<cdot>p, ssnd\<cdot>p:) = p"
   153 by (cases p, simp_all)
   154 
   155 lemma below_sprod: "(x \<sqsubseteq> y) = (sfst\<cdot>x \<sqsubseteq> sfst\<cdot>y \<and> ssnd\<cdot>x \<sqsubseteq> ssnd\<cdot>y)"
   156 by (simp add: Rep_sprod_simps sfst_def ssnd_def cont_Rep_sprod)
   157 
   158 lemma eq_sprod: "(x = y) = (sfst\<cdot>x = sfst\<cdot>y \<and> ssnd\<cdot>x = ssnd\<cdot>y)"
   159 by (auto simp add: po_eq_conv below_sprod)
   160 
   161 lemma sfst_below_iff: "sfst\<cdot>x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (:y, ssnd\<cdot>x:)"
   162 apply (cases "x = \<bottom>", simp, cases "y = \<bottom>", simp)
   163 apply (simp add: below_sprod)
   164 done
   165 
   166 lemma ssnd_below_iff: "ssnd\<cdot>x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (:sfst\<cdot>x, y:)"
   167 apply (cases "x = \<bottom>", simp, cases "y = \<bottom>", simp)
   168 apply (simp add: below_sprod)
   169 done
   170 
   171 subsection {* Compactness *}
   172 
   173 lemma compact_sfst: "compact x \<Longrightarrow> compact (sfst\<cdot>x)"
   174 by (rule compactI, simp add: sfst_below_iff)
   175 
   176 lemma compact_ssnd: "compact x \<Longrightarrow> compact (ssnd\<cdot>x)"
   177 by (rule compactI, simp add: ssnd_below_iff)
   178 
   179 lemma compact_spair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (:x, y:)"
   180 by (rule compact_sprod, simp add: Rep_sprod_spair seq_conv_if)
   181 
   182 lemma compact_spair_iff:
   183   "compact (:x, y:) = (x = \<bottom> \<or> y = \<bottom> \<or> (compact x \<and> compact y))"
   184 apply (safe elim!: compact_spair)
   185 apply (drule compact_sfst, simp)
   186 apply (drule compact_ssnd, simp)
   187 apply simp
   188 apply simp
   189 done
   190 
   191 subsection {* Properties of \emph{ssplit} *}
   192 
   193 lemma ssplit1 [simp]: "ssplit\<cdot>f\<cdot>\<bottom> = \<bottom>"
   194 by (simp add: ssplit_def)
   195 
   196 lemma ssplit2 [simp]: "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ssplit\<cdot>f\<cdot>(:x, y:) = f\<cdot>x\<cdot>y"
   197 by (simp add: ssplit_def)
   198 
   199 lemma ssplit3 [simp]: "ssplit\<cdot>spair\<cdot>z = z"
   200 by (cases z, simp_all)
   201 
   202 subsection {* Strict product preserves flatness *}
   203 
   204 instance sprod :: (flat, flat) flat
   205 proof
   206   fix x y :: "'a \<otimes> 'b"
   207   assume "x \<sqsubseteq> y" thus "x = \<bottom> \<or> x = y"
   208     apply (induct x, simp)
   209     apply (induct y, simp)
   210     apply (simp add: spair_below_iff flat_below_iff)
   211     done
   212 qed
   213 
   214 end