Theory Product_Type

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

section ‹Cartesian products›

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

subsection typbool is a datatype›

free_constructors (discs_sels) case_bool for True | False
  by auto

text ‹Avoid name clashes by prefixing the output of old_rep_datatype› with old›.›

setup Sign.mandatory_path "old"

old_rep_datatype True False by (auto intro: bool_induct)

setup Sign.parent_path

text ‹But erase the prefix for properties that are not generated by free_constructors›.›

setup Sign.mandatory_path "bool"

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

setup Sign.parent_path

declare case_split [cases type: bool]
  ― ‹prefer plain propositional version›

lemma [code]: "HOL.equal False P  ¬ P"
  and [code]: "HOL.equal True P  P"
  and [code]: "HOL.equal P False  ¬ P"
  and [code]: "HOL.equal P True  P"
  and [code nbe]: "HOL.equal P P  True"
  by (simp_all add: equal)

lemma If_case_cert:
  assumes "CASE  (λb. If b f g)"
  shows "(CASE True  f) &&& (CASE False  g)"
  using assms by simp_all

setup Code.declare_case_global @{thm If_case_cert}

code_printing
  constant "HOL.equal :: bool  bool  bool"  (Haskell) infix 4 "=="
| class_instance "bool" :: "equal"  (Haskell) -


subsection ‹The unit› type›

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 ‹
  Simplification procedure for @{thm [source] unit_eq}.  Cannot use
  this rule directly --- it loops!
›

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

free_constructors case_unit for "()"
  by auto

text ‹Avoid name clashes by prefixing the output of old_rep_datatype› with old›.›

setup Sign.mandatory_path "old"

old_rep_datatype "()" by simp

setup Sign.parent_path

text ‹But erase the prefix for properties that are not generated by free_constructors›.›

setup Sign.mandatory_path "unit"

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

setup Sign.parent_path

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 ‹
  This rewrite counters the effect of simproc unit_eq› on @{term
  [source] "λu::unit. f u"}, replacing it by @{term [source]
  f} rather than by @{term [source] "λu. f ()"}.
›

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  unit  bool"
  where "(_::unit)  _  True"

lemma less_eq_unit [iff]: "u  v" for u v :: unit
  by (simp add: less_eq_unit_def)

definition less_unit :: "unit  unit  bool"
  where "(_::unit) < _  False"

lemma less_unit [iff]: "¬ u < v" for u v :: unit
  by (simp_all add: less_eq_unit_def less_unit_def)

definition bot_unit :: unit
  where [code_unfold]: " = ()"

definition top_unit :: unit
  where [code_unfold]: " = ()"

definition inf_unit :: "unit  unit  unit"
  where [simp]: "_  _ = ()"

definition sup_unit :: "unit  unit  unit"
  where [simp]: "_  _ = ()"

definition Inf_unit :: "unit set  unit"
  where [simp]: "_ = ()"

definition Sup_unit :: "unit set  unit"
  where [simp]: "_ = ()"

definition uminus_unit :: "unit  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 v  True" for u v :: unit
  unfolding equal unit_eq [of u] unit_eq [of v] by (rule iffI TrueI refl)+

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

code_reserved SML
  unit

code_reserved OCaml
  unit

code_reserved Scala
  Unit


subsection ‹The product type›

subsubsection ‹Type definition›

definition Pair_Rep :: "'a  'b  'a  'b  bool"
  where "Pair_Rep a b = (λx y. x = a  y = b)"

definition "prod = {f. a b. f = Pair_Rep (a::'a) (b::'b)}"

typedef ('a, 'b) prod ("(_ ×/ _)" [21, 20] 20) = "prod :: ('a  'b  bool) set"
  unfolding prod_def by auto

type_notation (ASCII)
  prod  (infixr "*" 20)

definition Pair :: "'a  'b  'a × 'b"
  where "Pair a b = Abs_prod (Pair_Rep a b)"

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

free_constructors case_prod for Pair fst snd
proof -
  fix P :: bool and p :: "'a × 'b"
  show "(x1 x2. p = Pair x1 x2  P)  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  a = c  b = d"
    by (auto simp add: Pair_Rep_def fun_eq_iff)
  moreover have "Pair_Rep a b  prod" and "Pair_Rep c d  prod"
    by (auto simp add: prod_def)
  ultimately show "Pair a b = Pair c d  a = c  b = d"
    by (simp add: Pair_def Abs_prod_inject)
