| author | blanchet | 
| Mon, 02 Aug 2010 19:15:15 +0200 | |
| changeset 38168 | e5978befb951 | 
| parent 37808 | e604e5f9bb6a | 
| child 38715 | 6513ea67d95d | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Product_Type.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1992 University of Cambridge | |
| 11777 | 4 | *) | 
| 10213 | 5 | |
| 11838 | 6 | header {* Cartesian products *}
 | 
| 10213 | 7 | |
| 15131 | 8 | theory Product_Type | 
| 33959 
2afc55e8ed27
bootstrap datatype_rep_proofs in Datatype.thy (avoids unchecked dynamic name references)
 haftmann parents: 
33638diff
changeset | 9 | imports Typedef Inductive Fun | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 10 | uses | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 11 |   ("Tools/split_rule.ML")
 | 
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 12 |   ("Tools/inductive_codegen.ML")
 | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31667diff
changeset | 13 |   ("Tools/inductive_set.ML")
 | 
| 15131 | 14 | begin | 
| 11838 | 15 | |
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 16 | subsection {* @{typ bool} is a datatype *}
 | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 17 | |
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26975diff
changeset | 18 | rep_datatype True False by (auto intro: bool_induct) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 19 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 20 | declare case_split [cases type: bool] | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 21 | -- "prefer plain propositional version" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 22 | |
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 23 | lemma | 
| 28562 | 24 | shows [code]: "eq_class.eq False P \<longleftrightarrow> \<not> P" | 
| 25 | and [code]: "eq_class.eq True P \<longleftrightarrow> P" | |
| 26 | and [code]: "eq_class.eq P False \<longleftrightarrow> \<not> P" | |
| 27 | and [code]: "eq_class.eq P True \<longleftrightarrow> P" | |
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 28 | and [code nbe]: "eq_class.eq P P \<longleftrightarrow> True" | 
| 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 29 | by (simp_all add: eq) | 
| 25534 
d0b74fdd6067
simplified infrastructure for code generator operational equality
 haftmann parents: 
25511diff
changeset | 30 | |
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 31 | code_const "eq_class.eq \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool" | 
| 25534 
d0b74fdd6067
simplified infrastructure for code generator operational equality
 haftmann parents: 
25511diff
changeset | 32 | (Haskell infixl 4 "==") | 
| 
d0b74fdd6067
simplified infrastructure for code generator operational equality
 haftmann parents: 
25511diff
changeset | 33 | |
| 
d0b74fdd6067
simplified infrastructure for code generator operational equality
 haftmann parents: 
25511diff
changeset | 34 | code_instance bool :: eq | 
| 
d0b74fdd6067
simplified infrastructure for code generator operational equality
 haftmann parents: 
25511diff
changeset | 35 | (Haskell -) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 36 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 37 | |
| 37166 | 38 | subsection {* The @{text unit} type *}
 | 
| 11838 | 39 | |
| 40 | typedef unit = "{True}"
 | |
| 41 | proof | |
| 20588 | 42 | show "True : ?unit" .. | 
| 11838 | 43 | qed | 
| 44 | ||
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 45 | definition | 
| 11838 | 46 |   Unity :: unit    ("'(')")
 | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 47 | where | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 48 | "() = Abs_unit True" | 
| 11838 | 49 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 50 | lemma unit_eq [no_atp]: "u = ()" | 
| 11838 | 51 | by (induct u) (simp add: unit_def Unity_def) | 
| 52 | ||
| 53 | text {*
 | |
| 54 |   Simplification procedure for @{thm [source] unit_eq}.  Cannot use
 | |
| 55 | this rule directly --- it loops! | |
| 56 | *} | |
| 57 | ||
| 26480 | 58 | ML {*
 | 
| 13462 | 59 | val unit_eq_proc = | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 60 |     let val unit_meta_eq = mk_meta_eq @{thm unit_eq} in
 | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 61 |       Simplifier.simproc @{theory} "unit_eq" ["x::unit"]
 | 
| 15531 | 62 | (fn _ => fn _ => fn t => if HOLogic.is_unit t then NONE else SOME unit_meta_eq) | 
| 13462 | 63 | end; | 
| 11838 | 64 | |
| 65 | Addsimprocs [unit_eq_proc]; | |
| 66 | *} | |
| 67 | ||
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26975diff
changeset | 68 | rep_datatype "()" by simp | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 69 | |
| 11838 | 70 | lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()" | 
| 71 | by simp | |
| 72 | ||
| 73 | lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P" | |
| 74 | by (rule triv_forall_equality) | |
| 75 | ||
| 76 | text {*
 | |
| 77 |   This rewrite counters the effect of @{text unit_eq_proc} on @{term
 | |
| 78 |   [source] "%u::unit. f u"}, replacing it by @{term [source]
 | |
| 79 |   f} rather than by @{term [source] "%u. f ()"}.
 | |
| 80 | *} | |
| 81 | ||
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 82 | lemma unit_abs_eta_conv [simp,no_atp]: "(%u::unit. f ()) = f" | 
| 11838 | 83 | by (rule ext) simp | 
| 10213 | 84 | |
| 30924 | 85 | instantiation unit :: default | 
| 86 | begin | |
| 87 | ||
| 88 | definition "default = ()" | |
| 89 | ||
| 90 | instance .. | |
| 91 | ||
| 92 | end | |
| 10213 | 93 | |
| 28562 | 94 | lemma [code]: | 
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 95 | "eq_class.eq (u\<Colon>unit) v \<longleftrightarrow> True" unfolding eq unit_eq [of u] unit_eq [of v] by rule+ | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 96 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 97 | code_type unit | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 98 | (SML "unit") | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 99 | (OCaml "unit") | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 100 | (Haskell "()") | 
| 34886 | 101 | (Scala "Unit") | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 102 | |
| 37166 | 103 | code_const Unity | 
| 104 | (SML "()") | |
| 105 | (OCaml "()") | |
| 106 | (Haskell "()") | |
| 107 | (Scala "()") | |
| 108 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 109 | code_instance unit :: eq | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 110 | (Haskell -) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 111 | |
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 112 | code_const "eq_class.eq \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 113 | (Haskell infixl 4 "==") | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 114 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 115 | code_reserved SML | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 116 | unit | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 117 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 118 | code_reserved OCaml | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 119 | unit | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 120 | |
| 34886 | 121 | code_reserved Scala | 
| 122 | Unit | |
| 123 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 124 | |
| 37166 | 125 | subsection {* The product type *}
 | 
| 10213 | 126 | |
| 37166 | 127 | subsubsection {* Type definition *}
 | 
| 128 | ||
| 129 | definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 130 | "Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)" | 
| 10213 | 131 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 132 | typedef ('a, 'b) prod (infixr "*" 20)
 | 
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 133 |   = "{f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
 | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 134 | proof | 
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 135 | fix a b show "Pair_Rep a b \<in> ?prod" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 136 | by rule+ | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 137 | qed | 
| 10213 | 138 | |
| 35427 | 139 | type_notation (xsymbols) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 140 |   "prod"  ("(_ \<times>/ _)" [21, 20] 20)
 | 
| 35427 | 141 | type_notation (HTML output) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 142 |   "prod"  ("(_ \<times>/ _)" [21, 20] 20)
 | 
