| author | blanchet | 
| Wed, 20 Jan 2010 10:38:19 +0100 | |
| changeset 34937 | fb90752a9271 | 
| parent 33399 | 768b2bb9e66a | 
| child 35900 | aa5dfb03eb1e | 
| permissions | -rw-r--r-- | 
| 15600 | 1 | (* Title: HOLCF/Cprod.thy | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 2 | Author: Franz Regensburger | 
| 
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 cpo of cartesian products *}
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 6 | |
| 15577 | 7 | theory Cprod | 
| 25910 | 8 | imports Bifinite | 
| 15577 | 9 | begin | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 10 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 11 | defaultsort cpo | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 12 | |
| 16008 | 13 | subsection {* Type @{typ unit} is a pcpo *}
 | 
| 14 | ||
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 15 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 16 | unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 17 | "unit_when = (\<Lambda> a _. a)" | 
| 16008 | 18 | |
| 18289 | 19 | translations | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 20 | "\<Lambda>(). t" == "CONST unit_when\<cdot>t" | 
| 18289 | 21 | |
| 22 | lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a" | |
| 23 | by (simp add: unit_when_def) | |
| 24 | ||
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 25 | subsection {* Continuous versions of constants *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 26 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 27 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 28 |   cpair :: "'a \<rightarrow> 'b \<rightarrow> ('a * 'b)"  -- {* continuous pairing *}  where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 29 | "cpair = (\<Lambda> x y. (x, y))" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 30 | |
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 31 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 32 |   cfst :: "('a * 'b) \<rightarrow> 'a" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 33 | "cfst = (\<Lambda> p. fst p)" | 
| 17834 | 34 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 35 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 36 |   csnd :: "('a * 'b) \<rightarrow> 'b" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 37 | "csnd = (\<Lambda> p. snd p)" | 
| 17834 | 38 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 39 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 40 |   csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 41 | "csplit = (\<Lambda> f p. f\<cdot>(cfst\<cdot>p)\<cdot>(csnd\<cdot>p))" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 42 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 43 | syntax | 
| 17834 | 44 |   "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b"  ("(1<_,/ _>)")
 | 
| 45 | ||
| 46 | syntax (xsymbols) | |
| 47 |   "_ctuple" :: "['a, args] \<Rightarrow> 'a * 'b"  ("(1\<langle>_,/ _\<rangle>)")
 | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 48 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 49 | translations | 
| 18078 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
18077diff
changeset | 50 | "\<langle>x, y, z\<rangle>" == "\<langle>x, \<langle>y, z\<rangle>\<rangle>" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 51 | "\<langle>x, y\<rangle>" == "CONST cpair\<cdot>x\<cdot>y" | 
| 17834 | 52 | |
| 17816 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 huffman parents: 
16916diff
changeset | 53 | translations | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18289diff
changeset | 54 | "\<Lambda>(CONST cpair\<cdot>x\<cdot>y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)" | 
| 33399 | 55 | "\<Lambda>(CONST Pair x y). t" => "CONST csplit\<cdot>(\<Lambda> x y. t)" | 
| 17816 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 huffman parents: 
16916diff
changeset | 56 | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 57 | |
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 58 | subsection {* Convert all lemmas to the continuous versions *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 59 | |
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 60 | lemma cpair_eq_pair: "<x, y> = (x, y)" | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 61 | by (simp add: cpair_def cont_pair1 cont_pair2) | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 62 | |
| 25910 | 63 | lemma pair_eq_cpair: "(x, y) = <x, y>" | 
| 64 | by (simp add: cpair_def cont_pair1 cont_pair2) | |
| 65 | ||
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 66 | lemma inject_cpair: "<a,b> = <aa,ba> \<Longrightarrow> a = aa \<and> b = ba" | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 67 | by (simp add: cpair_eq_pair) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 68 | |
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 69 | lemma cpair_eq [iff]: "(<a, b> = <a', b'>) = (a = a' \<and> b = b')" | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 70 | by (simp add: cpair_eq_pair) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 71 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 72 | lemma cpair_below [iff]: "(<a, b> \<sqsubseteq> <a', b'>) = (a \<sqsubseteq> a' \<and> b \<sqsubseteq> b')" | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29530diff
changeset | 73 | by (simp add: cpair_eq_pair) | 
| 16057 
e23a61b3406f
added lemma cpair_eq, made cfst_strict and csnd_strict into simp rules
 huffman parents: 
16008diff
changeset | 74 | |
| 18077 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 huffman parents: 
17837diff
changeset | 75 | lemma cpair_defined_iff [iff]: "(<x, y> = \<bottom>) = (x = \<bottom> \<and> y = \<bottom>)" | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29530diff
changeset | 76 | by (simp add: cpair_eq_pair) | 
| 16916 | 77 | |
| 25913 | 78 | lemma cpair_strict [simp]: "\<langle>\<bottom>, \<bottom>\<rangle> = \<bottom>" | 
| 18077 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 huffman parents: 
17837diff
changeset | 79 | by simp | 
| 16210 
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
 huffman parents: 
16093diff
changeset | 80 | |
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 81 | lemma inst_cprod_pcpo2: "\<bottom> = <\<bottom>, \<bottom>>" | 
| 16916 | 82 | by (rule cpair_strict [symmetric]) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 83 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 84 | lemma defined_cpair_rev: | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 85 | "<a,b> = \<bottom> \<Longrightarrow> a = \<bottom> \<and> b = \<bottom>" | 
| 18077 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 huffman parents: 
17837diff
changeset | 86 | by simp | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 87 | |
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 88 | lemma Exh_Cprod2: "\<exists>a b. z = <a, b>" | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 89 | by (simp add: cpair_eq_pair) | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 90 | |
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 91 | lemma cprodE: "\<lbrakk>\<And>x y. p = <x, y> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 92 | by (cut_tac Exh_Cprod2, auto) | 
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 93 | |
| 16210 
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
 huffman parents: 
16093diff
changeset | 94 | lemma cfst_cpair [simp]: "cfst\<cdot><x, y> = x" | 
| 31113 
15cf300a742f
move bifinite instance for product type from Cprod.thy to Bifinite.thy
 huffman parents: 
31076diff
changeset | 95 | by (simp add: cpair_eq_pair cfst_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 96 | |
| 16210 
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
 huffman parents: 
16093diff
changeset | 97 | lemma csnd_cpair [simp]: "csnd\<cdot><x, y> = y" | 
| 31113 
15cf300a742f
move bifinite instance for product type from Cprod.thy to Bifinite.thy
 huffman parents: 
31076diff
changeset | 98 | by (simp add: cpair_eq_pair csnd_def) | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 99 | |
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 100 | lemma cfst_strict [simp]: "cfst\<cdot>\<bottom> = \<bottom>" | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29530diff
changeset | 101 | by (simp add: cfst_def) | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 102 | |
| 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 103 | lemma csnd_strict [simp]: "csnd\<cdot>\<bottom> = \<bottom>" | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29530diff
changeset | 104 | by (simp add: csnd_def) | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 105 | |
| 25910 | 106 | lemma cpair_cfst_csnd: "\<langle>cfst\<cdot>p, csnd\<cdot>p\<rangle> = p" | 
| 107 | by (cases p rule: cprodE, simp) | |
| 108 | ||
| 109 | lemmas surjective_pairing_Cprod2 = cpair_cfst_csnd | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 110 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 111 | lemma below_cprod: "x \<sqsubseteq> y = (cfst\<cdot>x \<sqsubseteq> cfst\<cdot>y \<and> csnd\<cdot>x \<sqsubseteq> csnd\<cdot>y)" | 
| 31113 
15cf300a742f
move bifinite instance for product type from Cprod.thy to Bifinite.thy
 huffman parents: 
31076diff
changeset | 112 | by (simp add: below_prod_def cfst_def csnd_def) | 
| 16315 | 113 | |
| 16750 | 114 | lemma eq_cprod: "(x = y) = (cfst\<cdot>x = cfst\<cdot>y \<and> csnd\<cdot>x = csnd\<cdot>y)" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 115 | by (auto simp add: po_eq_conv below_cprod) | 
| 16750 | 116 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 117 | lemma cfst_below_iff: "cfst\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <y, csnd\<cdot>x>" | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 118 | by (simp add: below_cprod) | 
| 25879 | 119 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 120 | lemma csnd_below_iff: "csnd\<cdot>x \<sqsubseteq> y = x \<sqsubseteq> <cfst\<cdot>x, y>" | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 121 | by (simp add: below_cprod) | 
| 25879 | 122 | |
| 123 | lemma compact_cfst: "compact x \<Longrightarrow> compact (cfst\<cdot>x)" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 124 | by (rule compactI, simp add: cfst_below_iff) | 
| 25879 | 125 | |
| 126 | lemma compact_csnd: "compact x \<Longrightarrow> compact (csnd\<cdot>x)" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29535diff
changeset | 127 | by (rule compactI, simp add: csnd_below_iff) | 
| 25879 | 128 | |
| 129 | lemma compact_cpair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact <x, y>" | |
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29530diff
changeset | 130 | by (simp add: cpair_eq_pair) | 
| 17837 | 131 | |
| 25879 | 132 | lemma compact_cpair_iff [simp]: "compact <x, y> = (compact x \<and> compact y)" | 
| 29535 
08824fad8879
add strictness and compactness lemmas to Product_Cpo.thy
 huffman parents: 
29530diff
changeset | 133 | by (simp add: cpair_eq_pair) | 
| 25905 | 134 | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 135 | lemma lub_cprod2: | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 136 | "chain S \<Longrightarrow> range S <<| <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>" | 
| 31113 
15cf300a742f
move bifinite instance for product type from Cprod.thy to Bifinite.thy
 huffman parents: 
31076diff
changeset | 137 | apply (simp add: cpair_eq_pair cfst_def csnd_def) | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 138 | apply (erule lub_cprod) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 139 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 140 | |
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 141 | lemma thelub_cprod2: | 
| 27413 | 142 | "chain S \<Longrightarrow> (\<Squnion>i. S i) = <\<Squnion>i. cfst\<cdot>(S i), \<Squnion>i. csnd\<cdot>(S i)>" | 
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 143 | by (rule lub_cprod2 [THEN thelubI]) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 144 | |
| 18077 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 huffman parents: 
17837diff
changeset | 145 | lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>" | 
| 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 huffman parents: 
17837diff
changeset | 146 | by (simp add: csplit_def) | 
| 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 huffman parents: 
17837diff
changeset | 147 | |
| 16081 
81a4b4a245b0
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
 huffman parents: 
16070diff
changeset | 148 | lemma csplit2 [simp]: "csplit\<cdot>f\<cdot><x,y> = f\<cdot>x\<cdot>y" | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 149 | by (simp add: csplit_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 150 | |
| 16553 | 151 | lemma csplit3 [simp]: "csplit\<cdot>cpair\<cdot>z = z" | 
| 25910 | 152 | by (simp add: csplit_def cpair_cfst_csnd) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 153 | |
| 33399 | 154 | lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y" | 
| 155 | by (simp add: csplit_def cfst_def csnd_def) | |
| 156 | ||
| 16210 
5d1b752cacc1
changed to use new contlubI, monofun_def; renamed cfst2, csnd2 to cfst_cpair, csnd_cpair; added lemma cpair_strict
 huffman parents: 
16093diff
changeset | 157 | lemmas Cprod_rews = cfst_cpair csnd_cpair csplit2 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 158 | |
| 25910 | 159 | subsection {* Product type is a bifinite domain *}
 | 
| 160 | ||
| 161 | lemma approx_cpair [simp]: | |
| 162 | "approx i\<cdot>\<langle>x, y\<rangle> = \<langle>approx i\<cdot>x, approx i\<cdot>y\<rangle>" | |
| 31113 
15cf300a742f
move bifinite instance for product type from Cprod.thy to Bifinite.thy
 huffman parents: 
31076diff
changeset | 163 | by (simp add: cpair_eq_pair) | 
| 25910 | 164 | |
| 165 | lemma cfst_approx: "cfst\<cdot>(approx i\<cdot>p) = approx i\<cdot>(cfst\<cdot>p)" | |
| 166 | by (cases p rule: cprodE, simp) | |
| 167 | ||
| 168 | lemma csnd_approx: "csnd\<cdot>(approx i\<cdot>p) = approx i\<cdot>(csnd\<cdot>p)" | |
| 169 | by (cases p rule: cprodE, simp) | |
| 170 | ||
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 171 | end |