src/HOL/Product_Type.thy
author wenzelm
Fri, 09 Oct 2015 20:26:03 +0200
changeset 61378 3e04c9ca001a
parent 61226 af7bed1360f3
child 61422 0dfcd0fb4172
permissions -rw-r--r--
discontinued specific HTML syntax;

(*  Title:      HOL/Product_Type.thy
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1992  University of Cambridge
*)

section \<open>Cartesian products\<close>

theory Product_Type
imports Typedef Inductive Fun
keywords "inductive_set" "coinductive_set" :: thy_decl
begin

subsection \<open>@{typ bool} is a datatype\<close>

free_constructors case_bool for True | False
  by auto

text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>

setup \<open>Sign.mandatory_path "old"\<close>

old_rep_datatype True False by (auto intro: bool_induct)

setup \<open>Sign.parent_path\<close>

text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>

setup \<open>Sign.mandatory_path "bool"\<close>

lemmas induct = old.bool.induct
lemmas inducts = old.bool.inducts
lemmas rec = old.bool.rec
lemmas simps = bool.distinct bool.case bool.rec

setup \<open>Sign.parent_path\<close>

declare case_split [cases type: bool]
  -- "prefer plain propositional version"

lemma
  shows [code]: "HOL.equal False P \<longleftrightarrow> \<not> P"
    and [code]: "HOL.equal True P \<longleftrightarrow> P" 
    and [code]: "HOL.equal P False \<longleftrightarrow> \<not> P"
    and [code]: "HOL.equal P True \<longleftrightarrow> P"
    and [code nbe]: "HOL.equal P P \<longleftrightarrow> True"
  by (simp_all add: equal)

lemma If_case_cert:
  assumes "CASE \<equiv> (\<lambda>b. If b f g)"
  shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)"
  using assms by simp_all

setup \<open>Code.add_case @{thm If_case_cert}\<close>

code_printing
  constant "HOL.equal :: bool \<Rightarrow> bool \<Rightarrow> bool" \<rightharpoonup> (Haskell) infix 4 "=="
| class_instance "bool" :: "equal" \<rightharpoonup> (Haskell) -


subsection \<open>The @{text unit} type\<close>

typedef unit = "{True}"
  by auto

definition Unity :: unit  ("'(')")
  where "() = Abs_unit True"

lemma unit_eq [no_atp]: "u = ()"
  by (induct u) (simp add: Unity_def)

text \<open>
  Simplification procedure for @{thm [source] unit_eq}.  Cannot use
  this rule directly --- it loops!
\<close>

simproc_setup unit_eq ("x::unit") = \<open>
  fn _ => fn _ => fn ct =>
    if HOLogic.is_unit (Thm.term_of ct) then NONE
    else SOME (mk_meta_eq @{thm unit_eq})
\<close>

free_constructors case_unit for "()"
  by auto

text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>

setup \<open>Sign.mandatory_path "old"\<close>

old_rep_datatype "()" by simp

setup \<open>Sign.parent_path\<close>

text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>

setup \<open>Sign.mandatory_path "unit"\<close>

lemmas induct = old.unit.induct
lemmas inducts = old.unit.inducts
lemmas rec = old.unit.rec
lemmas simps = unit.case unit.rec

setup \<open>Sign.parent_path\<close>

lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()"
  by simp

lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P"
  by (rule triv_forall_equality)

text \<open>
  This rewrite counters the effect of simproc @{text unit_eq} on @{term
  [source] "%u::unit. f u"}, replacing it by @{term [source]
  f} rather than by @{term [source] "%u. f ()"}.
\<close>

lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f"
  by (rule ext) simp

lemma UNIV_unit:
  "UNIV = {()}" by auto

instantiation unit :: default
begin

definition "default = ()"

instance ..

end

instantiation unit :: "{complete_boolean_algebra, complete_linorder, wellorder}"
begin

definition less_eq_unit :: "unit \<Rightarrow> unit \<Rightarrow> bool"
where
  "(_::unit) \<le> _ \<longleftrightarrow> True"

lemma less_eq_unit [iff]:
  "(u::unit) \<le> v"
  by (simp add: less_eq_unit_def)

definition less_unit :: "unit \<Rightarrow> unit \<Rightarrow> bool"
where
  "(_::unit) < _ \<longleftrightarrow> False"

lemma less_unit [iff]:
  "\<not> (u::unit) < v"
  by (simp_all add: less_eq_unit_def less_unit_def)

definition bot_unit :: unit
where
  [code_unfold]: "\<bottom> = ()"

definition top_unit :: unit
where
  [code_unfold]: "\<top> = ()"

definition inf_unit :: "unit \<Rightarrow> unit \<Rightarrow> unit"
where
  [simp]: "_ \<sqinter> _ = ()"

definition sup_unit :: "unit \<Rightarrow> unit \<Rightarrow> unit"
where
  [simp]: "_ \<squnion> _ = ()"

definition Inf_unit :: "unit set \<Rightarrow> unit"
where
  [simp]: "\<Sqinter>_ = ()"

definition Sup_unit :: "unit set \<Rightarrow> unit"
where
  [simp]: "\<Squnion>_ = ()"

definition uminus_unit :: "unit \<Rightarrow> unit"
where
  [simp]: "- _ = ()"

declare less_eq_unit_def [abs_def, code_unfold]
  less_unit_def [abs_def, code_unfold]
  inf_unit_def [abs_def, code_unfold]
  sup_unit_def [abs_def, code_unfold]
  Inf_unit_def [abs_def, code_unfold]
  Sup_unit_def [abs_def, code_unfold]
  uminus_unit_def [abs_def, code_unfold]

instance
  by intro_classes auto

end

lemma [code]:
  "HOL.equal (u::unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+

code_printing
  type_constructor unit \<rightharpoonup>
    (SML) "unit"
    and (OCaml) "unit"
    and (Haskell) "()"
    and (Scala) "Unit"
| constant Unity \<rightharpoonup>
    (SML) "()"
    and (OCaml) "()"
    and (Haskell) "()"
    and (Scala) "()"
