| author | huffman | 
| Thu, 14 Sep 2006 20:31:10 +0200 | |
| changeset 20539 | a7b2f90f698d | 
| parent 18078 | 20e5a6440790 | 
| child 25131 | 2c8caac48ade | 
| permissions | -rw-r--r-- | 
| 15600 | 1 | (* Title: HOLCF/Sprod.thy | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 2 | ID: $Id$ | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 3 | Author: Franz Regensburger and Brian Huffman | 
| 15576 
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 | Strict product with typedef. | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 6 | *) | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 7 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 8 | header {* The type of strict products *}
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 9 | |
| 15577 | 10 | theory Sprod | 
| 16699 | 11 | imports Cprod | 
| 15577 | 12 | begin | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 13 | |
| 16082 | 14 | defaultsort pcpo | 
| 15 | ||
| 15591 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 16 | subsection {* Definition of strict product type *}
 | 
| 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 17 | |
| 17817 | 18 | pcpodef (Sprod)  ('a, 'b) "**" (infixr "**" 20) =
 | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 19 |         "{p::'a \<times> 'b. p = \<bottom> \<or> (cfst\<cdot>p \<noteq> \<bottom> \<and> csnd\<cdot>p \<noteq> \<bottom>)}"
 | 
| 16699 | 20 | by simp | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 21 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 22 | syntax (xsymbols) | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 23 |   "**"		:: "[type, type] => type"	 ("(_ \<otimes>/ _)" [21,20] 20)
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 24 | syntax (HTML output) | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 25 |   "**"		:: "[type, type] => type"	 ("(_ \<otimes>/ _)" [21,20] 20)
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 26 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 27 | lemma spair_lemma: | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 28 | "<strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a> \<in> Sprod" | 
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 29 | by (simp add: Sprod_def strictify_conv_if cpair_strict) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 30 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 31 | subsection {* Definitions of constants *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 32 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 33 | consts | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 34 |   sfst :: "('a ** 'b) \<rightarrow> 'a"
 | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 35 |   ssnd :: "('a ** 'b) \<rightarrow> 'b"
 | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 36 |   spair :: "'a \<rightarrow> 'b \<rightarrow> ('a ** 'b)"
 | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 37 |   ssplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a ** 'b) \<rightarrow> 'c"
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 38 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 39 | defs | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 40 | sfst_def: "sfst \<equiv> \<Lambda> p. cfst\<cdot>(Rep_Sprod p)" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 41 | ssnd_def: "ssnd \<equiv> \<Lambda> p. csnd\<cdot>(Rep_Sprod p)" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 42 | spair_def: "spair \<equiv> \<Lambda> a b. Abs_Sprod | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 43 | <strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a>" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 44 | ssplit_def: "ssplit \<equiv> \<Lambda> f. strictify\<cdot>(\<Lambda> p. f\<cdot>(sfst\<cdot>p)\<cdot>(ssnd\<cdot>p))" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 45 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 46 | syntax | 
| 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 | 47 |   "@stuple" :: "['a, args] => 'a ** 'b"  ("(1'(:_,/ _:'))")
 | 
| 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: 
17837diff
changeset | 50 | "(:x, y, z:)" == "(:x, (:y, z:):)" | 
| 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 51 | "(:x, y:)" == "spair\<cdot>x\<cdot>y" | 
| 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 52 | |
| 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 53 | translations | 
| 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17837diff
changeset | 54 | "\<Lambda>(spair\<cdot>x\<cdot>y). t" == "ssplit\<cdot>(\<Lambda> x y. t)" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 55 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 56 | subsection {* Case analysis *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 57 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 58 | lemma spair_Abs_Sprod: | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 59 | "(:a, b:) = Abs_Sprod <strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a>" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 60 | apply (unfold spair_def) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 61 | apply (simp add: cont_Abs_Sprod spair_lemma) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 62 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 63 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 64 | lemma Exh_Sprod2: | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 65 | "z = \<bottom> \<or> (\<exists>a b. z = (:a, b:) \<and> a \<noteq> \<bottom> \<and> b \<noteq> \<bottom>)" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 66 | apply (rule_tac x=z in Abs_Sprod_cases) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 67 | apply (simp add: Sprod_def) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 68 | apply (erule disjE) | 
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 69 | apply (simp add: Abs_Sprod_strict) | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 70 | apply (rule disjI2) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 71 | apply (rule_tac x="cfst\<cdot>y" in exI) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 72 | apply (rule_tac x="csnd\<cdot>y" in exI) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 73 | apply (simp add: spair_Abs_Sprod Abs_Sprod_inject spair_lemma) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 74 | apply (simp add: surjective_pairing_Cprod2) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 75 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 76 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 77 | lemma sprodE: | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 78 | "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x y. \<lbrakk>p = (:x, y:); x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 79 | by (cut_tac z=p in Exh_Sprod2, auto) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 80 | |
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 81 | subsection {* Properties of @{term spair} *}
 | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 82 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 83 | lemma spair_strict1 [simp]: "(:\<bottom>, y:) = \<bottom>" | 
| 16920 | 84 | by (simp add: spair_Abs_Sprod strictify_conv_if cpair_strict Abs_Sprod_strict) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 85 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 86 | lemma spair_strict2 [simp]: "(:x, \<bottom>:) = \<bottom>" | 
| 16920 | 87 | by (simp add: spair_Abs_Sprod strictify_conv_if cpair_strict Abs_Sprod_strict) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 88 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 89 | lemma spair_strict: "x = \<bottom> \<or> y = \<bottom> \<Longrightarrow> (:x, y:) = \<bottom>" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 90 | by auto | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 91 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 92 | lemma spair_strict_rev: "(:x, y:) \<noteq> \<bottom> \<Longrightarrow> x \<noteq> \<bottom> \<and> y \<noteq> \<bottom>" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 93 | by (erule contrapos_np, auto) | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 94 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 95 | lemma spair_defined [simp]: | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 96 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<noteq> \<bottom>" | 
| 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 | 97 | by (simp add: spair_Abs_Sprod Abs_Sprod_defined Sprod_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 98 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 99 | lemma spair_defined_rev: "(:x, y:) = \<bottom> \<Longrightarrow> x = \<bottom> \<or> y = \<bottom>" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 100 | by (erule contrapos_pp, simp) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 101 | |
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 102 | lemma spair_eq: | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 103 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> ((:x, y:) = (:a, b:)) = (x = a \<and> y = b)" | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 104 | apply (simp add: spair_Abs_Sprod) | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 105 | apply (simp add: Abs_Sprod_inject [OF _ spair_lemma] Sprod_def) | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 106 | apply (simp add: strictify_conv_if) | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 107 | done | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 108 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 109 | lemma spair_inject: | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 110 | "\<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 | 111 | by (rule spair_eq [THEN iffD1]) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 112 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 113 | lemma inst_sprod_pcpo2: "UU = (:UU,UU:)" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 114 | by simp | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 115 | |
| 17837 | 116 | lemma Rep_Sprod_spair: | 
| 117 | "Rep_Sprod (:a, b:) = <strictify\<cdot>(\<Lambda> b. a)\<cdot>b, strictify\<cdot>(\<Lambda> a. b)\<cdot>a>" | |
| 118 | apply (unfold spair_def) | |
| 119 | apply (simp add: cont_Abs_Sprod Abs_Sprod_inverse spair_lemma) | |
| 120 | done | |
| 121 | ||
| 122 | lemma compact_spair: "\<lbrakk>compact x; compact y\<rbrakk> \<Longrightarrow> compact (:x, y:)" | |
| 123 | by (rule compact_Sprod, simp add: Rep_Sprod_spair strictify_conv_if) | |
| 124 | ||
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 125 | subsection {* Properties of @{term sfst} and @{term ssnd} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 126 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 127 | 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 | 128 | 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 | 129 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 130 | 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 | 131 | 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 | 132 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 133 | 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 | 134 | 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 | 135 | |
| 16212 
422f836f6b39
renamed strict, defined, and inject lemmas; renamed sfst2, ssnd2 to sfst_spair, ssnd_spair
 huffman parents: 
16082diff
changeset | 136 | 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 | 137 | 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 | 138 | |
| 16777 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 139 | lemma sfst_defined_iff [simp]: "(sfst\<cdot>p = \<bottom>) = (p = \<bottom>)" | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 140 | by (rule_tac p=p in sprodE, simp_all) | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 141 | |
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 142 | lemma ssnd_defined_iff [simp]: "(ssnd\<cdot>p = \<bottom>) = (p = \<bottom>)" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 143 | by (rule_tac p=p in sprodE, simp_all) | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 144 | |
| 16777 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 145 | 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 | 146 | by simp | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 147 | |
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 148 | 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 | 149 | by simp | 
| 
555c8951f05c
added lemmas sfst_defined_iff, ssnd_defined_iff, sfst_defined, ssnd_defined
 huffman parents: 
16751diff
changeset | 150 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 151 | lemma surjective_pairing_Sprod2: "(:sfst\<cdot>p, ssnd\<cdot>p:) = p" | 
| 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 152 | by (rule_tac p=p in sprodE, simp_all) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 153 | |
| 16751 | 154 | lemma less_sprod: "x \<sqsubseteq> y = (sfst\<cdot>x \<sqsubseteq> sfst\<cdot>y \<and> ssnd\<cdot>x \<sqsubseteq> ssnd\<cdot>y)" | 
| 16699 | 155 | apply (simp add: less_Sprod_def sfst_def ssnd_def cont_Rep_Sprod) | 
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 156 | apply (rule less_cprod) | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 157 | done | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 158 | |
| 16751 | 159 | lemma eq_sprod: "(x = y) = (sfst\<cdot>x = sfst\<cdot>y \<and> ssnd\<cdot>x = ssnd\<cdot>y)" | 
| 160 | by (auto simp add: po_eq_conv less_sprod) | |
| 161 | ||
| 16317 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 162 | lemma spair_less: | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 163 | "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> (:x, y:) \<sqsubseteq> (:a, b:) = (x \<sqsubseteq> a \<and> y \<sqsubseteq> b)" | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 164 | apply (case_tac "a = \<bottom>") | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 165 | apply (simp add: eq_UU_iff [symmetric]) | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 166 | apply (case_tac "b = \<bottom>") | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 167 | apply (simp add: eq_UU_iff [symmetric]) | 
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 168 | apply (simp add: less_sprod) | 
| 
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 | |
| 
868eddbcaf6e
added theorems less_sprod, spair_less, spair_eq, spair_inject
 huffman parents: 
16212diff
changeset | 171 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 172 | subsection {* Properties of @{term ssplit} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 173 | |
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 174 | 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 | 175 | by (simp add: ssplit_def) | 
| 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 176 | |
| 16920 | 177 | 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 | 178 | by (simp add: ssplit_def) | 
| 
50c3384ca6c4
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 179 | |
| 16553 | 180 | lemma ssplit3 [simp]: "ssplit\<cdot>spair\<cdot>z = z" | 
| 16059 
dab0d004732f
Simplified version of strict product theory, using TypedefPcpo
 huffman parents: 
15930diff
changeset | 181 | by (rule_tac p=z in sprodE, simp_all) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 182 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 183 | end |