| 10213 | 143 | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 144 | definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b" where | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 145 | "Pair a b = Abs_prod (Pair_Rep a b)" | 
| 37166 | 146 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 147 | rep_datatype Pair proof - | 
| 37166 | 148 | fix P :: "'a \<times> 'b \<Rightarrow> bool" and p | 
| 149 | assume "\<And>a b. P (Pair a b)" | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 150 | then show "P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) | 
| 37166 | 151 | next | 
| 152 | fix a c :: 'a and b d :: 'b | |
| 153 | have "Pair_Rep a b = Pair_Rep c d \<longleftrightarrow> a = c \<and> b = d" | |
| 154 | by (auto simp add: Pair_Rep_def expand_fun_eq) | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 155 | moreover have "Pair_Rep a b \<in> prod" and "Pair_Rep c d \<in> prod" | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 156 | by (auto simp add: prod_def) | 
| 37166 | 157 | ultimately show "Pair a b = Pair c d \<longleftrightarrow> a = c \<and> b = d" | 
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 158 | by (simp add: Pair_def Abs_prod_inject) | 
| 37166 | 159 | qed | 
| 160 | ||
| 37704 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 blanchet parents: 
37678diff
changeset | 161 | declare prod.simps(2) [nitpick_simp del] | 
| 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 blanchet parents: 
37678diff
changeset | 162 | |
| 37411 
c88c44156083
removed simplifier congruence rule of "prod_case"
 haftmann parents: 
37389diff
changeset | 163 | declare weak_case_cong [cong del] | 
| 
c88c44156083
removed simplifier congruence rule of "prod_case"
 haftmann parents: 
37389diff
changeset | 164 | |
| 37166 | 165 | |
| 166 | subsubsection {* Tuple syntax *}
 | |
| 167 | ||
| 37591 | 168 | abbreviation (input) split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
 | 
| 169 | "split \<equiv> prod_case" | |
| 19535 | 170 | |
| 11777 | 171 | text {*
 | 
| 172 |   Patterns -- extends pre-defined type @{typ pttrn} used in
 | |
| 173 | abstractions. | |
| 174 | *} | |
| 10213 | 175 | |
| 176 | nonterminals | |
| 177 | tuple_args patterns | |
| 178 | ||
| 179 | syntax | |
| 180 |   "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
 | |
| 181 |   "_tuple_arg"  :: "'a => tuple_args"                   ("_")
 | |
| 182 |   "_tuple_args" :: "'a => tuple_args => tuple_args"     ("_,/ _")
 | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 183 |   "_pattern"    :: "[pttrn, patterns] => pttrn"         ("'(_,/ _')")
 | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 184 |   ""            :: "pttrn => patterns"                  ("_")
 | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 185 |   "_patterns"   :: "[pttrn, patterns] => patterns"      ("_,/ _")
 | 
