src/HOL/Product_Type.thy
author hoelzl
Fri, 16 Nov 2012 18:45:57 +0100
changeset 50104 de19856feb54
parent 49897 cc69be3c8f87
child 50107 289181e3e524
permissions -rw-r--r--
move theorems to be more generally useable
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Product_Type.thy
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     3
    Copyright   1992  University of Cambridge
11777
wenzelm
parents: 11602
diff changeset
     4
*)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     5
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
     6
header {* Cartesian products *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14952
diff changeset
     8
theory Product_Type
33959
2afc55e8ed27 bootstrap datatype_rep_proofs in Datatype.thy (avoids unchecked dynamic name references)
haftmann
parents: 33638
diff changeset
     9
imports Typedef Inductive Fun
46950
d0181abdbdac declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents: 46630
diff changeset
    10
keywords "inductive_set" "coinductive_set" :: thy_decl
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14952
diff changeset
    11
begin
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    12
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    13
subsection {* @{typ bool} is a datatype *}
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    14
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
    15
rep_datatype True False by (auto intro: bool_induct)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    16
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    17
declare case_split [cases type: bool]
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    18
  -- "prefer plain propositional version"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    19
28346
b8390cd56b8f discontinued special treatment of op = vs. eq_class.eq
haftmann
parents: 28262
diff changeset
    20
lemma
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    21
  shows [code]: "HOL.equal False P \<longleftrightarrow> \<not> P"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    22
    and [code]: "HOL.equal True P \<longleftrightarrow> P" 
46630
3abc964cdc45 tuned whitespace
haftmann
parents: 46557
diff changeset
    23
    and [code]: "HOL.equal P False \<longleftrightarrow> \<not> P"
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    24
    and [code]: "HOL.equal P True \<longleftrightarrow> P"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    25
    and [code nbe]: "HOL.equal P P \<longleftrightarrow> True"
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    26
  by (simp_all add: equal)
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    27
43654
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    28
lemma If_case_cert:
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    29
  assumes "CASE \<equiv> (\<lambda>b. If b f g)"
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    30
  shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)"
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    31
  using assms by simp_all
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    32
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    33
setup {*
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    34
  Code.add_case @{thm If_case_cert}
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    35
*}
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    36
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    37
code_const "HOL.equal \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool"
39272
0b61951d2682 Haskell == is infix, not infixl
haftmann
parents: 39198
diff changeset
    38
  (Haskell infix 4 "==")
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    39
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    40
code_instance bool :: equal
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    41
  (Haskell -)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    42
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    43
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
    44
subsection {* The @{text unit} type *}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    45
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 49764
diff changeset
    46
typedef unit = "{True}"
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45662
diff changeset
    47
  by auto
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    48
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45662
diff changeset
    49
definition Unity :: unit  ("'(')")
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45662
diff changeset
    50
  where "() = Abs_unit True"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    51
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35427
diff changeset
    52
