| author | wenzelm | 
| Sat, 28 Aug 2010 20:24:41 +0200 | |
| changeset 38843 | d95522496593 | 
| parent 36452 | d37c6eed8117 | 
| child 39973 | c62b4ff97bfc | 
| permissions | -rw-r--r-- | 
| 15600 | 1 | (* Title: HOLCF/Sprod.thy | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 2 | Author: Franz Regensburger and Brian Huffman | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 3 | *) | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 4 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 5 | header {* The type of strict products *}
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 6 | |
| 15577 | 7 | theory Sprod | 
| 31114 | 8 | imports Bifinite | 
| 15577 | 9 | begin | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 10 | |
| 36452 | 11 | default_sort pcpo | 
| 16082 | 12 | |
| 15591 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 13 | subsection {* Definition of strict product type *}
 | 
| 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 14 | |
| 35525 | 15 | pcpodef (Sprod)  ('a, 'b) sprod (infixr "**" 20) =
 | 
| 31114 | 16 |         "{p::'a \<times> 'b. p = \<bottom> \<or> (fst p \<noteq> \<bottom> \<and> snd p \<noteq> \<bottom>)}"
 | 
| 29063 
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
 wenzelm parents: 
27310diff
changeset | 17 | by simp_all | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 18 | |
| 35525 | 19 | instance sprod :: ("{finite_po,pcpo}", "{finite_po,pcpo}") finite_po
 | 
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25757diff
changeset | 20 | by (rule typedef_finite_po [OF type_definition_Sprod]) | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25757diff
changeset | 21 | |
| 35525 | 22 | instance sprod :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin
 | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 23 | by (rule typedef_chfin [OF type_definition_Sprod below_Sprod_def]) | 
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25757diff
changeset | 24 | |
| 35427 | 25 | type_notation (xsymbols) | 
| 35547 | 26 |   sprod  ("(_ \<otimes>/ _)" [21,20] 20)
 | 