| 10213 | 186 | |
| 187 | translations | |
| 35115 | 188 | "(x, y)" == "CONST Pair x y" | 
| 10213 | 189 | "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" | 
| 37591 | 190 | "%(x, y, zs). b" == "CONST prod_case (%x (y, zs). b)" | 
| 191 | "%(x, y). b" == "CONST prod_case (%x y. b)" | |
| 35115 | 192 | "_abs (CONST Pair x y) t" => "%(x, y). t" | 
| 37166 | 193 |   -- {* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
 | 
| 194 | The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *} | |
| 10213 | 195 | |
| 35115 | 196 | (*reconstruct pattern from (nested) splits, avoiding eta-contraction of body; | 
| 197 | works best with enclosing "let", if "let" does not avoid eta-contraction*) | |
| 14359 | 198 | print_translation {*
 | 
| 35115 | 199 | let | 
| 200 | fun split_tr' [Abs (x, T, t as (Abs abs))] = | |
| 201 | (* split (%x y. t) => %(x,y) t *) | |
| 202 | let | |
| 203 | val (y, t') = atomic_abs_tr' abs; | |
| 204 | val (x', t'') = atomic_abs_tr' (x, T, t'); | |
| 205 | in | |
| 206 |           Syntax.const @{syntax_const "_abs"} $
 | |
| 207 |             (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
 | |
| 208 | end | |
| 37591 | 209 |     | split_tr' [Abs (x, T, (s as Const (@{const_syntax prod_case}, _) $ t))] =
 | 
| 35115 | 210 | (* split (%x. (split (%y z. t))) => %(x,y,z). t *) | 
| 211 | let | |
| 212 |           val Const (@{syntax_const "_abs"}, _) $
 | |
| 213 |             (Const (@{syntax_const "_pattern"}, _) $ y $ z) $ t' = split_tr' [t];
 | |
| 214 | val (x', t'') = atomic_abs_tr' (x, T, t'); | |
| 215 | in | |
| 216 |           Syntax.const @{syntax_const "_abs"} $
 | |
| 217 |             (Syntax.const @{syntax_const "_pattern"} $ x' $
 | |
| 218 |               (Syntax.const @{syntax_const "_patterns"} $ y $ z)) $ t''
 | |
| 219 | end | |
| 37591 | 220 |     | split_tr' [Const (@{const_syntax prod_case}, _) $ t] =
 | 
| 35115 | 221 | (* split (split (%x y z. t)) => %((x, y), z). t *) | 
| 222 | split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *) | |
| 223 |     | split_tr' [Const (@{syntax_const "_abs"}, _) $ x_y $ Abs abs] =
 | |
| 224 | (* split (%pttrn z. t) => %(pttrn,z). t *) | |
| 225 | let val (z, t) = atomic_abs_tr' abs in | |
| 226 |           Syntax.const @{syntax_const "_abs"} $
 | |
| 227 |             (Syntax.const @{syntax_const "_pattern"} $ x_y $ z) $ t
 | |
| 228 | end | |
| 229 | | split_tr' _ = raise Match; | |
| 37591 | 230 | in [(@{const_syntax prod_case}, split_tr')] end
 | 
| 14359 | 231 | *} | 
| 232 | ||
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 233 | (* print "split f" as "\<lambda>(x,y). f x y" and "split (\<lambda>x. f x)" as "\<lambda>(x,y). f x y" *) | 
| 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 234 | typed_print_translation {*
 | 
| 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 235 | let | 
| 35115 | 236 | fun split_guess_names_tr' _ T [Abs (x, _, Abs _)] = raise Match | 
| 237 | | split_guess_names_tr' _ T [Abs (x, xT, t)] = | |
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 238 | (case (head_of t) of | 
| 37591 | 239 |           Const (@{const_syntax prod_case}, _) => raise Match
 | 
| 35115 | 240 | | _ => | 
| 241 | let | |
| 242 | val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; | |
| 243 |             val (y, t') = atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0);
 | |
| 244 | val (x', t'') = atomic_abs_tr' (x, xT, t'); | |
| 245 | in | |
| 246 |             Syntax.const @{syntax_const "_abs"} $
 | |
| 247 |               (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
 | |
| 248 | end) | |
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 249 | | split_guess_names_tr' _ T [t] = | 
| 35115 | 250 | (case head_of t of | 
| 37591 | 251 |           Const (@{const_syntax prod_case}, _) => raise Match
 | 
| 35115 | 252 | | _ => | 
| 253 | let | |
| 254 | val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; | |
| 255 |             val (y, t') = atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0);
 | |
| 256 |             val (x', t'') = atomic_abs_tr' ("x", xT, t');
 | |
| 257 | in | |
| 258 |             Syntax.const @{syntax_const "_abs"} $
 | |
| 259 |               (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
 | |
| 260 | end) | |
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 261 | | split_guess_names_tr' _ _ _ = raise Match; | 
| 37591 | 262 | in [(@{const_syntax prod_case}, split_guess_names_tr')] end
 | 
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 263 | *} | 
| 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 264 | |
| 10213 | 265 | |
| 37166 | 266 | subsubsection {* Code generator setup *}
 | 
| 267 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 268 | code_type prod | 
| 37166 | 269 | (SML infix 2 "*") | 
| 270 | (OCaml infix 2 "*") | |
| 271 | (Haskell "!((_),/ (_))") | |
| 272 | (Scala "((_),/ (_))") | |
| 273 | ||
| 274 | code_const Pair | |
| 275 | (SML "!((_),/ (_))") | |
| 276 | (OCaml "!((_),/ (_))") | |
| 277 | (Haskell "!((_),/ (_))") | |
| 278 | (Scala "!((_),/ (_))") | |
| 279 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 280 | code_instance prod :: eq | 
| 37166 | 281 | (Haskell -) | 
| 282 | ||
| 283 | code_const "eq_class.eq \<Colon> 'a\<Colon>eq \<times> 'b\<Colon>eq \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool" | |
| 284 | (Haskell infixl 4 "==") | |
| 285 | ||
| 286 | types_code | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 287 |   "prod"     ("(_ */ _)")
 | 
| 37166 | 288 | attach (term_of) {*
 | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 289 | fun term_of_prod aF aT bF bT (x, y) = HOLogic.pair_const aT bT $ aF x $ bF y; | 
| 37166 | 290 | *} | 
| 291 | attach (test) {*
 | |
| 37808 
e604e5f9bb6a
correcting function name of generator for products of traditional code generator (introduced in 0040bafffdef)
 bulwahn parents: 
37765diff
changeset | 292 | fun gen_prod aG aT bG bT i = | 
| 37166 | 293 | let | 
| 294 | val (x, t) = aG i; | |
| 295 | val (y, u) = bG i | |
| 296 | in ((x, y), fn () => HOLogic.pair_const aT bT $ t () $ u ()) end; | |
| 297 | *} | |
| 298 | ||
| 299 | consts_code | |
| 300 |   "Pair"    ("(_,/ _)")
 | |
| 301 | ||
| 302 | setup {*
 | |
| 303 | let | |
| 304 | ||
| 305 | fun strip_abs_split 0 t = ([], t) | |
| 306 | | strip_abs_split i (Abs (s, T, t)) = | |
| 307 | let | |
| 308 | val s' = Codegen.new_name t s; | |
| 309 | val v = Free (s', T) | |
| 310 | in apfst (cons v) (strip_abs_split (i-1) (subst_bound (v, t))) end | |
| 37591 | 311 |   | strip_abs_split i (u as Const (@{const_name prod_case}, _) $ t) =
 | 
| 37166 | 312 | (case strip_abs_split (i+1) t of | 
| 313 | (v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u) | |
| 314 | | _ => ([], u)) | |
| 315 | | strip_abs_split i t = | |
| 316 |       strip_abs_split i (Abs ("x", hd (binder_types (fastype_of t)), t $ Bound 0));
 | |
| 317 | ||
| 318 | fun let_codegen thy defs dep thyname brack t gr = | |
| 319 | (case strip_comb t of | |
| 320 |     (t1 as Const (@{const_name Let}, _), t2 :: t3 :: ts) =>
 | |
| 321 | let | |
| 322 |       fun dest_let (l as Const (@{const_name Let}, _) $ t $ u) =
 | |
| 323 | (case strip_abs_split 1 u of | |
| 324 | ([p], u') => apfst (cons (p, t)) (dest_let u') | |
| 325 | | _ => ([], l)) | |
| 326 | | dest_let t = ([], t); | |
| 327 | fun mk_code (l, r) gr = | |
| 328 | let | |
| 329 | val (pl, gr1) = Codegen.invoke_codegen thy defs dep thyname false l gr; | |
| 330 | val (pr, gr2) = Codegen.invoke_codegen thy defs dep thyname false r gr1; | |
| 331 | in ((pl, pr), gr2) end | |
| 332 | in case dest_let (t1 $ t2 $ t3) of | |
| 333 | ([], _) => NONE | |
| 334 | | (ps, u) => | |
| 335 | let | |
| 336 | val (qs, gr1) = fold_map mk_code ps gr; | |
| 337 | val (pu, gr2) = Codegen.invoke_codegen thy defs dep thyname false u gr1; | |
| 338 | val (pargs, gr3) = fold_map | |
| 339 | (Codegen.invoke_codegen thy defs dep thyname true) ts gr2 | |
| 340 | in | |
| 341 | SOME (Codegen.mk_app brack | |
| 342 | (Pretty.blk (0, [Codegen.str "let ", Pretty.blk (0, flat | |
| 343 | (separate [Codegen.str ";", Pretty.brk 1] (map (fn (pl, pr) => | |
| 344 | [Pretty.block [Codegen.str "val ", pl, Codegen.str " =", | |
| 345 | Pretty.brk 1, pr]]) qs))), | |
| 346 | Pretty.brk 1, Codegen.str "in ", pu, | |
| 347 | Pretty.brk 1, Codegen.str "end"])) pargs, gr3) | |
| 348 | end | |
| 349 | end | |
| 350 | | _ => NONE); | |
| 351 | ||
| 352 | fun split_codegen thy defs dep thyname brack t gr = (case strip_comb t of | |
| 37591 | 353 |     (t1 as Const (@{const_name prod_case}, _), t2 :: ts) =>
 | 
| 37166 | 354 | let | 
| 355 | val ([p], u) = strip_abs_split 1 (t1 $ t2); | |
| 356 | val (q, gr1) = Codegen.invoke_codegen thy defs dep thyname false p gr; | |
| 357 | val (pu, gr2) = Codegen.invoke_codegen thy defs dep thyname false u gr1; | |
| 358 | val (pargs, gr3) = fold_map | |
| 359 | (Codegen.invoke_codegen thy defs dep thyname true) ts gr2 | |
| 360 | in | |
| 361 | SOME (Codegen.mk_app brack | |
| 362 | (Pretty.block [Codegen.str "(fn ", q, Codegen.str " =>", | |
| 363 | Pretty.brk 1, pu, Codegen.str ")"]) pargs, gr2) | |
| 364 | end | |
| 365 | | _ => NONE); | |
| 366 | ||
| 367 | in | |
| 368 | ||
| 369 | Codegen.add_codegen "let_codegen" let_codegen | |
| 370 | #> Codegen.add_codegen "split_codegen" split_codegen | |
| 371 | ||
| 372 | end | |
| 373 | *} | |
| 374 | ||
| 375 | ||
| 376 | subsubsection {* Fundamental operations and properties *}
 | |
| 11838 | 377 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 378 | lemma surj_pair [simp]: "EX x y. p = (x, y)" | 
| 37166 | 379 | by (cases p) simp | 
| 10213 | 380 | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 381 | definition fst :: "'a \<times> 'b \<Rightarrow> 'a" where | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 382 | "fst p = (case p of (a, b) \<Rightarrow> a)" | 
| 11838 | 383 | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 384 | definition snd :: "'a \<times> 'b \<Rightarrow> 'b" where | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 385 | "snd p = (case p of (a, b) \<Rightarrow> b)" | 
| 11838 | 386 | |
| 22886 | 387 | lemma fst_conv [simp, code]: "fst (a, b) = a" | 
| 37166 | 388 | unfolding fst_def by simp | 
| 11838 | 389 | |
| 22886 | 390 | lemma snd_conv [simp, code]: "snd (a, b) = b" | 
| 37166 | 391 | unfolding snd_def by simp | 
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 392 | |
| 37166 | 393 | code_const fst and snd | 
| 394 | (Haskell "fst" and "snd") | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 395 | |
| 37704 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 blanchet parents: 
37678diff
changeset | 396 | lemma prod_case_unfold [nitpick_def]: "prod_case = (%c p. c (fst p) (snd p))" | 
| 37166 | 397 | by (simp add: expand_fun_eq split: prod.split) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 398 | |
| 11838 | 399 | lemma fst_eqD: "fst (x, y) = a ==> x = a" | 
| 400 | by simp | |
| 401 | ||
| 402 | lemma snd_eqD: "snd (x, y) = a ==> y = a" | |
| 403 | by simp | |
| 404 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 405 | lemma pair_collapse [simp]: "(fst p, snd p) = p" | 
| 11838 | 406 | by (cases p) simp | 
| 407 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 408 | lemmas surjective_pairing = pair_collapse [symmetric] | 
| 11838 | 409 | |
| 37166 | 410 | lemma Pair_fst_snd_eq: "s = t \<longleftrightarrow> fst s = fst t \<and> snd s = snd t" | 
| 411 | by (cases s, cases t) simp | |
| 412 | ||
| 413 | lemma prod_eqI [intro?]: "fst p = fst q \<Longrightarrow> snd p = snd q \<Longrightarrow> p = q" | |
| 414 | by (simp add: Pair_fst_snd_eq) | |
| 415 | ||
| 416 | lemma split_conv [simp, code]: "split f (a, b) = f a b" | |
| 37591 | 417 | by (fact prod.cases) | 
| 37166 | 418 | |
| 419 | lemma splitI: "f a b \<Longrightarrow> split f (a, b)" | |
| 420 | by (rule split_conv [THEN iffD2]) | |
| 421 | ||
| 422 | lemma splitD: "split f (a, b) \<Longrightarrow> f a b" | |
| 423 | by (rule split_conv [THEN iffD1]) | |
| 424 | ||
| 425 | lemma split_Pair [simp]: "(\<lambda>(x, y). (x, y)) = id" | |
| 37591 | 426 | by (simp add: expand_fun_eq split: prod.split) | 
| 37166 | 427 | |
| 428 | lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f" | |
| 429 |   -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *}
 | |
| 37591 | 430 | by (simp add: expand_fun_eq split: prod.split) | 
| 37166 | 431 | |
| 432 | lemma split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)" | |
| 433 | by (cases x) simp | |
| 434 | ||
| 435 | lemma split_twice: "split f (split g p) = split (\<lambda>x y. split f (g x y)) p" | |
| 436 | by (cases p) simp | |
| 437 | ||
| 438 | lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))" | |
| 37591 | 439 | by (simp add: prod_case_unfold) | 
| 37166 | 440 | |
| 441 | lemma split_weak_cong: "p = q \<Longrightarrow> split c p = split c q" | |
| 442 |   -- {* Prevents simplification of @{term c}: much faster *}
 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 443 | by (fact weak_case_cong) | 
| 37166 | 444 | |
| 445 | lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" | |
| 446 | by (simp add: split_eta) | |
| 447 | ||
| 11838 | 448 | lemma split_paired_all: "(!!x. PROP P x) == (!!a b. PROP P (a, b))" | 
| 11820 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 449 | proof | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 450 | fix a b | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 451 | assume "!!x. PROP P x" | 
| 19535 | 452 | then show "PROP P (a, b)" . | 
| 11820 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 453 | next | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 454 | fix x | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 455 | assume "!!a b. PROP P (a, b)" | 
| 19535 | 456 | from `PROP P (fst x, snd x)` show "PROP P x" by simp | 
| 11820 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 457 | qed | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 458 | |
| 11838 | 459 | text {*
 | 
| 460 |   The rule @{thm [source] split_paired_all} does not work with the
 | |
| 461 | Simplifier because it also affects premises in congrence rules, | |
| 462 |   where this can lead to premises of the form @{text "!!a b. ... =
 | |
| 463 | ?P(a, b)"} which cannot be solved by reflexivity. | |
| 464 | *} | |
| 465 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 466 | lemmas split_tupled_all = split_paired_all unit_all_eq2 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 467 | |
| 26480 | 468 | ML {*
 | 
| 11838 | 469 | (* replace parameters of product type by individual component parameters *) | 
| 470 | val safe_full_simp_tac = generic_simp_tac true (true, false, false); | |
| 471 | local (* filtering with exists_paired_all is an essential optimization *) | |
| 16121 | 472 |     fun exists_paired_all (Const ("all", _) $ Abs (_, T, t)) =
 | 
| 11838 | 473 | can HOLogic.dest_prodT T orelse exists_paired_all t | 
| 474 | | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u | |
| 475 | | exists_paired_all (Abs (_, _, t)) = exists_paired_all t | |
| 476 | | exists_paired_all _ = false; | |
| 477 | val ss = HOL_basic_ss | |
| 26340 | 478 |       addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}]
 | 
| 11838 | 479 | addsimprocs [unit_eq_proc]; | 
| 480 | in | |
| 481 | val split_all_tac = SUBGOAL (fn (t, i) => | |
| 482 | if exists_paired_all t then safe_full_simp_tac ss i else no_tac); | |
| 483 | val unsafe_split_all_tac = SUBGOAL (fn (t, i) => | |
| 484 | if exists_paired_all t then full_simp_tac ss i else no_tac); | |
| 485 | fun split_all th = | |
| 26340 | 486 | if exists_paired_all (Thm.prop_of th) then full_simplify ss th else th; | 
| 11838 | 487 | end; | 
| 26340 | 488 | *} | 
| 11838 | 489 | |
| 26340 | 490 | declaration {* fn _ =>
 | 
| 491 |   Classical.map_cs (fn cs => cs addSbefore ("split_all_tac", split_all_tac))
 | |
| 16121 | 492 | *} | 
| 11838 | 493 | |
| 494 | lemma split_paired_All [simp]: "(ALL x. P x) = (ALL a b. P (a, b))" | |
| 495 |   -- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *}
 | |
| 496 | by fast | |
| 497 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 498 | lemma split_paired_Ex [simp]: "(EX x. P x) = (EX a b. P (a, b))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 499 | by fast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 500 | |
| 11838 | 501 | lemma split_paired_The: "(THE x. P x) = (THE (a, b). P (a, b))" | 
| 502 |   -- {* Can't be added to simpset: loops! *}
 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 503 | by (simp add: split_eta) | 
| 11838 | 504 | |
| 505 | text {*
 | |
| 506 |   Simplification procedure for @{thm [source] cond_split_eta}.  Using
 | |
| 507 |   @{thm [source] split_eta} as a rewrite rule is not general enough,
 | |
| 508 |   and using @{thm [source] cond_split_eta} directly would render some
 | |
| 509 |   existing proofs very inefficient; similarly for @{text
 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 510 | split_beta}. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 511 | *} | 
| 11838 | 512 | |
| 26480 | 513 | ML {*
 | 
| 11838 | 514 | local | 
| 35364 | 515 |   val cond_split_eta_ss = HOL_basic_ss addsimps @{thms cond_split_eta};
 | 
| 516 | fun Pair_pat k 0 (Bound m) = (m = k) | |
| 517 |     | Pair_pat k i (Const (@{const_name Pair},  _) $ Bound m $ t) =
 | |
| 518 | i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t | |
| 519 | | Pair_pat _ _ _ = false; | |
| 520 | fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t | |
| 521 | | no_args k i (t $ u) = no_args k i t andalso no_args k i u | |
| 522 | | no_args k i (Bound m) = m < k orelse m > k + i | |
| 523 | | no_args _ _ _ = true; | |
| 524 | fun split_pat tp i (Abs (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE | |
| 37591 | 525 |     | split_pat tp i (Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
 | 
| 35364 | 526 | | split_pat tp i _ = NONE; | 
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19656diff
changeset | 527 | fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] [] | 
| 35364 | 528 | (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))) | 
| 18328 | 529 | (K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1))); | 
| 11838 | 530 | |
| 35364 | 531 | fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t | 
| 532 | | beta_term_pat k i (t $ u) = | |
| 533 | Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u) | |
| 534 | | beta_term_pat k i t = no_args k i t; | |
| 535 | fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg | |
| 536 | | eta_term_pat _ _ _ = false; | |
| 11838 | 537 | fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) | 
| 35364 | 538 | | subst arg k i (t $ u) = | 
| 539 | if Pair_pat k i (t $ u) then incr_boundvars k arg | |
| 540 | else (subst arg k i t $ subst arg k i u) | |
| 541 | | subst arg k i t = t; | |
| 37591 | 542 |   fun beta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t) $ arg) =
 | 
| 11838 | 543 | (case split_pat beta_term_pat 1 t of | 
| 35364 | 544 | SOME (i, f) => SOME (metaeq ss s (subst arg 0 i f)) | 
| 15531 | 545 | | NONE => NONE) | 
| 35364 | 546 | | beta_proc _ _ = NONE; | 
| 37591 | 547 |   fun eta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t)) =
 | 
| 11838 | 548 | (case split_pat eta_term_pat 1 t of | 
| 35364 | 549 | SOME (_, ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end)) | 
| 15531 | 550 | | NONE => NONE) | 
| 35364 | 551 | | eta_proc _ _ = NONE; | 
| 11838 | 552 | in | 
| 32010 | 553 |   val split_beta_proc = Simplifier.simproc @{theory} "split_beta" ["split f z"] (K beta_proc);
 | 
| 554 |   val split_eta_proc = Simplifier.simproc @{theory} "split_eta" ["split f"] (K eta_proc);
 | |
| 11838 | 555 | end; | 
| 556 | ||
| 557 | Addsimprocs [split_beta_proc, split_eta_proc]; | |
| 558 | *} | |
| 559 | ||
| 26798 
a9134a089106
split_beta is now declared as monotonicity rule, to allow bounded
 berghofe parents: 
26588diff
changeset | 560 | lemma split_beta [mono]: "(%(x, y). P x y) z = P (fst z) (snd z)" | 
| 11838 | 561 | by (subst surjective_pairing, rule split_conv) | 
| 562 | ||
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 563 | lemma split_split [no_atp]: "R(split c p) = (ALL x y. p = (x, y) --> R(c x y))" | 
| 11838 | 564 |   -- {* For use with @{text split} and the Simplifier. *}
 | 
| 15481 | 565 | by (insert surj_pair [of p], clarify, simp) | 
| 11838 | 566 | |
| 567 | text {*
 | |
| 568 |   @{thm [source] split_split} could be declared as @{text "[split]"}
 | |
| 569 | done after the Splitter has been speeded up significantly; | |
| 570 | precompute the constants involved and don't do anything unless the | |
| 571 | current goal contains one of those constants. | |
| 572 | *} | |
| 573 | ||
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 574 | lemma split_split_asm [no_atp]: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))" | 
| 14208 | 575 | by (subst split_split, simp) | 
| 11838 | 576 | |
| 577 | text {*
 | |
| 578 |   \medskip @{term split} used as a logical connective or set former.
 | |
| 579 | ||
| 580 |   \medskip These rules are for use with @{text blast}; could instead
 | |
| 581 |   call @{text simp} using @{thm [source] split} as rewrite. *}
 | |
| 582 | ||
| 583 | lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p" | |
| 584 | apply (simp only: split_tupled_all) | |
| 585 | apply (simp (no_asm_simp)) | |
| 586 | done | |
| 587 | ||
| 588 | lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x" | |
| 589 | apply (simp only: split_tupled_all) | |
| 590 | apply (simp (no_asm_simp)) | |
| 591 | done | |
| 592 | ||
| 593 | lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" | |
| 37591 | 594 | by (induct p) auto | 
| 11838 | 595 | |
| 596 | lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" | |
| 37591 | 597 | by (induct p) auto | 
| 11838 | 598 | |
| 599 | lemma splitE2: | |
| 600 | "[| Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R" | |
| 601 | proof - | |
| 602 | assume q: "Q (split P z)" | |
| 603 | assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R" | |
| 604 | show R | |
| 605 | apply (rule r surjective_pairing)+ | |
| 606 | apply (rule split_beta [THEN subst], rule q) | |
| 607 | done | |
| 608 | qed | |
| 609 | ||
| 610 | lemma splitD': "split R (a,b) c ==> R a b c" | |
| 611 | by simp | |
| 612 | ||
| 613 | lemma mem_splitI: "z: c a b ==> z: split c (a, b)" | |
| 614 | by simp | |
| 615 | ||
| 616 | lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p" | |
| 14208 | 617 | by (simp only: split_tupled_all, simp) | 
| 11838 | 618 | |
| 18372 | 619 | lemma mem_splitE: | 
| 37166 | 620 | assumes major: "z \<in> split c p" | 
| 621 | and cases: "\<And>x y. p = (x, y) \<Longrightarrow> z \<in> c x y \<Longrightarrow> Q" | |
| 18372 | 622 | shows Q | 
| 37591 | 623 | by (rule major [unfolded prod_case_unfold] cases surjective_pairing)+ | 
| 11838 | 624 | |
| 625 | declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!] | |
| 626 | declare mem_splitE [elim!] splitE' [elim!] splitE [elim!] | |
| 627 | ||
| 26340 | 628 | ML {*
 | 
| 11838 | 629 | local (* filtering with exists_p_split is an essential optimization *) | 
| 37591 | 630 |   fun exists_p_split (Const (@{const_name prod_case},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
 | 
| 11838 | 631 | | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u | 
| 632 | | exists_p_split (Abs (_, _, t)) = exists_p_split t | |
| 633 | | exists_p_split _ = false; | |
| 35364 | 634 |   val ss = HOL_basic_ss addsimps @{thms split_conv};
 | 
| 11838 | 635 | in | 
| 636 | val split_conv_tac = SUBGOAL (fn (t, i) => | |
| 637 | if exists_p_split t then safe_full_simp_tac ss i else no_tac); | |
| 638 | end; | |
| 26340 | 639 | *} | 
| 640 | ||
| 11838 | 641 | (* This prevents applications of splitE for already splitted arguments leading | 
| 642 | to quite time-consuming computations (in particular for nested tuples) *) | |
| 26340 | 643 | declaration {* fn _ =>
 | 
| 644 |   Classical.map_cs (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac))
 | |
| 16121 | 645 | *} | 
| 11838 | 646 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 647 | lemma split_eta_SetCompr [simp,no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" | 
| 18372 | 648 | by (rule ext) fast | 
| 11838 | 649 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 650 | lemma split_eta_SetCompr2 [simp,no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P" | 
| 18372 | 651 | by (rule ext) fast | 
| 11838 | 652 | |
| 653 | lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)" | |
| 654 |   -- {* Allows simplifications of nested splits in case of independent predicates. *}
 | |
| 18372 | 655 | by (rule ext) blast | 
| 11838 | 656 | |
| 14337 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 657 | (* Do NOT make this a simp rule as it | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 658 | a) only helps in special situations | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 659 | b) can lead to nontermination in the presence of split_def | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 660 | *) | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 661 | lemma split_comp_eq: | 
| 20415 | 662 | fixes f :: "'a => 'b => 'c" and g :: "'d => 'a" | 
| 663 | shows "(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))" | |
| 18372 | 664 | by (rule ext) auto | 
| 14101 | 665 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 666 | lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 667 | apply (rule_tac x = "(a, b)" in image_eqI) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 668 | apply auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 669 | done | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 670 | |
| 11838 | 671 | lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)" | 
| 672 | by blast | |
| 673 | ||
| 674 | (* | |
| 675 | the following would be slightly more general, | |
| 676 | but cannot be used as rewrite rule: | |
| 677 | ### Cannot add premise as rewrite rule because it contains (type) unknowns: | |
| 678 | ### ?y = .x | |
| 679 | Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" | |
| 14208 | 680 | by (rtac some_equality 1) | 
| 681 | by ( Simp_tac 1) | |
| 682 | by (split_all_tac 1) | |
| 683 | by (Asm_full_simp_tac 1) | |
| 11838 | 684 | qed "The_split_eq"; | 
| 685 | *) | |
| 686 | ||
| 687 | text {*
 | |
| 688 |   Setup of internal @{text split_rule}.
 | |
| 689 | *} | |
| 690 | ||
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 691 | lemmas prod_caseI = prod.cases [THEN iffD2, standard] | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 692 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 693 | lemma prod_caseI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> prod_case c p" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 694 | by (fact splitI2) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 695 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 696 | lemma prod_caseI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> prod_case c p x" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 697 | by (fact splitI2') | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 698 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 699 | lemma prod_caseE: "prod_case c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 700 | by (fact splitE) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 701 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 702 | lemma prod_caseE': "prod_case c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 703 | by (fact splitE') | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 704 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 705 | declare prod_caseI [intro!] | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 706 | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
25885diff
changeset | 707 | lemma prod_case_beta: | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
25885diff
changeset | 708 | "prod_case f p = f (fst p) (snd p)" | 
| 37591 | 709 | by (fact split_beta) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
25885diff
changeset | 710 | |
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 711 | lemma prod_cases3 [cases type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 712 | obtains (fields) a b c where "y = (a, b, c)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 713 | by (cases y, case_tac b) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 714 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 715 | lemma prod_induct3 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 716 | "(!!a b c. P (a, b, c)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 717 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 718 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 719 | lemma prod_cases4 [cases type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 720 | obtains (fields) a b c d where "y = (a, b, c, d)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 721 | by (cases y, case_tac c) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 722 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 723 | lemma prod_induct4 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 724 | "(!!a b c d. P (a, b, c, d)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 725 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 726 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 727 | lemma prod_cases5 [cases type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 728 | obtains (fields) a b c d e where "y = (a, b, c, d, e)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 729 | by (cases y, case_tac d) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 730 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 731 | lemma prod_induct5 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 732 | "(!!a b c d e. P (a, b, c, d, e)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 733 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 734 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 735 | lemma prod_cases6 [cases type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 736 | obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 737 | by (cases y, case_tac e) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 738 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 739 | lemma prod_induct6 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 740 | "(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 741 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 742 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 743 | lemma prod_cases7 [cases type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 744 | obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 745 | by (cases y, case_tac f) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 746 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 747 | lemma prod_induct7 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 748 | "(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 749 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 750 | |
| 37166 | 751 | lemma split_def: | 
| 752 | "split = (\<lambda>c p. c (fst p) (snd p))" | |
| 37591 | 753 | by (fact prod_case_unfold) | 
| 37166 | 754 | |
| 755 | definition internal_split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
 | |
| 756 | "internal_split == split" | |
| 757 | ||
| 758 | lemma internal_split_conv: "internal_split c (a, b) = c a b" | |
| 759 | by (simp only: internal_split_def split_conv) | |
| 760 | ||
| 761 | use "Tools/split_rule.ML" | |
| 762 | setup Split_Rule.setup | |
| 763 | ||
| 764 | hide_const internal_split | |
| 765 | ||
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 766 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 767 | subsubsection {* Derived operations *}
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 768 | |
| 37387 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
37278diff
changeset | 769 | definition curry    :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c" where
 | 
| 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
37278diff
changeset | 770 | "curry = (\<lambda>c x y. c (x, y))" | 
| 37166 | 771 | |
| 772 | lemma curry_conv [simp, code]: "curry f a b = f (a, b)" | |
| 773 | by (simp add: curry_def) | |
| 774 | ||
| 775 | lemma curryI [intro!]: "f (a, b) \<Longrightarrow> curry f a b" | |
| 776 | by (simp add: curry_def) | |
| 777 | ||
| 778 | lemma curryD [dest!]: "curry f a b \<Longrightarrow> f (a, b)" | |
| 779 | by (simp add: curry_def) | |
| 780 | ||
| 781 | lemma curryE: "curry f a b \<Longrightarrow> (f (a, b) \<Longrightarrow> Q) \<Longrightarrow> Q" | |
| 782 | by (simp add: curry_def) | |
| 783 | ||
| 784 | lemma curry_split [simp]: "curry (split f) = f" | |
| 785 | by (simp add: curry_def split_def) | |
| 786 | ||
| 787 | lemma split_curry [simp]: "split (curry f) = f" | |
| 788 | by (simp add: curry_def split_def) | |
| 789 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 790 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 791 | The composition-uncurry combinator. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 792 | *} | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 793 | |
| 37751 | 794 | notation fcomp (infixl "\<circ>>" 60) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 795 | |
| 37751 | 796 | definition scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd" (infixl "\<circ>\<rightarrow>" 60) where
 | 
| 797 | "f \<circ>\<rightarrow> g = (\<lambda>x. prod_case g (f x))" | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 798 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 799 | lemma scomp_unfold: "scomp = (\<lambda>f g x. g (fst (f x)) (snd (f x)))" | 
| 37751 | 800 | by (simp add: expand_fun_eq scomp_def prod_case_unfold) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 801 | |
| 37751 | 802 | lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = prod_case g (f x)" | 
| 803 | by (simp add: scomp_unfold prod_case_unfold) | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 804 | |
| 37751 | 805 | lemma Pair_scomp: "Pair x \<circ>\<rightarrow> f = f x" | 
| 26588 
d83271bfaba5
removed syntax from monad combinators; renamed mbind to scomp
 haftmann parents: 
26480diff
changeset | 806 | by (simp add: expand_fun_eq scomp_apply) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 807 | |
| 37751 | 808 | lemma scomp_Pair: "x \<circ>\<rightarrow> Pair = x" | 
| 26588 
d83271bfaba5
removed syntax from monad combinators; renamed mbind to scomp
 haftmann parents: 
26480diff
changeset | 809 | by (simp add: expand_fun_eq scomp_apply) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 810 | |
| 37751 | 811 | lemma scomp_scomp: "(f \<circ>\<rightarrow> g) \<circ>\<rightarrow> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>\<rightarrow> h)" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 812 | by (simp add: expand_fun_eq scomp_unfold) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 813 | |
| 37751 | 814 | lemma scomp_fcomp: "(f \<circ>\<rightarrow> g) \<circ>> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>> h)" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 815 | by (simp add: expand_fun_eq scomp_unfold fcomp_def) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 816 | |
| 37751 | 817 | lemma fcomp_scomp: "(f \<circ>> g) \<circ>\<rightarrow> h = f \<circ>> (g \<circ>\<rightarrow> h)" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 818 | by (simp add: expand_fun_eq scomp_unfold fcomp_apply) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 819 | |
| 31202 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
30924diff
changeset | 820 | code_const scomp | 
| 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
30924diff
changeset | 821 | (Eval infixl 3 "#->") | 
| 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
30924diff
changeset | 822 | |
| 37751 | 823 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 824 | no_notation scomp (infixl "\<circ>\<rightarrow>" 60) | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 825 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 826 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 827 |   @{term prod_fun} --- action of the product functor upon
 | 
| 36664 
6302f9ad7047
repaired comments where SOMEthing went utterly wrong (cf. 2b04504fcb69)
 krauss parents: 
36622diff
changeset | 828 | functions. | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 829 | *} | 
| 21195 | 830 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 831 | definition prod_fun :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
 | 
| 37765 | 832 | "prod_fun f g = (\<lambda>(x, y). (f x, g y))" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 833 | |
| 28562 | 834 | lemma prod_fun [simp, code]: "prod_fun f g (a, b) = (f a, g b)" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 835 | by (simp add: prod_fun_def) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 836 | |
| 37278 | 837 | lemma fst_prod_fun[simp]: "fst (prod_fun f g x) = f (fst x)" | 
| 838 | by (cases x, auto) | |
| 839 | ||
| 840 | lemma snd_prod_fun[simp]: "snd (prod_fun f g x) = g (snd x)" | |
| 841 | by (cases x, auto) | |
| 842 | ||
| 843 | lemma fst_comp_prod_fun[simp]: "fst \<circ> prod_fun f g = f \<circ> fst" | |
| 844 | by (rule ext) auto | |
| 845 | ||
| 846 | lemma snd_comp_prod_fun[simp]: "snd \<circ> prod_fun f g = g \<circ> snd" | |
| 847 | by (rule ext) auto | |
| 848 | ||
| 849 | ||
| 850 | lemma prod_fun_compose: | |
| 851 | "prod_fun (f1 o f2) (g1 o g2) = (prod_fun f1 g1 o prod_fun f2 g2)" | |
| 852 | by (rule ext) auto | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 853 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 854 | lemma prod_fun_ident [simp]: "prod_fun (%x. x) (%y. y) = (%z. z)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 855 | by (rule ext) auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 856 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 857 | lemma prod_fun_imageI [intro]: "(a, b) : r ==> (f a, g b) : prod_fun f g ` r" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 858 | apply (rule image_eqI) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 859 | apply (rule prod_fun [symmetric], assumption) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 860 | done | 
| 21195 | 861 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 862 | lemma prod_fun_imageE [elim!]: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 863 | assumes major: "c: (prod_fun f g)`r" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 864 | and cases: "!!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 865 | shows P | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 866 | apply (rule major [THEN imageE]) | 
| 37166 | 867 | apply (case_tac x) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 868 | apply (rule cases) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 869 | apply (blast intro: prod_fun) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 870 | apply blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 871 | done | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 872 | |
| 37278 | 873 | |
| 37166 | 874 | definition apfst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b" where
 | 
| 875 | "apfst f = prod_fun f id" | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 876 | |
| 37166 | 877 | definition apsnd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c" where
 | 
| 878 | "apsnd f = prod_fun id f" | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 879 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 880 | lemma apfst_conv [simp, code]: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 881 | "apfst f (x, y) = (f x, y)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 882 | by (simp add: apfst_def) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 883 | |
| 33638 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 884 | lemma apsnd_conv [simp, code]: | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 885 | "apsnd f (x, y) = (x, f y)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 886 | by (simp add: apsnd_def) | 
| 21195 | 887 | |
| 33594 | 888 | lemma fst_apfst [simp]: | 
| 889 | "fst (apfst f x) = f (fst x)" | |
| 890 | by (cases x) simp | |
| 891 | ||
| 892 | lemma fst_apsnd [simp]: | |
| 893 | "fst (apsnd f x) = fst x" | |
| 894 | by (cases x) simp | |
| 895 | ||
| 896 | lemma snd_apfst [simp]: | |
| 897 | "snd (apfst f x) = snd x" | |
| 898 | by (cases x) simp | |
| 899 | ||
| 900 | lemma snd_apsnd [simp]: | |
| 901 | "snd (apsnd f x) = f (snd x)" | |
| 902 | by (cases x) simp | |
| 903 | ||
| 904 | lemma apfst_compose: | |
| 905 | "apfst f (apfst g x) = apfst (f \<circ> g) x" | |
| 906 | by (cases x) simp | |
| 907 | ||
| 908 | lemma apsnd_compose: | |
| 909 | "apsnd f (apsnd g x) = apsnd (f \<circ> g) x" | |
| 910 | by (cases x) simp | |
| 911 | ||
| 912 | lemma apfst_apsnd [simp]: | |
| 913 | "apfst f (apsnd g x) = (f (fst x), g (snd x))" | |
| 914 | by (cases x) simp | |
| 915 | ||
| 916 | lemma apsnd_apfst [simp]: | |
| 917 | "apsnd f (apfst g x) = (g (fst x), f (snd x))" | |
| 918 | by (cases x) simp | |
| 919 | ||
| 920 | lemma apfst_id [simp] : | |
| 921 | "apfst id = id" | |
| 922 | by (simp add: expand_fun_eq) | |
| 923 | ||
| 924 | lemma apsnd_id [simp] : | |
| 925 | "apsnd id = id" | |
| 926 | by (simp add: expand_fun_eq) | |
| 927 | ||
| 928 | lemma apfst_eq_conv [simp]: | |
| 929 | "apfst f x = apfst g x \<longleftrightarrow> f (fst x) = g (fst x)" | |
| 930 | by (cases x) simp | |
| 931 | ||
| 932 | lemma apsnd_eq_conv [simp]: | |
| 933 | "apsnd f x = apsnd g x \<longleftrightarrow> f (snd x) = g (snd x)" | |
| 934 | by (cases x) simp | |
| 935 | ||
| 33638 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 936 | lemma apsnd_apfst_commute: | 
| 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 937 | "apsnd f (apfst g p) = apfst g (apsnd f p)" | 
| 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 938 | by simp | 
| 21195 | 939 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 940 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 941 | Disjoint union of a family of sets -- Sigma. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 942 | *} | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 943 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 944 | definition  Sigma :: "['a set, 'a => 'b set] => ('a \<times> 'b) set" where
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 945 |   Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 946 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 947 | abbreviation | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 948 |   Times :: "['a set, 'b set] => ('a * 'b) set"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 949 | (infixr "<*>" 80) where | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 950 | "A <*> B == Sigma A (%_. B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 951 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 952 | notation (xsymbols) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 953 | Times (infixr "\<times>" 80) | 
| 15394 | 954 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 955 | notation (HTML output) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 956 | Times (infixr "\<times>" 80) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 957 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 958 | syntax | 
| 35115 | 959 |   "_Sigma" :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"  ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
 | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 960 | translations | 
| 35115 | 961 | "SIGMA x:A. B" == "CONST Sigma A (%x. B)" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 962 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 963 | lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 964 | by (unfold Sigma_def) blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 965 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 966 | lemma SigmaE [elim!]: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 967 | "[| c: Sigma A B; | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 968 | !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 969 | |] ==> P" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 970 |   -- {* The general elimination rule. *}
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 971 | by (unfold Sigma_def) blast | 
| 20588 | 972 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 973 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 974 |   Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 975 | eigenvariables. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 976 | *} | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 977 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 978 | lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 979 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 980 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 981 | lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 982 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 983 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 984 | lemma SigmaE2: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 985 | "[| (a, b) : Sigma A B; | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 986 | [| a:A; b:B(a) |] ==> P | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 987 | |] ==> P" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 988 | by blast | 
| 20588 | 989 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 990 | lemma Sigma_cong: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 991 | "\<lbrakk>A = B; !!x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk> | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 992 | \<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 993 | by auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 994 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 995 | lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 996 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 997 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 998 | lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 999 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1000 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1001 | lemma Sigma_empty2 [simp]: "A <*> {} = {}"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1002 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1003 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1004 | lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1005 | by auto | 
| 21908 | 1006 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1007 | lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1008 | by auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1009 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1010 | lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1011 | by auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1012 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1013 | lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1014 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1015 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1016 | lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1017 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1018 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1019 | lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1020 | by (blast elim: equalityE) | 
| 20588 | 1021 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1022 | lemma SetCompr_Sigma_eq: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1023 | "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1024 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1025 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1026 | lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1027 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1028 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1029 | lemma UN_Times_distrib: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1030 | "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1031 |   -- {* Suggested by Pierre Chartier *}
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1032 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1033 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 1034 | lemma split_paired_Ball_Sigma [simp,no_atp]: | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1035 | "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1036 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1037 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 1038 | lemma split_paired_Bex_Sigma [simp,no_atp]: | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1039 | "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1040 | by blast | 
| 21908 | 1041 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1042 | lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1043 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1044 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1045 | lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1046 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1047 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1048 | lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1049 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1050 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1051 | lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1052 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1053 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1054 | lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1055 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1056 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1057 | lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1058 | by blast | 
| 21908 | 1059 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1060 | lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1061 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1062 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1063 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1064 | Non-dependent versions are needed to avoid the need for higher-order | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1065 | matching, especially when the rules are re-oriented. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1066 | *} | 
| 21908 | 1067 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1068 | lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)" | 
| 28719 | 1069 | by blast | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1070 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1071 | lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)" | 
| 28719 | 1072 | by blast | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1073 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1074 | lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)" | 
| 28719 | 1075 | by blast | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1076 | |
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1077 | lemma Times_empty[simp]: "A \<times> B = {} \<longleftrightarrow> A = {} \<or> B = {}"
 | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1078 | by auto | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1079 | |
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1080 | lemma fst_image_times[simp]: "fst ` (A \<times> B) = (if B = {} then {} else A)"
 | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1081 | by (auto intro!: image_eqI) | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1082 | |
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1083 | lemma snd_image_times[simp]: "snd ` (A \<times> B) = (if A = {} then {} else B)"
 | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1084 | by (auto intro!: image_eqI) | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1085 | |
| 28719 | 1086 | lemma insert_times_insert[simp]: | 
| 1087 | "insert a A \<times> insert b B = | |
| 1088 | insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)" | |
| 1089 | by blast | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1090 | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: 
33089diff
changeset | 1091 | lemma vimage_Times: "f -` (A \<times> B) = ((fst \<circ> f) -` A) \<inter> ((snd \<circ> f) -` B)" | 
| 37166 | 1092 | by (auto, case_tac "f x", auto) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: 
33089diff
changeset | 1093 | |
| 37278 | 1094 | text{* The following @{const prod_fun} lemmas are due to Joachim Breitner: *}
 | 