| class_instance unit :: equal \<rightharpoonup>
    (Haskell) -
| constant "HOL.equal :: unit \<Rightarrow> unit \<Rightarrow> bool" \<rightharpoonup>
    (Haskell) infix 4 "=="

code_reserved SML
  unit

code_reserved OCaml
  unit

code_reserved Scala
  Unit


subsection \<open>The product type\<close>

subsubsection \<open>Type definition\<close>

definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
  "Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"

definition "prod = {f. \<exists>a b. f = Pair_Rep (a::'a) (b::'b)}"

typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \<Rightarrow> 'b \<Rightarrow> bool) set"
  unfolding prod_def by auto

type_notation (xsymbols)
  "prod"  ("(_ \<times>/ _)" [21, 20] 20)

definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b" where
  "Pair a b = Abs_prod (Pair_Rep a b)"

lemma prod_cases: "(\<And>a b. P (Pair a b)) \<Longrightarrow> P p"
  by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)

free_constructors uncurry for Pair fst snd
proof -
  fix P :: bool and p :: "'a \<times> 'b"
  show "(\<And>x1 x2. p = Pair x1 x2 \<Longrightarrow> P) \<Longrightarrow> P"
    by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
next
  fix a c :: 'a and b d :: 'b
  have "Pair_Rep a b = Pair_Rep c d \<longleftrightarrow> a = c \<and> b = d"
    by (auto simp add: Pair_Rep_def fun_eq_iff)
  moreover have "Pair_Rep a b \<in> prod" and "Pair_Rep c d \<in> prod"
    by (auto simp add: prod_def)
  ultimately show "Pair a b = Pair c d \<longleftrightarrow> a = c \<and> b = d"
    by (simp add: Pair_def Abs_prod_inject)
qed

text \<open>Avoid name clashes by prefixing the output of @{text old_rep_datatype} with @{text old}.\<close>

setup \<open>Sign.mandatory_path "old"\<close>

old_rep_datatype Pair
by (erule prod_cases) (rule prod.inject)

setup \<open>Sign.parent_path\<close>

text \<open>But erase the prefix for properties that are not generated by @{text free_constructors}.\<close>

setup \<open>Sign.mandatory_path "prod"\<close>

declare
  old.prod.inject[iff del]

lemmas induct = old.prod.induct
lemmas inducts = old.prod.inducts
lemmas rec = old.prod.rec
lemmas simps = prod.inject prod.case prod.rec

setup \<open>Sign.parent_path\<close>

declare prod.case [nitpick_simp del]
declare prod.case_cong_weak [cong del]


subsubsection \<open>Tuple syntax\<close>

text \<open>
  Patterns -- extends pre-defined type @{typ pttrn} used in
  abstractions.
\<close>

nonterminal tuple_args and patterns

syntax
  "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
  "_tuple_arg"  :: "'a => tuple_args"                   ("_")
  "_tuple_args" :: "'a => tuple_args => tuple_args"     ("_,/ _")
  "_pattern"    :: "[pttrn, patterns] => pttrn"         ("'(_,/ _')")
  ""            :: "pttrn => patterns"                  ("_")
  "_patterns"   :: "[pttrn, patterns] => patterns"      ("_,/ _")

translations
  "(x, y)" \<rightleftharpoons> "CONST Pair x y"
  "_pattern x y" \<rightleftharpoons> "CONST Pair x y"
  "_patterns x y" \<rightleftharpoons> "CONST Pair x y"
  "_tuple x (_tuple_args y z)" \<rightleftharpoons> "_tuple x (_tuple_arg (_tuple y z))"
  "\<lambda>(x, y, zs). b" \<rightleftharpoons> "CONST uncurry (\<lambda>x (y, zs). b)"
  "\<lambda>(x, y). b" \<rightleftharpoons> "CONST uncurry (\<lambda>x y. b)"
  "_abs (CONST Pair x y) t" \<rightharpoonup> "\<lambda>(x, y). t"
  -- \<open>This rule accommodates tuples in @{text "case C \<dots> (x, y) \<dots> \<Rightarrow> \<dots>"}:
     The @{text "(x, y)"} is parsed as @{text "Pair x y"} because it is @{text logic},
     not @{text pttrn}.\<close>

text \<open>print @{term "uncurry f"} as @{term "\<lambda>(x, y). f x y"} and
  @{term "uncurry (\<lambda>x. f x)"} as @{term "\<lambda>(x, y). f x y"}\<close>