lemma unit_eq [no_atp]: "u = ()"
40590
b994d855dbd2 typedef (open) unit
huffman
parents: 39302
diff changeset
    53
  by (induct u) (simp add: Unity_def)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    54
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    55
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    56
  Simplification procedure for @{thm [source] unit_eq}.  Cannot use
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    57
  this rule directly --- it loops!
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    58
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    59
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    60
simproc_setup unit_eq ("x::unit") = {*
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    61
  fn _ => fn _ => fn ct =>
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    62
    if HOLogic.is_unit (term_of ct) then NONE
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    63
    else SOME (mk_meta_eq @{thm unit_eq})
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    64
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    65
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26975
diff changeset
    66
rep_datatype "()" by simp
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    67
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    68
lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    69
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    70
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    71
lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    72
  by (rule triv_forall_equality)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    73
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    74
text {*
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    75
  This rewrite counters the effect of simproc @{text unit_eq} on @{term
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    76
  [source] "%u::unit. f u"}, replacing it by @{term [source]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    77
  f} rather than by @{term [source] "%u. f ()"}.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    78
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    79
43866
8a50dc70cbff moving UNIV = ... equations to their proper theories
haftmann
parents: 43654
diff changeset
    80
lemma unit_abs_eta_conv [simp, no_atp]: "(%u::unit. f ()) = f"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    81
  by (rule ext) simp
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    82
43866
8a50dc70cbff moving UNIV = ... equations to their proper theories
haftmann
parents: 43654
diff changeset
    83
lemma UNIV_unit [no_atp]:
8a50dc70cbff moving UNIV = ... equations to their proper theories
haftmann
parents: 43654
diff changeset
    84
  "UNIV = {()}" by auto
8a50dc70cbff moving UNIV = ... equations to their proper theories
haftmann
parents: 43654
diff changeset
    85
30924
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    86
instantiation unit :: default
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    87
begin
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    88
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    89
definition "default = ()"
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    90
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    91
instance ..
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    92
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
    93
end
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    94
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    95
lemma [code]:
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    96
  "HOL.equal (u\<Colon>unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    97
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    98
code_type unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    99
  (SML "unit")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   100
  (OCaml "unit")
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   101
  (Haskell "()")
34886
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   102
  (Scala "Unit")
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   103
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   104
code_const Unity
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   105
  (SML "()")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   106
  (OCaml "()")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   107
  (Haskell "()")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   108
  (Scala "()")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   109
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
   110
code_instance unit :: equal
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   111
  (Haskell -)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   112
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
   113
code_const "HOL.equal \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool"
39272
0b61951d2682 Haskell == is infix, not infixl
haftmann
parents: 39198
diff changeset
   114
  (Haskell infix 4 "==")
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   115
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   116
code_reserved SML
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   117
  unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   118
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   119
code_reserved OCaml
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   120
  unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   121
34886
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   122
code_reserved Scala
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   123
  Unit
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   124
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   125
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   126
subsection {* The product type *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   127
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   128
subsubsection {* Type definition *}
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   129
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   130
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: 26340
diff changeset
   131
  "Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   132
45696
476ad865f125 prefer typedef without alternative name;
wenzelm
parents: 45694
diff changeset
   133
definition "prod = {f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
476ad865f125 prefer typedef without alternative name;
wenzelm
parents: 45694
diff changeset
   134
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 49764
diff changeset
   135
typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \<Rightarrow> 'b \<Rightarrow> bool) set"
45696
476ad865f125 prefer typedef without alternative name;
wenzelm
parents: 45694
diff changeset
   136
  unfolding prod_def by auto
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   137
35427
ad039d29e01c proper (type_)notation;
wenzelm
parents: 35365
diff changeset
   138
type_notation (xsymbols)
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   139
  "prod"  ("(_ \<times>/ _)" [21, 20] 20)
35427
ad039d29e01c proper (type_)notation;
wenzelm
parents: 35365
diff changeset
   140
type_notation (HTML output)
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   141
  "prod"  ("(_ \<times>/ _)" [21, 20] 20)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   142
37389
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   143
definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b" where
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   144
  "Pair a b = Abs_prod (Pair_Rep a b)"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   145
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   146
rep_datatype Pair proof -
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   147
  fix P :: "'a \<times> 'b \<Rightarrow> bool" and p
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   148
  assume "\<And>a b. P (Pair a b)"
37389
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   149
  then show "P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   150
next
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   151
  fix a c :: 'a and b d :: 'b
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   152
  have "Pair_Rep a b = Pair_Rep c d \<longleftrightarrow> a = c \<and> b = d"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   153
    by (auto simp add: Pair_Rep_def fun_eq_iff)
37389
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   154
  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: 37387
diff changeset
   155
    by (auto simp add: prod_def)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   156
  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: 37387
diff changeset
   157
    by (simp add: Pair_def Abs_prod_inject)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   158
qed
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   159
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   160
declare prod.simps(2) [nitpick_simp del]
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   161
40929
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40702
diff changeset
   162
declare prod.weak_case_cong [cong del]
37411
c88c44156083 removed simplifier congruence rule of "prod_case"
haftmann
parents: 37389
diff changeset
   163
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   164
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   165
subsubsection {* Tuple syntax *}
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   166
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   167
abbreviation (input) split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   168
  "split \<equiv> prod_case"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   169
11777
wenzelm
parents: 11602
diff changeset
   170
text {*
wenzelm
parents: 11602
diff changeset
   171
  Patterns -- extends pre-defined type @{typ pttrn} used in
wenzelm
parents: 11602
diff changeset
   172
  abstractions.
wenzelm
parents: 11602
diff changeset
   173
*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   174
41229
d797baa3d57c replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents: 40968
diff changeset
   175
nonterminal tuple_args and patterns
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   176
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   177
syntax
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   178
  "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   179
  "_tuple_arg"  :: "'a => tuple_args"                   ("_")
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   180
  "_tuple_args" :: "'a => tuple_args => tuple_args"     ("_,/ _")
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   181
  "_pattern"    :: "[pttrn, patterns] => pttrn"         ("'(_,/ _')")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   182
  ""            :: "pttrn => patterns"                  ("_")
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   183
  "_patterns"   :: "[pttrn, patterns] => patterns"      ("_,/ _")
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   184
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   185
translations
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   186
  "(x, y)" == "CONST Pair x y"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   187
  "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   188
  "%(x, y, zs). b" == "CONST prod_case (%x (y, zs). b)"
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   189
  "%(x, y). b" == "CONST prod_case (%x y. b)"
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   190
  "_abs (CONST Pair x y) t" => "%(x, y). t"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   191
  -- {* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   192
     The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   193
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   194
(*reconstruct pattern from (nested) splits, avoiding eta-contraction of body;
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   195
  works best with enclosing "let", if "let" does not avoid eta-contraction*)
14359
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   196
print_translation {*
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   197
let
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   198
  fun split_tr' [Abs (x, T, t as (Abs abs))] =
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   199
        (* split (%x y. t) => %(x,y) t *)
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   200
        let
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   201
          val (y, t') = Syntax_Trans.atomic_abs_tr' abs;
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   202
          val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t');
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   203
        in
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   204
          Syntax.const @{syntax_const "_abs"} $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   205
            (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   206
        end
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   207
    | split_tr' [Abs (x, T, (s as Const (@{const_syntax prod_case}, _) $ t))] =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   208
        (* split (%x. (split (%y z. t))) => %(x,y,z). t *)
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   209
        let
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   210
          val Const (@{syntax_const "_abs"}, _) $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   211
            (Const (@{syntax_const "_pattern"}, _) $ y $ z) $ t' = split_tr' [t];
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   212
          val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t');
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   213
        in
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   214
          Syntax.const @{syntax_const "_abs"} $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   215
            (Syntax.const @{syntax_const "_pattern"} $ x' $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   216
              (Syntax.const @{syntax_const "_patterns"} $ y $ z)) $ t''
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   217
        end
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   218
    | split_tr' [Const (@{const_syntax prod_case}, _) $ t] =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   219
        (* split (split (%x y z. t)) => %((x, y), z). t *)
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   220
        split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *)
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   221
    | split_tr' [Const (@{syntax_const "_abs"}, _) $ x_y $ Abs abs] =
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   222
        (* split (%pttrn z. t) => %(pttrn,z). t *)
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   223
        let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   224
          Syntax.const @{syntax_const "_abs"} $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   225
            (Syntax.const @{syntax_const "_pattern"} $ x_y $ z) $ t
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   226
        end
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   227
    | split_tr' _ = raise Match;
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   228
in [(@{const_syntax prod_case}, split_tr')] end
14359
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   229
*}
3d9948163018 Added print translation for pairs
schirmer
parents: 14337
diff changeset
   230
15422
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   231
(* 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: 15404
diff changeset
   232
typed_print_translation {*
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   233
let
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42083
diff changeset
   234
  fun split_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42083
diff changeset
   235
    | 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: 15404
diff changeset
   236
        (case (head_of t) of
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   237
          Const (@{const_syntax prod_case}, _) => raise Match
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   238
        | _ =>
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   239
          let 
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   240
            val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   241
            val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0);
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   242
            val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t');
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   243
          in
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   244
            Syntax.const @{syntax_const "_abs"} $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   245
              (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   246
          end)
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42083
diff changeset
   247
    | split_guess_names_tr' T [t] =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   248
        (case head_of t of
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   249
          Const (@{const_syntax prod_case}, _) => raise Match
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   250
        | _ =>
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   251
          let
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   252
            val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   253
            val (y, t') =
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   254
              Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0);
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   255
            val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t');
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   256
          in
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   257
            Syntax.const @{syntax_const "_abs"} $
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   258
              (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   259
          end)
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42083
diff changeset
   260
    | split_guess_names_tr' _ _ = raise Match;
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   261
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: 15404
diff changeset
   262
*}
cbdddc0efadf added print translation for split: split f --> %(x,y). f x y
schirmer
parents: 15404
diff changeset
   263
42059
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   264
(* Force eta-contraction for terms of the form "Q A (%p. prod_case P p)"
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   265
   where Q is some bounded quantifier or set operator.
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   266
   Reason: the above prints as "Q p : A. case p of (x,y) \<Rightarrow> P x y"
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   267
   whereas we want "Q (x,y):A. P x y".
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   268
   Otherwise prevent eta-contraction.
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   269
*)
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   270
print_translation {*
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   271
let
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   272
  fun contract Q f ts =
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   273
    case ts of
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   274
      [A, Abs(_, _, (s as Const (@{const_syntax prod_case},_) $ t) $ Bound 0)]
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 42059
diff changeset
   275
      => if Term.is_dependent t then f ts else Syntax.const Q $ A $ s
42059
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   276
    | _ => f ts;
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   277
  fun contract2 (Q,f) = (Q, contract Q f);
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   278
  val pairs =
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   279
    [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Ball} @{syntax_const "_Ball"},
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   280
     Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Bex} @{syntax_const "_Bex"},
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   281
     Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42247
diff changeset
   282
     Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
42059
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   283
in map contract2 pairs end
83f3dc509068 fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
nipkow
parents: 41792
diff changeset
   284
*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   285
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   286
subsubsection {* Code generator setup *}
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   287
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   288
code_type prod
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   289
  (SML infix 2 "*")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   290
  (OCaml infix 2 "*")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   291
  (Haskell "!((_),/ (_))")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   292
  (Scala "((_),/ (_))")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   293
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   294
code_const Pair
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   295
  (SML "!((_),/ (_))")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   296
  (OCaml "!((_),/ (_))")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   297
  (Haskell "!((_),/ (_))")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   298
  (Scala "!((_),/ (_))")
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   299
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
   300
code_instance prod :: equal
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   301
  (Haskell -)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   302
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
   303
code_const "HOL.equal \<Colon> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
39272
0b61951d2682 Haskell == is infix, not infixl
haftmann
parents: 39198
diff changeset
   304
  (Haskell infix 4 "==")
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   305
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   306
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   307
subsubsection {* Fundamental operations and properties *}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   308
49897
cc69be3c8f87 moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
bulwahn
parents: 49834
diff changeset
   309
lemma Pair_inject:
cc69be3c8f87 moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
bulwahn
parents: 49834
diff changeset
   310
  assumes "(a, b) = (a', b')"
cc69be3c8f87 moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
bulwahn
parents: 49834
diff changeset
   311
    and "a = a' ==> b = b' ==> R"
cc69be3c8f87 moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
bulwahn
parents: 49834
diff changeset
   312
  shows R
cc69be3c8f87 moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
bulwahn
parents: 49834
diff changeset
   313
  using assms by simp
cc69be3c8f87 moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
bulwahn
parents: 49834
diff changeset
   314
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   315
lemma surj_pair [simp]: "EX x y. p = (x, y)"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   316
  by (cases p) simp
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   317
37389
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   318
definition fst :: "'a \<times> 'b \<Rightarrow> 'a" where
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   319
  "fst p = (case p of (a, b) \<Rightarrow> a)"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   320
37389
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   321
definition snd :: "'a \<times> 'b \<Rightarrow> 'b" where
09467cdfa198 qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents: 37387
diff changeset
   322
  "snd p = (case p of (a, b) \<Rightarrow> b)"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   323
22886
cdff6ef76009 moved recfun_codegen.ML to Code_Generator.thy
haftmann
parents: 22838
diff changeset
   324
lemma fst_conv [simp, code]: "fst (a, b) = a"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   325
  unfolding fst_def by simp
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   326
22886
cdff6ef76009 moved recfun_codegen.ML to Code_Generator.thy
haftmann
parents: 22838
diff changeset
   327
lemma snd_conv [simp, code]: "snd (a, b) = b"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   328
  unfolding snd_def by simp
11025
a70b796d9af8 converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents: 10289
diff changeset
   329
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   330
code_const fst and snd
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   331
  (Haskell "fst" and "snd")
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   332
41792
ff3cb0c418b7 renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents: 41505
diff changeset
   333
lemma prod_case_unfold [nitpick_unfold]: "prod_case = (%c p. c (fst p) (snd p))"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   334
  by (simp add: fun_eq_iff split: prod.split)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   335
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   336
lemma fst_eqD: "fst (x, y) = a ==> x = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   337
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   338
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   339
lemma snd_eqD: "snd (x, y) = a ==> y = a"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   340
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   341
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   342
lemma pair_collapse [simp]: "(fst p, snd p) = p"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   343
  by (cases p) simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   344
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   345
lemmas surjective_pairing = pair_collapse [symmetric]
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   346
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 43866
diff changeset
   347
lemma prod_eq_iff: "s = t \<longleftrightarrow> fst s = fst t \<and> snd s = snd t"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   348
  by (cases s, cases t) simp
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   349
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   350
lemma prod_eqI [intro?]: "fst p = fst q \<Longrightarrow> snd p = snd q \<Longrightarrow> p = q"
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 43866
diff changeset
   351
  by (simp add: prod_eq_iff)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   352
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   353
lemma split_conv [simp, code]: "split f (a, b) = f a b"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   354
  by (fact prod.cases)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   355
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   356
lemma splitI: "f a b \<Longrightarrow> split f (a, b)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   357
  by (rule split_conv [THEN iffD2])
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   358
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   359
lemma splitD: "split f (a, b) \<Longrightarrow> f a b"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   360
  by (rule split_conv [THEN iffD1])
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   361
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   362
lemma split_Pair [simp]: "(\<lambda>(x, y). (x, y)) = id"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   363
  by (simp add: fun_eq_iff split: prod.split)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   364
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   365
lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   366
  -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *}
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   367
  by (simp add: fun_eq_iff split: prod.split)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   368
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   369
lemma split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   370
  by (cases x) simp
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   371
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   372
lemma split_twice: "split f (split g p) = split (\<lambda>x y. split f (g x y)) p"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   373
  by (cases p) simp
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   374
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   375
lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   376
  by (simp add: prod_case_unfold)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   377
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   378
lemma split_weak_cong: "p = q \<Longrightarrow> split c p = split c q"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   379
  -- {* Prevents simplification of @{term c}: much faster *}
40929
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40702
diff changeset
   380
  by (fact prod.weak_case_cong)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   381
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   382
lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   383
  by (simp add: split_eta)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   384
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   385
lemma split_paired_all [no_atp]: "(!!x. PROP P x) == (!!a b. PROP P (a, b))"
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   386
proof
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   387
  fix a b
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   388
  assume "!!x. PROP P x"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   389
  then show "PROP P (a, b)" .
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   390
next
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   391
  fix x
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   392
  assume "!!a b. PROP P (a, b)"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   393
  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: 11777
diff changeset
   394
qed
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   395
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   396
lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   397
  by (cases x) simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   398
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   399
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   400
  The rule @{thm [source] split_paired_all} does not work with the
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   401
  Simplifier because it also affects premises in congrence rules,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   402
  where this can lead to premises of the form @{text "!!a b. ... =
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   403
  ?P(a, b)"} which cannot be solved by reflexivity.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   404
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   405
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   406
lemmas split_tupled_all = split_paired_all unit_all_eq2
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   407
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26358
diff changeset
   408
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   409
  (* replace parameters of product type by individual component parameters *)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   410
  val safe_full_simp_tac = generic_simp_tac true (true, false, false);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   411
  local (* filtering with exists_paired_all is an essential optimization *)
16121
wenzelm
parents: 15570
diff changeset
   412
    fun exists_paired_all (Const ("all", _) $ Abs (_, T, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   413
          can HOLogic.dest_prodT T orelse exists_paired_all t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   414
      | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   415
      | exists_paired_all (Abs (_, _, t)) = exists_paired_all t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   416
      | exists_paired_all _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   417
    val ss = HOL_basic_ss
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   418
      addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}]
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
   419
      addsimprocs [@{simproc unit_eq}];
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   420
  in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   421
    val split_all_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   422
      if exists_paired_all t then safe_full_simp_tac ss i else no_tac);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   423
    val unsafe_split_all_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   424
      if exists_paired_all t then full_simp_tac ss i else no_tac);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   425
    fun split_all th =
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   426
   if exists_paired_all (Thm.prop_of th) then full_simplify ss th else th;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   427
  end;
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   428
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   429
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   430
declaration {* fn _ =>
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   431
  Classical.map_cs (fn cs => cs addSbefore ("split_all_tac", split_all_tac))
16121
wenzelm
parents: 15570
diff changeset
   432
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   433
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   434
lemma split_paired_All [simp, no_atp]: "(ALL x. P x) = (ALL a b. P (a, b))"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   435
  -- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   436
  by fast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   437
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   438
lemma split_paired_Ex [simp, no_atp]: "(EX x. P x) = (EX a b. P (a, b))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   439
  by fast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   440
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   441
lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   442
  -- {* Can't be added to simpset: loops! *}
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   443
  by (simp add: split_eta)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   444
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   445
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   446
  Simplification procedure for @{thm [source] cond_split_eta}.  Using
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   447
  @{thm [source] split_eta} as a rewrite rule is not general enough,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   448
  and using @{thm [source] cond_split_eta} directly would render some
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   449
  existing proofs very inefficient; similarly for @{text
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   450
  split_beta}.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   451
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   452
26480
544cef16045b replaced 'ML_setup' by 'ML';
wenzelm
parents: 26358
diff changeset
   453
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   454
local
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   455
  val cond_split_eta_ss = HOL_basic_ss addsimps @{thms cond_split_eta};
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   456
  fun Pair_pat k 0 (Bound m) = (m = k)
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   457
    | Pair_pat k i (Const (@{const_name Pair},  _) $ Bound m $ t) =
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   458
        i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   459
    | Pair_pat _ _ _ = false;
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   460
  fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   461
    | no_args k i (t $ u) = no_args k i t andalso no_args k i u
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   462
    | no_args k i (Bound m) = m < k orelse m > k + i
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   463
    | no_args _ _ _ = true;
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   464
  fun split_pat tp i (Abs  (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   465
    | split_pat tp i (Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   466
    | split_pat tp i _ = NONE;
20044
92cc2f4c7335 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19656
diff changeset
   467
  fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] []
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   468
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
18328
841261f303a1 simprocs: static evaluation of simpset;
wenzelm
parents: 18220
diff changeset
   469
        (K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1)));
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   470
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   471
  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   472
    | beta_term_pat k i (t $ u) =
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   473
        Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u)
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   474
    | beta_term_pat k i t = no_args k i t;
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   475
  fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   476
    | eta_term_pat _ _ _ = false;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   477
  fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t)
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   478
    | subst arg k i (t $ u) =
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   479
        if Pair_pat k i (t $ u) then incr_boundvars k arg
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   480
        else (subst arg k i t $ subst arg k i u)
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   481
    | subst arg k i t = t;
43595
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 43594
diff changeset
   482
in
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   483
  fun beta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t) $ arg) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   484
        (case split_pat beta_term_pat 1 t of
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   485
          SOME (i, f) => SOME (metaeq ss s (subst arg 0 i f))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   486
        | NONE => NONE)
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   487
    | beta_proc _ _ = NONE;
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   488
  fun eta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   489
        (case split_pat eta_term_pat 1 t of
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   490
          SOME (_, ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   491
        | NONE => NONE)
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   492
    | eta_proc _ _ = NONE;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   493
end;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   494
*}
43595
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 43594
diff changeset
   495
simproc_setup split_beta ("split f z") = {* fn _ => fn ss => fn ct => beta_proc ss (term_of ct) *}
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 43594
diff changeset
   496
simproc_setup split_eta ("split f") = {* fn _ => fn ss => fn ct => eta_proc ss (term_of ct) *}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   497
26798
a9134a089106 split_beta is now declared as monotonicity rule, to allow bounded
berghofe
parents: 26588
diff changeset
   498
lemma split_beta [mono]: "(%(x, y). P x y) z = P (fst z) (snd z)"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   499
  by (subst surjective_pairing, rule split_conv)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   500
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   501
lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   502
  by (auto simp: fun_eq_iff)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   503
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   504
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35427
diff changeset
   505
lemma split_split [no_atp]: "R(split c p) = (ALL x y. p = (x, y) --> R(c x y))"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   506
  -- {* For use with @{text split} and the Simplifier. *}
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15422
diff changeset
   507
  by (insert surj_pair [of p], clarify, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   508
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   509
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   510
  @{thm [source] split_split} could be declared as @{text "[split]"}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   511
  done after the Splitter has been speeded up significantly;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   512
  precompute the constants involved and don't do anything unless the
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   513
  current goal contains one of those constants.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   514
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   515
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35427
diff changeset
   516
lemma split_split_asm [no_atp]: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   517
by (subst split_split, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   518
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   519
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   520
  \medskip @{term split} used as a logical connective or set former.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   521
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   522
  \medskip These rules are for use with @{text blast}; could instead
40929
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40702
diff changeset
   523
  call @{text simp} using @{thm [source] prod.split} as rewrite. *}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   524
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   525
lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   526
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   527
  apply (simp (no_asm_simp))
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   528
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   529
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   530
lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   531
  apply (simp only: split_tupled_all)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   532
  apply (simp (no_asm_simp))
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   533
  done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   534
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   535
lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   536
  by (induct p) auto
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   537
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   538
lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   539
  by (induct p) auto
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   540
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   541
lemma splitE2:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   542
  "[| Q (split P z);  !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   543
proof -
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   544
  assume q: "Q (split P z)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   545
  assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   546
  show R
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   547
    apply (rule r surjective_pairing)+
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   548
    apply (rule split_beta [THEN subst], rule q)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   549
    done
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   550
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   551
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   552
lemma splitD': "split R (a,b) c ==> R a b c"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   553
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   554
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   555
lemma mem_splitI: "z: c a b ==> z: split c (a, b)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   556
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   557
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   558
lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   559
by (simp only: split_tupled_all, simp)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   560
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   561
lemma mem_splitE:
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   562
  assumes major: "z \<in> split c p"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   563
    and cases: "\<And>x y. p = (x, y) \<Longrightarrow> z \<in> c x y \<Longrightarrow> Q"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   564
  shows Q
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   565
  by (rule major [unfolded prod_case_unfold] cases surjective_pairing)+
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   566
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   567
declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   568
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!]
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   569
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   570
ML {*
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   571
local (* filtering with exists_p_split is an essential optimization *)
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   572
  fun exists_p_split (Const (@{const_name prod_case},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   573
    | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   574
    | exists_p_split (Abs (_, _, t)) = exists_p_split t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   575
    | exists_p_split _ = false;
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   576
  val ss = HOL_basic_ss addsimps @{thms split_conv};
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   577
in
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   578
val split_conv_tac = SUBGOAL (fn (t, i) =>
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   579
    if exists_p_split t then safe_full_simp_tac ss i else no_tac);
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   580
end;
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   581
*}
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   582
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   583
(* This prevents applications of splitE for already splitted arguments leading
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   584
   to quite time-consuming computations (in particular for nested tuples) *)
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   585
declaration {* fn _ =>
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   586
  Classical.map_cs (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac))
16121
wenzelm
parents: 15570
diff changeset
   587
*}
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   588
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35427
diff changeset
   589
lemma split_eta_SetCompr [simp,no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   590
  by (rule ext) fast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   591
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35427
diff changeset
   592
lemma split_eta_SetCompr2 [simp,no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   593
  by (rule ext) fast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   594
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   595
lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   596
  -- {* Allows simplifications of nested splits in case of independent predicates. *}
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   597
  by (rule ext) blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   598
14337
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   599
(* 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: 14208
diff changeset
   600
   a) only helps in special situations
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   601
   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: 14208
diff changeset
   602
*)
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   603
lemma split_comp_eq: 
20415
e3d2d7b01279 explicit type variables prevent empty sorts
paulson
parents: 20380
diff changeset
   604
  fixes f :: "'a => 'b => 'c" and g :: "'d => 'a"
e3d2d7b01279 explicit type variables prevent empty sorts
paulson
parents: 20380
diff changeset
   605
  shows "(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   606
  by (rule ext) auto
14101
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   607
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   608
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: 26340
diff changeset
   609
  apply (rule_tac x = "(a, b)" in image_eqI)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   610
   apply auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   611
  done
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   612
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   613
lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)"
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   614
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   615
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   616
(*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   617
the following  would be slightly more general,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   618
but cannot be used as rewrite rule:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   619
### Cannot add premise as rewrite rule because it contains (type) unknowns:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   620
### ?y = .x
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   621
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)"
14208
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   622
by (rtac some_equality 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   623
by ( Simp_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   624
by (split_all_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   625
by (Asm_full_simp_tac 1)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   626
qed "The_split_eq";
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   627
*)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   628
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   629
text {*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   630
  Setup of internal @{text split_rule}.
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   631
*}
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   632
45607
16b4f5774621 eliminated obsolete "standard";
wenzelm
parents: 45205
diff changeset
   633
lemmas prod_caseI = prod.cases [THEN iffD2]
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   634
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   635
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: 37591
diff changeset
   636
  by (fact splitI2)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   637
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   638
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: 37591
diff changeset
   639
  by (fact splitI2')
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   640
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   641
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: 37591
diff changeset
   642
  by (fact splitE)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   643
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   644
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: 37591
diff changeset
   645
  by (fact splitE')
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   646
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   647
declare prod_caseI [intro!]
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   648
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   649
lemma prod_case_beta:
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   650
  "prod_case f p = f (fst p) (snd p)"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   651
  by (fact split_beta)
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   652
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   653
lemma prod_cases3 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   654
  obtains (fields) a b c where "y = (a, b, c)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   655
  by (cases y, case_tac b) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   656
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   657
lemma prod_induct3 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   658
    "(!!a b c. P (a, b, c)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   659
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   660
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   661
lemma prod_cases4 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   662
  obtains (fields) a b c d where "y = (a, b, c, d)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   663
  by (cases y, case_tac c) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   664
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   665
lemma prod_induct4 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   666
    "(!!a b c d. P (a, b, c, d)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   667
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   668
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   669
lemma prod_cases5 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   670
  obtains (fields) a b c d e where "y = (a, b, c, d, e)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   671
  by (cases y, case_tac d) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   672
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   673
lemma prod_induct5 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   674
    "(!!a b c d e. P (a, b, c, d, e)) ==> P x"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   675
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   676
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   677
lemma prod_cases6 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   678
  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: 24286
diff changeset
   679
  by (cases y, case_tac e) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   680
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   681
lemma prod_induct6 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   682
    "(!!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: 24286
diff changeset
   683
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   684
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   685
lemma prod_cases7 [cases type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   686
  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: 24286
diff changeset
   687
  by (cases y, case_tac f) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   688
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   689
lemma prod_induct7 [case_names fields, induct type]:
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   690
    "(!!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: 24286
diff changeset
   691
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   692
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   693
lemma split_def:
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   694
  "split = (\<lambda>c p. c (fst p) (snd p))"
37591
d3daea901123 merged constants "split" and "prod_case"
haftmann
parents: 37411
diff changeset
   695
  by (fact prod_case_unfold)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   696
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   697
definition internal_split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   698
  "internal_split == split"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   699
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   700
lemma internal_split_conv: "internal_split c (a, b) = c a b"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   701
  by (simp only: internal_split_def split_conv)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   702
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47988
diff changeset
   703
ML_file "Tools/split_rule.ML"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   704
setup Split_Rule.setup
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   705
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   706
hide_const internal_split
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   707
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   708
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   709
subsubsection {* Derived operations *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   710
37387
3581483cca6c qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
haftmann
parents: 37278
diff changeset
   711
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: 37278
diff changeset
   712
  "curry = (\<lambda>c x y. c (x, y))"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   713
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   714
lemma curry_conv [simp, code]: "curry f a b = f (a, b)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   715
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   716
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   717
lemma curryI [intro!]: "f (a, b) \<Longrightarrow> curry f a b"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   718
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   719
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   720
lemma curryD [dest!]: "curry f a b \<Longrightarrow> f (a, b)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   721
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   722
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   723
lemma curryE: "curry f a b \<Longrightarrow> (f (a, b) \<Longrightarrow> Q) \<Longrightarrow> Q"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   724
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   725
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   726
lemma curry_split [simp]: "curry (split f) = f"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   727
  by (simp add: curry_def split_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   728
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   729
lemma split_curry [simp]: "split (curry f) = f"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   730
  by (simp add: curry_def split_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   731
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   732
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   733
  The composition-uncurry combinator.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   734
*}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   735
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   736
notation fcomp (infixl "\<circ>>" 60)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   737
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   738
definition scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd" (infixl "\<circ>\<rightarrow>" 60) where
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   739
  "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: 26340
diff changeset
   740
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   741
lemma scomp_unfold: "scomp = (\<lambda>f g x. g (fst (f x)) (snd (f x)))"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   742
  by (simp add: fun_eq_iff scomp_def prod_case_unfold)
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   743
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   744
lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = prod_case g (f x)"
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   745
  by (simp add: scomp_unfold prod_case_unfold)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   746
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   747
lemma Pair_scomp: "Pair x \<circ>\<rightarrow> f = f x"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   748
  by (simp add: fun_eq_iff)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   749
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   750
lemma scomp_Pair: "x \<circ>\<rightarrow> Pair = x"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   751
  by (simp add: fun_eq_iff)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   752
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   753
lemma scomp_scomp: "(f \<circ>\<rightarrow> g) \<circ>\<rightarrow> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>\<rightarrow> h)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   754
  by (simp add: fun_eq_iff scomp_unfold)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   755
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   756
lemma scomp_fcomp: "(f \<circ>\<rightarrow> g) \<circ>> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>> h)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   757
  by (simp add: fun_eq_iff scomp_unfold fcomp_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   758
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   759
lemma fcomp_scomp: "(f \<circ>> g) \<circ>\<rightarrow> h = f \<circ>> (g \<circ>\<rightarrow> h)"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   760
  by (simp add: fun_eq_iff scomp_unfold)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   761
31202
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   762
code_const scomp
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   763
  (Eval infixl 3 "#->")
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   764
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   765
no_notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   766
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   767
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   768
text {*
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   769
  @{term map_pair} --- action of the product functor upon
36664
6302f9ad7047 repaired comments where SOMEthing went utterly wrong (cf. 2b04504fcb69)
krauss
parents: 36622
diff changeset
   770
  functions.
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   771
*}
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   772
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   773
definition map_pair :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   774
  "map_pair f g = (\<lambda>(x, y). (f x, g y))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   775
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   776
lemma map_pair_simp [simp, code]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   777
  "map_pair f g (a, b) = (f a, g b)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   778
  by (simp add: map_pair_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   779
41505
6d19301074cf "enriched_type" replaces less specific "type_lifting"
haftmann
parents: 41372
diff changeset
   780
enriched_type map_pair: map_pair
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   781
  by (auto simp add: split_paired_all)
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   782
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   783
lemma fst_map_pair [simp]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   784
  "fst (map_pair f g x) = f (fst x)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   785
  by (cases x) simp_all
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   786
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   787
lemma snd_prod_fun [simp]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   788
  "snd (map_pair f g x) = g (snd x)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   789
  by (cases x) simp_all
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   790
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   791
lemma fst_comp_map_pair [simp]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   792
  "fst \<circ> map_pair f g = f \<circ> fst"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   793
  by (rule ext) simp_all
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   794
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   795
lemma snd_comp_map_pair [simp]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   796
  "snd \<circ> map_pair f g = g \<circ> snd"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   797
  by (rule ext) simp_all
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   798
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   799
lemma map_pair_compose:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   800
  "map_pair (f1 o f2) (g1 o g2) = (map_pair f1 g1 o map_pair f2 g2)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   801
  by (rule ext) (simp add: map_pair.compositionality comp_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   802
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   803
lemma map_pair_ident [simp]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   804
  "map_pair (%x. x) (%y. y) = (%z. z)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   805
  by (rule ext) (simp add: map_pair.identity)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   806
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   807
lemma map_pair_imageI [intro]:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   808
  "(a, b) \<in> R \<Longrightarrow> (f a, g b) \<in> map_pair f g ` R"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   809
  by (rule image_eqI) simp_all
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   810
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   811
lemma prod_fun_imageE [elim!]:
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   812
  assumes major: "c \<in> map_pair f g ` R"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   813
    and cases: "\<And>x y. c = (f x, g y) \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> P"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   814
  shows P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   815
  apply (rule major [THEN imageE])
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   816
  apply (case_tac x)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   817
  apply (rule cases)
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   818
  apply simp_all
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   819
  done
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   820
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   821
definition apfst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b" where
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   822
  "apfst f = map_pair f id"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   823
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   824
definition apsnd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c" where
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   825
  "apsnd f = map_pair id f"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   826
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   827
lemma apfst_conv [simp, code]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   828
  "apfst f (x, y) = (f x, y)" 
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   829
  by (simp add: apfst_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   830
33638
548a34929e98 Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
hoelzl
parents: 33594
diff changeset
   831
lemma apsnd_conv [simp, code]:
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   832
  "apsnd f (x, y) = (x, f y)" 
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   833
  by (simp add: apsnd_def)
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   834
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   835
lemma fst_apfst [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   836
  "fst (apfst f x) = f (fst x)"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   837
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   838
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   839
lemma fst_apsnd [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   840
  "fst (apsnd f x) = fst x"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   841
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   842
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   843
lemma snd_apfst [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   844
  "snd (apfst f x) = snd x"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   845
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   846
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   847
lemma snd_apsnd [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   848
  "snd (apsnd f x) = f (snd x)"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   849
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   850
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   851
lemma apfst_compose:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   852
  "apfst f (apfst g x) = apfst (f \<circ> g) x"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   853
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   854
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   855
lemma apsnd_compose:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   856
  "apsnd f (apsnd g x) = apsnd (f \<circ> g) x"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   857
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   858
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   859
lemma apfst_apsnd [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   860
  "apfst f (apsnd g x) = (f (fst x), g (snd x))"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   861
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   862
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   863
lemma apsnd_apfst [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   864
  "apsnd f (apfst g x) = (g (fst x), f (snd x))"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   865
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   866
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   867
lemma apfst_id [simp] :
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   868
  "apfst id = id"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   869
  by (simp add: fun_eq_iff)
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   870
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   871
lemma apsnd_id [simp] :
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   872
  "apsnd id = id"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   873
  by (simp add: fun_eq_iff)
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   874
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   875
lemma apfst_eq_conv [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   876
  "apfst f x = apfst g x \<longleftrightarrow> f (fst x) = g (fst x)"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   877
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   878
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   879
lemma apsnd_eq_conv [simp]:
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   880
  "apsnd f x = apsnd g x \<longleftrightarrow> f (snd x) = g (snd x)"
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   881
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   882
33638
548a34929e98 Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
hoelzl
parents: 33594
diff changeset
   883
lemma apsnd_apfst_commute:
548a34929e98 Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
hoelzl
parents: 33594
diff changeset
   884
  "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: 33594
diff changeset
   885
  by simp
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   886
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   887
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   888
  Disjoint union of a family of sets -- Sigma.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   889
*}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   890
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45696
diff changeset
   891
definition Sigma :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<times> 'b) set" where
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   892
  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: 26340
diff changeset
   893
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   894
abbreviation
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45696
diff changeset
   895
  Times :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   896
    (infixr "<*>" 80) where
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   897
  "A <*> B == Sigma A (%_. B)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   898
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   899
notation (xsymbols)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   900
  Times  (infixr "\<times>" 80)
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   901
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   902
notation (HTML output)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   903
  Times  (infixr "\<times>" 80)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   904
45662
4f7c05990420 Hide Product_Type.Times - too precious an identifier
nipkow
parents: 45607
diff changeset
   905
hide_const (open) Times
4f7c05990420 Hide Product_Type.Times - too precious an identifier
nipkow
parents: 45607
diff changeset
   906
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   907
syntax
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   908
  "_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: 26340
diff changeset
   909
translations
35115
446c5063e4fd modernized translations;
wenzelm
parents: 34900
diff changeset
   910
  "SIGMA x:A. B" == "CONST Sigma A (%x. B)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   911
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   912
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: 26340
diff changeset
   913
  by (unfold Sigma_def) blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   914
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   915
lemma SigmaE [elim!]:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   916
    "[| c: Sigma A B;
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   917
        !!x y.[| x:A;  y:B(x);  c=(x,y) |] ==> P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   918
     |] ==> P"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   919
  -- {* The general elimination rule. *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   920
  by (unfold Sigma_def) blast
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   921
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   922
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   923
  Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   924
  eigenvariables.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   925
*}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   926
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   927
lemma SigmaD1: "(a, b) : Sigma A B ==> a : A"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   928
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   929
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   930
lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   931
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   932
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   933
lemma SigmaE2:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   934
    "[| (a, b) : Sigma A B;
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   935
        [| a:A;  b:B(a) |] ==> P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   936
     |] ==> P"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   937
  by blast
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   938
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   939
lemma Sigma_cong:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   940
     "\<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: 26340
diff changeset
   941
      \<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   942
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   943
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   944
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: 26340
diff changeset
   945
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   946
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   947
lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   948
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   949
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   950
lemma Sigma_empty2 [simp]: "A <*> {} = {}"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   951
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   952
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   953
lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   954
  by auto
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   955
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   956
lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   957
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   958
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   959
lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   960
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   961
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   962
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: 26340
diff changeset
   963
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   964
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   965
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: 26340
diff changeset
   966
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   967
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   968
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: 26340
diff changeset
   969
  by (blast elim: equalityE)
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   970
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   971
lemma SetCompr_Sigma_eq:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   972
    "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: 26340
diff changeset
   973
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   974
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   975
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: 26340
diff changeset
   976
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   977
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   978
lemma UN_Times_distrib:
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   979
  "(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: 26340
diff changeset
   980
  -- {* Suggested by Pierre Chartier *}
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   981
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   982
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   983
lemma split_paired_Ball_Sigma [simp, no_atp]:
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   984
    "(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: 26340
diff changeset
   985
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   986
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   987
lemma split_paired_Bex_Sigma [simp, no_atp]:
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   988
    "(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: 26340
diff changeset
   989
  by blast
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
   990
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   991
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: 26340
diff changeset
   992
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   993
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   994
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: 26340
diff changeset
   995
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   996
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   997
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: 26340
diff changeset
   998
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   999
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1000
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: 26340
diff changeset
  1001
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1002
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1003
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: 26340
diff changeset
  1004
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1005
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1006
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: 26340
diff changeset
  1007
  by blast
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1008
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1009
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: 26340
diff changeset
  1010
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1011
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1012
text {*
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1013
  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: 26340
diff changeset
  1014
  matching, especially when the rules are re-oriented.
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1015
*}
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1016
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1017
lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)"
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1018
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1019
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1020
lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)"
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1021
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1022
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1023
lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)"
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1024
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1025
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: 36176
diff changeset
  1026
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: 36176
diff changeset
  1027
  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: 36176
diff changeset
  1028
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1029
lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1030
  by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1031
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: 36176
diff changeset
  1032
lemma fst_image_times[simp]: "fst ` (A \<times> B) = (if B = {} then {} else A)"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
  1033
  by force
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: 36176
diff changeset
  1034
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: 36176
diff changeset
  1035
lemma snd_image_times[simp]: "snd ` (A \<times> B) = (if A = {} then {} else B)"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
  1036
  by force
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: 36176
diff changeset
  1037
28719
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1038
lemma insert_times_insert[simp]:
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1039
  "insert a A \<times> insert b B =
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1040
   insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)"
01e04e41cc7b added lemma
nipkow
parents: 28562
diff changeset
  1041
by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1042
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33089
diff changeset
  1043
lemma vimage_Times: "f -` (A \<times> B) = ((fst \<circ> f) -` A) \<inter> ((snd \<circ> f) -` B)"
47988
e4b69e10b990 tuned proofs;
wenzelm
parents: 47740
diff changeset
  1044
  apply auto
e4b69e10b990 tuned proofs;
wenzelm
parents: 47740
diff changeset
  1045
  apply (case_tac "f x")
e4b69e10b990 tuned proofs;
wenzelm
parents: 47740
diff changeset
  1046
  apply auto
e4b69e10b990 tuned proofs;
wenzelm
parents: 47740
diff changeset
  1047
  done
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33089
diff changeset
  1048
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1049
lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1050
  by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1051
35822
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1052
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: 36176
diff changeset
  1053
  "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: 36176
diff changeset
  1054
  by (auto intro!: inj_onI)
35822
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1055
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1056
lemma swap_product:
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1057
  "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1058
  by (simp add: split_def image_def) blast
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1059
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: 36176
diff changeset
  1060
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: 36176
diff changeset
  1061
  "(\<lambda>x. (f x, g x)) ` A = Sigma (f ` A) (\<lambda>x. g ` (f -` {x} \<inter> A))"
46128
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1062
proof (safe intro!: imageI)
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: 36176
diff changeset
  1063
  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: 36176
diff changeset
  1064
  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: 36176
diff changeset
  1065
    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: 36176
diff changeset
  1066
qed simp_all
35822
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1067
46128
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1068
definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1069
  [code_abbrev]: "product A B = A \<times> B"
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1070
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1071
hide_const (open) product
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1072
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1073
lemma member_product:
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1074
  "x \<in> Product_Type.product A B \<longleftrightarrow> x \<in> A \<times> B"
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1075
  by (simp add: product_def)
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1076
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1077
text {* The following @{const map_pair} lemmas are due to Joachim Breitner: *}
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1078
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1079
lemma map_pair_inj_on:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1080
  assumes "inj_on f A" and "inj_on g B"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1081
  shows "inj_on (map_pair f g) (A \<times> B)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1082
proof (rule inj_onI)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1083
  fix x :: "'a \<times> 'c" and y :: "'a \<times> 'c"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1084
  assume "x \<in> A \<times> B" hence "fst x \<in> A" and "snd x \<in> B" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1085
  assume "y \<in> A \<times> B" hence "fst y \<in> A" and "snd y \<in> B" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1086
  assume "map_pair f g x = map_pair f g y"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1087
  hence "fst (map_pair f g x) = fst (map_pair f g y)" by (auto)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1088
  hence "f (fst x) = f (fst y)" by (cases x,cases y,auto)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1089
  with `inj_on f A` and `fst x \<in> A` and `fst y \<in> A`
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1090
  have "fst x = fst y" by (auto dest:dest:inj_onD)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1091
  moreover from `map_pair f g x = map_pair f g y`
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1092
  have "snd (map_pair f g x) = snd (map_pair f g y)" by (auto)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1093
  hence "g (snd x) = g (snd y)" by (cases x,cases y,auto)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1094
  with `inj_on g B` and `snd x \<in> B` and `snd y \<in> B`
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1095
  have "snd x = snd y" by (auto dest:dest:inj_onD)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1096
  ultimately show "x = y" by(rule prod_eqI)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1097
qed
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1098
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1099
lemma map_pair_surj:
40702
cf26dd7395e4 Replace surj by abbreviation; remove surj_on.
hoelzl
parents: 40607
diff changeset
  1100
  fixes f :: "'a \<Rightarrow> 'b" and g :: "'c \<Rightarrow> 'd"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1101
  assumes "surj f" and "surj g"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1102
  shows "surj (map_pair f g)"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1103
unfolding surj_def
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1104
proof
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1105
  fix y :: "'b \<times> 'd"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1106
  from `surj f` obtain a where "fst y = f a" by (auto elim:surjE)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1107
  moreover
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1108
  from `surj g` obtain b where "snd y = g b" by (auto elim:surjE)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1109
  ultimately have "(fst y, snd y) = map_pair f g (a,b)" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1110
  thus "\<exists>x. y = map_pair f g x" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1111
qed
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1112
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1113
lemma map_pair_surj_on:
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1114
  assumes "f ` A = A'" and "g ` B = B'"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1115
  shows "map_pair f g ` (A \<times> B) = A' \<times> B'"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1116
unfolding image_def
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1117
proof(rule set_eqI,rule iffI)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1118
  fix x :: "'a \<times> 'c"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1119
  assume "x \<in> {y\<Colon>'a \<times> 'c. \<exists>x\<Colon>'b \<times> 'd\<in>A \<times> B. y = map_pair f g x}"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1120
  then obtain y where "y \<in> A \<times> B" and "x = map_pair f g y" by blast
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1121
  from `image f A = A'` and `y \<in> A \<times> B` have "f (fst y) \<in> A'" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1122
  moreover from `image g B = B'` and `y \<in> A \<times> B` have "g (snd y) \<in> B'" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1123
  ultimately have "(f (fst y), g (snd y)) \<in> (A' \<times> B')" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1124
  with `x = map_pair f g y` show "x \<in> A' \<times> B'" by (cases y, auto)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1125
next
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1126
  fix x :: "'a \<times> 'c"
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1127
  assume "x \<in> A' \<times> B'" hence "fst x \<in> A'" and "snd x \<in> B'" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1128
  from `image f A = A'` and `fst x \<in> A'` have "fst x \<in> image f A" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1129
  then obtain a where "a \<in> A" and "fst x = f a" by (rule imageE)
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1130
  moreover from `image g B = B'` and `snd x \<in> B'`
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1131
  obtain b where "b \<in> B" and "snd x = g b" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1132
  ultimately have "(fst x, snd x) = map_pair f g (a,b)" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1133
  moreover from `a \<in> A` and  `b \<in> B` have "(a , b) \<in> A \<times> B" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1134
  ultimately have "\<exists>y \<in> A \<times> B. x = map_pair f g y" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1135
  thus "x \<in> {x. \<exists>y \<in> A \<times> B. x = map_pair f g y}" by auto
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1136
qed
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1137
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1138
49764
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1139
subsection {* Simproc for rewriting a set comprehension into a pointfree expression *}
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1140
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1141
ML_file "Tools/set_comprehension_pointfree.ML"
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1142
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1143
setup {*
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1144
  Code_Preproc.map_pre (fn ss => ss addsimprocs
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1145
    [Raw_Simplifier.make_simproc {name = "set comprehension", lhss = [@{cpat "Collect ?P"}],
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1146
    proc = K Set_Comprehension_Pointfree.code_simproc, identifier = []}])
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1147
*}
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1148
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1149
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1150
subsection {* Inductively defined sets *}
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1151
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47988
diff changeset
  1152
ML_file "Tools/inductive_set.ML"
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31667
diff changeset
  1153
setup Inductive_Set.setup
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1154
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1155
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1156
subsection {* Legacy theorem bindings and duplicates *}
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1157
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1158
lemma PairE:
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1159
  obtains x y where "p = (x, y)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1160
  by (fact prod.exhaust)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1161
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1162
lemmas Pair_eq = prod.inject
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1163
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1164
lemmas split = split_conv  -- {* for backwards compatibility *}
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1165
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 43866
diff changeset
  1166
lemmas Pair_fst_snd_eq = prod_eq_iff
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 43866
diff changeset
  1167
45204
5e4a1270c000 hide typedef-generated constants Product_Type.prod and Sum_Type.sum
huffman
parents: 44921
diff changeset
  1168
hide_const (open) prod
5e4a1270c000 hide typedef-generated constants Product_Type.prod and Sum_Type.sum
huffman
parents: 44921
diff changeset
  1169
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
  1170
end