qed

text ‹Avoid name clashes by prefixing the output of old_rep_datatype› with old›.›

setup Sign.mandatory_path "old"

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

setup Sign.parent_path

text ‹But erase the prefix for properties that are not generated by free_constructors›.›

setup Sign.mandatory_path "prod"

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 Sign.parent_path

declare prod.case [nitpick_simp del]
declare old.prod.case_cong_weak [cong del]
declare prod.case_eq_if [mono]
declare prod.split [no_atp]
declare prod.split_asm [no_atp]

text @{thm [source] prod.split} could be declared as [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.
›


subsubsection ‹Tuple syntax›

text ‹
  Patterns -- extends pre-defined type typpttrn used in
  abstractions.
›

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"      ("_,/ _")
  "_unit"       :: pttrn                                ("'(')")
translations
  "(x, y)"  "CONST Pair x y"
  "_pattern x y"  "CONST Pair x y"
  "_patterns x y"  "CONST Pair x y"
  "_tuple x (_tuple_args y z)"  "_tuple x (_tuple_arg (_tuple y z))"
  "λ(x, y, zs). b"  "CONST case_prod (λx (y, zs). b)"
  "λ(x, y). b"  "CONST case_prod (λx y. b)"
  "_abs (CONST Pair x y) t"  "λ(x, y). t"
  ― ‹This rule accommodates tuples in case C … (x, y) … ⇒ …›:
     The (x, y)› is parsed as Pair x y› because it is logic›,
     not pttrn›.›
  "λ(). b"  "CONST case_unit b"
  "_abs (CONST Unity) t"  "λ(). t"

text ‹print termcase_prod f as termλ(x, y). f x y and
  termcase_prod (λx. f x) as termλ(x, y). f x y

typed_print_translation let
    fun case_prod_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match
      | case_prod_guess_names_tr' T [Abs (x, xT, t)] =
          (case (head_of t) of
            Const (const_syntaxcase_prod, _) => 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)
      | case_prod_guess_names_tr' T [t] =
          (case head_of t of
            Const (const_syntaxcase_prod, _) => 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)
      | case_prod_guess_names_tr' _ _ = raise Match;
  in [(const_syntaxcase_prod, K case_prod_guess_names_tr')] end

text ‹Reconstruct pattern from (nested) constcase_prods,
  avoiding eta-contraction of body; required for enclosing "let",
  if "let" does not avoid eta-contraction, which has been observed to occur.›

print_translation let
    fun case_prod_tr' [Abs (x, T, t as (Abs abs))] =
          (* case_prod (λx y. t) ⇒ λ(x, y) t *)
          let
            val (y, t') = Syntax_Trans.atomic_abs_tr' abs;
            val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t');
          in
            Syntax.const syntax_const‹_abs› $
              (Syntax.const syntax_const‹_pattern› $ x' $ y) $ t''
          end
      | case_prod_tr' [Abs (x, T, (s as Const (const_syntaxcase_prod, _) $ t))] =
          (* case_prod (λx. (case_prod (λy z. t))) ⇒ λ(x, y, z). t *)
          let
            val Const (syntax_const‹_abs›, _) $
              (Const (syntax_const‹_pattern›, _) $ y $ z) $ t' =
                case_prod_tr' [t];
            val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t');
          in
            Syntax.const syntax_const‹_abs› $
              (Syntax.const syntax_const‹_pattern› $ x' $
                (Syntax.const syntax_const‹_patterns› $ y $ z)) $ t''
          end
      | case_prod_tr' [Const (const_syntaxcase_prod, _) $ t] =
          (* case_prod (case_prod (λx y z. t)) ⇒ λ((x, y), z). t *)
          case_prod_tr' [(case_prod_tr' [t])]
            (* inner case_prod_tr' creates next pattern *)
      | case_prod_tr' [Const (syntax_const‹_abs›, _) $ x_y $ Abs abs] =
          (* case_prod (λpttrn z. t) ⇒ λ(pttrn, z). t *)
          let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in
            Syntax.const syntax_const‹_abs› $
              (Syntax.const syntax_const‹_pattern› $ x_y $ z) $ t
          end
      | case_prod_tr' _ = raise Match;
  in [(const_syntaxcase_prod, K case_prod_tr')] end


subsubsection ‹Code generator setup›

code_printing
  type_constructor prod 
    (SML) infix 2 "*"
    and (OCaml) infix 2 "*"
    and (Haskell) "!((_),/ (_))"
    and (Scala) "((_),/ (_))"
| constant Pair 
    (SML) "!((_),/ (_))"
    and (OCaml) "!((_),/ (_))"
    and (Haskell) "!((_),/ (_))"
    and (Scala) "!((_),/ (_))"
| class_instance  prod :: equal 
    (Haskell) -
| constant "HOL.equal :: 'a × 'b  'a × 'b  bool" 
    (Haskell) infix 4 "=="
| constant fst  (Haskell) "fst"
| constant snd  (Haskell) "snd"


subsubsection ‹Fundamental operations and properties›

lemma Pair_inject: "(a, b) = (a', b')  (a = a'  b = b'  R)  R"
  by simp

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

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

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

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

lemma case_prod_conv [simp, code]: "(case (a, b) of (c, d)  f c d) = f a b"
  by (fact prod.case)

lemmas surjective_pairing = prod.collapse [symmetric]

lemma prod_eq_iff: "s = t  fst s = fst t  snd s = snd t"
  by (cases s, cases t) simp

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

lemma case_prodI: "f a b  case (a, b) of (c, d)  f c d"
  by (rule prod.case [THEN iffD2])

lemma case_prodD: "(case (a, b) of (c, d)  f c d)  f a b"
  by (rule prod.case [THEN iffD1])

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

lemma case_prod_eta: "(λ(x, y). f (x, y)) = f"
  ― ‹Subsumes the old split_Pair› when termf is the identity function.›
  by (simp add: fun_eq_iff split: prod.split)

(* This looks like a sensible simp-rule but appears to do more harm than good:
lemma case_prod_const [simp]: "(λ(_,_). c) = (λ_. c)"
by(rule case_prod_eta)
*)

lemma case_prod_comp: "(case x of (a, b)  (f  g) a b) = f (g (fst x)) (snd x)"
  by (cases x) simp

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

lemma cond_case_prod_eta: "(x y. f x y = g (x, y))  (λ(x, y). f x y) = g"
  by (simp add: case_prod_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 PROP P (fst x, snd x) show "PROP P x" by simp
qed

text ‹
  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 ⋀a b. … = ?P(a, b)›
  which cannot be solved by reflexivity.
›

lemmas split_tupled_all = split_paired_all unit_all_eq2

ML (* 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_namePure.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 [simprocunit_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;

setup map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))

lemma split_paired_All [simp, no_atp]: "(x. P x)  (a b. P (a, b))"
  ― ‹[iff]› is not a good idea because it makes blast› loop›
  by fast

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

lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))"
  ― ‹Can't be added to simpset: loops!›
  by (simp add: case_prod_eta)

text ‹
  Simplification procedure for @{thm [source] cond_case_prod_eta}.  Using
  @{thm [source] case_prod_eta} as a rewrite rule is not general enough,
  and using @{thm [source] cond_case_prod_eta} directly would render some
  existing proofs very inefficient; similarly for prod.case_eq_if›.
›

ML local
  val cond_case_prod_eta_ss =
    simpset_of (put_simpset HOL_basic_ss context addsimps @{thms cond_case_prod_eta});
  fun Pair_pat k 0 (Bound m) = (m = k)
    | Pair_pat k i (Const (const_namePair,  _) $ 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_namecase_prod, _) $ 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_case_prod_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_namecase_prod, _) $ 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_namecase_prod, _) $ 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;
simproc_setup case_prod_beta ("case_prod f z") =
  K (fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct))
simproc_setup case_prod_eta ("case_prod f") =
  K (fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct))

lemma case_prod_beta': "(λ(x,y). f x y) = (λx. f (fst x) (snd x))"
  by (auto simp: fun_eq_iff)

text  constcase_prod used as a logical connective or set former.

   These rules are for use with blast›; could instead
  call simp› using @{thm [source] prod.split} as rewrite.›

lemma case_prodI2:
  "p. (a b. p = (a, b)  c a b)  case p of (a, b)  c a b"
  by (simp add: split_tupled_all)

lemma case_prodI2':
  "p. (a b. (a, b) = p  c a b x)  (case p of (a, b)  c a b) x"
  by (simp add: split_tupled_all)

lemma case_prodE [elim!]:
  "(case p of (a, b)  c a b)  (x y. p = (x, y)  c x y  Q)  Q"
  by (induct p) simp

lemma case_prodE' [elim!]:
  "(case p of (a, b)  c a b) z  (x y. p = (x, y)  c x y z  Q)  Q"
  by (induct p) simp

lemma case_prodE2:
  assumes q: "Q (case z of (a, b)  P a b)"
    and r: "x y. z = (x, y)  Q (P x y)  R"
  shows R
proof (rule r)
  show "z = (fst z, snd z)" by simp
  then show "Q (P (fst z) (snd z))"
    using q by (simp add: case_prod_unfold)
qed

lemma case_prodD': "(case (a, b) of (c, d)  R c d) c  R a b c"
  by simp

lemma mem_case_prodI: "z  c a b  z  (case (a, b) of (d, e)  c d e)"
  by simp

lemma mem_case_prodI2 [intro!]:
  "p. (a b. p = (a, b)  z  c a b)  z  (case p of (a, b)  c a b)"
  by (simp only: split_tupled_all) simp

declare mem_case_prodI [intro!] ― ‹postponed to maintain traditional declaration order!›
declare case_prodI2' [intro!] ― ‹postponed to maintain traditional declaration order!›
declare case_prodI2 [intro!] ― ‹postponed to maintain traditional declaration order!›
declare case_prodI [intro!] ― ‹postponed to maintain traditional declaration order!›

lemma mem_case_prodE [elim!]:
  assumes "z  case_prod c p"
  obtains x y where "p = (x, y)" and "z  c x y"
  using assms by (rule case_prodE2)

ML local (* filtering with exists_p_split is an essential optimization *)
  fun exists_p_split (Const (const_namecase_prod,_) $ _ $ (Const (const_namePair,_)$_$_)) = 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 case_prod_conv}) i
    else no_tac);