typed_print_translation \<open>
  let
    fun uncurry_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match
      | uncurry_guess_names_tr' T [Abs (x, xT, t)] =
          (case (head_of t) of
            Const (@{const_syntax uncurry}, _) => raise Match
          | _ =>
            let 
              val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
              val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0);
              val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t');
            in
              Syntax.const @{syntax_const "_abs"} $
                (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
            end)
      | uncurry_guess_names_tr' T [t] =
          (case head_of t of
            Const (@{const_syntax uncurry}, _) => raise Match
          | _ =>
            let
              val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match;
              val (y, t') =
                Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0);
              val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t');
            in
              Syntax.const @{syntax_const "_abs"} $
                (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
            end)
      | uncurry_guess_names_tr' _ _ = raise Match;
  in [(@{const_syntax uncurry}, K uncurry_guess_names_tr')] end
\<close>


subsubsection \<open>Code generator setup\<close>

code_printing
  type_constructor prod \<rightharpoonup>
    (SML) infix 2 "*"
    and (OCaml) infix 2 "*"
    and (Haskell) "!((_),/ (_))"
    and (Scala) "((_),/ (_))"
| constant Pair \<rightharpoonup>
    (SML) "!((_),/ (_))"
    and (OCaml) "!((_),/ (_))"
    and (Haskell) "!((_),/ (_))"
    and (Scala) "!((_),/ (_))"
| class_instance  prod :: equal \<rightharpoonup>
    (Haskell) -
| constant "HOL.equal :: 'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool" \<rightharpoonup>
    (Haskell) infix 4 "=="


subsubsection \<open>Fundamental operations and properties\<close>

lemma Pair_inject:
  assumes "(a, b) = (a', b')"
    and "a = a' ==> b = b' ==> R"
  shows R
  using assms by simp

lemma surj_pair [simp]: "EX x y. p = (x, y)"
  by (cases p) simp

code_printing
  constant fst \<rightharpoonup> (Haskell) "fst"
| constant snd \<rightharpoonup> (Haskell) "snd"

lemma case_prod_unfold [nitpick_unfold]: "uncurry = (%c p. c (fst p) (snd p))"
  by (simp add: fun_eq_iff split: prod.split)

lemma fst_eqD: "fst (x, y) = a ==> x = a"
  by simp

lemma snd_eqD: "snd (x, y) = a ==> y = a"
  by simp

lemmas surjective_pairing = prod.collapse [symmetric]

lemma prod_eq_iff: "s = t \<longleftrightarrow> fst s = fst t \<and> snd s = snd t"
  by (cases s, cases t) simp

lemma prod_eqI [intro?]: "fst p = fst q \<Longrightarrow> snd p = snd q \<Longrightarrow> p = q"
  by (simp add: prod_eq_iff)

lemma split_conv [simp, code]: "(case (a, b) of (c, d) \<Rightarrow> f c d) = f a b"
  by (fact prod.case)

lemma splitI: "f a b \<Longrightarrow> case (a, b) of (c, d) \<Rightarrow> f c d"
  by (rule split_conv [THEN iffD2])

lemma splitD: "(case (a, b) of (c, d) \<Rightarrow> f c d) \<Longrightarrow> f a b"
  by (rule split_conv [THEN iffD1])

lemma split_Pair [simp]: "uncurry Pair = id"
  by (simp add: fun_eq_iff split: prod.split)

lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f"
  -- \<open>Subsumes the old @{text split_Pair} when @{term f} is the identity function.\<close>
  by (simp add: fun_eq_iff split: prod.split)

lemma split_comp: "(case x of (a, b) \<Rightarrow> (f \<circ> g) a b) = f (g (fst x)) (snd x)"
  by (cases x) simp

lemma split_twice: "uncurry f (uncurry g p) = uncurry (\<lambda>x y. uncurry f (g x y)) p"
  by (fact prod.case_distrib)

lemma The_split: "The (uncurry P) = (THE xy. P (fst xy) (snd xy))"
  by (simp add: case_prod_unfold)

lemmas split_weak_cong = prod.case_cong_weak
  -- \<open>Prevents simplification of @{term c}: much faster\<close>

lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g"
  by (simp add: split_eta)

lemma split_paired_all [no_atp]: "(!!x. PROP P x) == (!!a b. PROP P (a, b))"
proof
  fix a b
  assume "!!x. PROP P x"
  then show "PROP P (a, b)" .
next
  fix x
  assume "!!a b. PROP P (a, b)"
  from \<open>PROP P (fst x, snd x)\<close> show "PROP P x" by simp
qed

lemma uncurry_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
  by (cases x) simp

text \<open>
  The rule @{thm [source] split_paired_all} does not work with the
  Simplifier because it also affects premises in congrence rules,
  where this can lead to premises of the form @{text "!!a b. ... =
  ?P(a, b)"} which cannot be solved by reflexivity.
\<close>

lemmas split_tupled_all = split_paired_all unit_all_eq2

ML \<open>
  (* replace parameters of product type by individual component parameters *)
  local (* filtering with exists_paired_all is an essential optimization *)
    fun exists_paired_all (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) =
          can HOLogic.dest_prodT T orelse exists_paired_all t
      | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u
      | exists_paired_all (Abs (_, _, t)) = exists_paired_all t
      | exists_paired_all _ = false;
    val ss =
      simpset_of
       (put_simpset HOL_basic_ss @{context}
        addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}]
        addsimprocs [@{simproc unit_eq}]);
  in
    fun split_all_tac ctxt = SUBGOAL (fn (t, i) =>
      if exists_paired_all t then safe_full_simp_tac (put_simpset ss ctxt) i else no_tac);

    fun unsafe_split_all_tac ctxt = SUBGOAL (fn (t, i) =>
      if exists_paired_all t then full_simp_tac (put_simpset ss ctxt) i else no_tac);

    fun split_all ctxt th =
      if exists_paired_all (Thm.prop_of th)
      then full_simplify (put_simpset ss ctxt) th else th;
  end;
\<close>

setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\<close>

lemma split_paired_All [simp, no_atp]: "(ALL x. P x) = (ALL a b. P (a, b))"
  -- \<open>@{text "[iff]"} is not a good idea because it makes @{text blast} loop\<close>
  by fast

lemma split_paired_Ex [simp, no_atp]: "(EX x. P x) = (EX a b. P (a, b))"
  by fast

lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))"
  -- \<open>Can't be added to simpset: loops!\<close>
  by (simp add: split_eta)

text \<open>
  Simplification procedure for @{thm [source] cond_split_eta}.  Using
  @{thm [source] split_eta} as a rewrite rule is not general enough,
  and using @{thm [source] cond_split_eta} directly would render some
  existing proofs very inefficient; similarly for @{text
  split_beta}.
\<close>

