src/HOL/Product_Type.thy
author wenzelm
Tue, 05 Nov 2019 14:16:16 +0100
changeset 71046 b8aeeedf7e68
parent 70044 da5857dbcbb9
child 72581 de581f98a3a1
permissions -rw-r--r--
support for Linux user management;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58469
66ddc5ad4f63 corrected white-space accident
haftmann
parents: 58468
diff changeset
     1
(*  Title:      HOL/Product_Type.thy
10213
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
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
     6
section \<open>Cartesian products\<close>
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
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
     9
  imports Typedef Inductive Fun
69913
ca515cf61651 more specific keyword kinds;
wenzelm
parents: 69605
diff changeset
    10
  keywords "inductive_set" "coinductive_set" :: thy_defn
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
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
    13
subsection \<open>\<^typ>\<open>bool\<close> is a datatype\<close>
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    14
62594
75452e3eda14 generate theorems like 'bool.split_sel'
blanchet
parents: 62379
diff changeset
    15
free_constructors (discs_sels) case_bool for True | False
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 57983
diff changeset
    16
  by auto
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    17
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
    18
text \<open>Avoid name clashes by prefixing the output of \<open>old_rep_datatype\<close> with \<open>old\<close>.\<close>
55442
blanchet
parents: 55417
diff changeset
    19
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    20
setup \<open>Sign.mandatory_path "old"\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    21
58306
117ba6cbe414 renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents: 58292
diff changeset
    22
old_rep_datatype True False by (auto intro: bool_induct)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    23
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    24
setup \<open>Sign.parent_path\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    25
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
    26
text \<open>But erase the prefix for properties that are not generated by \<open>free_constructors\<close>.\<close>
55442
blanchet
parents: 55417
diff changeset
    27
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    28
setup \<open>Sign.mandatory_path "bool"\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    29
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    30
lemmas induct = old.bool.induct
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    31
lemmas inducts = old.bool.inducts
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55469
diff changeset
    32
lemmas rec = old.bool.rec
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55469
diff changeset
    33
lemmas simps = bool.distinct bool.case bool.rec
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    34
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    35
setup \<open>Sign.parent_path\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    36
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    37
declare case_split [cases type: bool]
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 66251
diff changeset
    38
  \<comment> \<open>prefer plain propositional version\<close>
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    39
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
    40
lemma [code]: "HOL.equal False P \<longleftrightarrow> \<not> P"
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
    41
  and [code]: "HOL.equal True P \<longleftrightarrow> P"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
    42
  and [code]: "HOL.equal P False \<longleftrightarrow> \<not> P"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
    43
  and [code]: "HOL.equal P True \<longleftrightarrow> P"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
    44
  and [code nbe]: "HOL.equal P P \<longleftrightarrow> True"
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38715
diff changeset
    45
  by (simp_all add: equal)
25534
d0b74fdd6067 simplified infrastructure for code generator operational equality
haftmann
parents: 25511
diff changeset
    46
43654
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    47
lemma If_case_cert:
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    48
  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
    49
  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
    50
  using assms by simp_all
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    51
66251
cd935b7cb3fb proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents: 63575
diff changeset
    52
setup \<open>Code.declare_case_global @{thm If_case_cert}\<close>
43654
3f1a44c2d645 install case certificate for If after code_datatype declaration for bool
haftmann
parents: 43595
diff changeset
    53
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
    54
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
    55
  constant "HOL.equal :: bool \<Rightarrow> bool \<Rightarrow> bool" \<rightharpoonup> (Haskell) infix 4 "=="
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
    56
| class_instance "bool" :: "equal" \<rightharpoonup> (Haskell) -
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    57
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
    58
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
    59
subsection \<open>The \<open>unit\<close> type\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    60
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 49764
diff changeset
    61
typedef unit = "{True}"
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45662
diff changeset
    62
  by auto
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    63
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45662
diff changeset
    64
definition Unity :: unit  ("'(')")
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 45662
diff changeset
    65
  where "() = Abs_unit True"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    66
35828
46cfc4b8112e now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents: 35427
diff changeset
    67
lemma unit_eq [no_atp]: "u = ()"
40590
b994d855dbd2 typedef (open) unit
huffman
parents: 39302
diff changeset
    68
  by (induct u) (simp add: Unity_def)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    69
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    70
text \<open>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    71
  Simplification procedure for @{thm [source] unit_eq}.  Cannot use
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    72
  this rule directly --- it loops!
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    73
\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    74
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    75
simproc_setup unit_eq ("x::unit") = \<open>
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    76
  fn _ => fn _ => fn ct =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    77
    if HOLogic.is_unit (Thm.term_of ct) then NONE
43594
ef1ddc59b825 modernized some simproc setup;
wenzelm
parents: 42411
diff changeset
    78
    else SOME (mk_meta_eq @{thm unit_eq})
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    79
\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
    80
55469
b19dd108f0c2 aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
blanchet
parents: 55468
diff changeset
    81
free_constructors case_unit for "()"
58189
9d714be4f028 added 'plugins' option to control which hooks are enabled
blanchet
parents: 57983
diff changeset
    82
  by auto
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    83
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
    84
text \<open>Avoid name clashes by prefixing the output of \<open>old_rep_datatype\<close> with \<open>old\<close>.\<close>
55442
blanchet
parents: 55417
diff changeset
    85
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    86
setup \<open>Sign.mandatory_path "old"\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    87
58306
117ba6cbe414 renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents: 58292
diff changeset
    88
old_rep_datatype "()" by simp
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
    89
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    90
setup \<open>Sign.parent_path\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    91
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
    92
text \<open>But erase the prefix for properties that are not generated by \<open>free_constructors\<close>.\<close>
55442
blanchet
parents: 55417
diff changeset
    93
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
    94
setup \<open>Sign.mandatory_path "unit"\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    95
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    96
lemmas induct = old.unit.induct
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
    97
lemmas inducts = old.unit.inducts
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55469
diff changeset
    98
lemmas rec = old.unit.rec
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55469
diff changeset
    99
lemmas simps = unit.case unit.rec
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   100
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   101
setup \<open>Sign.parent_path\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   102
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   103
lemma unit_all_eq1: "(\<And>x::unit. PROP P x) \<equiv> PROP P ()"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   104
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   105
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   106
lemma unit_all_eq2: "(\<And>x::unit. PROP P) \<equiv> PROP P"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   107
  by (rule triv_forall_equality)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   108
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   109
text \<open>
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   110
  This rewrite counters the effect of simproc \<open>unit_eq\<close> on @{term
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   111
  [source] "\<lambda>u::unit. f u"}, replacing it by @{term [source]
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   112
  f} rather than by @{term [source] "\<lambda>u. f ()"}.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   113
\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   114
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   115
lemma unit_abs_eta_conv [simp]: "(\<lambda>u::unit. f ()) = f"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   116
  by (rule ext) simp
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   117
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   118
lemma UNIV_unit: "UNIV = {()}"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   119
  by auto
43866
8a50dc70cbff moving UNIV = ... equations to their proper theories
haftmann
parents: 43654
diff changeset
   120
30924
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   121
instantiation unit :: default
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   122
begin
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   123
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   124
definition "default = ()"
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   125
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   126
instance ..
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   127
c1ed09f3fbfe default instantiation for unit type
haftmann
parents: 30604
diff changeset
   128
end
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   129
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   130
instantiation unit :: "{complete_boolean_algebra,complete_linorder,wellorder}"
57016
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   131
begin
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   132
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   133
definition less_eq_unit :: "unit \<Rightarrow> unit \<Rightarrow> bool"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   134
  where "(_::unit) \<le> _ \<longleftrightarrow> True"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   135
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   136
lemma less_eq_unit [iff]: "u \<le> v" for u v :: unit
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   137
  by (simp add: less_eq_unit_def)
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   138
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   139
definition less_unit :: "unit \<Rightarrow> unit \<Rightarrow> bool"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   140
  where "(_::unit) < _ \<longleftrightarrow> False"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   141
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   142
lemma less_unit [iff]: "\<not> u < v" for u v :: unit
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   143
  by (simp_all add: less_eq_unit_def less_unit_def)
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   144
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   145
definition bot_unit :: unit
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   146
  where [code_unfold]: "\<bottom> = ()"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   147
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   148
definition top_unit :: unit
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   149
  where [code_unfold]: "\<top> = ()"
57016
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   150
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   151
definition inf_unit :: "unit \<Rightarrow> unit \<Rightarrow> unit"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   152
  where [simp]: "_ \<sqinter> _ = ()"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   153
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   154
definition sup_unit :: "unit \<Rightarrow> unit \<Rightarrow> unit"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   155
  where [simp]: "_ \<squnion> _ = ()"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   156
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   157
definition Inf_unit :: "unit set \<Rightarrow> unit"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   158
  where [simp]: "\<Sqinter>_ = ()"
57016
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   159
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   160
definition Sup_unit :: "unit set \<Rightarrow> unit"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   161
  where [simp]: "\<Squnion>_ = ()"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   162
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   163
definition uminus_unit :: "unit \<Rightarrow> unit"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   164
  where [simp]: "- _ = ()"
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   165
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   166
declare less_eq_unit_def [abs_def, code_unfold]
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   167
  less_unit_def [abs_def, code_unfold]
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   168
  inf_unit_def [abs_def, code_unfold]
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   169
  sup_unit_def [abs_def, code_unfold]
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   170
  Inf_unit_def [abs_def, code_unfold]
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   171
  Sup_unit_def [abs_def, code_unfold]
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   172
  uminus_unit_def [abs_def, code_unfold]
57016
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   173
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   174
instance
57233
8fcbfce2a2a9 uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
haftmann
parents: 57201
diff changeset
   175
  by intro_classes auto
57016
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   176
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   177
end
c44ce6f4067d added unit :: linorder
nipkow
parents: 56626
diff changeset
   178
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   179
lemma [code]: "HOL.equal u v \<longleftrightarrow> True" for u v :: unit
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   180
  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
   181
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   182
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   183
  type_constructor unit \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   184
    (SML) "unit"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   185
    and (OCaml) "unit"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   186
    and (Haskell) "()"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   187
    and (Scala) "Unit"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   188
| constant Unity \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   189
    (SML) "()"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   190
    and (OCaml) "()"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   191
    and (Haskell) "()"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   192
    and (Scala) "()"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   193
| class_instance unit :: equal \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   194
    (Haskell) -
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   195
| constant "HOL.equal :: unit \<Rightarrow> unit \<Rightarrow> bool" \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   196
    (Haskell) infix 4 "=="
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   197
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   198
code_reserved SML
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   199
  unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   200
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   201
code_reserved OCaml
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   202
  unit
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   203
34886
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   204
code_reserved Scala
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   205
  Unit
873c31d9f10d some syntax setup for Scala
haftmann
parents: 33959
diff changeset
   206
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   207
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   208
subsection \<open>The product type\<close>
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   209
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   210
subsubsection \<open>Type definition\<close>
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   211
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   212
definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   213
  where "Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   214
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 61032
diff changeset
   215
definition "prod = {f. \<exists>a b. f = Pair_Rep (a::'a) (b::'b)}"
45696
476ad865f125 prefer typedef without alternative name;
wenzelm
parents: 45694
diff changeset
   216
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61943
diff changeset
   217
typedef ('a, 'b) prod ("(_ \<times>/ _)" [21, 20] 20) = "prod :: ('a \<Rightarrow> 'b \<Rightarrow> bool) set"
45696
476ad865f125 prefer typedef without alternative name;
wenzelm
parents: 45694
diff changeset
   218
  unfolding prod_def by auto
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   219
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61943
diff changeset
   220
type_notation (ASCII)
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61943
diff changeset
   221
  prod  (infixr "*" 20)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   222
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   223
definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   224
  where "Pair a b = Abs_prod (Pair_Rep a b)"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   225
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   226
lemma prod_cases: "(\<And>a b. P (Pair a b)) \<Longrightarrow> P p"
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   227
  by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   228
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   229
free_constructors case_prod for Pair fst snd
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   230
proof -
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   231
  fix P :: bool and p :: "'a \<times> 'b"
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   232
  show "(\<And>x1 x2. p = Pair x1 x2 \<Longrightarrow> P) \<Longrightarrow> P"
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   233
    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
   234
next
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   235
  fix a c :: 'a and b d :: 'b
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   236
  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
   237
    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
   238
  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
   239
    by (auto simp add: prod_def)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   240
  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
   241
    by (simp add: Pair_def Abs_prod_inject)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   242
qed
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   243
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   244
text \<open>Avoid name clashes by prefixing the output of \<open>old_rep_datatype\<close> with \<open>old\<close>.\<close>
55442
blanchet
parents: 55417
diff changeset
   245
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   246
setup \<open>Sign.mandatory_path "old"\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   247
58306
117ba6cbe414 renamed 'rep_datatype' to 'old_rep_datatype' (HOL)
blanchet
parents: 58292
diff changeset
   248
old_rep_datatype Pair
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   249
  by (erule prod_cases) (rule prod.inject)
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   250
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   251
setup \<open>Sign.parent_path\<close>
37704
c6161bee8486 adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents: 37678
diff changeset
   252
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   253
text \<open>But erase the prefix for properties that are not generated by \<open>free_constructors\<close>.\<close>
55442
blanchet
parents: 55417
diff changeset
   254
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   255
setup \<open>Sign.mandatory_path "prod"\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   256
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   257
declare old.prod.inject [iff del]
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   258
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   259
lemmas induct = old.prod.induct
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   260
lemmas inducts = old.prod.inducts
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55469
diff changeset
   261
lemmas rec = old.prod.rec
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55469
diff changeset
   262
lemmas simps = prod.inject prod.case prod.rec
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   263
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   264
setup \<open>Sign.parent_path\<close>
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   265
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   266
declare prod.case [nitpick_simp del]
63566
e5abbdee461a more accurate cong del;
wenzelm
parents: 63400
diff changeset
   267
declare old.prod.case_cong_weak [cong del]
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   268
declare prod.case_eq_if [mono]
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   269
declare prod.split [no_atp]
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   270
declare prod.split_asm [no_atp]
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   271
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   272
text \<open>
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   273
  @{thm [source] prod.split} could be declared as \<open>[split]\<close>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   274
  done after the Splitter has been speeded up significantly;
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   275
  precompute the constants involved and don't do anything unless the
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   276
  current goal contains one of those constants.
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   277
\<close>
37411
c88c44156083 removed simplifier congruence rule of "prod_case"
haftmann
parents: 37389
diff changeset
   278
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   279
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   280
subsubsection \<open>Tuple syntax\<close>
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   281
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   282
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   283
  Patterns -- extends pre-defined type \<^typ>\<open>pttrn\<close> used in
11777
wenzelm
parents: 11602
diff changeset
   284
  abstractions.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   285
\<close>
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   286
41229
d797baa3d57c replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents: 40968
diff changeset
   287
nonterminal tuple_args and patterns
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   288
syntax
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   289
  "_tuple"      :: "'a \<Rightarrow> tuple_args \<Rightarrow> 'a \<times> 'b"        ("(1'(_,/ _'))")
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   290
  "_tuple_arg"  :: "'a \<Rightarrow> tuple_args"                   ("_")
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   291
  "_tuple_args" :: "'a \<Rightarrow> tuple_args \<Rightarrow> tuple_args"     ("_,/ _")
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   292
  "_pattern"    :: "pttrn \<Rightarrow> patterns \<Rightarrow> pttrn"         ("'(_,/ _')")
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   293
  ""            :: "pttrn \<Rightarrow> patterns"                  ("_")
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   294
  "_patterns"   :: "pttrn \<Rightarrow> patterns \<Rightarrow> patterns"      ("_,/ _")
63237
3e908f762817 conventional syntax for unit abstractions
haftmann
parents: 63007
diff changeset
   295
  "_unit"       :: pttrn                                ("'(')")
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   296
translations
61124
haftmann
parents: 61123
diff changeset
   297
  "(x, y)" \<rightleftharpoons> "CONST Pair x y"
haftmann
parents: 61123
diff changeset
   298
  "_pattern x y" \<rightleftharpoons> "CONST Pair x y"
haftmann
parents: 61123
diff changeset
   299
  "_patterns x y" \<rightleftharpoons> "CONST Pair x y"
haftmann
parents: 61123
diff changeset
   300
  "_tuple x (_tuple_args y z)" \<rightleftharpoons> "_tuple x (_tuple_arg (_tuple y z))"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   301
  "\<lambda>(x, y, zs). b" \<rightleftharpoons> "CONST case_prod (\<lambda>x (y, zs). b)"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   302
  "\<lambda>(x, y). b" \<rightleftharpoons> "CONST case_prod (\<lambda>x y. b)"
61124
haftmann
parents: 61123
diff changeset
   303
  "_abs (CONST Pair x y) t" \<rightharpoonup> "\<lambda>(x, y). t"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   304
  \<comment> \<open>This rule accommodates tuples in \<open>case C \<dots> (x, y) \<dots> \<Rightarrow> \<dots>\<close>:
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   305
     The \<open>(x, y)\<close> is parsed as \<open>Pair x y\<close> because it is \<open>logic\<close>,
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   306
     not \<open>pttrn\<close>.\<close>
63237
3e908f762817 conventional syntax for unit abstractions
haftmann
parents: 63007
diff changeset
   307
  "\<lambda>(). b" \<rightleftharpoons> "CONST case_unit b"
3e908f762817 conventional syntax for unit abstractions
haftmann
parents: 63007
diff changeset
   308
  "_abs (CONST Unity) t" \<rightharpoonup> "\<lambda>(). t"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   309
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   310
text \<open>print \<^term>\<open>case_prod f\<close> as \<^term>\<open>\<lambda>(x, y). f x y\<close> and
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   311
  \<^term>\<open>case_prod (\<lambda>x. f x)\<close> as \<^term>\<open>\<lambda>(x, y). f x y\<close>\<close>
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   312
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   313
typed_print_translation \<open>
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   314
  let
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   315
    fun case_prod_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   316
      | case_prod_guess_names_tr' T [Abs (x, xT, t)] =
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   317
          (case (head_of t) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   318
            Const (\<^const_syntax>\<open>case_prod\<close>, _) => raise Match
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   319
          | _ =>
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   320
            let
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   321
              val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   322
              val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0);
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   323
              val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t');
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   324
            in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   325
              Syntax.const \<^syntax_const>\<open>_abs\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   326
                (Syntax.const \<^syntax_const>\<open>_pattern\<close> $ x' $ y) $ t''
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   327
            end)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   328
      | case_prod_guess_names_tr' T [t] =
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   329
          (case head_of t of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   330
            Const (\<^const_syntax>\<open>case_prod\<close>, _) => raise Match
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   331
          | _ =>
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   332
            let
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   333
              val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   334
              val (y, t') =
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   335
                Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0);
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   336
              val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t');
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   337
            in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   338
              Syntax.const \<^syntax_const>\<open>_abs\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   339
                (Syntax.const \<^syntax_const>\<open>_pattern\<close> $ x' $ y) $ t''
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   340
            end)
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   341
      | case_prod_guess_names_tr' _ _ = raise Match;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   342
  in [(\<^const_syntax>\<open>case_prod\<close>, K case_prod_guess_names_tr')] end
61226
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   343
\<close>
af7bed1360f3 effective revert of e6b1236f9b3d: spontaneous eta-contraction happens on the print translation level and can only be suppressed on the print translation level
haftmann
parents: 61144
diff changeset
   344
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   345
text \<open>Reconstruct pattern from (nested) \<^const>\<open>case_prod\<close>s,
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   346
  avoiding eta-contraction of body; required for enclosing "let",
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   347
  if "let" does not avoid eta-contraction, which has been observed to occur.\<close>
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   348
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   349
print_translation \<open>
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   350
  let
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   351
    fun case_prod_tr' [Abs (x, T, t as (Abs abs))] =
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   352
          (* case_prod (\<lambda>x y. t) \<Rightarrow> \<lambda>(x, y) t *)
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   353
          let
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   354
            val (y, t') = Syntax_Trans.atomic_abs_tr' abs;
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   355
            val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t');
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   356
          in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   357
            Syntax.const \<^syntax_const>\<open>_abs\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   358
              (Syntax.const \<^syntax_const>\<open>_pattern\<close> $ x' $ y) $ t''
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   359
          end
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   360
      | case_prod_tr' [Abs (x, T, (s as Const (\<^const_syntax>\<open>case_prod\<close>, _) $ t))] =
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   361
          (* case_prod (\<lambda>x. (case_prod (\<lambda>y z. t))) \<Rightarrow> \<lambda>(x, y, z). t *)
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   362
          let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   363
            val Const (\<^syntax_const>\<open>_abs\<close>, _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   364
              (Const (\<^syntax_const>\<open>_pattern\<close>, _) $ y $ z) $ t' =
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   365
                case_prod_tr' [t];
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   366
            val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t');
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   367
          in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   368
            Syntax.const \<^syntax_const>\<open>_abs\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   369
              (Syntax.const \<^syntax_const>\<open>_pattern\<close> $ x' $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   370
                (Syntax.const \<^syntax_const>\<open>_patterns\<close> $ y $ z)) $ t''
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   371
          end
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   372
      | case_prod_tr' [Const (\<^const_syntax>\<open>case_prod\<close>, _) $ t] =
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   373
          (* case_prod (case_prod (\<lambda>x y z. t)) \<Rightarrow> \<lambda>((x, y), z). t *)
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   374
          case_prod_tr' [(case_prod_tr' [t])]
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   375
            (* inner case_prod_tr' creates next pattern *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   376
      | case_prod_tr' [Const (\<^syntax_const>\<open>_abs\<close>, _) $ x_y $ Abs abs] =
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   377
          (* case_prod (\<lambda>pttrn z. t) \<Rightarrow> \<lambda>(pttrn, z). t *)
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   378
          let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   379
            Syntax.const \<^syntax_const>\<open>_abs\<close> $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   380
              (Syntax.const \<^syntax_const>\<open>_pattern\<close> $ x_y $ z) $ t
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   381
          end
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   382
      | case_prod_tr' _ = raise Match;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   383
  in [(\<^const_syntax>\<open>case_prod\<close>, K case_prod_tr')] end
61425
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   384
\<close>
fb95d06fb21f restored print translation from a1141fb798ff, to prevent a printing misfit observable using "thm divmod_nat_if" in theory "Divides", with a meagure indication in the comment
haftmann
parents: 61424
diff changeset
   385
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   386
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   387
subsubsection \<open>Code generator setup\<close>
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   388
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   389
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   390
  type_constructor prod \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   391
    (SML) infix 2 "*"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   392
    and (OCaml) infix 2 "*"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   393
    and (Haskell) "!((_),/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   394
    and (Scala) "((_),/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   395
| constant Pair \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   396
    (SML) "!((_),/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   397
    and (OCaml) "!((_),/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   398
    and (Haskell) "!((_),/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   399
    and (Scala) "!((_),/ (_))"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   400
| class_instance  prod :: equal \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   401
    (Haskell) -
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   402
| constant "HOL.equal :: 'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool" \<rightharpoonup>
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   403
    (Haskell) infix 4 "=="
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   404
| constant fst \<rightharpoonup> (Haskell) "fst"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   405
| constant snd \<rightharpoonup> (Haskell) "snd"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   406
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   407
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   408
subsubsection \<open>Fundamental operations and properties\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   409
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   410
lemma Pair_inject: "(a, b) = (a', b') \<Longrightarrow> (a = a' \<Longrightarrow> b = b' \<Longrightarrow> R) \<Longrightarrow> R"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   411
  by simp
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
   412
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   413
lemma surj_pair [simp]: "\<exists>x y. p = (x, y)"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   414
  by (cases p) simp
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   415
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   416
lemma fst_eqD: "fst (x, y) = a \<Longrightarrow> x = a"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   417
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   418
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   419
lemma snd_eqD: "snd (x, y) = a \<Longrightarrow> y = a"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   420
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   421
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   422
lemma case_prod_unfold [nitpick_unfold]: "case_prod = (\<lambda>c p. c (fst p) (snd p))"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   423
  by (simp add: fun_eq_iff split: prod.split)
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   425
lemma case_prod_conv [simp, code]: "(case (a, b) of (c, d) \<Rightarrow> f c d) = f a b"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   426
  by (fact prod.case)
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   427
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
   428
lemmas surjective_pairing = prod.collapse [symmetric]
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   429
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 43866
diff changeset
   430
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
   431
  by (cases s, cases t) simp
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   432
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   433
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
   434
  by (simp add: prod_eq_iff)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   435
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   436
lemma case_prodI: "f a b \<Longrightarrow> case (a, b) of (c, d) \<Rightarrow> f c d"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   437
  by (rule prod.case [THEN iffD2])
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   438
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   439
lemma case_prodD: "(case (a, b) of (c, d) \<Rightarrow> f c d) \<Longrightarrow> f a b"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   440
  by (rule prod.case [THEN iffD1])
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   441
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   442
lemma case_prod_Pair [simp]: "case_prod Pair = 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
   443
  by (simp add: fun_eq_iff split: prod.split)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   444
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   445
lemma case_prod_eta: "(\<lambda>(x, y). f (x, y)) = f"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   446
  \<comment> \<open>Subsumes the old \<open>split_Pair\<close> when \<^term>\<open>f\<close> is the identity function.\<close>
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39272
diff changeset
   447
  by (simp add: fun_eq_iff split: prod.split)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   448
67575
nipkow
parents: 67443
diff changeset
   449
(* This looks like a sensible simp-rule but appears to do more harm than good:
nipkow
parents: 67443
diff changeset
   450
lemma case_prod_const [simp]: "(\<lambda>(_,_). c) = (\<lambda>_. c)"
nipkow
parents: 67443
diff changeset
   451
by(rule case_prod_eta)
nipkow
parents: 67443
diff changeset
   452
*)
nipkow
parents: 67443
diff changeset
   453
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   454
lemma case_prod_comp: "(case x of (a, b) \<Rightarrow> (f \<circ> g) a b) = f (g (fst x)) (snd x)"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   455
  by (cases x) simp
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   456
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   457
lemma The_case_prod: "The (case_prod P) = (THE xy. P (fst xy) (snd xy))"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55403
diff changeset
   458
  by (simp add: case_prod_unfold)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   459
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   460
lemma cond_case_prod_eta: "(\<And>x y. f x y = g (x, y)) \<Longrightarrow> (\<lambda>(x, y). f x y) = g"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   461
  by (simp add: case_prod_eta)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   462
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   463
lemma split_paired_all [no_atp]: "(\<And>x. PROP P x) \<equiv> (\<And>a b. PROP P (a, b))"
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   464
proof
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   465
  fix a b
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   466
  assume "\<And>x. PROP P x"
19535
e4fdeb32eadf replaced syntax/translations by abbreviation;
wenzelm
parents: 19179
diff changeset
   467
  then show "PROP P (a, b)" .
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   468
next
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   469
  fix x
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   470
  assume "\<And>a b. PROP P (a, b)"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   471
  from \<open>PROP P (fst x, snd x)\<close> show "PROP P x" by simp
11820
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   472
qed
015a82d4ee96 proper proof of split_paired_all (presently unused);
wenzelm
parents: 11777
diff changeset
   473
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   474
text \<open>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   475
  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
   476
  Simplifier because it also affects premises in congrence rules,
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   477
  where this can lead to premises of the form \<open>\<And>a b. \<dots> = ?P(a, b)\<close>
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   478
  which cannot be solved by reflexivity.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   479
\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   480
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   481
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
   482
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   483
ML \<open>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   484
  (* replace parameters of product type by individual component parameters *)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   485
  local (* filtering with exists_paired_all is an essential optimization *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   486
    fun exists_paired_all (Const (\<^const_name>\<open>Pure.all\<close>, _) $ Abs (_, T, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   487
          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
   488
      | 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
   489
      | exists_paired_all (Abs (_, _, t)) = exists_paired_all t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   490
      | exists_paired_all _ = false;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   491
    val ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   492
      simpset_of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   493
       (put_simpset HOL_basic_ss \<^context>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   494
        addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}]
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   495
        addsimprocs [\<^simproc>\<open>unit_eq\<close>]);
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   496
  in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   497
    fun split_all_tac ctxt = SUBGOAL (fn (t, i) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   498
      if exists_paired_all t then safe_full_simp_tac (put_simpset ss ctxt) i else no_tac);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   499
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   500
    fun unsafe_split_all_tac ctxt = SUBGOAL (fn (t, i) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   501
      if exists_paired_all t then full_simp_tac (put_simpset ss ctxt) i else no_tac);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   502
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   503
    fun split_all ctxt th =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   504
      if exists_paired_all (Thm.prop_of th)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   505
      then full_simplify (put_simpset ss ctxt) th else th;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   506
  end;
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   507
\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   508
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   509
setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   510
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   511
lemma split_paired_All [simp, no_atp]: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>a b. P (a, b))"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   512
  \<comment> \<open>\<open>[iff]\<close> is not a good idea because it makes \<open>blast\<close> loop\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   513
  by fast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   514
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   515
lemma split_paired_Ex [simp, no_atp]: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>a b. P (a, b))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   516
  by fast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   517
47740
a8989fe9a3a5 added "no_atp"s for extremely prolific, useless facts for ATPs
blanchet
parents: 46950
diff changeset
   518
lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   519
  \<comment> \<open>Can't be added to simpset: loops!\<close>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   520
  by (simp add: case_prod_eta)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   521
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   522
text \<open>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   523
  Simplification procedure for @{thm [source] cond_case_prod_eta}.  Using
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   524
  @{thm [source] case_prod_eta} as a rewrite rule is not general enough,
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   525
  and using @{thm [source] cond_case_prod_eta} directly would render some
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   526
  existing proofs very inefficient; similarly for \<open>prod.case_eq_if\<close>.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   527
\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   528
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   529
ML \<open>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   530
local
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   531
  val cond_case_prod_eta_ss =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   532
    simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms cond_case_prod_eta});
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   533
  fun Pair_pat k 0 (Bound m) = (m = k)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   534
    | Pair_pat k i (Const (\<^const_name>\<open>Pair\<close>,  _) $ Bound m $ t) =
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   535
        i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   536
    | Pair_pat _ _ _ = false;
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   537
  fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   538
    | 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
   539
    | no_args k i (Bound m) = m < k orelse m > k + i
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   540
    | no_args _ _ _ = true;
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   541
  fun split_pat tp i (Abs  (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   542
    | split_pat tp i (Const (\<^const_name>\<open>case_prod\<close>, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   543
    | split_pat tp i _ = NONE;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   544
  fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] []
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   545
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   546
        (K (simp_tac (put_simpset cond_case_prod_eta_ss ctxt) 1)));
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   547
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   548
  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   549
    | beta_term_pat k i (t $ u) =
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   550
        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
   551
    | beta_term_pat k i t = no_args k i t;
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   552
  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
   553
    | eta_term_pat _ _ _ = false;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   554
  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
   555
    | subst arg k i (t $ u) =
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   556
        if Pair_pat k i (t $ u) then incr_boundvars k arg
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   557
        else (subst arg k i t $ subst arg k i u)
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   558
    | subst arg k i t = t;
43595
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 43594
diff changeset
   559
in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   560
  fun beta_proc ctxt (s as Const (\<^const_name>\<open>case_prod\<close>, _) $ Abs (_, _, t) $ arg) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   561
        (case split_pat beta_term_pat 1 t of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   562
          SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   563
        | NONE => NONE)
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   564
    | beta_proc _ _ = NONE;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   565
  fun eta_proc ctxt (s as Const (\<^const_name>\<open>case_prod\<close>, _) $ Abs (_, _, t)) =
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   566
        (case split_pat eta_term_pat 1 t of
58468
haftmann
parents: 58389
diff changeset
   567
          SOME (_, ft) => SOME (metaeq ctxt s (let val f $ _ = ft in f end))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
   568
        | NONE => NONE)
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 35115
diff changeset
   569
    | eta_proc _ _ = NONE;
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   570
end;
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   571
\<close>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   572
simproc_setup case_prod_beta ("case_prod f z") =
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   573
  \<open>fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct)\<close>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   574
simproc_setup case_prod_eta ("case_prod f") =
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   575
  \<open>fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct)\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   576
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   577
lemma case_prod_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   578
  by (auto simp: fun_eq_iff)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
   579
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   580
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   581
  \<^medskip> \<^const>\<open>case_prod\<close> used as a logical connective or set former.
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   582
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   583
  \<^medskip> These rules are for use with \<open>blast\<close>; could instead
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   584
  call \<open>simp\<close> using @{thm [source] prod.split} as rewrite.\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   585
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   586
lemma case_prodI2:
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   587
  "\<And>p. (\<And>a b. p = (a, b) \<Longrightarrow> c a b) \<Longrightarrow> case p of (a, b) \<Rightarrow> c a b"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   588
  by (simp add: split_tupled_all)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   589
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   590
lemma case_prodI2':
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   591
  "\<And>p. (\<And>a b. (a, b) = p \<Longrightarrow> c a b x) \<Longrightarrow> (case p of (a, b) \<Rightarrow> c a b) x"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   592
  by (simp add: split_tupled_all)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   593
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   594
lemma case_prodE [elim!]:
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   595
  "(case p of (a, b) \<Rightarrow> c a b) \<Longrightarrow> (\<And>x y. p = (x, y) \<Longrightarrow> c x y \<Longrightarrow> Q) \<Longrightarrow> Q"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   596
  by (induct p) simp
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   597
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   598
lemma case_prodE' [elim!]:
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   599
  "(case p of (a, b) \<Rightarrow> c a b) z \<Longrightarrow> (\<And>x y. p = (x, y) \<Longrightarrow> c x y z \<Longrightarrow> Q) \<Longrightarrow> Q"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   600
  by (induct p) simp
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   601
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   602
lemma case_prodE2:
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   603
  assumes q: "Q (case z of (a, b) \<Rightarrow> P a b)"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   604
    and r: "\<And>x y. z = (x, y) \<Longrightarrow> Q (P x y) \<Longrightarrow> R"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   605
  shows R
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   606
proof (rule r)
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   607
  show "z = (fst z, snd z)" by simp
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   608
  then show "Q (P (fst z) (snd z))"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   609
    using q by (simp add: case_prod_unfold)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   610
qed
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   611
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   612
lemma case_prodD': "(case (a, b) of (c, d) \<Rightarrow> R c d) c \<Longrightarrow> R a b c"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   613
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   614
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   615
lemma mem_case_prodI: "z \<in> c a b \<Longrightarrow> z \<in> (case (a, b) of (d, e) \<Rightarrow> c d e)"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   616
  by simp
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   617
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   618
lemma mem_case_prodI2 [intro!]:
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
   619
  "\<And>p. (\<And>a b. p = (a, b) \<Longrightarrow> z \<in> c a b) \<Longrightarrow> z \<in> (case p of (a, b) \<Rightarrow> c a b)"
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
   620
  by (simp only: split_tupled_all) simp
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   621
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   622
declare mem_case_prodI [intro!] \<comment> \<open>postponed to maintain traditional declaration order!\<close>
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   623
declare case_prodI2' [intro!] \<comment> \<open>postponed to maintain traditional declaration order!\<close>
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   624
declare case_prodI2 [intro!] \<comment> \<open>postponed to maintain traditional declaration order!\<close>
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   625
declare case_prodI [intro!] \<comment> \<open>postponed to maintain traditional declaration order!\<close>
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   626
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   627
lemma mem_case_prodE [elim!]:
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   628
  assumes "z \<in> case_prod c p"
58468
haftmann
parents: 58389
diff changeset
   629
  obtains x y where "p = (x, y)" and "z \<in> c x y"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   630
  using assms by (rule case_prodE2)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   631
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   632
ML \<open>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   633
local (* filtering with exists_p_split is an essential optimization *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   634
  fun exists_p_split (Const (\<^const_name>\<open>case_prod\<close>,_) $ _ $ (Const (\<^const_name>\<open>Pair\<close>,_)$_$_)) = true
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   635
    | 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
   636
    | exists_p_split (Abs (_, _, t)) = exists_p_split t
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   637
    | exists_p_split _ = false;
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   638
in
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   639
  fun split_conv_tac ctxt = SUBGOAL (fn (t, i) =>
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   640
    if exists_p_split t
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   641
    then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms case_prod_conv}) i
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   642
    else no_tac);
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   643
end;
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   644
\<close>
26340
a85fe32e7b2f more antiquotations;
wenzelm
parents: 26143
diff changeset
   645
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   646
(* 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
   647
   to quite time-consuming computations (in particular for nested tuples) *)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   648
setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac))\<close>
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   649
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   650
lemma split_eta_SetCompr [simp, no_atp]: "(\<lambda>u. \<exists>x y. u = (x, y) \<and> P (x, y)) = P"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   651
  by (rule ext) fast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   652
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   653
lemma split_eta_SetCompr2 [simp, no_atp]: "(\<lambda>u. \<exists>x y. u = (x, y) \<and> P x y) = case_prod P"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   654
  by (rule ext) fast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   655
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   656
lemma split_part [simp]: "(\<lambda>(a,b). P \<and> Q a b) = (\<lambda>ab. P \<and> case_prod Q ab)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   657
  \<comment> \<open>Allows simplifications of nested splits in case of independent predicates.\<close>
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   658
  by (rule ext) blast
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   659
14337
e13731554e50 undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents: 14208
diff changeset
   660
(* 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
   661
   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
   662
   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
   663
*)
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   664
lemma split_comp_eq:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   665
  fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   666
    and g :: "'d \<Rightarrow> 'a"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   667
  shows "(\<lambda>u. f (g (fst u)) (snd u)) = case_prod (\<lambda>x. f (g x))"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 18334
diff changeset
   668
  by (rule ext) auto
14101
d25c23e46173 added upd_fst, upd_snd, some thms
oheimb
parents: 13480
diff changeset
   669
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   670
lemma pair_imageI [intro]: "(a, b) \<in> A \<Longrightarrow> f a b \<in> (\<lambda>(a, b). f a b) ` A"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   671
  by (rule image_eqI [where x = "(a, b)"]) auto
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   672
68457
517aa9076fc9 added simp rule
nipkow
parents: 67575
diff changeset
   673
lemma Collect_const_case_prod[simp]: "{(a,b). P} = (if P then UNIV else {})"
517aa9076fc9 added simp rule
nipkow
parents: 67575
diff changeset
   674
by auto
517aa9076fc9 added simp rule
nipkow
parents: 67575
diff changeset
   675
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   676
lemma The_split_eq [simp]: "(THE (x',y'). x = x' \<and> y = y') = (x, y)"
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   677
  by blast
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   678
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   679
(*
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   680
the following  would be slightly more general,
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   681
but cannot be used as rewrite rule:
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   682
### 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
   683
### ?y = .x
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   684
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
   685
by (rtac some_equality 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   686
by ( Simp_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   687
by (split_all_tac 1)
144f45277d5a misc tidying
paulson
parents: 14190
diff changeset
   688
by (Asm_full_simp_tac 1)
11838
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   689
qed "The_split_eq";
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   690
*)
02d75712061d got rid of ML proof scripts for Product_Type;
wenzelm
parents: 11820
diff changeset
   691
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   692
lemma case_prod_beta: "case_prod f p = f (fst p) (snd p)"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   693
  by (fact prod.case_eq_if)
26143
314c0bcb7df7 Added useful general lemmas from the work with the HeapMonad
bulwahn
parents: 25885
diff changeset
   694
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
   695
lemma prod_cases3 [cases type]:
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   696
  obtains (fields) a b c where "y = (a, b, c)"
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   697
  by (cases y, case_tac b) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   698
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   699
lemma prod_induct3 [case_names fields, induct type]:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   700
  "(\<And>a b c. P (a, b, c)) \<Longrightarrow> P x"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   701
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   702
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
   703
lemma prod_cases4 [cases type]:
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   704
  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
   705
  by (cases y, case_tac c) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   706
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   707
lemma prod_induct4 [case_names fields, induct type]:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   708
  "(\<And>a b c d. P (a, b, c, d)) \<Longrightarrow> P x"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   709
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   710
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
   711
lemma prod_cases5 [cases type]:
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   712
  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
   713
  by (cases y, case_tac d) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   714
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   715
lemma prod_induct5 [case_names fields, induct type]:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   716
  "(\<And>a b c d e. P (a, b, c, d, e)) \<Longrightarrow> P x"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   717
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   718
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
   719
lemma prod_cases6 [cases type]:
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   720
  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
   721
  by (cases y, case_tac e) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   722
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   723
lemma prod_induct6 [case_names fields, induct type]:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   724
  "(\<And>a b c d e f. P (a, b, c, d, e, f)) \<Longrightarrow> P x"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   725
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   726
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 55414
diff changeset
   727
lemma prod_cases7 [cases type]:
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   728
  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
   729
  by (cases y, case_tac f) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   730
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   731
lemma prod_induct7 [case_names fields, induct type]:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   732
  "(\<And>a b c d e f g. P (a, b, c, d, e, f, g)) \<Longrightarrow> P x"
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   733
  by (cases x) blast
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   734
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   735
definition internal_case_prod :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   736
  where "internal_case_prod \<equiv> case_prod"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   737
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   738
lemma internal_case_prod_conv: "internal_case_prod c (a, b) = c a b"
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   739
  by (simp only: internal_case_prod_def case_prod_conv)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   740
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   741
ML_file \<open>Tools/split_rule.ML\<close>
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   742
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   743
hide_const internal_case_prod
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   744
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
   745
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   746
subsubsection \<open>Derived operations\<close>
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   747
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   748
definition curry :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   749
  where "curry = (\<lambda>c x y. c (x, y))"
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   750
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   751
lemma curry_conv [simp, code]: "curry f a b = f (a, b)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   752
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   753
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   754
lemma curryI [intro!]: "f (a, b) \<Longrightarrow> curry f a b"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   755
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   756
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   757
lemma curryD [dest!]: "curry f a b \<Longrightarrow> f (a, b)"
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   758
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   759
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   760
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
   761
  by (simp add: curry_def)
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   762
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   763
lemma curry_case_prod [simp]: "curry (case_prod f) = f"
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 60758
diff changeset
   764
  by (simp add: curry_def case_prod_unfold)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   765
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   766
lemma case_prod_curry [simp]: "case_prod (curry f) = f"
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 60758
diff changeset
   767
  by (simp add: curry_def case_prod_unfold)
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   768
54630
9061af4d5ebc restrict admissibility to non-empty chains to allow more syntax-directed proof rules
Andreas Lochbihler
parents: 54147
diff changeset
   769
lemma curry_K: "curry (\<lambda>x. c) = (\<lambda>x y. c)"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   770
  by (simp add: fun_eq_iff)
54630
9061af4d5ebc restrict admissibility to non-empty chains to allow more syntax-directed proof rules
Andreas Lochbihler
parents: 54147
diff changeset
   771
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   772
text \<open>The composition-uncurry combinator.\<close>
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   773
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   774
notation fcomp (infixl "\<circ>>" 60)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   775
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   776
definition scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd"  (infixl "\<circ>\<rightarrow>" 60)
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   777
  where "f \<circ>\<rightarrow> g = (\<lambda>x. case_prod g (f x))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   778
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   779
lemma scomp_unfold: "scomp = (\<lambda>f g x. g (fst (f x)) (snd (f x)))"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55403
diff changeset
   780
  by (simp add: fun_eq_iff scomp_def case_prod_unfold)
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37591
diff changeset
   781
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
   782
lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = case_prod g (f x)"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55403
diff changeset
   783
  by (simp add: scomp_unfold case_prod_unfold)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   784
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   785
lemma Pair_scomp: "Pair x \<circ>\<rightarrow> f = f x"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   786
  by (simp add: fun_eq_iff)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   787
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   788
lemma scomp_Pair: "x \<circ>\<rightarrow> Pair = x"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   789
  by (simp add: fun_eq_iff)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   790
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   791
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
   792
  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
   793
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   794
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
   795
  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
   796
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   797
lemma fcomp_scomp: "(f \<circ>> g) \<circ>\<rightarrow> h = f \<circ>> (g \<circ>\<rightarrow> h)"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   798
  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
   799
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   800
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 52143
diff changeset
   801
  constant scomp \<rightharpoonup> (Eval) infixl 3 "#->"
31202
52d332f8f909 pretty printing of functional combinators for evaluation code
haftmann
parents: 30924
diff changeset
   802
37751
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   803
no_notation fcomp (infixl "\<circ>>" 60)
89e16802b6cc nicer xsymbol syntax for fcomp and scomp
haftmann
parents: 37704
diff changeset
   804
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   805
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   806
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   807
  \<^term>\<open>map_prod\<close> --- action of the product functor upon functions.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   808
\<close>
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   809
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   810
definition map_prod :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   811
  where "map_prod 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
   812
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   813
lemma map_prod_simp [simp, code]: "map_prod f g (a, b) = (f a, g b)"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
   814
  by (simp add: map_prod_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   815
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
   816
functor map_prod: map_prod
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
   817
  by (auto simp add: split_paired_all)
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   818
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   819
lemma fst_map_prod [simp]: "fst (map_prod f g x) = f (fst x)"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   820
  by (cases x) simp_all
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   821
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   822
lemma snd_map_prod [simp]: "snd (map_prod f g x) = g (snd x)"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   823
  by (cases x) simp_all
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   824
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   825
lemma fst_comp_map_prod [simp]: "fst \<circ> map_prod f g = f \<circ> fst"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   826
  by (rule ext) simp_all
37278
307845cc7f51 added lemmas
nipkow
parents: 37166
diff changeset
   827
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   828
lemma snd_comp_map_prod [simp]: "snd \<circ> map_prod f g = g \<circ> snd"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   829
  by (rule ext) simp_all
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   830
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   831
lemma map_prod_compose: "map_prod (f1 \<circ> f2) (g1 \<circ> g2) = (map_prod f1 g1 \<circ> map_prod f2 g2)"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
   832
  by (rule ext) (simp add: map_prod.compositionality comp_def)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   833
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   834
lemma map_prod_ident [simp]: "map_prod (\<lambda>x. x) (\<lambda>y. y) = (\<lambda>z. z)"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
   835
  by (rule ext) (simp add: map_prod.identity)
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   836
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   837
lemma map_prod_imageI [intro]: "(a, b) \<in> R \<Longrightarrow> (f a, g b) \<in> map_prod f g ` R"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   838
  by (rule image_eqI) simp_all
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   839
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   840
lemma prod_fun_imageE [elim!]:
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
   841
  assumes major: "c \<in> map_prod f g ` R"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
   842
    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
   843
  shows P
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   844
  apply (rule major [THEN imageE])
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
   845
  apply (case_tac x)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   846
  apply (rule cases)
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   847
   apply simp_all
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   848
  done
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   849
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   850
definition apfst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   851
  where "apfst f = map_prod f id"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   852
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   853
definition apsnd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   854
  where "apsnd f = map_prod id f"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   855
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   856
lemma apfst_conv [simp, code]: "apfst f (x, y) = (f x, y)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   857
  by (simp add: apfst_def)
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   858
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
   859
lemma apsnd_conv [simp, code]: "apsnd f (x, y) = (x, f y)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   860
  by (simp add: apsnd_def)
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   861
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   862
lemma fst_apfst [simp]: "fst (apfst f x) = f (fst x)"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   863
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   864
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   865
lemma fst_comp_apfst [simp]: "fst \<circ> apfst f = f \<circ> fst"
51173
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   866
  by (simp add: fun_eq_iff)
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   867
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   868
lemma fst_apsnd [simp]: "fst (apsnd f x) = fst x"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   869
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   870
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   871
lemma fst_comp_apsnd [simp]: "fst \<circ> apsnd f = fst"
51173
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   872
  by (simp add: fun_eq_iff)
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   873
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   874
lemma snd_apfst [simp]: "snd (apfst f x) = snd x"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   875
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   876
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   877
lemma snd_comp_apfst [simp]: "snd \<circ> apfst f = snd"
51173
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   878
  by (simp add: fun_eq_iff)
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   879
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   880
lemma snd_apsnd [simp]: "snd (apsnd f x) = f (snd x)"
33594
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
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   883
lemma snd_comp_apsnd [simp]: "snd \<circ> apsnd f = f \<circ> snd"
51173
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   884
  by (simp add: fun_eq_iff)
3cbb4e95a565 Sieve of Eratosthenes
haftmann
parents: 50107
diff changeset
   885
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   886
lemma apfst_compose: "apfst f (apfst g x) = apfst (f \<circ> g) x"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   887
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   888
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   889
lemma apsnd_compose: "apsnd f (apsnd g x) = apsnd (f \<circ> g) x"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   890
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   891
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   892
lemma apfst_apsnd [simp]: "apfst f (apsnd g x) = (f (fst x), g (snd x))"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   893
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   894
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   895
lemma apsnd_apfst [simp]: "apsnd f (apfst g x) = (g (fst x), f (snd x))"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   896
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   897
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   898
lemma apfst_id [simp]: "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
   899
  by (simp add: fun_eq_iff)
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   900
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   901
lemma apsnd_id [simp]: "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
   902
  by (simp add: fun_eq_iff)
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   903
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   904
lemma apfst_eq_conv [simp]: "apfst f x = apfst g x \<longleftrightarrow> f (fst x) = g (fst x)"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   905
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   906
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   907
lemma apsnd_eq_conv [simp]: "apsnd f x = apsnd g x \<longleftrightarrow> f (snd x) = g (snd x)"
33594
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   908
  by (cases x) simp
357f74e0090c lemmas about apfst and apsnd
haftmann
parents: 33275
diff changeset
   909
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   910
lemma apsnd_apfst_commute: "apsnd f (apfst g p) = apfst g (apsnd f p)"
33638
548a34929e98 Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
hoelzl
parents: 33594
diff changeset
   911
  by simp
21195
0cca8d19557d two further lemmas on split
haftmann
parents: 21046
diff changeset
   912
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   913
context
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   914
begin
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   915
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   916
local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")\<close>
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   917
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   918
definition swap :: "'a \<times> 'b \<Rightarrow> 'b \<times> 'a"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   919
  where "swap p = (snd p, fst p)"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   920
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   921
end
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   922
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   923
lemma swap_simp [simp]: "prod.swap (x, y) = (y, x)"
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   924
  by (simp add: prod.swap_def)
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   925
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   926
lemma swap_swap [simp]: "prod.swap (prod.swap p) = p"
58195
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   927
  by (cases p) simp
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   928
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   929
lemma swap_comp_swap [simp]: "prod.swap \<circ> prod.swap = id"
58195
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   930
  by (simp add: fun_eq_iff)
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   931
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   932
lemma pair_in_swap_image [simp]: "(y, x) \<in> prod.swap ` A \<longleftrightarrow> (x, y) \<in> A"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   933
  by (auto intro!: image_eqI)
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   934
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   935
lemma inj_swap [simp]: "inj_on prod.swap A"
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   936
  by (rule inj_onI) auto
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   937
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   938
lemma swap_inj_on: "inj_on (\<lambda>(i, j). (j, i)) A"
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
   939
  by (rule inj_onI) auto
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   940
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   941
lemma surj_swap [simp]: "surj prod.swap"
58195
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   942
  by (rule surjI [of _ prod.swap]) simp
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   943
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   944
lemma bij_swap [simp]: "bij prod.swap"
58195
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   945
  by (simp add: bij_def)
1fee63e0377d added various facts
haftmann
parents: 58189
diff changeset
   946
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   947
lemma case_swap [simp]: "(case prod.swap p of (y, x) \<Rightarrow> f x y) = (case p of (x, y) \<Rightarrow> f x y)"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   948
  by (cases p) simp
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
   949
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 62101
diff changeset
   950
lemma fst_swap [simp]: "fst (prod.swap x) = snd x"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   951
  by (cases x) simp
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 62101
diff changeset
   952
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 62101
diff changeset
   953
lemma snd_swap [simp]: "snd (prod.swap x) = fst x"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   954
  by (cases x) simp
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 62101
diff changeset
   955
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   956
text \<open>Disjoint union of a family of sets -- Sigma.\<close>
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   957
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   958
definition Sigma :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<times> 'b) set"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   959
  where "Sigma A B \<equiv> \<Union>x\<in>A. \<Union>y\<in>B x. {Pair x y}"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   960
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   961
abbreviation Times :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"  (infixr "\<times>" 80)
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   962
  where "A \<times> B \<equiv> Sigma A (\<lambda>_. B)"
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
   963
45662
4f7c05990420 Hide Product_Type.Times - too precious an identifier
nipkow
parents: 45607
diff changeset
   964
hide_const (open) Times
4f7c05990420 Hide Product_Type.Times - too precious an identifier
nipkow
parents: 45607
diff changeset
   965
68467
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   966
bundle no_Set_Product_syntax begin
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   967
no_notation Product_Type.Times (infixr "\<times>" 80)
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   968
end
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   969
bundle Set_Product_syntax begin
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   970
notation Product_Type.Times (infixr "\<times>" 80)
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   971
end
44ffc5b9cd76 fixing overloading problems involving vector cross products
paulson <lp15@cam.ac.uk>
parents: 68457
diff changeset
   972
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   973
syntax
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   974
  "_Sigma" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"  ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   975
translations
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   976
  "SIGMA x:A. B" \<rightleftharpoons> "CONST Sigma A (\<lambda>x. B)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   977
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   978
lemma SigmaI [intro!]: "a \<in> A \<Longrightarrow> b \<in> B a \<Longrightarrow> (a, b) \<in> Sigma A B"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   979
  unfolding Sigma_def by blast
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   980
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   981
lemma SigmaE [elim!]: "c \<in> Sigma A B \<Longrightarrow> (\<And>x y. x \<in> A \<Longrightarrow> y \<in> B x \<Longrightarrow> c = (x, y) \<Longrightarrow> P) \<Longrightarrow> P"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
   982
  \<comment> \<open>The general elimination rule.\<close>
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   983
  unfolding Sigma_def by blast
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   984
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   985
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
   986
  Elimination of \<^term>\<open>(a, b) \<in> A \<times> B\<close> -- introduces no
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   987
  eigenvariables.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
   988
\<close>
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   989
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   990
lemma SigmaD1: "(a, b) \<in> Sigma A B \<Longrightarrow> a \<in> A"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   991
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   992
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   993
lemma SigmaD2: "(a, b) \<in> Sigma A B \<Longrightarrow> b \<in> B a"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   994
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   995
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   996
lemma SigmaE2: "(a, b) \<in> Sigma A B \<Longrightarrow> (a \<in> A \<Longrightarrow> b \<in> B a \<Longrightarrow> P) \<Longrightarrow> P"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
   997
  by blast
20588
c847c56edf0c added operational equality
haftmann
parents: 20415
diff changeset
   998
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
   999
lemma Sigma_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (SIGMA x:A. C x) = (SIGMA x:B. D x)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1000
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1001
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1002
lemma Sigma_mono: "A \<subseteq> C \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> D x) \<Longrightarrow> Sigma A B \<subseteq> Sigma C D"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1003
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1004
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1005
lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1006
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1007
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 61799
diff changeset
  1008
lemma Sigma_empty2 [simp]: "A \<times> {} = {}"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1009
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1010
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 61799
diff changeset
  1011
lemma UNIV_Times_UNIV [simp]: "UNIV \<times> UNIV = UNIV"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1012
  by auto
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1013
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 61799
diff changeset
  1014
lemma Compl_Times_UNIV1 [simp]: "- (UNIV \<times> A) = UNIV \<times> (-A)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1015
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1016
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 61799
diff changeset
  1017
lemma Compl_Times_UNIV2 [simp]: "- (A \<times> UNIV) = (-A) \<times> UNIV"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1018
  by auto
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1019
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1020
lemma mem_Sigma_iff [iff]: "(a, b) \<in> Sigma A B \<longleftrightarrow> a \<in> A \<and> b \<in> B a"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1021
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1022
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61955
diff changeset
  1023
lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61955
diff changeset
  1024
  by (induct x) simp
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61955
diff changeset
  1025
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58916
diff changeset
  1026
lemma Sigma_empty_iff: "(SIGMA i:I. X i) = {} \<longleftrightarrow> (\<forall>i\<in>I. X i = {})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58916
diff changeset
  1027
  by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58916
diff changeset
  1028
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1029
lemma Times_subset_cancel2: "x \<in> C \<Longrightarrow> A \<times> C \<subseteq> B \<times> C \<longleftrightarrow> A \<subseteq> B"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1030
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1031
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1032
lemma Times_eq_cancel2: "x \<in> C \<Longrightarrow> A \<times> C = B \<times> C \<longleftrightarrow> A = B"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1033
  by (blast elim: equalityE)
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1034
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1035
lemma Collect_case_prod_Sigma: "{(x, y). P x \<and> Q x y} = (SIGMA x:Collect P. Collect (Q x))"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1036
  by blast
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1037
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1038
lemma Collect_case_prod [simp]: "{(a, b). P a \<and> Q b} = Collect P \<times> Collect Q "
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1039
  by (fact Collect_case_prod_Sigma)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1040
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1041
lemma Collect_case_prodD: "x \<in> Collect (case_prod A) \<Longrightarrow> A (fst x) (snd x)"
61422
0dfcd0fb4172 moved lemmas
haftmann
parents: 61378
diff changeset
  1042
  by auto
0dfcd0fb4172 moved lemmas
haftmann
parents: 61378
diff changeset
  1043
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1044
lemma Collect_case_prod_mono: "A \<le> B \<Longrightarrow> Collect (case_prod A) \<subseteq> Collect (case_prod B)"
61422
0dfcd0fb4172 moved lemmas
haftmann
parents: 61378
diff changeset
  1045
  by auto (auto elim!: le_funE)
0dfcd0fb4172 moved lemmas
haftmann
parents: 61378
diff changeset
  1046
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1047
lemma Collect_split_mono_strong:
61422
0dfcd0fb4172 moved lemmas
haftmann
parents: 61378
diff changeset
  1048
  "X = fst ` A \<Longrightarrow> Y = snd ` A \<Longrightarrow> \<forall>a\<in>X. \<forall>b \<in> Y. P a b \<longrightarrow> Q a b
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1049
    \<Longrightarrow> A \<subseteq> Collect (case_prod P) \<Longrightarrow> A \<subseteq> Collect (case_prod Q)"
61422
0dfcd0fb4172 moved lemmas
haftmann
parents: 61378
diff changeset
  1050
  by fastforce
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1051
69275
9bbd5497befd clarified status of legacy input abbreviations
haftmann
parents: 69144
diff changeset
  1052
lemma UN_Times_distrib: "(\<Union>(a, b)\<in>A \<times> B. E a \<times> F b) = \<Union>(E ` A) \<times> \<Union>(F ` B)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61630
diff changeset
  1053
  \<comment> \<open>Suggested by Pierre Chartier\<close>
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1054
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1055
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1056
lemma split_paired_Ball_Sigma [simp, no_atp]: "(\<forall>z\<in>Sigma A B. P z) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B x. P (x, y))"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1057
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1058
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1059
lemma split_paired_Bex_Sigma [simp, no_atp]: "(\<exists>z\<in>Sigma A B. P z) \<longleftrightarrow> (\<exists>x\<in>A. \<exists>y\<in>B x. P (x, y))"
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1060
  by blast
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1061
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1062
lemma Sigma_Un_distrib1: "Sigma (I \<union> J) C = Sigma I C \<union> Sigma J C"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1063
  by blast
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1064
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1065
lemma Sigma_Un_distrib2: "(SIGMA i:I. A i \<union> B i) = Sigma I A \<union> Sigma I B"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1066
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1067
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1068
lemma Sigma_Int_distrib1: "Sigma (I \<inter> J) C = Sigma I C \<inter> Sigma J C"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1069
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1070
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1071
lemma Sigma_Int_distrib2: "(SIGMA i:I. A i \<inter> B i) = Sigma I A \<inter> Sigma I B"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1072
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1073
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1074
lemma Sigma_Diff_distrib1: "Sigma (I - J) C = Sigma I C - Sigma J C"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1075
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1076
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1077
lemma Sigma_Diff_distrib2: "(SIGMA i:I. A i - B i) = Sigma I A - Sigma I B"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1078
  by blast
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1079
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1080
lemma Sigma_Union: "Sigma (\<Union>X) B = (\<Union>A\<in>X. Sigma A B)"
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1081
  by blast
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1082
61630
608520e0e8e2 add various lemmas
Andreas Lochbihler
parents: 61425
diff changeset
  1083
lemma Pair_vimage_Sigma: "Pair x -` Sigma A f = (if x \<in> A then f x else {})"
608520e0e8e2 add various lemmas
Andreas Lochbihler
parents: 61425
diff changeset
  1084
  by auto
608520e0e8e2 add various lemmas
Andreas Lochbihler
parents: 61425
diff changeset
  1085
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1086
text \<open>
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1087
  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
  1088
  matching, especially when the rules are re-oriented.
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1089
\<close>
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1090
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1091
lemma Times_Un_distrib1: "(A \<union> B) \<times> C = A \<times> C \<union> B \<times> C "
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1092
  by (fact Sigma_Un_distrib1)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1093
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1094
lemma Times_Int_distrib1: "(A \<inter> B) \<times> C = A \<times> C \<inter> B \<times> C "
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1095
  by (fact Sigma_Int_distrib1)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1096
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1097
lemma Times_Diff_distrib1: "(A - B) \<times> C = A \<times> C - B \<times> C "
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1098
  by (fact Sigma_Diff_distrib1)
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1099
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1100
lemma Times_empty [simp]: "A \<times> B = {} \<longleftrightarrow> A = {} \<or> B = {}"
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
  1101
  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
  1102
69144
f13b82281715 new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents: 68467
diff changeset
  1103
lemma times_subset_iff: "A \<times> C \<subseteq> B \<times> D \<longleftrightarrow> A={} \<or> C={} \<or> A \<subseteq> B \<and> C \<subseteq> D"
f13b82281715 new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents: 68467
diff changeset
  1104
  by blast
f13b82281715 new theory Abstract_Topology with lots of stuff from HOL Light's metric.sml
paulson <lp15@cam.ac.uk>
parents: 68467
diff changeset
  1105
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1106
lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> (A = {} \<or> B = {}) \<and> (C = {} \<or> D = {})"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1107
  by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1108
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1109
lemma fst_image_times [simp]: "fst ` (A \<times> B) = (if B = {} then {} else A)"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
  1110
  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
  1111
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1112
lemma snd_image_times [simp]: "snd ` (A \<times> B) = (if A = {} then {} else B)"
44921
58eef4843641 tuned proofs
huffman
parents: 44066
diff changeset
  1113
  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
  1114
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1115
lemma fst_image_Sigma: "fst ` (Sigma A B) = {x \<in> A. B(x) \<noteq> {}}"
62379
340738057c8c An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents: 62139
diff changeset
  1116
  by force
340738057c8c An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents: 62139
diff changeset
  1117
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1118
lemma snd_image_Sigma: "snd ` (Sigma A B) = (\<Union> x \<in> A. B x)"
62379
340738057c8c An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents: 62139
diff changeset
  1119
  by force
340738057c8c An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
paulson <lp15@cam.ac.uk>
parents: 62139
diff changeset
  1120
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1121
lemma vimage_fst: "fst -` A = A \<times> UNIV"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1122
  by auto
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1123
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1124
lemma vimage_snd: "snd -` A = UNIV \<times> A"
56545
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1125
  by auto
8f1e7596deb7 more operations and lemmas
haftmann
parents: 56512
diff changeset
  1126
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69913
diff changeset
  1127
lemma insert_Times_insert [simp]:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1128
  "insert a A \<times> insert b B = insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)"
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1129
  by blast
26358
d6a508c16908 Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
haftmann
parents: 26340
diff changeset
  1130
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1131
lemma vimage_Times: "f -` (A \<times> B) = (fst \<circ> f) -` A \<inter> (snd \<circ> f) -` B"
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1132
proof (rule set_eqI)
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1133
  show "x \<in> f -` (A \<times> B) \<longleftrightarrow> x \<in> (fst \<circ> f) -` A \<inter> (snd \<circ> f) -` B" for x
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1134
    by (cases "f x") (auto split: prod.split)
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1135
qed
33271
7be66dee1a5a New theory Probability, which contains a development of measure theory
paulson
parents: 33089
diff changeset
  1136
69939
812ce526da33 new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  1137
lemma Times_Int_Times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1138
  by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49897
diff changeset
  1139
70044
da5857dbcbb9 More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
  1140
lemma image_paired_Times:
da5857dbcbb9 More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
  1141
   "(\<lambda>(x,y). (f x, g y)) ` (A \<times> B) = (f ` A) \<times> (g ` B)"
da5857dbcbb9 More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
  1142
  by auto
da5857dbcbb9 More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents: 69994
diff changeset
  1143
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1144
lemma product_swap: "prod.swap ` (A \<times> B) = B \<times> A"
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
  1145
  by (auto simp add: set_eq_iff)
35822
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1146
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1147
lemma swap_product: "(\<lambda>(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
56626
6532efd66a70 swap with qualifier;
haftmann
parents: 56545
diff changeset
  1148
  by (auto simp add: set_eq_iff)
35822
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1149
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1150
lemma image_split_eq_Sigma: "(\<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
  1151
proof (safe intro!: imageI)
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1152
  fix a b
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1153
  assume *: "a \<in> A" "b \<in> A" and eq: "f a = f b"
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
  1154
  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
  1155
    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
  1156
qed simp_all
35822
67e4de90d2c2 lemma swap_inj_on, swap_product
haftmann
parents: 35427
diff changeset
  1157
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62913
diff changeset
  1158
lemma subset_fst_snd: "A \<subseteq> (fst ` A \<times> snd ` A)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62913
diff changeset
  1159
  by force
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62913
diff changeset
  1160
60057
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1161
lemma inj_on_apfst [simp]: "inj_on (apfst f) (A \<times> UNIV) \<longleftrightarrow> inj_on f A"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1162
  by (auto simp add: inj_on_def)
60057
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1163
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1164
lemma inj_apfst [simp]: "inj (apfst f) \<longleftrightarrow> inj f"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1165
  using inj_on_apfst[of f UNIV] by simp
60057
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1166
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1167
lemma inj_on_apsnd [simp]: "inj_on (apsnd f) (UNIV \<times> A) \<longleftrightarrow> inj_on f A"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1168
  by (auto simp add: inj_on_def)
60057
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1169
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1170
lemma inj_apsnd [simp]: "inj (apsnd f) \<longleftrightarrow> inj f"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1171
  using inj_on_apsnd[of f UNIV] by simp
60057
86fa63ce8156 add lemmas
Andreas Lochbihler
parents: 59880
diff changeset
  1172
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1173
context
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1174
begin
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1175
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1176
qualified definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1177
  where [code_abbrev]: "product A B = A \<times> B"
46128
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1178
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1179
lemma member_product: "x \<in> Product_Type.product A B \<longleftrightarrow> x \<in> A \<times> B"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1180
  by (simp add: product_def)
46128
53e7cc599f58 interaction of set operations for execution and membership predicate
haftmann
parents: 46028
diff changeset
  1181
61127
76cd7f1ec257 tuned notation, proofs, namespace
haftmann
parents: 61126
diff changeset
  1182
end
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1183
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
  1184
text \<open>The following \<^const>\<open>map_prod\<close> lemmas are due to Joachim Breitner:\<close>
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1185
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1186
lemma map_prod_inj_on:
63575
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1187
  assumes "inj_on f A"
b9bd9e61fd63 misc tuning and modernization;
wenzelm
parents: 63566
diff changeset
  1188
    and "inj_on g B"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1189
  shows "inj_on (map_prod f g) (A \<times> B)"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1190
proof (rule inj_onI)
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1191
  fix x :: "'a \<times> 'c"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1192
  fix y :: "'a \<times> 'c"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1193
  assume "x \<in> A \<times> B"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1194
  then have "fst x \<in> A" and "snd x \<in> B" by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1195
  assume "y \<in> A \<times> B"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1196
  then have "fst y \<in> A" and "snd y \<in> B" by auto
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1197
  assume "map_prod f g x = map_prod f g y"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1198
  then have "fst (map_prod f g x) = fst (map_prod f g y)" by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1199
  then have "f (fst x) = f (fst y)" by (cases x, cases y) auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1200
  with \<open>inj_on f A\<close> and \<open>fst x \<in> A\<close> and \<open>fst y \<in> A\<close> have "fst x = fst y"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1201
    by (auto dest: inj_onD)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1202
  moreover from \<open>map_prod f g x = map_prod f g y\<close>
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1203
  have "snd (map_prod f g x) = snd (map_prod f g y)" by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1204
  then have "g (snd x) = g (snd y)" by (cases x, cases y) auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1205
  with \<open>inj_on g B\<close> and \<open>snd x \<in> B\<close> and \<open>snd y \<in> B\<close> have "snd x = snd y"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1206
    by (auto dest: inj_onD)
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1207
  ultimately show "x = y" by (rule prod_eqI)
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1208
qed
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1209
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1210
lemma map_prod_surj:
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1211
  fixes f :: "'a \<Rightarrow> 'b"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1212
    and g :: "'c \<Rightarrow> 'd"
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1213
  assumes "surj f" and "surj g"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1214
  shows "surj (map_prod f g)"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1215
  unfolding surj_def
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1216
proof
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1217
  fix y :: "'b \<times> 'd"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1218
  from \<open>surj f\<close> obtain a where "fst y = f a"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1219
    by (auto elim: surjE)
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1220
  moreover
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1221
  from \<open>surj g\<close> obtain b where "snd y = g b"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1222
    by (auto elim: surjE)
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1223
  ultimately have "(fst y, snd y) = map_prod f g (a,b)"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1224
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1225
  then show "\<exists>x. y = map_prod f g x"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1226
    by auto
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1227
qed
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1228
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1229
lemma map_prod_surj_on:
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1230
  assumes "f ` A = A'" and "g ` B = B'"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55642
diff changeset
  1231
  shows "map_prod f g ` (A \<times> B) = A' \<times> B'"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1232
  unfolding image_def
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1233
proof (rule set_eqI, rule iffI)
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1234
  fix x :: "'a \<times> 'c"
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 61032
diff changeset
  1235
  assume "x \<in> {y::'a \<times> 'c. \<exists>x::'b \<times> 'd\<in>A \<times> B. y = map_prod f g x}"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1236
  then obtain y where "y \<in> A \<times> B" and "x = map_prod f g y"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1237
    by blast
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1238
  from \<open>image f A = A'\<close> and \<open>y \<in> A \<times> B\<close> have "f (fst y) \<in> A'"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1239
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1240
  moreover from \<open>image g B = B'\<close> and \<open>y \<in> A \<times> B\<close> have "g (snd y) \<in> B'"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1241
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1242
  ultimately have "(f (fst y), g (snd y)) \<in> (A' \<times> B')"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1243
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1244
  with \<open>x = map_prod f g y\<close> show "x \<in> A' \<times> B'"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1245
    by (cases y) auto
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1246
next
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1247
  fix x :: "'a \<times> 'c"
63400
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1248
  assume "x \<in> A' \<times> B'"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1249
  then have "fst x \<in> A'" and "snd x \<in> B'"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1250
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1251
  from \<open>image f A = A'\<close> and \<open>fst x \<in> A'\<close> have "fst x \<in> image f A"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1252
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1253
  then obtain a where "a \<in> A" and "fst x = f a"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1254
    by (rule imageE)
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1255
  moreover from \<open>image g B = B'\<close> and \<open>snd x \<in> B'\<close> obtain b where "b \<in> B" and "snd x = g b"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1256
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1257
  ultimately have "(fst x, snd x) = map_prod f g (a, b)"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1258
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1259
  moreover from \<open>a \<in> A\<close> and  \<open>b \<in> B\<close> have "(a , b) \<in> A \<times> B"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1260
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1261
  ultimately have "\<exists>y \<in> A \<times> B. x = map_prod f g y"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1262
    by auto
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1263
  then show "x \<in> {x. \<exists>y \<in> A \<times> B. x = map_prod f g y}"
249fa34faba2 misc tuning and modernization;
wenzelm
parents: 63399
diff changeset
  1264
    by auto
40607
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1265
qed
30d512bf47a7 map_pair replaces prod_fun
haftmann
parents: 40590
diff changeset
  1266
21908
d02ba728cd56 moved code generator product setup here
haftmann
parents: 21454
diff changeset
  1267
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1268
subsection \<open>Simproc for rewriting a set comprehension into a pointfree expression\<close>
49764
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1269
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
  1270
ML_file \<open>Tools/set_comprehension_pointfree.ML\<close>
49764
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1271
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1272
setup \<open>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
  1273
  Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
  1274
    [Simplifier.make_simproc \<^context> "set comprehension"
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
  1275
      {lhss = [\<^term>\<open>Collect P\<close>],
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 62594
diff changeset
  1276
       proc = K Set_Comprehension_Pointfree.code_simproc}])
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1277
\<close>
49764
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1278
69994
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1279
subsection \<open>Lemmas about disjointness\<close>
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1280
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1281
lemma disjnt_Times1_iff [simp]: "disjnt (C \<times> A) (C \<times> B) \<longleftrightarrow> C = {} \<or> disjnt A B"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1282
  by (auto simp: disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1283
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1284
lemma disjnt_Times2_iff [simp]: "disjnt (A \<times> C) (B \<times> C) \<longleftrightarrow> C = {} \<or> disjnt A B"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1285
  by (auto simp: disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1286
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1287
lemma disjnt_Sigma_iff: "disjnt (Sigma A C) (Sigma B C) \<longleftrightarrow> (\<forall>i \<in> A\<inter>B. C i = {}) \<or> disjnt A B"
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1288
  by (auto simp: disjnt_def)
cf7150ab1075 more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents: 69939
diff changeset
  1289
49764
9979d64b8016 moving simproc from Finite_Set to more appropriate Product_Type theory
bulwahn
parents: 48891
diff changeset
  1290
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1291
subsection \<open>Inductively defined sets\<close>
15394
a2c34e6ca4f8 New code generator for let and split.
berghofe
parents: 15140
diff changeset
  1292
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1293
(* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1294
simproc_setup Collect_mem ("Collect t") = \<open>
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1295
  fn _ => fn ctxt => fn ct =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1296
    (case Thm.term_of ct of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
  1297
      S as Const (\<^const_name>\<open>Collect\<close>, Type (\<^type_name>\<open>fun\<close>, [_, T])) $ t =>
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1298
        let val (u, _, ps) = HOLogic.strip_ptupleabs t in
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1299
          (case u of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
  1300
            (c as Const (\<^const_name>\<open>Set.member\<close>, _)) $ q $ S' =>
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1301
              (case try (HOLogic.strip_ptuple ps) q of
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1302
                NONE => NONE
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1303
              | SOME ts =>
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1304
                  if not (Term.is_open S') andalso
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1305
                    ts = map Bound (length ps downto 0)
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1306
                  then
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1307
                    let val simp =
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1308
                      full_simp_tac (put_simpset HOL_basic_ss ctxt
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1309
                        addsimps [@{thm split_paired_all}, @{thm case_prod_conv}]) 1
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1310
                    in
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1311
                      SOME (Goal.prove ctxt [] []
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69275
diff changeset
  1312
                        (Const (\<^const_name>\<open>Pure.eq\<close>, T --> T --> propT) $ S $ S')
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1313
                        (K (EVERY
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59000
diff changeset
  1314
                          [resolve_tac ctxt [eq_reflection] 1,
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59000
diff changeset
  1315
                           resolve_tac ctxt @{thms subset_antisym} 1,
63399
d1742d1b7f0f more antiquotations;
wenzelm
parents: 63237
diff changeset
  1316
                           resolve_tac ctxt @{thms subsetI} 1,
d1742d1b7f0f more antiquotations;
wenzelm
parents: 63237
diff changeset
  1317
                           dresolve_tac ctxt @{thms CollectD} 1, simp,
d1742d1b7f0f more antiquotations;
wenzelm
parents: 63237
diff changeset
  1318
                           resolve_tac ctxt @{thms subsetI} 1,
d1742d1b7f0f more antiquotations;
wenzelm
parents: 63237
diff changeset
  1319
                           resolve_tac ctxt @{thms CollectI} 1, simp])))
56512
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1320
                    end
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1321
                  else NONE)
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1322
          | _ => NONE)
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1323
        end
9276da80f7c3 modernized simproc_setup;
wenzelm
parents: 56245
diff changeset
  1324
    | _ => NONE)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1325
\<close>
58389
ee1f45ca0d73 made new 'primrec' bootstrapping-capable
blanchet
parents: 58306
diff changeset
  1326
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
  1327
ML_file \<open>Tools/inductive_set.ML\<close>
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24286
diff changeset
  1328
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1329
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60057
diff changeset
  1330
subsection \<open>Legacy theorem bindings and duplicates\<close>
37166
e8400e31528a more coherent theory structure; tuned headings
haftmann
parents: 37136
diff changeset
  1331
55393
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
  1332
lemmas fst_conv = prod.sel(1)
ce5cebfaedda se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
blanchet
parents: 54630
diff changeset
  1333
lemmas snd_conv = prod.sel(2)
61032
b57df8eecad6 standardized some occurences of ancient "split" alias
haftmann
parents: 60758
diff changeset
  1334
lemmas split_def = case_prod_unfold
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1335
lemmas split_beta' = case_prod_beta'
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1336
lemmas split_beta = prod.case_eq_if
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1337
lemmas split_conv = case_prod_conv
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61422
diff changeset
  1338
lemmas split = case_prod_conv
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 43866
diff changeset
  1339
45204
5e4a1270c000 hide typedef-generated constants Product_Type.prod and Sum_Type.sum
huffman
parents: 44921
diff changeset
  1340
hide_const (open) prod
5e4a1270c000 hide typedef-generated constants Product_Type.prod and Sum_Type.sum
huffman
parents: 44921
diff changeset
  1341
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
  1342
end