end;

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

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

lemma split_eta_SetCompr2 [simp, no_atp]: "(λu. x y. u = (x, y)  P x y) = case_prod P"
  by (rule ext) fast

lemma split_part [simp]: "(λ(a,b). P  Q a b) = (λab. P  case_prod Q ab)"
  ― ‹Allows simplifications of nested splits in case of independent predicates.›
  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)) = case_prod (λx. f (g x))"
  by (rule ext) auto

lemma pair_imageI [intro]: "(a, b)  A  f a b  (λ(a, b). f a b) ` A"
  by (rule image_eqI [where x = "(a, b)"]) auto

lemma Collect_const_case_prod[simp]: "{(a,b). P} = (if P then UNIV else {})"
by auto

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";
*)

lemma case_prod_beta: "case_prod f p = f (fst p) (snd p)"
  by (fact prod.case_eq_if)

lemma prod_cases3 [cases type]:
  obtains (fields) a b c where "y = (a, b, c)"
proof (cases y)
  case (Pair a b)
  with that show ?thesis
    by (cases b) blast
qed


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)"
proof (cases y)
  case (fields a b c)
  with that show ?thesis
    by (cases c) blast
qed

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)"
proof (cases y)
  case (fields a b c d)
  with that show ?thesis
    by (cases d) blast
