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