| 1095 | ||
| 1096 | lemma prod_fun_inj_on: | |
| 1097 | assumes "inj_on f A" and "inj_on g B" | |
| 1098 | shows "inj_on (prod_fun f g) (A \<times> B)" | |
| 1099 | proof (rule inj_onI) | |
| 1100 | fix x :: "'a \<times> 'c" and y :: "'a \<times> 'c" | |
| 1101 | assume "x \<in> A \<times> B" hence "fst x \<in> A" and "snd x \<in> B" by auto | |
| 1102 | assume "y \<in> A \<times> B" hence "fst y \<in> A" and "snd y \<in> B" by auto | |
| 1103 | assume "prod_fun f g x = prod_fun f g y" | |
| 1104 | hence "fst (prod_fun f g x) = fst (prod_fun f g y)" by (auto) | |
| 1105 | hence "f (fst x) = f (fst y)" by (cases x,cases y,auto) | |
| 1106 | with `inj_on f A` and `fst x \<in> A` and `fst y \<in> A` | |
| 1107 | have "fst x = fst y" by (auto dest:dest:inj_onD) | |
| 1108 | moreover from `prod_fun f g x = prod_fun f g y` | |
| 1109 | have "snd (prod_fun f g x) = snd (prod_fun f g y)" by (auto) | |
| 1110 | hence "g (snd x) = g (snd y)" by (cases x,cases y,auto) | |
| 1111 | with `inj_on g B` and `snd x \<in> B` and `snd y \<in> B` | |
| 1112 | have "snd x = snd y" by (auto dest:dest:inj_onD) | |
| 1113 | ultimately show "x = y" by(rule prod_eqI) | |
| 1114 | qed | |
| 1115 | ||
| 1116 | lemma prod_fun_surj: | |
| 1117 | assumes "surj f" and "surj g" | |
| 1118 | shows "surj (prod_fun f g)" | |
| 1119 | unfolding surj_def | |
| 1120 | proof | |
| 1121 | fix y :: "'b \<times> 'd" | |
| 1122 | from `surj f` obtain a where "fst y = f a" by (auto elim:surjE) | |
| 1123 | moreover | |
| 1124 | from `surj g` obtain b where "snd y = g b" by (auto elim:surjE) | |
| 1125 | ultimately have "(fst y, snd y) = prod_fun f g (a,b)" by auto | |
| 1126 | thus "\<exists>x. y = prod_fun f g x" by auto | |
| 1127 | qed | |
| 1128 | ||
| 1129 | lemma prod_fun_surj_on: | |
| 1130 | assumes "f ` A = A'" and "g ` B = B'" | |
| 1131 | shows "prod_fun f g ` (A \<times> B) = A' \<times> B'" | |
| 1132 | unfolding image_def | |
| 1133 | proof(rule set_ext,rule iffI) | |
| 1134 | fix x :: "'a \<times> 'c" | |
| 1135 |   assume "x \<in> {y\<Colon>'a \<times> 'c. \<exists>x\<Colon>'b \<times> 'd\<in>A \<times> B. y = prod_fun f g x}"
 | |