qed

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)"
proof (cases y)
  case (fields a b c d e)
  with that show ?thesis
    by (cases e) blast
qed

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)"
proof (cases y)
  case (fields a b c d e f)
  with that show ?thesis
    by (cases f) blast
qed


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_case_prod :: "('a  'b  'c)  'a × 'b  'c"
  where "internal_case_prod  case_prod"

lemma internal_case_prod_conv: "internal_case_prod c (a, b) = c a b"
  by (simp only: internal_case_prod_def case_prod_conv)

ML_file ‹Tools/split_rule.ML›

hide_const internal_case_prod


subsubsection ‹Derived operations›

definition curry :: "('a × 'b  'c)  'a  'b  'c"
  where "curry = (λ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)  curry f a b"
  by (simp add: curry_def)

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

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

lemma curry_case_prod [simp]: "curry (case_prod f) = f"
  by (simp add: curry_def case_prod_unfold)

lemma case_prod_curry [simp]: "case_prod (curry f) = f"
  by (simp add: curry_def case_prod_unfold)

lemma curry_K: "curry (λx. c) = (λx y. c)"
  by (simp add: fun_eq_iff)

text ‹The composition-uncurry combinator.›

definition scomp :: "('a  'b × 'c)  ('b  'c  'd)  'a  'd"  (infixl "∘→" 60)
  where "f ∘→ g = (λx. case_prod g (f x))"