| 35427 | 27 | type_notation (HTML output) | 
| 35547 | 28 |   sprod  ("(_ \<otimes>/ _)" [21,20] 20)
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 29 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 30 | lemma spair_lemma: | 
| 31114 | 31 | "(strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a) \<in> Sprod" | 
| 25914 | 32 | by (simp add: Sprod_def strictify_conv_if) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 33 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 34 | subsection {* Definitions of constants *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 35 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 36 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 37 |   sfst :: "('a ** 'b) \<rightarrow> 'a" where
 | 
| 31114 | 38 | "sfst = (\<Lambda> p. fst (Rep_Sprod p))" | 
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 39 | |
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 40 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 41 |   ssnd :: "('a ** 'b) \<rightarrow> 'b" where
 | 
| 31114 | 42 | "ssnd = (\<Lambda> p. snd (Rep_Sprod p))" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 43 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 44 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 45 |   spair :: "'a \<rightarrow> 'b \<rightarrow> ('a ** 'b)" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 46 | "spair = (\<Lambda> a b. Abs_Sprod | 
| 31114 | 47 | (strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a))" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 48 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 49 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 50 |   ssplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a ** 'b) \<rightarrow> 'c" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 51 | "ssplit = (\<Lambda> f. strictify\<cdot>(\<Lambda> p. f\<cdot>(sfst\<cdot>p)\<cdot>(ssnd\<cdot>p)))" | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 52 | |
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
25131diff
changeset | 53 | syntax | 
| 35115 | 54 |   "_stuple" :: "['a, args] => 'a ** 'b"  ("(1'(:_,/ _:'))")
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 55 | translations | 
| 18078 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 56 | "(:x, y, z:)" == "(:x, (:y, z:):)" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18078diff
changeset | 57 | "(:x, y:)" == "CONST spair\<cdot>x\<cdot>y" | 
| 18078 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 58 | |
| 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 59 | translations | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18078diff
changeset | 60 | "\<Lambda>(CONST spair\<cdot>x\<cdot>y). t" == "CONST ssplit\<cdot>(\<Lambda> x y. t)" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 61 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 62 | subsection {* Case analysis *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 63 | |
| 25914 | 64 | lemma Rep_Sprod_spair: | 
| 31114 | 65 | "Rep_Sprod (:a, b:) = (strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a)" | 
| 25914 | 66 | unfolding spair_def | 
| 67 | by (simp add: cont_Abs_Sprod Abs_Sprod_inverse spair_lemma) | |
| 68 | ||
| 69 | lemmas Rep_Sprod_simps = | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 70 | Rep_Sprod_inject [symmetric] below_Sprod_def | 
| 25914 | 71 | Rep_Sprod_strict Rep_Sprod_spair | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 72 | |
| 27310 | 73 | lemma Exh_Sprod: | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 74 | "z = \<bottom> \<or> (\<exists>a b. z = (:a, b:) \<and> a \<noteq> \<bottom> \<and> b \<noteq> \<bottom>)" | 
| 25914 | 75 | apply (insert Rep_Sprod [of z]) | 
| 31114 | 76 | apply (simp add: Rep_Sprod_simps Pair_fst_snd_eq) | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 77 | apply (simp add: Sprod_def) | 
| 25914 | 78 | apply (erule disjE, simp) | 
| 79 | apply (simp add: strictify_conv_if) | |
| 80 | apply fast | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 81 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 82 | |
| 35783 | 83 | lemma sprodE [case_names bottom spair, cases type: sprod]: | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 84 | "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x y. \<lbrakk>p = (:x, y:); x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 35783 | 85 | using Exh_Sprod [of p] by auto | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 86 | |
| 35783 | 87 | lemma sprod_induct [case_names bottom spair, induct type: sprod]: | 
| 25757 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 88 | "\<lbrakk>P \<bottom>; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> P (:x, y:)\<rbrakk> \<Longrightarrow> P x" | 
| 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 89 | by (cases x, simp_all) | 
| 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 90 | |
| 35900 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 huffman parents: 
35783diff
changeset | 91 | subsection {* Properties of \emph{spair} *}
 | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 92 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 93 | lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>" | 
| 25914 | 94 | by (simp add: Rep_Sprod_simps strictify_conv_if) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 95 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 96 | lemma spair_strict2 [simp]: "(:x, \<bottom>:) = \<bottom>" | 
| 25914 | 97 | by (simp add: Rep_Sprod_simps strictify_conv_if) | 
| 98 | ||
| 99 | lemma spair_strict_iff [simp]: "((:x, y:) = \<bottom>) = (x = \<bottom> \<or> y = \<bottom>)" | |
| 100 | by (simp add: Rep_Sprod_simps strictify_conv_if) | |
| 101 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 102 | lemma spair_below_iff: | 
| 25914 | 103 | "((:a, b:) \<sqsubseteq> (:c, d:)) = (a = \<bottom> \<or> b = \<bottom> \<or> (a \<sqsubseteq> c \<and> b \<sqsubseteq> d))" | 
| 104 | by (simp add: Rep_Sprod_simps strictify_conv_if) | |
| 105 | ||
| 106 | lemma spair_eq_iff: | |
| 107 | "((:a, b:) = (:c, d:)) = | |
| 108 | (a = c \<and> b = d \<or> (a = \<bottom> \<or> b = \<bottom>) \<and> (c = \<bottom> \<or> d = \<bottom>))" | |
| 109 | by (simp add: Rep_Sprod_simps strictify_conv_if) | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 110 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 111 | lemma spair_strict: "x = \<bottom> \<or> y = \<bottom> \<Longrightarrow> (:x, y:) = \<bottom>" | 
| 25914 | 112 | by simp | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 113 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 114 | lemma spair_strict_rev: "(:x, y:) \<noteq> \<bottom> \<Longrightarrow> x \<noteq> \<bottom> \<and> y \<noteq> \<bottom>" | 
| 25914 | 115 | by simp | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 116 | |
| 25914 | 117 | lemma spair_defined: "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<noteq> \<bottom>" | 
| 118 | by simp | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 119 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 120 | lemma spair_defined_rev: "(:x, y:) = \<bottom> \<Longrightarrow> x = \<bottom> \<or> y = \<bottom>" | 
| 25914 | 121 | by simp | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 122 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 123 | lemma spair_eq: | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 124 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ((:x, y:) = (:a, b:)) = (x = a \<and> y = b)" | 
| 25914 | 125 | by (simp add: spair_eq_iff) | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 126 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 127 | lemma spair_inject: | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 128 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; (:x, y:) = (:a, b:)\<rbrakk> \<Longrightarrow> x = a \<and> y = b" | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 129 | by (rule spair_eq [THEN iffD1]) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 130 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 131 | lemma inst_sprod_pcpo2: "UU = (:UU,UU:)" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 132 | by simp | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 133 | |
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 134 | lemma sprodE2: "(\<And>x y. p = (:x, y:) \<Longrightarrow> Q) \<Longrightarrow> Q" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 135 | by (cases p, simp only: inst_sprod_pcpo2, simp) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 136 | |
| 35900 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 huffman parents: 
35783diff
changeset | 137 | subsection {* Properties of \emph{sfst} and \emph{ssnd} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 138 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 139 | lemma sfst_strict [simp]: "sfst\<cdot>\<bottom> = \<bottom>" | 
| 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 140 | by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_strict) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 141 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 142 | lemma ssnd_strict [simp]: "ssnd\<cdot>\<bottom> = \<bottom>" | 
| 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 143 | by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_strict) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 144 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 145 | lemma sfst_spair [simp]: "y \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>(:x, y:) = x" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 146 | by (simp add: sfst_def cont_Rep_Sprod Rep_Sprod_spair) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 147 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 148 | lemma ssnd_spair [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>(:x, y:) = y" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 149 | by (simp add: ssnd_def cont_Rep_Sprod Rep_Sprod_spair) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 150 | |
| 16777 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 151 | lemma sfst_defined_iff [simp]: "(sfst\<cdot>p = \<bottom>) = (p = \<bottom>)" | 
| 25757 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 152 | by (cases p, simp_all) | 
| 16777 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 153 | |
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 154 | lemma ssnd_defined_iff [simp]: "(ssnd\<cdot>p = \<bottom>) = (p = \<bottom>)" | 
| 25757 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 155 | by (cases p, simp_all) | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 156 | |
| 16777 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 157 | lemma sfst_defined: "p \<noteq> \<bottom> \<Longrightarrow> sfst\<cdot>p \<noteq> \<bottom>" | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 158 | by simp | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 159 | |
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 160 | lemma ssnd_defined: "p \<noteq> \<bottom> \<Longrightarrow> ssnd\<cdot>p \<noteq> \<bottom>" | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 161 | by simp | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 162 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 163 | lemma surjective_pairing_Sprod2: "(:sfst\<cdot>p, ssnd\<cdot>p:) = p" | 
| 25757 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 164 | by (cases p, simp_all) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 165 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 166 | lemma below_sprod: "x \<sqsubseteq> y = (sfst\<cdot>x \<sqsubseteq> sfst\<cdot>y \<and> ssnd\<cdot>x \<sqsubseteq> ssnd\<cdot>y)" | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 167 | apply (simp add: below_Sprod_def sfst_def ssnd_def cont_Rep_Sprod) | 
| 31114 | 168 | apply (simp only: below_prod_def) | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 169 | done | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 170 | |
| 16751 | 171 | lemma eq_sprod: "(x = y) = (sfst\<cdot>x = sfst\<cdot>y \<and> ssnd\<cdot>x = ssnd\<cdot>y)" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 172 | by (auto simp add: po_eq_conv below_sprod) | 
| 16751 | 173 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 174 | lemma spair_below: | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 175 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<sqsubseteq> (:a, b:) = (x \<sqsubseteq> a \<and> y \<sqsubseteq> b)" | 
| 25757 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 176 | apply (cases "a = \<bottom>", simp) | 
| 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 177 | apply (cases "b = \<bottom>", simp) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 178 | apply (simp add: below_sprod) | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 179 | done | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 180 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 181 | lemma sfst_below_iff: "sfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> (:y, ssnd\<cdot>x:)" | 
| 25881 | 182 | apply (cases "x = \<bottom>", simp, cases "y = \<bottom>", simp) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 183 | apply (simp add: below_sprod) | 
| 25881 | 184 | done | 
| 185 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 186 | lemma ssnd_below_iff: "ssnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> (:sfst\<cdot>x, y:)" | 
| 25881 | 187 | apply (cases "x = \<bottom>", simp, cases "y = \<bottom>", simp) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 188 | apply (simp add: below_sprod) | 
| 25881 | 189 | done | 
| 190 | ||
| 191 | subsection {* Compactness *}
 | |
| 192 | ||
| 193 | lemma compact_sfst: "compact x \<Longrightarrow> compact (sfst\<cdot>x)" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 194 | by (rule compactI, simp add: sfst_below_iff) | 
| 25881 | 195 | |
| 196 | lemma compact_ssnd: "compact x \<Longrightarrow> compact (ssnd\<cdot>x)" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 197 | by (rule compactI, simp add: ssnd_below_iff) | 
| 25881 | 198 | |
| 199 | lemma compact_spair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (:x, y:)" | |
| 200 | by (rule compact_Sprod, simp add: Rep_Sprod_spair strictify_conv_if) | |
| 201 | ||
| 202 | lemma compact_spair_iff: | |
| 203 | "compact (:x, y:) = (x = \<bottom> \<or> y = \<bottom> \<or> (compact x \<and> compact y))" | |
| 204 | apply (safe elim!: compact_spair) | |
| 205 | apply (drule compact_sfst, simp) | |
| 206 | apply (drule compact_ssnd, simp) | |
| 207 | apply simp | |
| 208 | apply simp | |
| 209 | done | |
| 210 | ||
| 35900 
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
 huffman parents: 
35783diff
changeset | 211 | subsection {* Properties of \emph{ssplit} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 212 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 213 | lemma ssplit1 [simp]: "ssplit\<cdot>f\<cdot>\<bottom> = \<bottom>" | 
| 15591 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 214 | by (simp add: ssplit_def) | 
| 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 215 | |
| 16920 | 216 | lemma ssplit2 [simp]: "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ssplit\<cdot>f\<cdot>(:x, y:) = f\<cdot>x\<cdot>y" | 
| 15591 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 217 | by (simp add: ssplit_def) | 
| 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 218 | |
| 16553 | 219 | lemma ssplit3 [simp]: "ssplit\<cdot>spair\<cdot>z = z" | 
| 25757 
5957e3d72fec
declare sprodE as cases rule; new induction rule sprod_induct
 huffman parents: 
25135diff
changeset | 220 | by (cases z, simp_all) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 221 | |
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25757diff
changeset | 222 | subsection {* Strict product preserves flatness *}
 | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25757diff
changeset | 223 | |
| 35525 | 224 | instance sprod :: (flat, flat) flat | 
| 27310 | 225 | proof | 
| 226 | fix x y :: "'a \<otimes> 'b" | |
| 227 | assume "x \<sqsubseteq> y" thus "x = \<bottom> \<or> x = y" | |
| 228 | apply (induct x, simp) | |
| 229 | apply (induct y, simp) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29138diff
changeset | 230 | apply (simp add: spair_below_iff flat_below_iff) | 
| 27310 | 231 | done | 
| 232 | qed | |
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25757diff
changeset | 233 | |
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 234 | subsection {* Map function for strict products *}
 | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 235 | |
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 236 | definition | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 237 |   sprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<otimes> 'c \<rightarrow> 'b \<otimes> 'd"
 | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 238 | where | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 239 | "sprod_map = (\<Lambda> f g. ssplit\<cdot>(\<Lambda> x y. (:f\<cdot>x, g\<cdot>y:)))" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 240 | |
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 241 | lemma sprod_map_strict [simp]: "sprod_map\<cdot>a\<cdot>b\<cdot>\<bottom> = \<bottom>" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 242 | unfolding sprod_map_def by simp | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 243 | |
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 244 | lemma sprod_map_spair [simp]: | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 245 | "x \<noteq> \<bottom> \<Longrightarrow> y \<noteq> \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 246 | by (simp add: sprod_map_def) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 247 | |
| 35491 | 248 | lemma sprod_map_spair': | 
| 249 | "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)" | |
| 250 | by (cases "x = \<bottom> \<or> y = \<bottom>") auto | |
| 251 | ||
| 33808 | 252 | lemma sprod_map_ID: "sprod_map\<cdot>ID\<cdot>ID = ID" | 
| 253 | unfolding sprod_map_def by (simp add: expand_cfun_eq eta_cfun) | |
| 254 | ||
| 33587 | 255 | lemma sprod_map_map: | 
| 256 | "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow> | |
| 257 | sprod_map\<cdot>f1\<cdot>g1\<cdot>(sprod_map\<cdot>f2\<cdot>g2\<cdot>p) = | |
| 258 | sprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p" | |
| 259 | apply (induct p, simp) | |
| 260 | apply (case_tac "f2\<cdot>x = \<bottom>", simp) | |
| 261 | apply (case_tac "g2\<cdot>y = \<bottom>", simp) | |
| 262 | apply simp | |
| 263 | done | |
| 264 | ||
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 265 | lemma ep_pair_sprod_map: | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 266 | assumes "ep_pair e1 p1" and "ep_pair e2 p2" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 267 | shows "ep_pair (sprod_map\<cdot>e1\<cdot>e2) (sprod_map\<cdot>p1\<cdot>p2)" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 268 | proof | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 269 | interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 270 | interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 271 | fix x show "sprod_map\<cdot>p1\<cdot>p2\<cdot>(sprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 272 | by (induct x) simp_all | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 273 | fix y show "sprod_map\<cdot>e1\<cdot>e2\<cdot>(sprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 274 | apply (induct y, simp) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 275 | apply (case_tac "p1\<cdot>x = \<bottom>", simp, case_tac "p2\<cdot>y = \<bottom>", simp) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 276 | apply (simp add: monofun_cfun e1p1.e_p_below e2p2.e_p_below) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 277 | done | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 278 | qed | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 279 | |
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 280 | lemma deflation_sprod_map: | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 281 | assumes "deflation d1" and "deflation d2" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 282 | shows "deflation (sprod_map\<cdot>d1\<cdot>d2)" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 283 | proof | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 284 | interpret d1: deflation d1 by fact | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 285 | interpret d2: deflation d2 by fact | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 286 | fix x | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 287 | show "sprod_map\<cdot>d1\<cdot>d2\<cdot>(sprod_map\<cdot>d1\<cdot>d2\<cdot>x) = sprod_map\<cdot>d1\<cdot>d2\<cdot>x" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 288 | apply (induct x, simp) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 289 | apply (case_tac "d1\<cdot>x = \<bottom>", simp, case_tac "d2\<cdot>y = \<bottom>", simp) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 290 | apply (simp add: d1.idem d2.idem) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 291 | done | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 292 | show "sprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 293 | apply (induct x, simp) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 294 | apply (simp add: monofun_cfun d1.below d2.below) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 295 | done | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 296 | qed | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 297 | |
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 298 | lemma finite_deflation_sprod_map: | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 299 | assumes "finite_deflation d1" and "finite_deflation d2" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 300 | shows "finite_deflation (sprod_map\<cdot>d1\<cdot>d2)" | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 301 | proof (intro finite_deflation.intro finite_deflation_axioms.intro) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 302 | interpret d1: finite_deflation d1 by fact | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 303 | interpret d2: finite_deflation d2 by fact | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 304 | have "deflation d1" and "deflation d2" by fact+ | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 305 | thus "deflation (sprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_sprod_map) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 306 |   have "{x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq> insert \<bottom>
 | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 307 |         ((\<lambda>(x, y). (:x, y:)) ` ({x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}))"
 | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 308 | by (rule subsetI, case_tac x, auto simp add: spair_eq_iff) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 309 |   thus "finite {x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
 | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 310 | by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes) | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 311 | qed | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 312 | |
| 25914 | 313 | subsection {* Strict product is a bifinite domain *}
 | 
| 314 | ||
| 35525 | 315 | instantiation sprod :: (bifinite, bifinite) bifinite | 
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
25914diff
changeset | 316 | begin | 
| 25914 | 317 | |
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
25914diff
changeset | 318 | definition | 
| 25914 | 319 | approx_sprod_def: | 
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 320 | "approx = (\<lambda>n. sprod_map\<cdot>(approx n)\<cdot>(approx n))" | 
| 25914 | 321 | |
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
25914diff
changeset | 322 | instance proof | 
| 25914 | 323 | fix i :: nat and x :: "'a \<otimes> 'b" | 
| 27310 | 324 | show "chain (approx :: nat \<Rightarrow> 'a \<otimes> 'b \<rightarrow> 'a \<otimes> 'b)" | 
| 25914 | 325 | unfolding approx_sprod_def by simp | 
| 326 | show "(\<Squnion>i. approx i\<cdot>x) = x" | |
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 327 | unfolding approx_sprod_def sprod_map_def | 
| 25914 | 328 | by (simp add: lub_distribs eta_cfun) | 
| 329 | show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" | |
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 330 | unfolding approx_sprod_def sprod_map_def | 
| 25914 | 331 | by (simp add: ssplit_def strictify_conv_if) | 
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 332 |   show "finite {x::'a \<otimes> 'b. approx i\<cdot>x = x}"
 | 
| 25914 | 333 | unfolding approx_sprod_def | 
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 334 | by (intro finite_deflation.finite_fixes | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 335 | finite_deflation_sprod_map | 
| 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 336 | finite_deflation_approx) | 
| 25914 | 337 | qed | 
| 338 | ||
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
25914diff
changeset | 339 | end | 
| 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
25914diff
changeset | 340 | |
| 25914 | 341 | lemma approx_spair [simp]: | 
| 342 | "approx i\<cdot>(:x, y:) = (:approx i\<cdot>x, approx i\<cdot>y:)" | |
| 33504 
b4210cc3ac97
map functions for various types, with ep_pair/deflation/finite_deflation lemmas
 huffman parents: 
32960diff
changeset | 343 | unfolding approx_sprod_def sprod_map_def | 
| 25914 | 344 | by (simp add: ssplit_def strictify_conv_if) | 
| 345 | ||
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 346 | end |