| 1136 | then obtain y where "y \<in> A \<times> B" and "x = prod_fun f g y" by blast | |
| 1137 | from `image f A = A'` and `y \<in> A \<times> B` have "f (fst y) \<in> A'" by auto | |
| 1138 | moreover from `image g B = B'` and `y \<in> A \<times> B` have "g (snd y) \<in> B'" by auto | |
| 1139 | ultimately have "(f (fst y), g (snd y)) \<in> (A' \<times> B')" by auto | |
| 1140 | with `x = prod_fun f g y` show "x \<in> A' \<times> B'" by (cases y, auto) | |
| 1141 | next | |
| 1142 | fix x :: "'a \<times> 'c" | |
| 1143 | assume "x \<in> A' \<times> B'" hence "fst x \<in> A'" and "snd x \<in> B'" by auto | |
| 1144 | from `image f A = A'` and `fst x \<in> A'` have "fst x \<in> image f A" by auto | |
| 1145 | then obtain a where "a \<in> A" and "fst x = f a" by (rule imageE) | |
| 1146 | moreover from `image g B = B'` and `snd x \<in> B'` | |
| 1147 | obtain b where "b \<in> B" and "snd x = g b" by auto | |
| 1148 | ultimately have "(fst x, snd x) = prod_fun f g (a,b)" by auto | |
| 1149 | moreover from `a \<in> A` and `b \<in> B` have "(a , b) \<in> A \<times> B" by auto | |
| 1150 | ultimately have "\<exists>y \<in> A \<times> B. x = prod_fun f g y" by auto | |
| 1151 |   thus "x \<in> {x. \<exists>y \<in> A \<times> B. x = prod_fun f g y}" by auto
 | |