ML \<open>
local
  val cond_split_eta_ss =
    simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms cond_split_eta});
  fun Pair_pat k 0 (Bound m) = (m = k)
    | Pair_pat k i (Const (@{const_name Pair},  _) $ Bound m $ t) =
        i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t
    | Pair_pat _ _ _ = false;
  fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t
    | no_args k i (t $ u) = no_args k i t andalso no_args k i u
    | no_args k i (Bound m) = m < k orelse m > k + i
    | no_args _ _ _ = true;
  fun split_pat tp i (Abs  (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE
    | split_pat tp i (Const (@{const_name uncurry}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
    | split_pat tp i _ = NONE;
  fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] []
        (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
        (K (simp_tac (put_simpset cond_split_eta_ss ctxt) 1)));

  fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t
    | beta_term_pat k i (t $ u) =
        Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u)
    | beta_term_pat k i t = no_args k i t;
  fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg
    | eta_term_pat _ _ _ = false;
  fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t)
    | subst arg k i (t $ u) =
        if Pair_pat k i (t $ u) then incr_boundvars k arg
        else (subst arg k i t $ subst arg k i u)
    | subst arg k i t = t;
in
  fun beta_proc ctxt (s as Const (@{const_name uncurry}, _) $ Abs (_, _, t) $ arg) =
        (case split_pat beta_term_pat 1 t of
          SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f))
        | NONE => NONE)
    | beta_proc _ _ = NONE;
  fun eta_proc ctxt (s as Const (@{const_name uncurry}, _) $ Abs (_, _, t)) =
        (case split_pat eta_term_pat 1 t of
          SOME (_, ft) => SOME (metaeq ctxt s (let val f $ _ = ft in f end))
        | NONE => NONE)
    | eta_proc _ _ = NONE;
end;
\<close>
simproc_setup split_beta ("split f z") =
  \<open>fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct)\<close>
simproc_setup split_eta ("split f") =
  \<open>fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct)\<close>

lemmas split_beta [mono] = prod.case_eq_if

lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))"
  by (auto simp: fun_eq_iff)

lemmas split_split [no_atp] = prod.split
  -- \<open>For use with @{text split} and the Simplifier.\<close>

text \<open>
  @{thm [source] split_split} could be declared as @{text "[split]"}
  done after the Splitter has been speeded up significantly;
  precompute the constants involved and don't do anything unless the
  current goal contains one of those constants.
\<close>

lemmas split_split_asm [no_atp] = prod.split_asm

text \<open>
  \medskip @{const uncurry} used as a logical connective or set former.

  \medskip These rules are for use with @{text blast}; could instead
  call @{text simp} using @{thm [source] prod.split} as rewrite.\<close>

lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> case p of (a, b) \<Rightarrow> c a b"
  apply (simp only: split_tupled_all)
  apply (simp (no_asm_simp))
  done

lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> (case p of (a, b) \<Rightarrow> c a b) x"
  apply (simp only: split_tupled_all)
  apply (simp (no_asm_simp))
  done

lemma splitE: "(case p of (a, b) \<Rightarrow> c a b) ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
  by (induct p) auto

lemma splitE': "(case p of (a, b) \<Rightarrow> c a b) z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
  by (induct p) auto

lemma splitE2:
  "[| Q (case z of (a, b) \<Rightarrow> P a b);  !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R"
proof -
  assume q: "Q (uncurry P z)"
  assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R"
  show R
    apply (rule r surjective_pairing)+
    apply (rule split_beta [THEN subst], rule q)
    done
qed

lemma splitD':
  "(case (a, b) of (c, d) \<Rightarrow> R c d) c \<Longrightarrow> R a b c"
  by simp

lemma mem_splitI:
  "z \<in> c a b \<Longrightarrow> z \<in> (case (a, b) of (d, e) \<Rightarrow> c d e)"
  by simp

lemma mem_splitI2:
  "\<And>p. (\<And>a b. p = (a, b) \<Longrightarrow> z \<in> c a b) \<Longrightarrow> z \<in> (case p of (a, b) \<Rightarrow> c a b)"
  by (simp only: split_tupled_all) simp

lemma mem_splitE:
  assumes "z \<in> uncurry c p"
  obtains x y where "p = (x, y)" and "z \<in> c x y"
  using assms by (rule splitE2)

declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!]
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!]

ML \<open>
local (* filtering with exists_p_split is an essential optimization *)
  fun exists_p_split (Const (@{const_name uncurry},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
    | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u
    | exists_p_split (Abs (_, _, t)) = exists_p_split t
    | exists_p_split _ = false;
in
fun split_conv_tac ctxt = SUBGOAL (fn (t, i) =>
  if exists_p_split t
  then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_conv}) i
  else no_tac);
end;
\<close>

(* This prevents applications of splitE for already splitted arguments leading
   to quite time-consuming computations (in particular for nested tuples) *)
setup \<open>map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac))\<close>

lemma split_eta_SetCompr [simp, no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
  by (rule ext) fast

lemma split_eta_SetCompr2 [simp, no_atp]: "(%u. EX x y. u = (x, y) & P x y) = uncurry P"
  by (rule ext) fast

lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & uncurry Q ab)"
  -- \<open>Allows simplifications of nested splits in case of independent predicates.\<close>
  by (rule ext) blast

(* Do NOT make this a simp rule as it
   a) only helps in special situations
   b) can lead to nontermination in the presence of split_def
*)
lemma split_comp_eq: 
  fixes f :: "'a => 'b => 'c" and g :: "'d => 'a"
  shows "(%u. f (g (fst u)) (snd u)) = (uncurry (%x. f (g x)))"
  by (rule ext) auto

lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A"
  apply (rule_tac x = "(a, b)" in image_eqI)
   apply auto
  done

lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)"
  by blast

(*
the following  would be slightly more general,
but cannot be used as rewrite rule:
### Cannot add premise as rewrite rule because it contains (type) unknowns:
### ?y = .x
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)"
by (rtac some_equality 1)
by ( Simp_tac 1)
by (split_all_tac 1)
by (Asm_full_simp_tac 1)
qed "The_split_eq";
*)

text \<open>
  Setup of internal @{text split_rule}.
\<close>

lemmas case_prodI = prod.case [THEN iffD2]

lemma case_prodI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> uncurry c p"
  by (fact splitI2)

