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