no_notation scomp (infixl "∘→" 60)

bundle state_combinator_syntax
begin

notation fcomp (infixl "∘>" 60)
notation scomp (infixl "∘→" 60)

end

context
  includes state_combinator_syntax
begin

lemma scomp_unfold: "(∘→) = (λ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 ∘→ g) x = case_prod g (f x)"
  by (simp add: scomp_unfold case_prod_unfold)

lemma Pair_scomp: "Pair x ∘→ f = f x"
  by (simp add: fun_eq_iff)

lemma scomp_Pair: "x ∘→ Pair = x"
  by (simp add: fun_eq_iff)

lemma scomp_scomp: "(f ∘→ g) ∘→ h = f ∘→ (λx. g x ∘→ h)"
  by (simp add: fun_eq_iff scomp_unfold)

lemma scomp_fcomp: "(f ∘→ g) ∘> h = f ∘→ (λx. g x ∘> h)"
  by (simp add: fun_eq_iff scomp_unfold fcomp_def)

lemma fcomp_scomp: "(f ∘> g) ∘→ h = f ∘> (g ∘→ h)"
  by (simp add: fun_eq_iff scomp_unfold)

end

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

text termmap_prod --- action of the product functor upon functions.
›

definition map_prod :: "('a  'c)  ('b  'd)  'a × 'b  'c × 'd"
  where "map_prod f g = (λ(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  map_prod f g = f  fst"
  by (rule ext) simp_all

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

lemma map_prod_compose: "map_prod (f1  f2) (g1  g2) = (map_prod f1 g1  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)  R  (f a, g b)  map_prod f g ` R"
  by (rule image_eqI) simp_all

lemma prod_fun_imageE [elim!]:
  assumes major: "c  map_prod f g ` R"
    and cases: "x y. c = (f x, g y)  (x, y)  R  P"
  shows P
proof (rule major [THEN imageE])
  fix x
  assume "c = map_prod f g x" "x  R"
  then show P
    using cases by (cases x) simp
qed

definition apfst :: "('a  'c)  'a × 'b  'c × 'b"
  where "apfst f = map_prod f id"

definition apsnd :: "('b  'c)  'a × 'b  'a × '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  apfst f = f  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  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  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  apsnd f = f  snd"
  by (simp add: fun_eq_iff)

lemma apfst_compose: "apfst f (apfst g x) = apfst (f  g) x"
  by (cases x) simp

lemma apsnd_compose: "apsnd f (apsnd g x) = apsnd (f  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  f (fst x) = g (fst x)"
  by (cases x) simp

lemma apsnd_eq_conv [simp]: "apsnd f x = apsnd g x  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 Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")

definition swap :: "'a × 'b  'b × '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  prod.swap = id"
  by (simp add: fun_eq_iff)

lemma pair_in_swap_image [simp]: "(y, x)  prod.swap ` A  (x, y)  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 (λ(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)  f x y) = (case p of (x, y)  f x y)"
  by (cases p) simp

lemma fst_swap [simp]: "fst (prod.swap x) = snd x"
  by (cases x) simp

lemma snd_swap [simp]: "snd (prod.swap x) = fst x"
  by (cases x) simp

text ‹Disjoint union of a family of sets -- Sigma.›

definition Sigma :: "'a set  ('a  'b set)  ('a × 'b) set"
  where "Sigma A B  xA. yB x. {Pair x y}"

abbreviation Times :: "'a set  'b set  ('a × 'b) set"  (infixr "×" 80)
  where "A × B  Sigma A (λ_. B)"

hide_const (open) Times

bundle no_Set_Product_syntax begin
no_notation Product_Type.Times (infixr "×" 80)
end
bundle Set_Product_syntax begin
notation Product_Type.Times (infixr "×" 80)
end

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"
  unfolding Sigma_def by blast

lemma SigmaE [elim!]: "c  Sigma A B  (x y. x  A  y  B x  c = (x, y)  P)  P"
  ― ‹The general elimination rule.›
  unfolding Sigma_def by blast

text ‹
  Elimination of term(a, b)  A × B -- introduces no
  eigenvariables.
›

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: "A = B  (x. x  B  C x = D x)  (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 mem_Times_iff: "x  A × B  fst x  A  snd x  B"
  by (induct x) simp