lemma case_prodI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> uncurry c p x"
  by (fact splitI2')

lemma case_prodE: "uncurry c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q"
  by (fact splitE)

lemma case_prodE': "uncurry c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q"
  by (fact splitE')

declare case_prodI [intro!]

lemma case_prod_beta:
  "uncurry f p = f (fst p) (snd p)"
  by (fact split_beta)

lemma prod_cases3 [cases type]:
  obtains (fields) a b c where "y = (a, b, c)"
  by (cases y, case_tac b) blast

lemma prod_induct3 [case_names fields, induct type]:
    "(!!a b c. P (a, b, c)) ==> P x"
  by (cases x) blast

lemma prod_cases4 [cases type]:
  obtains (fields) a b c d where "y = (a, b, c, d)"
  by (cases y, case_tac c) blast

lemma prod_induct4 [case_names fields, induct type]:
    "(!!a b c d. P (a, b, c, d)) ==> P x"
  by (cases x) blast

lemma prod_cases5 [cases type]:
  obtains (fields) a b c d e where "y = (a, b, c, d, e)"
  by (cases y, case_tac d) blast

lemma prod_induct5 [case_names fields, induct type]:
    "(!!a b c d e. P (a, b, c, d, e)) ==> P x"
  by (cases x) blast

lemma prod_cases6 [cases type]:
  obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)"
  by (cases y, case_tac e) blast

lemma prod_induct6 [case_names fields, induct type]:
    "(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x"
  by (cases x) blast

lemma prod_cases7 [cases type]:
  obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)"
  by (cases y, case_tac f) blast

lemma prod_induct7 [case_names fields, induct type]:
    "(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x"
  by (cases x) blast

definition internal_split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
  "internal_split == uncurry"

lemma internal_split_conv: "internal_split c (a, b) = c a b"
  by (simp only: internal_split_def split_conv)

ML_file "Tools/split_rule.ML"

hide_const internal_split


subsubsection \<open>Derived operations\<close>

definition curry    :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c" where
  "curry = (\<lambda>c x y. c (x, y))"

lemma curry_conv [simp, code]: "curry f a b = f (a, b)"
  by (simp add: curry_def)

lemma curryI [intro!]: "f (a, b) \<Longrightarrow> curry f a b"
  by (simp add: curry_def)

lemma curryD [dest!]: "curry f a b \<Longrightarrow> f (a, b)"
  by (simp add: curry_def)

lemma curryE: "curry f a b \<Longrightarrow> (f (a, b) \<Longrightarrow> Q) \<Longrightarrow> Q"
  by (simp add: curry_def)

lemma curry_split [simp]: "curry (uncurry f) = f"
  by (simp add: curry_def case_prod_unfold)

lemma split_curry [simp]: "uncurry (curry f) = f"
  by (simp add: curry_def case_prod_unfold)

lemma curry_K: "curry (\<lambda>x. c) = (\<lambda>x y. c)"
by(simp add: fun_eq_iff)

text \<open>
  The composition-uncurry combinator.
\<close>

notation fcomp (infixl "\<circ>>" 60)

definition scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd" (infixl "\<circ>\<rightarrow>" 60) where
  "f \<circ>\<rightarrow> g = (\<lambda>x. uncurry g (f x))"

lemma scomp_unfold: "scomp = (\<lambda>f g x. g (fst (f x)) (snd (f x)))"
  by (simp add: fun_eq_iff scomp_def case_prod_unfold)

lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = uncurry g (f x)"
  by (simp add: scomp_unfold case_prod_unfold)

lemma Pair_scomp: "Pair x \<circ>\<rightarrow> f = f x"
  by (simp add: fun_eq_iff)

lemma scomp_Pair: "x \<circ>\<rightarrow> Pair = x"
  by (simp add: fun_eq_iff)

lemma scomp_scomp: "(f \<circ>\<rightarrow> g) \<circ>\<rightarrow> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>\<rightarrow> h)"
  by (simp add: fun_eq_iff scomp_unfold)

lemma scomp_fcomp: "(f \<circ>\<rightarrow> g) \<circ>> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>> h)"
  by (simp add: fun_eq_iff scomp_unfold fcomp_def)

lemma fcomp_scomp: "(f \<circ>> g) \<circ>\<rightarrow> h = f \<circ>> (g \<circ>\<rightarrow> h)"
  by (simp add: fun_eq_iff scomp_unfold)

code_printing
  constant scomp \<rightharpoonup> (Eval) infixl 3 "#->"

no_notation fcomp (infixl "\<circ>>" 60)
no_notation scomp (infixl "\<circ>\<rightarrow>" 60)

text \<open>
  @{term map_prod} --- action of the product functor upon
  functions.
\<close>

definition map_prod :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
  "map_prod f g = (\<lambda>(x, y). (f x, g y))"

lemma map_prod_simp [simp, code]:
  "map_prod f g (a, b) = (f a, g b)"
  by (simp add: map_prod_def)

functor map_prod: map_prod
  by (auto simp add: split_paired_all)

lemma fst_map_prod [simp]:
  "fst (map_prod f g x) = f (fst x)"
  by (cases x) simp_all

lemma snd_map_prod [simp]:
  "snd (map_prod f g x) = g (snd x)"
  by (cases x) simp_all

lemma fst_comp_map_prod [simp]:
  "fst \<circ> map_prod f g = f \<circ> fst"
  by (rule ext) simp_all

lemma snd_comp_map_prod [simp]:
  "snd \<circ> map_prod f g = g \<circ> snd"
  by (rule ext) simp_all

lemma map_prod_compose:
  "map_prod (f1 o f2) (g1 o g2) = (map_prod f1 g1 o map_prod f2 g2)"
  by (rule ext) (simp add: map_prod.compositionality comp_def)

lemma map_prod_ident [simp]:
  "map_prod (%x. x) (%y. y) = (%z. z)"
  by (rule ext) (simp add: map_prod.identity)

lemma map_prod_imageI [intro]:
  "(a, b) \<in> R \<Longrightarrow> (f a, g b) \<in> map_prod f g ` R"
  by (rule image_eqI) simp_all

lemma prod_fun_imageE [elim!]:
  assumes major: "c \<in> map_prod f g ` R"
    and cases: "\<And>x y. c = (f x, g y) \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> P"
  shows P
  apply (rule major [THEN imageE])
  apply (case_tac x)
  apply (rule cases)
  apply simp_all
  done

