simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
1 (* Title: HOLCF/Cprod.thy
2 Author: Franz Regensburger
5 header {* The cpo of cartesian products *}
13 subsection {* Type @{typ unit} is a pcpo *}
16 unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where
17 "unit_when = (\<Lambda> a _. a)"
20 "\<Lambda>(). t" == "CONST unit_when\<cdot>t"
22 lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a"
23 by (simp add: unit_when_def)
25 subsection {* Continuous versions of constants *}
28 cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)" -- {* continuous pairing *} where
29 "cpair = (\<Lambda> x y. (x, y))"
32 cfst :: "('a * 'b) \<rightarrow> 'a" where
33 "cfst = (\<Lambda> p. fst p)"
36 csnd :: "('a * 'b) \<rightarrow> 'b" where
37 "csnd = (\<Lambda> p. snd p)"
40 csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where
41 "csplit = (\<Lambda> f p. f\<cdot>(cfst\<cdot>p)\<cdot>(csnd\<cdot>p))"
44 "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1<_,/ _>)")
47 "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(1\<langle>_,/ _\<rangle>)")
50 "\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>"
51 "\<langle>x, y\<rangle>" == "CONST cpair\<cdot>x\<cdot>y"
54 "\<Lambda>(CONST cpair\<cdot>x\<cdot>y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)"
57 subsection {* Convert all lemmas to the continuous versions *}
59 lemma cpair_eq_pair: "<x, y> = (x, y)"
60 by (simp add: cpair_def cont_pair1 cont_pair2)
62 lemma pair_eq_cpair: "(x, y) = <x, y>"
63 by (simp add: cpair_def cont_pair1 cont_pair2)
65 lemma inject_cpair: "<a,b> = <aa,ba> \<Longrightarrow> a = aa \<and> b = ba"
66 by (simp add: cpair_eq_pair)
68 lemma cpair_eq [iff]: "(<a, b> = <a', b'>) = (a = a' \<and> b = b')"
69 by (simp add: cpair_eq_pair)
71 lemma cpair_below [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')"
72 by (simp add: cpair_eq_pair)
74 lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)"
75 by (simp add: cpair_eq_pair)
77 lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>"
80 lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>"
81 by (rule cpair_strict [symmetric])
83 lemma defined_cpair_rev:
84 "<a,b> = \<bottom> \<Longrightarrow> a = \<bottom> \<and> b = \<bottom>"
87 lemma Exh_Cprod2: "\<exists>a b. z = <a, b>"
88 by (simp add: cpair_eq_pair)
90 lemma cprodE: "\<lbrakk>\<And>x y. p = <x, y> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
91 by (cut_tac Exh_Cprod2, auto)
93 lemma cfst_cpair [simp]: "cfst\<cdot><x, y> = x"
94 by (simp add: cpair_eq_pair cfst_def cont_fst)
96 lemma csnd_cpair [simp]: "csnd\<cdot><x, y> = y"
97 by (simp add: cpair_eq_pair csnd_def cont_snd)
99 lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>"
100 by (simp add: cfst_def)
102 lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>"
103 by (simp add: csnd_def)
105 lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p"
106 by (cases p rule: cprodE, simp)
108 lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd
110 lemma below_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)"
111 by (simp add: below_prod_def cfst_def csnd_def cont_fst cont_snd)
113 lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)"
114 by (auto simp add: po_eq_conv below_cprod)
116 lemma cfst_below_iff: "cfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd\<cdot>x>"
117 by (simp add: below_cprod)
119 lemma csnd_below_iff: "csnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <cfst\<cdot>x, y>"
120 by (simp add: below_cprod)
122 lemma compact_cfst: "compact x \<Longrightarrow> compact (cfst\<cdot>x)"
123 by (rule compactI, simp add: cfst_below_iff)
125 lemma compact_csnd: "compact x \<Longrightarrow> compact (csnd\<cdot>x)"
126 by (rule compactI, simp add: csnd_below_iff)
128 lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>"
129 by (simp add: cpair_eq_pair)
131 lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)"
132 by (simp add: cpair_eq_pair)
135 "chain S \<Longrightarrow> range S <<| <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
136 apply (simp add: cpair_eq_pair cfst_def csnd_def cont_fst cont_snd)
137 apply (erule lub_cprod)
141 "chain S \<Longrightarrow> (\<Squnion>i. S i) = <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>"
142 by (rule lub_cprod2 [THEN thelubI])
144 lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>"
145 by (simp add: csplit_def)
147 lemma csplit2 [simp]: "csplit\<cdot>f\<cdot><x,y> = f\<cdot>x\<cdot>y"
148 by (simp add: csplit_def)
150 lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z"
151 by (simp add: csplit_def cpair_cfst_csnd)
153 lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2
155 subsection {* Product type is a bifinite domain *}
157 instantiation "*" :: (profinite, profinite) profinite
162 "approx = (\<lambda>n. \<Lambda>\<langle>x, y\<rangle>. \<langle>approx n\<cdot>x, approx n\<cdot>y\<rangle>)"
165 fix i :: nat and x :: "'a \<times> 'b"
166 show "chain (approx :: nat \<Rightarrow> 'a \<times> 'b \<rightarrow> 'a \<times> 'b)"
167 unfolding approx_cprod_def by simp
168 show "(\<Squnion>i. approx i\<cdot>x) = x"
169 unfolding approx_cprod_def
170 by (simp add: lub_distribs eta_cfun)
171 show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
172 unfolding approx_cprod_def csplit_def by simp
173 have "{x::'a \<times> 'b. approx i\<cdot>x = x} \<subseteq>
174 {x::'a. approx i\<cdot>x = x} \<times> {x::'b. approx i\<cdot>x = x}"
175 unfolding approx_cprod_def
176 by (clarsimp simp add: pair_eq_cpair)
177 thus "finite {x::'a \<times> 'b. approx i\<cdot>x = x}"
178 by (rule finite_subset,
179 intro finite_cartesian_product finite_fixes_approx)
184 instance "*" :: (bifinite, bifinite) bifinite ..
186 lemma approx_cpair [simp]:
187 "approx i\<cdot>\<langle>x, y\<rangle> = \<langle>approx i\<cdot>x, approx i\<cdot>y\<rangle>"
188 unfolding approx_cprod_def by simp
190 lemma cfst_approx: "cfst\<cdot>(approx i\<cdot>p) = approx i\<cdot>(cfst\<cdot>p)"
191 by (cases p rule: cprodE, simp)
193 lemma csnd_approx: "csnd\<cdot>(approx i\<cdot>p) = approx i\<cdot>(csnd\<cdot>p)"
194 by (cases p rule: cprodE, simp)