| 1152 | qed | |
| 1153 | ||
| 35822 | 1154 | lemma swap_inj_on: | 
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1155 | "inj_on (\<lambda>(i, j). (j, i)) A" | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1156 | by (auto intro!: inj_onI) | 
| 35822 | 1157 | |
| 1158 | lemma swap_product: | |
| 1159 | "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A" | |
| 1160 | by (simp add: split_def image_def) blast | |
| 1161 | ||
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1162 | lemma image_split_eq_Sigma: | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1163 |   "(\<lambda>x. (f x, g x)) ` A = Sigma (f ` A) (\<lambda>x. g ` (f -` {x} \<inter> A))"
 | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1164 | proof (safe intro!: imageI vimageI) | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1165 | fix a b assume *: "a \<in> A" "b \<in> A" and eq: "f a = f b" | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1166 | show "(f b, g a) \<in> (\<lambda>x. (f x, g x)) ` A" | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1167 | using * eq[symmetric] by auto | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1168 | qed simp_all | 
| 35822 | 1169 | |
| 21908 | 1170 | |
| 37166 | 1171 | subsection {* Inductively defined sets *}
 | 
| 15394 | 1172 | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 1173 | use "Tools/inductive_codegen.ML" | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 1174 | setup Inductive_Codegen.setup | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 1175 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31667diff
changeset | 1176 | use "Tools/inductive_set.ML" | 
| 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31667diff
changeset | 1177 | setup Inductive_Set.setup | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 1178 | |
| 37166 | 1179 | |
| 1180 | subsection {* Legacy theorem bindings and duplicates *}
 | |
| 1181 | ||
| 1182 | lemma PairE: | |
| 1183 | obtains x y where "p = (x, y)" | |
| 1184 | by (fact prod.exhaust) | |
| 1185 | ||
| 1186 | lemma Pair_inject: | |
| 1187 | assumes "(a, b) = (a', b')" | |
| 1188 | and "a = a' ==> b = b' ==> R" | |
| 1189 | shows R | |
| 1190 | using assms by simp | |
| 1191 | ||
| 1192 | lemmas Pair_eq = prod.inject | |
| 1193 | ||
| 1194 | lemmas split = split_conv  -- {* for backwards compatibility *}
 | |
| 1195 | ||
| 10213 | 1196 | end |