definition apfst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b" where
  "apfst f = map_prod f id"

definition apsnd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c" where
  "apsnd f = map_prod id f"

lemma apfst_conv [simp, code]:
  "apfst f (x, y) = (f x, y)" 
  by (simp add: apfst_def)

lemma apsnd_conv [simp, code]:
  "apsnd f (x, y) = (x, f y)" 
  by (simp add: apsnd_def)

lemma fst_apfst [simp]:
  "fst (apfst f x) = f (fst x)"
  by (cases x) simp

lemma fst_comp_apfst [simp]:
  "fst \<circ> apfst f = f \<circ> fst"
  by (simp add: fun_eq_iff)

lemma fst_apsnd [simp]:
  "fst (apsnd f x) = fst x"
  by (cases x) simp

lemma fst_comp_apsnd [simp]:
  "fst \<circ> apsnd f = fst"
  by (simp add: fun_eq_iff)

lemma snd_apfst [simp]:
  "snd (apfst f x) = snd x"
  by (cases x) simp

lemma snd_comp_apfst [simp]:
  "snd \<circ> apfst f = snd"
  by (simp add: fun_eq_iff)

lemma snd_apsnd [simp]:
  "snd (apsnd f x) = f (snd x)"
  by (cases x) simp

lemma snd_comp_apsnd [simp]:
  "snd \<circ> apsnd f = f \<circ> snd"
  by (simp add: fun_eq_iff)

lemma apfst_compose:
  "apfst f (apfst g x) = apfst (f \<circ> g) x"
  by (cases x) simp

lemma apsnd_compose:
  "apsnd f (apsnd g x) = apsnd (f \<circ> g) x"
  by (cases x) simp

lemma apfst_apsnd [simp]:
  "apfst f (apsnd g x) = (f (fst x), g (snd x))"
  by (cases x) simp

lemma apsnd_apfst [simp]:
  "apsnd f (apfst g x) = (g (fst x), f (snd x))"
  by (cases x) simp

lemma apfst_id [simp] :
  "apfst id = id"
  by (simp add: fun_eq_iff)

lemma apsnd_id [simp] :
  "apsnd id = id"
  by (simp add: fun_eq_iff)

lemma apfst_eq_conv [simp]:
  "apfst f x = apfst g x \<longleftrightarrow> f (fst x) = g (fst x)"
  by (cases x) simp

lemma apsnd_eq_conv [simp]:
  "apsnd f x = apsnd g x \<longleftrightarrow> f (snd x) = g (snd x)"
  by (cases x) simp

lemma apsnd_apfst_commute:
  "apsnd f (apfst g p) = apfst g (apsnd f p)"
  by simp

context
begin

local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")\<close>

definition swap :: "'a \<times> 'b \<Rightarrow> 'b \<times> 'a"
where
  "swap p = (snd p, fst p)"

end

lemma swap_simp [simp]:
  "prod.swap (x, y) = (y, x)"
  by (simp add: prod.swap_def)

lemma swap_swap [simp]:
  "prod.swap (prod.swap p) = p"
  by (cases p) simp

lemma swap_comp_swap [simp]:
  "prod.swap \<circ> prod.swap = id"
  by (simp add: fun_eq_iff)

lemma pair_in_swap_image [simp]:
  "(y, x) \<in> prod.swap ` A \<longleftrightarrow> (x, y) \<in> A"
  by (auto intro!: image_eqI)

lemma inj_swap [simp]:
  "inj_on prod.swap A"
  by (rule inj_onI) auto

lemma swap_inj_on:
  "inj_on (\<lambda>(i, j). (j, i)) A"
  by (rule inj_onI) auto

lemma surj_swap [simp]:
  "surj prod.swap"
  by (rule surjI [of _ prod.swap]) simp

lemma bij_swap [simp]:
  "bij prod.swap"
  by (simp add: bij_def)

lemma case_swap [simp]:
  "(case prod.swap p of (y, x) \<Rightarrow> f x y) = (case p of (x, y) \<Rightarrow> f x y)"
  by (cases p) simp

text \<open>
  Disjoint union of a family of sets -- Sigma.
\<close>

definition Sigma :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<times> 'b) set" where
  Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"

abbreviation
  Times :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"
    (infixr "<*>" 80) where
  "A <*> B == Sigma A (%_. B)"

notation (xsymbols)
  Times  (infixr "\<times>" 80)

hide_const (open) Times

syntax
  "_Sigma" :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"  ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
translations
  "SIGMA x:A. B" == "CONST Sigma A (%x. B)"

lemma SigmaI [intro!]: "[| a:A;  b:B(a) |] ==> (a,b) : Sigma A B"
  by (unfold Sigma_def) blast

lemma SigmaE [elim!]:
    "[| c: Sigma A B;
        !!x y.[| x:A;  y:B(x);  c=(x,y) |] ==> P
     |] ==> P"
  -- \<open>The general elimination rule.\<close>
  by (unfold Sigma_def) blast

text \<open>
  Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
  eigenvariables.
\<close>

lemma SigmaD1: "(a, b) : Sigma A B ==> a : A"
  by blast

lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a"
  by blast

lemma SigmaE2:
    "[| (a, b) : Sigma A B;
        [| a:A;  b:B(a) |] ==> P
     |] ==> P"
  by blast

lemma Sigma_cong:
     "\<lbrakk>A = B; !!x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk>
      \<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)"
  by auto

lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D"
  by blast

lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
  by blast

lemma Sigma_empty2 [simp]: "A <*> {} = {}"
  by blast

lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV"
  by auto

lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)"
  by auto

lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV"
  by auto

lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))"
  by blast

lemma Sigma_empty_iff: "(SIGMA i:I. X i) = {} \<longleftrightarrow> (\<forall>i\<in>I. X i = {})"
  by auto

lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)"
  by blast

lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)"
  by (blast elim: equalityE)

lemma SetCompr_Sigma_eq:
  "{(x, y). P x \<and> Q x y} = (SIGMA x:Collect P. Collect (Q x))"
  by blast