lemma Sigma_empty_iff: "(SIGMA i:I. X i) = {}  (iI. 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 Collect_case_prod_Sigma: "{(x, y). P x  Q x y} = (SIGMA x:Collect P. Collect (Q x))"
  by blast

lemma Collect_case_prod [simp]: "{(a, b). P a  Q b} = Collect P × Collect Q "
  by (fact Collect_case_prod_Sigma)

lemma Collect_case_prodD: "x  Collect (case_prod A)  A (fst x) (snd x)"
  by auto

lemma Collect_case_prod_mono: "A  B  Collect (case_prod A)  Collect (case_prod B)"
  by auto (auto elim!: le_funE)

lemma Collect_split_mono_strong:
  "X = fst ` A  Y = snd ` A  aX. b  Y. P a b  Q a b
     A  Collect (case_prod P)  A  Collect (case_prod Q)"
  by fastforce

lemma UN_Times_distrib: "((a, b)A × B. E a × F b) = (E ` A) × (F ` B)"
  ― ‹Suggested by Pierre Chartier›
  by blast

lemma split_paired_Ball_Sigma [simp, no_atp]: "(zSigma A B. P z)  (xA. yB x. P (x, y))"
  by blast

lemma split_paired_Bex_Sigma [simp, no_atp]: "(zSigma A B. P z)  (xA. yB x. P (x, y))"
  by blast

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

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

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

lemma Sigma_Int_distrib2: "(SIGMA i:I. A i  B i) = Sigma I A  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 (X) B = (AX. Sigma A B)"
  by blast

lemma Pair_vimage_Sigma: "Pair x -` Sigma A f = (if x  A then f x else {})"
  by auto

text ‹
  Non-dependent versions are needed to avoid the need for higher-order
  matching, especially when the rules are re-oriented.
›

lemma Times_Un_distrib1: "(A  B) × C = A × C  B × C "
  by (fact Sigma_Un_distrib1)

lemma Times_Int_distrib1: "(A  B) × C = A × C  B × C "
  by (fact Sigma_Int_distrib1)

lemma Times_Diff_distrib1: "(A - B) × C = A × C - B × C "
  by (fact Sigma_Diff_distrib1)

lemma Times_empty [simp]: "A × B = {}  A = {}  B = {}"
  by auto

lemma times_subset_iff: "A × C  B × D  A={}  C={}  A  B  C  D"
  by blast

lemma times_eq_iff: "A × B = C × D  A = C  B = D  (A = {}  B = {})  (C = {}  D = {})"
  by auto

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

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

lemma fst_image_Sigma: "fst ` (Sigma A B) = {x  A. B(x)  {}}"
  by force

lemma snd_image_Sigma: "snd ` (Sigma A B) = ( x  A. B x)"
  by force

lemma vimage_fst: "fst -` A = A × UNIV"
  by auto

lemma vimage_snd: "snd -` A = UNIV × A"
  by auto

lemma insert_Times_insert [simp]:
  "insert a A × insert b B = insert (a,b) (A × insert b B  {a} × B)"
  by blast

lemma vimage_Times: "f -` (A × B) = (fst  f) -` A  (snd  f) -` B"
proof (rule set_eqI)
  show "x  f -` (A × B)  x  (fst  f) -` A  (snd  f) -` B" for x
    by (cases "f x") (auto split: prod.split)
qed

lemma Times_Int_Times: "A × B  C × D = (A  C) × (B  D)"
  by auto

lemma image_paired_Times:
   "(λ(x,y). (f x, g y)) ` (A × B) = (f ` A) × (g ` B)"
  by auto

lemma product_swap: "prod.swap ` (A × B) = B × A"
  by (auto simp add: set_eq_iff)

lemma swap_product: "(λ(i, j). (j, i)) ` (A × B) = B × A"
  by (auto simp add: set_eq_iff)

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

lemma subset_fst_snd: "A  (fst ` A × snd ` A)"
  by force

lemma inj_on_apfst [simp]: "inj_on (apfst f) (A × UNIV)  inj_on f A"
  by (auto simp add: inj_on_def)

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

lemma inj_on_apsnd [simp]: "inj_on (apsnd f) (UNIV × A)  inj_on f A"
  by (auto simp add: inj_on_def)

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

context
begin

qualified definition product :: "'a set  'b set  ('a × 'b) set"
  where [code_abbrev]: "product A B = A × B"