lemma Collect_split [simp]:
  "{(a, b). P a \<and> Q b} = Collect P \<times> Collect Q "
  by (fact SetCompr_Sigma_eq)

lemma UN_Times_distrib:
  "(\<Union>(a, b)\<in>A \<times> B. E a \<times> F b) = UNION A E \<times> UNION B F"
  -- \<open>Suggested by Pierre Chartier\<close>
  by blast

lemma split_paired_Ball_Sigma [simp, no_atp]:
  "(\<forall>z\<in>Sigma A B. P z) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B x. P (x, y))"
  by blast

lemma split_paired_Bex_Sigma [simp, no_atp]:
  "(\<exists>z\<in>Sigma A B. P z) \<longleftrightarrow> (\<exists>x\<in>A. \<exists>y\<in>B x. P (x, y))"
  by blast

lemma Sigma_Un_distrib1:
  "Sigma (I \<union> J) C = Sigma I C \<union> Sigma J C"
  by blast

lemma Sigma_Un_distrib2:
  "(SIGMA i:I. A i \<union> B i) = Sigma I A \<union> Sigma I B"
  by blast

lemma Sigma_Int_distrib1:
  "Sigma (I \<inter> J) C = Sigma I C \<inter> Sigma J C"
  by blast

lemma Sigma_Int_distrib2:
  "(SIGMA i:I. A i \<inter> B i) = Sigma I A \<inter> Sigma I B"
  by blast

lemma Sigma_Diff_distrib1:
  "Sigma (I - J) C = Sigma I C - Sigma J C"
  by blast

lemma Sigma_Diff_distrib2:
  "(SIGMA i:I. A i - B i) = Sigma I A - Sigma I B"
  by blast

lemma Sigma_Union:
  "Sigma (\<Union>X) B = (\<Union>A\<in>X. Sigma A B)"
  by blast

text \<open>
  Non-dependent versions are needed to avoid the need for higher-order
  matching, especially when the rules are re-oriented.
\<close>

lemma Times_Un_distrib1:
  "(A \<union> B) \<times> C = A \<times> C \<union> B \<times> C "
  by (fact Sigma_Un_distrib1)

lemma Times_Int_distrib1:
  "(A \<inter> B) \<times> C = A \<times> C \<inter> B \<times> C "
  by (fact Sigma_Int_distrib1)

lemma Times_Diff_distrib1:
  "(A - B) \<times> C = A \<times> C - B \<times> C "
  by (fact Sigma_Diff_distrib1)

lemma Times_empty [simp]:
  "A \<times> B = {} \<longleftrightarrow> A = {} \<or> B = {}"
  by auto

lemma times_eq_iff:
  "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> (A = {} \<or> B = {}) \<and> (C = {} \<or> D = {})"
  by auto

lemma fst_image_times [simp]:
  "fst ` (A \<times> B) = (if B = {} then {} else A)"
  by force

lemma snd_image_times [simp]:
  "snd ` (A \<times> B) = (if A = {} then {} else B)"
  by force

lemma vimage_fst:
  "fst -` A = A \<times> UNIV"
  by auto

lemma vimage_snd:
  "snd -` A = UNIV \<times> A"
  by auto

lemma insert_times_insert[simp]:
  "insert a A \<times> insert b B =
   insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)"
  by blast

lemma vimage_Times:
  "f -` (A \<times> B) = (fst \<circ> f) -` A \<inter> (snd \<circ> f) -` B"
proof (rule set_eqI)
  fix x
  show "x \<in> f -` (A \<times> B) \<longleftrightarrow> x \<in> (fst \<circ> f) -` A \<inter> (snd \<circ> f) -` B"
    by (cases "f x") (auto split: prod.split)
qed

lemma times_Int_times:
  "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
  by auto

lemma product_swap:
  "prod.swap ` (A \<times> B) = B \<times> A"
  by (auto simp add: set_eq_iff)

lemma swap_product:
  "(\<lambda>(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
  by (auto simp add: set_eq_iff)

lemma image_split_eq_Sigma:
  "(\<lambda>x. (f x, g x)) ` A = Sigma (f ` A) (\<lambda>x. g ` (f -` {x} \<inter> A))"
proof (safe intro!: imageI)
  fix a b assume *: "a \<in> A" "b \<in> A" and eq: "f a = f b"
  show "(f b, g a) \<in> (\<lambda>x. (f x, g x)) ` A"
    using * eq[symmetric] by auto
qed simp_all

lemma inj_on_apfst [simp]: "inj_on (apfst f) (A \<times> UNIV) \<longleftrightarrow> inj_on f A"
by(auto simp add: inj_on_def)

lemma inj_apfst [simp]: "inj (apfst f) \<longleftrightarrow> inj f"
using inj_on_apfst[of f UNIV] by simp

lemma inj_on_apsnd [simp]: "inj_on (apsnd f) (UNIV \<times> A) \<longleftrightarrow> inj_on f A"
by(auto simp add: inj_on_def)

lemma inj_apsnd [simp]: "inj (apsnd f) \<longleftrightarrow> inj f"
using inj_on_apsnd[of f UNIV] by simp

context
begin

qualified definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
  [code_abbrev]: "product A B = A \<times> B"

lemma member_product:
  "x \<in> Product_Type.product A B \<longleftrightarrow> x \<in> A \<times> B"
  by (simp add: Product_Type.product_def)

end
  
text \<open>The following @{const map_prod} lemmas are due to Joachim Breitner:\<close>

lemma map_prod_inj_on:
  assumes "inj_on f A" and "inj_on g B"
  shows "inj_on (map_prod f g) (A \<times> B)"
proof (rule inj_onI)
  fix x :: "'a \<times> 'c" and y :: "'a \<times> 'c"
  assume "x \<in> A \<times> B" hence "fst x \<in> A" and "snd x \<in> B" by auto
  assume "y \<in> A \<times> B" hence "fst y \<in> A" and "snd y \<in> B" by auto
  assume "map_prod f g x = map_prod f g y"
  hence "fst (map_prod f g x) = fst (map_prod f g y)" by (auto)
  hence "f (fst x) = f (fst y)" by (cases x,cases y,auto)
  with \<open>inj_on f A\<close> and \<open>fst x \<in> A\<close> and \<open>fst y \<in> A\<close>
  have "fst x = fst y" by (auto dest:dest:inj_onD)
  moreover from \<open>map_prod f g x = map_prod f g y\<close>
  have "snd (map_prod f g x) = snd (map_prod f g y)" by (auto)
  hence "g (snd x) = g (snd y)" by (cases x,cases y,auto)
  with \<open>inj_on g B\<close> and \<open>snd x \<in> B\<close> and \<open>snd y \<in> B\<close>
  have "snd x = snd y" by (auto dest:dest:inj_onD)
  ultimately show "x = y" by(rule prod_eqI)
qed

lemma map_prod_surj:
  fixes f :: "'a \<Rightarrow> 'b" and g :: "'c \<Rightarrow> 'd"
  assumes "surj f" and "surj g"
  shows "surj (map_prod f g)"
unfolding surj_def
proof
  fix y :: "'b \<times> 'd"
  from \<open>surj f\<close> obtain a where "fst y = f a" by (auto elim:surjE)
  moreover
  from \<open>surj g\<close> obtain b where "snd y = g b" by (auto elim:surjE)
  ultimately have "(fst y, snd y) = map_prod f g (a,b)" by auto
  thus "\<exists>x. y = map_prod f g x" by auto
qed

lemma map_prod_surj_on:
  assumes "f ` A = A'" and "g ` B = B'"
  shows "map_prod f g ` (A \<times> B) = A' \<times> B'"
unfolding image_def
proof(rule set_eqI,rule iffI)
  fix x :: "'a \<times> 'c"
  assume "x \<in> {y::'a \<times> 'c. \<exists>x::'b \<times> 'd\<in>A \<times> B. y = map_prod f g x}"
  then obtain y where "y \<in> A \<times> B" and "x = map_prod f g y" by blast
  from \<open>image f A = A'\<close> and \<open>y \<in> A \<times> B\<close> have "f (fst y) \<in> A'" by auto
  moreover from \<open>image g B = B'\<close> and \<open>y \<in> A \<times> B\<close> have "g (snd y) \<in> B'" by auto
  ultimately have "(f (fst y), g (snd y)) \<in> (A' \<times> B')" by auto
  with \<open>x = map_prod f g y\<close> show "x \<in> A' \<times> B'" by (cases y, auto)
next
  fix x :: "'a \<times> 'c"
  assume "x \<in> A' \<times> B'" hence "fst x \<in> A'" and "snd x \<in> B'" by auto
  from \<open>image f A = A'\<close> and \<open>fst x \<in> A'\<close> have "fst x \<in> image f A" by auto
  then obtain a where "a \<in> A" and "fst x = f a" by (rule imageE)
  moreover from \<open>image g B = B'\<close> and \<open>snd x \<in> B'\<close>
  obtain b where "b \<in> B" and "snd x = g b" by auto
  ultimately have "(fst x, snd x) = map_prod f g (a,b)" by auto
  moreover from \<open>a \<in> A\<close> and  \<open>b \<in> B\<close> have "(a , b) \<in> A \<times> B" by auto
  ultimately have "\<exists>y \<in> A \<times> B. x = map_prod f g y" by auto
  thus "x \<in> {x. \<exists>y \<in> A \<times> B. x = map_prod f g y}" by auto
qed


subsection \<open>Simproc for rewriting a set comprehension into a pointfree expression\<close>

ML_file "Tools/set_comprehension_pointfree.ML"

setup \<open>
  Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs
    [Simplifier.make_simproc @{context} "set comprehension"
      {lhss = [@{term "Collect P"}],
       proc = K Set_Comprehension_Pointfree.code_simproc,
       identifier = []}])
\<close>


subsection \<open>Inductively defined sets\<close>

(* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
simproc_setup Collect_mem ("Collect t") = \<open>
  fn _ => fn ctxt => fn ct =>
    (case Thm.term_of ct of
      S as Const (@{const_name Collect}, Type (@{type_name fun}, [_, T])) $ t =>
        let val (u, _, ps) = HOLogic.strip_psplits t in
          (case u of
            (c as Const (@{const_name Set.member}, _)) $ q $ S' =>
              (case try (HOLogic.strip_ptuple ps) q of
                NONE => NONE
              | SOME ts =>
                  if not (Term.is_open S') andalso
                    ts = map Bound (length ps downto 0)
                  then
                    let val simp =
                      full_simp_tac (put_simpset HOL_basic_ss ctxt
                        addsimps [@{thm split_paired_all}, @{thm split_conv}]) 1
                    in
                      SOME (Goal.prove ctxt [] []
                        (Const (@{const_name Pure.eq}, T --> T --> propT) $ S $ S')
                        (K (EVERY
                          [resolve_tac ctxt [eq_reflection] 1,
                           resolve_tac ctxt @{thms subset_antisym} 1,
                           resolve_tac ctxt [subsetI] 1, dresolve_tac ctxt [CollectD] 1, simp,
                           resolve_tac ctxt [subsetI] 1, resolve_tac ctxt [CollectI] 1, simp])))
                    end
                  else NONE)
          | _ => NONE)
        end
    | _ => NONE)
\<close>

ML_file "Tools/inductive_set.ML"


subsection \<open>Legacy theorem bindings and duplicates\<close>

abbreviation (input) case_prod :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
  "case_prod \<equiv> uncurry"

abbreviation (input) split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
  "split \<equiv> uncurry"

lemmas PairE = prod.exhaust
lemmas Pair_eq = prod.inject
lemmas fst_conv = prod.sel(1)
lemmas snd_conv = prod.sel(2)
lemmas pair_collapse = prod.collapse
lemmas split = split_conv
lemmas Pair_fst_snd_eq = prod_eq_iff
lemmas split_def = case_prod_unfold

hide_const (open) prod

end