lemma member_product: "x  Product_Type.product A B  x  A × B"
  by (simp add: product_def)

end

text ‹The following constmap_prod lemmas are due to Joachim Breitner:›

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

lemma map_prod_surj:
  fixes f :: "'a  'b"
    and g :: "'c  'd"
  assumes "surj f" and "surj g"
  shows "surj (map_prod f g)"
  unfolding surj_def
proof
  fix y :: "'b × 'd"
  from surj f obtain a where "fst y = f a"
    by (auto elim: surjE)
  moreover
  from surj g 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
  then show "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 × B) = A' × B'"
  unfolding image_def
proof (rule set_eqI, rule iffI)
  fix x :: "'a × 'c"
  assume "x  {y::'a × 'c. x::'b × 'dA × B. y = map_prod f g x}"
  then obtain y where "y  A × B" and "x = map_prod f g y"
    by blast
  from image f A = A' and y  A × B have "f (fst y)  A'"
    by auto
  moreover from image g B = B' and y  A × B have "g (snd y)  B'"
    by auto
  ultimately have "(f (fst y), g (snd y))  (A' × B')"
    by auto
  with x = map_prod f g y show "x  A' × B'"
    by (cases y) auto
next
  fix x :: "'a × 'c"
  assume "x  A' × B'"
  then have "fst x  A'" and "snd x  B'"
    by auto
  from image f A = A' and fst x  A' have "fst x  image f A"
    by auto
  then obtain a where "a  A" and "fst x = f a"
    by (rule imageE)
  moreover from image g B = B' and snd x  B' obtain b where "b  B" and "snd x = g b"
    by auto
  ultimately have "(fst x, snd x) = map_prod f g (a, b)"
    by auto
  moreover from a  A and  b  B have "(a , b)  A × B"
    by auto
  ultimately have "y  A × B. x = map_prod f g y"
    by auto
  then show "x  {x. y  A × B. x = map_prod f g y}"
    by auto
qed


subsection ‹Simproc for rewriting a set comprehension into a pointfree expression›

ML_file ‹Tools/set_comprehension_pointfree.ML›

setup Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs
    [Simplifier.make_simproc context "set comprehension"
      {lhss = [termCollect P],
       proc = K Set_Comprehension_Pointfree.code_simproc}])

subsection ‹Lemmas about disjointness›

lemma disjnt_Times1_iff [simp]: "disjnt (C × A) (C × B)  C = {}  disjnt A B"
  by (auto simp: disjnt_def)

lemma disjnt_Times2_iff [simp]: "disjnt (A × C) (B × C)  C = {}  disjnt A B"
  by (auto simp: disjnt_def)

lemma disjnt_Sigma_iff: "disjnt (Sigma A C) (Sigma B C)  (i  AB. C i = {})  disjnt A B"
  by (auto simp: disjnt_def)


subsection ‹Inductively defined sets›

(* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
simproc_setup Collect_mem ("Collect t") = K (fn ctxt => fn ct =>
    (case Thm.term_of ct of
      S as Const (const_nameCollect, Type (type_namefun, [_, T])) $ t =>
        let val (u, _, ps) = HOLogic.strip_ptupleabs t in
          (case u of
            (c as Const (const_nameSet.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 case_prod_conv}]) 1
                    in
                      SOME (Goal.prove ctxt [] []
                        (Const (const_namePure.eq, T --> T --> propT) $ S $ S')
                        (K (EVERY
                          [resolve_tac ctxt [eq_reflection] 1,
                           resolve_tac ctxt @{thms subset_antisym} 1,
                           resolve_tac ctxt @{thms subsetI} 1,
                           dresolve_tac ctxt @{thms CollectD} 1, simp,
                           resolve_tac ctxt @{thms subsetI} 1,
                           resolve_tac ctxt @{thms CollectI} 1, simp])))
                    end
                  else NONE)
          | _ => NONE)
        end
    | _ => NONE))

ML_file ‹Tools/inductive_set.ML›


subsection ‹Legacy theorem bindings and duplicates›

lemmas fst_conv = prod.sel(1)
lemmas snd_conv = prod.sel(2)
lemmas split_def = case_prod_unfold
lemmas split_beta' = case_prod_beta'
lemmas split_beta = prod.case_eq_if
lemmas split_conv = case_prod_conv
lemmas split = case_prod_conv

hide_const (open) prod

end