Theory Relation

(*  Title:      HOL/Relation.thy
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Author:     Stefan Berghofer, TU Muenchen
    Author:     Martin Desharnais, MPI-INF Saarbruecken
*)

section ‹Relations -- as sets of pairs, and binary predicates›

theory Relation
  imports Product_Type Sum_Type Fields
begin

text ‹A preliminary: classical rules for reasoning on predicates›

declare predicate1I [Pure.intro!, intro!]
declare predicate1D [Pure.dest, dest]
declare predicate2I [Pure.intro!, intro!]
declare predicate2D [Pure.dest, dest]
declare bot1E [elim!]
declare bot2E [elim!]
declare top1I [intro!]
declare top2I [intro!]
declare inf1I [intro!]
declare inf2I [intro!]
declare inf1E [elim!]
declare inf2E [elim!]
declare sup1I1 [intro?]
declare sup2I1 [intro?]
declare sup1I2 [intro?]
declare sup2I2 [intro?]
declare sup1E [elim!]
declare sup2E [elim!]
declare sup1CI [intro!]
declare sup2CI [intro!]
declare Inf1_I [intro!]
declare INF1_I [intro!]
declare Inf2_I [intro!]
declare INF2_I [intro!]
declare Inf1_D [elim]
declare INF1_D [elim]
declare Inf2_D [elim]
declare INF2_D [elim]
declare Inf1_E [elim]
declare INF1_E [elim]
declare Inf2_E [elim]
declare INF2_E [elim]
declare Sup1_I [intro]
declare SUP1_I [intro]
declare Sup2_I [intro]
declare SUP2_I [intro]
declare Sup1_E [elim!]
declare SUP1_E [elim!]
declare Sup2_E [elim!]
declare SUP2_E [elim!]


subsection ‹Fundamental›

subsubsection ‹Relations as sets of pairs›

type_synonym 'a rel = "('a × 'a) set"

lemma subrelI: "(x y. (x, y)  r  (x, y)  s)  r  s"
  ― ‹Version of @{thm [source] subsetI} for binary relations›
  by auto

lemma lfp_induct2:
  "(a, b)  lfp f  mono f 
    (a b. (a, b)  f (lfp f  {(x, y). P x y})  P a b)  P a b"
  ― ‹Version of @{thm [source] lfp_induct} for binary relations›
  using lfp_induct_set [of "(a, b)" f "case_prod P"] by auto


subsubsection ‹Conversions between set and predicate relations›

lemma pred_equals_eq [pred_set_conv]: "(λx. x  R) = (λx. x  S)  R = S"
  by (simp add: set_eq_iff fun_eq_iff)

lemma pred_equals_eq2 [pred_set_conv]: "(λx y. (x, y)  R) = (λx y. (x, y)  S)  R = S"
  by (simp add: set_eq_iff fun_eq_iff)

lemma pred_subset_eq [pred_set_conv]: "(λx. x  R)  (λx. x  S)  R  S"
  by (simp add: subset_iff le_fun_def)

lemma pred_subset_eq2 [pred_set_conv]: "(λx y. (x, y)  R)  (λx y. (x, y)  S)  R  S"
  by (simp add: subset_iff le_fun_def)

lemma bot_empty_eq [pred_set_conv]: " = (λx. x  {})"
  by (auto simp add: fun_eq_iff)

lemma bot_empty_eq2 [pred_set_conv]: " = (λx y. (x, y)  {})"
  by (auto simp add: fun_eq_iff)

lemma top_empty_eq: " = (λx. x  UNIV)"
  by (auto simp add: fun_eq_iff)

lemma top_empty_eq2: " = (λx y. (x, y)  UNIV)"
  by (auto simp add: fun_eq_iff)

lemma inf_Int_eq [pred_set_conv]: "(λx. x  R)  (λx. x  S) = (λx. x  R  S)"
  by (simp add: inf_fun_def)

lemma inf_Int_eq2 [pred_set_conv]: "(λx y. (x, y)  R)  (λx y. (x, y)  S) = (λx y. (x, y)  R  S)"
  by (simp add: inf_fun_def)

lemma sup_Un_eq [pred_set_conv]: "(λx. x  R)  (λx. x  S) = (λx. x  R  S)"
  by (simp add: sup_fun_def)

lemma sup_Un_eq2 [pred_set_conv]: "(λx y. (x, y)  R)  (λx y. (x, y)  S) = (λx y. (x, y)  R  S)"
  by (simp add: sup_fun_def)

lemma INF_INT_eq [pred_set_conv]: "(iS. (λx. x  r i)) = (λx. x  (iS. r i))"
  by (simp add: fun_eq_iff)

lemma INF_INT_eq2 [pred_set_conv]: "(iS. (λx y. (x, y)  r i)) = (λx y. (x, y)  (iS. r i))"
  by (simp add: fun_eq_iff)

lemma SUP_UN_eq [pred_set_conv]: "(iS. (λx. x  r i)) = (λx. x  (iS. r i))"
  by (simp add: fun_eq_iff)

lemma SUP_UN_eq2 [pred_set_conv]: "(iS. (λx y. (x, y)  r i)) = (λx y. (x, y)  (iS. r i))"
  by (simp add: fun_eq_iff)

lemma Inf_INT_eq [pred_set_conv]: "S = (λx. x  ((Collect ` S)))"
  by (simp add: fun_eq_iff)

lemma INF_Int_eq [pred_set_conv]: "(iS. (λx. x  i)) = (λx. x  S)"
  by (simp add: fun_eq_iff)

lemma Inf_INT_eq2 [pred_set_conv]: "S = (λx y. (x, y)  ((Collect ` case_prod ` S)))"
  by (simp add: fun_eq_iff)

lemma INF_Int_eq2 [pred_set_conv]: "(iS. (λx y. (x, y)  i)) = (λx y. (x, y)  S)"
  by (simp add: fun_eq_iff)

lemma Sup_SUP_eq [pred_set_conv]: "S = (λx. x  (Collect ` S))"
  by (simp add: fun_eq_iff)

lemma SUP_Sup_eq [pred_set_conv]: "(iS. (λx. x  i)) = (λx. x  S)"
  by (simp add: fun_eq_iff)

lemma Sup_SUP_eq2 [pred_set_conv]: "S = (λx y. (x, y)  ((Collect ` case_prod ` S)))"
  by (simp add: fun_eq_iff)

lemma SUP_Sup_eq2 [pred_set_conv]: "(iS. (λx y. (x, y)  i)) = (λx y. (x, y)  S)"
  by (simp add: fun_eq_iff)


subsection ‹Properties of relations›

subsubsection ‹Reflexivity›

definition refl_on :: "'a set  'a rel  bool"
  where "refl_on A r  r  A × A  (xA. (x, x)  r)"

abbreviation refl :: "'a rel  bool" ― ‹reflexivity over a type›
  where "refl  refl_on UNIV"

definition reflp_on :: "'a set  ('a  'a  bool)  bool"
  where "reflp_on A R  (xA. R x x)"

abbreviation reflp :: "('a  'a  bool)  bool"
  where "reflp  reflp_on UNIV"

lemma reflp_def[no_atp]: "reflp R  (x. R x x)"
  by (simp add: reflp_on_def)

text @{thm [source] reflp_def} is for backward compatibility.›

lemma reflp_refl_eq [pred_set_conv]: "reflp (λx y. (x, y)  r)  refl r"
  by (simp add: refl_on_def reflp_def)

lemma refl_onI [intro?]: "r  A × A  (x. x  A  (x, x)  r)  refl_on A r"
  unfolding refl_on_def by (iprover intro!: ballI)

lemma reflI: "(x. (x, x)  r)  refl r"
  by (auto intro: refl_onI)

lemma reflp_onI:
  "(x. x  A  R x x)  reflp_on A R"
  by (simp add: reflp_on_def)

lemma reflpI[intro?]: "(x. R x x)  reflp R"
  by (rule reflp_onI)

lemma refl_onD: "refl_on A r  a  A  (a, a)  r"
  unfolding refl_on_def by blast

lemma refl_onD1: "refl_on A r  (x, y)  r  x  A"
  unfolding refl_on_def by blast

lemma refl_onD2: "refl_on A r  (x, y)  r  y  A"
  unfolding refl_on_def by blast

lemma reflD: "refl r  (a, a)  r"
  unfolding refl_on_def by blast

lemma reflp_onD:
  "reflp_on A R  x  A  R x x"
  by (simp add: reflp_on_def)

lemma reflpD[dest?]: "reflp R  R x x"
  by (simp add: reflp_onD)

lemma reflpE:
  assumes "reflp r"
  obtains "r x x"
  using assms by (auto dest: refl_onD simp add: reflp_def)

lemma reflp_on_subset: "reflp_on A R  B  A  reflp_on B R"
  by (auto intro: reflp_onI dest: reflp_onD)

lemma reflp_on_image: "reflp_on (f ` A) R  reflp_on A (λa b. R (f a) (f b))"
  by (simp add: reflp_on_def)

lemma refl_on_Int: "refl_on A r  refl_on B s  refl_on (A  B) (r  s)"
  unfolding refl_on_def by blast

lemma reflp_on_inf: "reflp_on A R  reflp_on B S  reflp_on (A  B) (R  S)"
  by (auto intro: reflp_onI dest: reflp_onD)

lemma reflp_inf: "reflp r  reflp s  reflp (r  s)"
  by (rule reflp_on_inf[of UNIV _ UNIV, unfolded Int_absorb])

lemma refl_on_Un: "refl_on A r  refl_on B s  refl_on (A  B) (r  s)"
  unfolding refl_on_def by blast

lemma reflp_on_sup: "reflp_on A R  reflp_on B S  reflp_on (A  B) (R  S)"
  by (auto intro: reflp_onI dest: reflp_onD)

lemma reflp_sup: "reflp r  reflp s  reflp (r  s)"
  by (rule reflp_on_sup[of UNIV _ UNIV, unfolded Un_absorb])

lemma refl_on_INTER: "xS. refl_on (A x) (r x)  refl_on ((A ` S)) ((r ` S))"
  unfolding refl_on_def by fast

lemma reflp_on_Inf: "xS. reflp_on (A x) (R x)  reflp_on ((A ` S)) ((R ` S))"
  by (auto intro: reflp_onI dest: reflp_onD)

lemma refl_on_UNION: "xS. refl_on (A x) (r x)  refl_on ((A ` S)) ((r ` S))"
  unfolding refl_on_def by blast

lemma reflp_on_Sup: "xS. reflp_on (A x) (R x)  reflp_on ((A ` S)) ((R ` S))"
  by (auto intro: reflp_onI dest: reflp_onD)

lemma refl_on_empty [simp]: "refl_on {} {}"
  by (simp add: refl_on_def)

lemma reflp_on_empty [simp]: "reflp_on {} R"
  by (auto intro: reflp_onI)

lemma refl_on_singleton [simp]: "refl_on {x} {(x, x)}"
by (blast intro: refl_onI)

lemma refl_on_def' [nitpick_unfold, code]:
  "refl_on A r  ((x, y)  r. x  A  y  A)  (x  A. (x, x)  r)"
  by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)

lemma reflp_on_equality [simp]: "reflp_on A (=)"
  by (simp add: reflp_on_def)

lemma reflp_on_mono:
  "reflp_on A R  (x y. x  A  y  A  R x y  Q x y)  reflp_on A Q"
  by (auto intro: reflp_onI dest: reflp_onD)

lemma reflp_mono: "reflp R  (x y. R x y  Q x y)  reflp Q"
  by (rule reflp_on_mono[of UNIV R Q]) simp_all

lemma (in preorder) reflp_on_le[simp]: "reflp_on A (≤)"
  by (simp add: reflp_onI)

lemma (in preorder) reflp_on_ge[simp]: "reflp_on A (≥)"
  by (simp add: reflp_onI)


subsubsection ‹Irreflexivity›

definition irrefl_on :: "'a set  'a rel  bool" where
  "irrefl_on A r  (a  A. (a, a)  r)"

abbreviation irrefl :: "'a rel  bool" where
  "irrefl  irrefl_on UNIV"

definition irreflp_on :: "'a set  ('a  'a  bool)  bool" where
  "irreflp_on A R  (a  A. ¬ R a a)"

abbreviation irreflp :: "('a  'a  bool)  bool" where
  "irreflp  irreflp_on UNIV"

lemma irrefl_def[no_atp]: "irrefl r  (a. (a, a)  r)"
  by (simp add: irrefl_on_def)

lemma irreflp_def[no_atp]: "irreflp R  (a. ¬ R a a)"
  by (simp add: irreflp_on_def)

text @{thm [source] irrefl_def} and @{thm [source] irreflp_def} are for backward compatibility.›

lemma irreflp_on_irrefl_on_eq [pred_set_conv]: "irreflp_on A (λa b. (a, b)  r)  irrefl_on A r"
  by (simp add: irrefl_on_def irreflp_on_def)

lemmas irreflp_irrefl_eq = irreflp_on_irrefl_on_eq[of UNIV]

lemma irrefl_onI: "(a. a  A  (a, a)  r)  irrefl_on A r"
  by (simp add: irrefl_on_def)

lemma irreflI[intro?]: "(a. (a, a)  r)  irrefl r"
  by (rule irrefl_onI[of UNIV, simplified])

lemma irreflp_onI: "(a. a  A  ¬ R a a)  irreflp_on A R"
  by (rule irrefl_onI[to_pred])

lemma irreflpI[intro?]: "(a. ¬ R a a)  irreflp R"
  by (rule irreflI[to_pred])

lemma irrefl_onD: "irrefl_on A r  a  A  (a, a)  r"
  by (simp add: irrefl_on_def)

lemma irreflD: "irrefl r  (x, x)  r"
  by (rule irrefl_onD[of UNIV, simplified])

lemma irreflp_onD: "irreflp_on A R  a  A  ¬ R a a"
  by (rule irrefl_onD[to_pred])

lemma irreflpD: "irreflp R  ¬ R x x"
  by (rule irreflD[to_pred])

lemma irrefl_on_distinct [code]: "irrefl_on A r  ((a, b)  r. a  A  b  A  a  b)"
  by (auto simp add: irrefl_on_def)

lemmas irrefl_distinct = irrefl_on_distinct ― ‹For backward compatibility›

lemma irrefl_on_subset: "irrefl_on A r  B  A  irrefl_on B r"
  by (auto simp: irrefl_on_def)

lemma irreflp_on_subset: "irreflp_on A R  B  A  irreflp_on B R"
  by (auto simp: irreflp_on_def)

lemma irreflp_on_image: "irreflp_on (f ` A) R  irreflp_on A (λa b. R (f a) (f b))"
  by (simp add: irreflp_on_def)

lemma (in preorder) irreflp_on_less[simp]: "irreflp_on A (<)"
  by (simp add: irreflp_onI)

lemma (in preorder) irreflp_on_greater[simp]: "irreflp_on A (>)"
  by (simp add: irreflp_onI)


subsubsection ‹Asymmetry›

definition asym_on :: "'a set  'a rel  bool" where
  "asym_on A r  (x  A. y  A. (x, y)  r  (y, x)  r)"

abbreviation asym :: "'a rel  bool" where
  "asym  asym_on UNIV"

definition asymp_on :: "'a set  ('a  'a  bool)  bool" where
  "asymp_on A R  (x  A. y  A. R x y  ¬ R y x)"

abbreviation asymp :: "('a  'a  bool)  bool" where
  "asymp  asymp_on UNIV"

lemma asymp_on_asym_on_eq[pred_set_conv]: "asymp_on A (λx y. (x, y)  r)  asym_on A r"
  by (simp add: asymp_on_def asym_on_def)

lemmas asymp_asym_eq = asymp_on_asym_on_eq[of UNIV] ― ‹For backward compatibility›

lemma asym_onI[intro]:
  "(x y. x  A  y  A  (x, y)  r  (y, x)  r)  asym_on A r"
  by (simp add: asym_on_def)

lemma asymI[intro]: "(x y. (x, y)  r  (y, x)  r)  asym r"
  by (simp add: asym_onI)

lemma asymp_onI[intro]:
  "(x y. x  A  y  A  R x y  ¬ R y x)  asymp_on A R"
  by (rule asym_onI[to_pred])

lemma asympI[intro]: "(x y. R x y  ¬ R y x)  asymp R"
  by (rule asymI[to_pred])

lemma asym_onD: "asym_on A r  x  A  y  A  (x, y)  r  (y, x)  r"
  by (simp add: asym_on_def)

lemma asymD: "asym r  (x, y)  r  (y, x)  r"
  by (simp add: asym_onD)

lemma asymp_onD: "asymp_on A R  x  A  y  A  R x y  ¬ R y x"
  by (rule asym_onD[to_pred])

lemma asympD: "asymp R  R x y  ¬ R y x"
  by (rule asymD[to_pred])

lemma asym_iff: "asym r  (x y. (x,y)  r  (y,x)  r)"
  by (blast dest: asymD)

lemma asym_on_subset: "asym_on A r  B  A  asym_on B r"
  by (auto simp: asym_on_def)

lemma asymp_on_subset: "asymp_on A R  B  A  asymp_on B R"
  by (auto simp: asymp_on_def)

lemma asymp_on_image: "asymp_on (f ` A) R  asymp_on A (λa b. R (f a) (f b))"
  by (simp add: asymp_on_def)

lemma irrefl_on_if_asym_on[simp]: "asym_on A r  irrefl_on A r"
  by (auto intro: irrefl_onI dest: asym_onD)

lemma irreflp_on_if_asymp_on[simp]: "asymp_on A r  irreflp_on A r"
  by (rule irrefl_on_if_asym_on[to_pred])

lemma (in preorder) asymp_on_less[simp]: "asymp_on A (<)"
  by (auto intro: dual_order.asym)

lemma (in preorder) asymp_on_greater[simp]: "asymp_on A (>)"
  by (auto intro: dual_order.asym)


subsubsection ‹Symmetry›

definition sym_on :: "'a set  'a rel  bool" where
  "sym_on A r  (x  A. y  A. (x, y)  r  (y, x)  r)"

abbreviation sym :: "'a rel  bool" where
  "sym  sym_on UNIV"

definition symp_on :: "'a set  ('a  'a  bool)  bool" where
  "symp_on A R  (x  A. y  A. R x y  R y x)"

abbreviation symp :: "('a  'a  bool)  bool" where
  "symp  symp_on UNIV"

lemma sym_def[no_atp]: "sym r  (x y. (x, y)  r  (y, x)  r)"
  by (simp add: sym_on_def)

lemma symp_def[no_atp]: "symp R  (x y. R x y  R y x)"
  by (simp add: symp_on_def)

text @{thm [source] sym_def} and @{thm [source] symp_def} are for backward compatibility.›

lemma symp_on_sym_on_eq[pred_set_conv]: "symp_on A (λx y. (x, y)  r)  sym_on A r"
  by (simp add: sym_on_def symp_on_def)

lemmas symp_sym_eq = symp_on_sym_on_eq[of UNIV] ― ‹For backward compatibility›

lemma sym_on_subset: "sym_on A r  B  A  sym_on B r"
  by (auto simp: sym_on_def)

lemma symp_on_subset: "symp_on A R  B  A  symp_on B R"
  by (auto simp: symp_on_def)

lemma symp_on_image: "symp_on (f ` A) R  symp_on A (λa b. R (f a) (f b))"
  by (simp add: symp_on_def)

lemma sym_onI: "(x y. x  A  y  A  (x, y)  r  (y, x)  r)  sym_on A r"
  by (simp add: sym_on_def)

lemma symI [intro?]: "(x y. (x, y)  r  (y, x)  r)  sym r"
  by (simp add: sym_onI)

lemma symp_onI: "(x y. x  A  y  A  R x y  R y x)  symp_on A R"
  by (rule sym_onI[to_pred])

lemma sympI [intro?]: "(x y. R x y  R y x)  symp R"
  by (rule symI[to_pred])

lemma symE:
  assumes "sym r" and "(b, a)  r"
  obtains "(a, b)  r"
  using assms by (simp add: sym_def)

lemma sympE:
  assumes "symp r" and "r b a"
  obtains "r a b"
  using assms by (rule symE [to_pred])

lemma sym_onD: "sym_on A r  x  A  y  A  (x, y)  r  (y, x)  r"
  by (simp add: sym_on_def)

lemma symD [dest?]: "sym r  (x, y)  r  (y, x)  r"
  by (simp add: sym_onD)

lemma symp_onD: "symp_on A R  x  A  y  A  R x y  R y x"
  by (rule sym_onD[to_pred])

lemma sympD [dest?]: "symp R  R x y  R y x"
  by (rule symD[to_pred])

lemma sym_Int: "sym r  sym s  sym (r  s)"
  by (fast intro: symI elim: symE)

lemma symp_inf: "symp r  symp s  symp (r  s)"
  by (fact sym_Int [to_pred])

lemma sym_Un: "sym r  sym s  sym (r  s)"
  by (fast intro: symI elim: symE)

lemma symp_sup: "symp r  symp s  symp (r  s)"
  by (fact sym_Un [to_pred])

lemma sym_INTER: "xS. sym (r x)  sym ((r ` S))"
  by (fast intro: symI elim: symE)

lemma symp_INF: "xS. symp (r x)  symp ((r ` S))"
  by (fact sym_INTER [to_pred])

lemma sym_UNION: "xS. sym (r x)  sym ((r ` S))"
  by (fast intro: symI elim: symE)

lemma symp_SUP: "xS. symp (r x)  symp ((r ` S))"
  by (fact sym_UNION [to_pred])


subsubsection ‹Antisymmetry›

definition antisym_on :: "'a set  'a rel  bool" where
  "antisym_on A r  (x  A. y  A. (x, y)  r  (y, x)  r  x = y)"

abbreviation antisym :: "'a rel  bool" where
  "antisym  antisym_on UNIV"

definition antisymp_on :: "'a set  ('a  'a  bool)  bool" where
  "antisymp_on A R  (x  A. y  A. R x y  R y x  x = y)"

abbreviation antisymp :: "('a  'a  bool)  bool" where
  "antisymp  antisymp_on UNIV"

lemma antisym_def[no_atp]: "antisym r  (x y. (x, y)  r  (y, x)  r  x = y)"
  by (simp add: antisym_on_def)

lemma antisymp_def[no_atp]: "antisymp R  (x y. R x y  R y x  x = y)"
  by (simp add: antisymp_on_def)

text @{thm [source] antisym_def} and @{thm [source] antisymp_def} are for backward compatibility.›

lemma antisymp_on_antisym_on_eq[pred_set_conv]:
  "antisymp_on A (λx y. (x, y)  r)  antisym_on A r"
  by (simp add: antisym_on_def antisymp_on_def)

lemmas antisymp_antisym_eq = antisymp_on_antisym_on_eq[of UNIV] ― ‹For backward compatibility›

lemma antisym_on_subset: "antisym_on A r  B  A  antisym_on B r"
  by (auto simp: antisym_on_def)

lemma antisymp_on_subset: "antisymp_on A R  B  A  antisymp_on B R"
  by (auto simp: antisymp_on_def)

lemma antisymp_on_image:
  assumes "inj_on f A"
  shows "antisymp_on (f ` A) R  antisymp_on A (λa b. R (f a) (f b))"
  using assms by (auto simp: antisymp_on_def inj_on_def)

lemma antisym_onI:
  "(x y. x  A  y  A  (x, y)  r  (y, x)  r  x = y)  antisym_on A r"
  unfolding antisym_on_def by simp

lemma antisymI [intro?]:
  "(x y. (x, y)  r  (y, x)  r  x = y)  antisym r"
  by (simp add: antisym_onI)

lemma antisymp_onI:
  "(x y. x  A  y  A  R x y  R y x  x = y)  antisymp_on A R"
  by (rule antisym_onI[to_pred])

lemma antisympI [intro?]:
  "(x y. R x y  R y x  x = y)  antisymp R"
  by (rule antisymI[to_pred])
    
lemma antisym_onD:
  "antisym_on A r  x  A  y  A  (x, y)  r  (y, x)  r  x = y"
  by (simp add: antisym_on_def)

lemma antisymD [dest?]:
  "antisym r  (x, y)  r  (y, x)  r  x = y"
  by (simp add: antisym_onD)

lemma antisymp_onD:
  "antisymp_on A R  x  A  y  A  R x y  R y x  x = y"
  by (rule antisym_onD[to_pred])

lemma antisympD [dest?]:
  "antisymp R  R x y  R y x  x = y"
  by (rule antisymD[to_pred])

lemma antisym_subset:
  "r  s  antisym s  antisym r"
  unfolding antisym_def by blast

lemma antisymp_less_eq:
  "r  s  antisymp s  antisymp r"
  by (fact antisym_subset [to_pred])
    
lemma antisym_empty [simp]:
  "antisym {}"
  unfolding antisym_def by blast

lemma antisym_bot [simp]:
  "antisymp "
  by (fact antisym_empty [to_pred])
    
lemma antisymp_equality [simp]:
  "antisymp HOL.eq"
  by (auto intro: antisympI)

lemma antisym_singleton [simp]:
  "antisym {x}"
  by (blast intro: antisymI)

lemma antisym_on_if_asym_on: "asym_on A r  antisym_on A r"
  by (auto intro: antisym_onI dest: asym_onD)

lemma antisymp_on_if_asymp_on: "asymp_on A R  antisymp_on A R"
  by (rule antisym_on_if_asym_on[to_pred])

lemma (in preorder) antisymp_on_less[simp]: "antisymp_on A (<)"
  by (rule antisymp_on_if_asymp_on[OF asymp_on_less])

lemma (in preorder) antisymp_on_greater[simp]: "antisymp_on A (>)"
  by (rule antisymp_on_if_asymp_on[OF asymp_on_greater])

lemma (in order) antisymp_on_le[simp]: "antisymp_on A (≤)"
  by (simp add: antisymp_onI)

lemma (in order) antisymp_on_ge[simp]: "antisymp_on A (≥)"
  by (simp add: antisymp_onI)


subsubsection ‹Transitivity›

definition trans_on :: "'a set  'a rel  bool" where
  "trans_on A r  (x  A. y  A. z  A. (x, y)  r  (y, z)  r  (x, z)  r)"

abbreviation trans :: "'a rel  bool" where
  "trans  trans_on UNIV"

definition transp_on :: "'a set  ('a  'a  bool)  bool" where
  "transp_on A R  (x  A. y  A. z  A. R x y  R y z  R x z)"

abbreviation transp :: "('a  'a  bool)  bool" where
  "transp  transp_on UNIV"

lemma trans_def[no_atp]: "trans r  (x y z. (x, y)  r  (y, z)  r  (x, z)  r)"
  by (simp add: trans_on_def)

lemma transp_def: "transp R  (x y z. R x y  R y z  R x z)"
  by (simp add: transp_on_def)

text @{thm [source] trans_def} and @{thm [source] transp_def} are for backward compatibility.›

lemma transp_on_trans_on_eq[pred_set_conv]: "transp_on A (λx y. (x, y)  r)  trans_on A r"
  by (simp add: trans_on_def transp_on_def)

lemmas transp_trans_eq = transp_on_trans_on_eq[of UNIV] ― ‹For backward compatibility›

lemma trans_onI:
  "(x y z. x  A  y  A  z  A  (x, y)  r  (y, z)  r  (x, z)  r) 
  trans_on A r"
  unfolding trans_on_def
  by (intro ballI) iprover

lemma transI [intro?]: "(x y z. (x, y)  r  (y, z)  r  (x, z)  r)  trans r"
  by (rule trans_onI)

lemma transp_onI:
  "(x y z. x  A  y  A  z  A  R x y  R y z  R x z)  transp_on A R"
  by (rule trans_onI[to_pred])

lemma transpI [intro?]: "(x y z. R x y  R y z  R x z)  transp R"
  by (rule transI[to_pred])

lemma transE:
  assumes "trans r" and "(x, y)  r" and "(y, z)  r"
  obtains "(x, z)  r"
  using assms by (unfold trans_def) iprover

lemma transpE:
  assumes "transp r" and "r x y" and "r y z"
  obtains "r x z"
  using assms by (rule transE [to_pred])

lemma trans_onD:
  "trans_on A r  x  A  y  A  z  A  (x, y)  r  (y, z)  r  (x, z)  r"
  unfolding trans_on_def
  by (elim ballE) iprover+

lemma transD[dest?]: "trans r  (x, y)  r  (y, z)  r  (x, z)  r"
  by (simp add: trans_onD[of UNIV r x y z])

lemma transp_onD: "transp_on A R  x  A  y  A  z  A  R x y  R y z  R x z"
  by (rule trans_onD[to_pred])

lemma transpD[dest?]: "transp R  R x y  R y z  R x z"
  by (rule transD[to_pred])

lemma trans_on_subset: "trans_on A r  B  A  trans_on B r"
  by (auto simp: trans_on_def)

lemma transp_on_subset: "transp_on A R  B  A  transp_on B R"
  by (auto simp: transp_on_def)

lemma transp_on_image: "transp_on (f ` A) R  transp_on A (λa b. R (f a) (f b))"
  by (simp add: transp_on_def)

lemma trans_Int: "trans r  trans s  trans (r  s)"
  by (fast intro: transI elim: transE)

lemma transp_inf: "transp r  transp s  transp (r  s)"
  by (fact trans_Int [to_pred])

lemma trans_INTER: "xS. trans (r x)  trans ((r ` S))"
  by (fast intro: transI elim: transD)

lemma transp_INF: "xS. transp (r x)  transp ((r ` S))"
  by (fact trans_INTER [to_pred])

lemma trans_on_join [code]:
  "trans_on A r  ((x, y1)  r. x  A  y1  A 
    ((y2, z)  r. y1 = y2  z  A  (x, z)  r))"
  by (auto simp: trans_on_def)

lemma trans_join: "trans r  ((x, y1)  r. (y2, z)  r. y1 = y2  (x, z)  r)"
  by (auto simp add: trans_def)

lemma transp_trans: "transp r  trans {(x, y). r x y}"
  by (simp add: trans_def transp_def)

lemma transp_equality [simp]: "transp (=)"
  by (auto intro: transpI)

lemma trans_empty [simp]: "trans {}"
  by (blast intro: transI)

lemma transp_empty [simp]: "transp (λx y. False)"
  using trans_empty[to_pred] by (simp add: bot_fun_def)

lemma trans_singleton [simp]: "trans {(a, a)}"
  by (blast intro: transI)

lemma transp_singleton [simp]: "transp (λx y. x = a  y = a)"
  by (simp add: transp_def)

lemma asym_on_iff_irrefl_on_if_trans_on: "trans_on A r  asym_on A r  irrefl_on A r"
  by (auto intro: irrefl_on_if_asym_on dest: trans_onD irrefl_onD)

lemma asymp_on_iff_irreflp_on_if_transp_on: "transp_on A R  asymp_on A R  irreflp_on A R"
  by (rule asym_on_iff_irrefl_on_if_trans_on[to_pred])

lemma (in preorder) transp_on_le[simp]: "transp_on A (≤)"
  by (auto intro: transp_onI order_trans)

lemma (in preorder) transp_on_less[simp]: "transp_on A (<)"
  by (auto intro: transp_onI less_trans)

lemma (in preorder) transp_on_ge[simp]: "transp_on A (≥)"
  by (auto intro: transp_onI order_trans)

lemma (in preorder) transp_on_greater[simp]: "transp_on A (>)"
  by (auto intro: transp_onI less_trans)


subsubsection ‹Totality›

definition total_on :: "'a set  'a rel  bool" where
  "total_on A r  (xA. yA. x  y  (x, y)  r  (y, x)  r)"

abbreviation total :: "'a rel  bool" where
  "total  total_on UNIV"

definition totalp_on :: "'a set  ('a  'a  bool)  bool" where
  "totalp_on A R  (x  A. y  A. x  y  R x y  R y x)"

abbreviation totalp :: "('a  'a  bool)  bool" where
  "totalp  totalp_on UNIV"

lemma totalp_on_total_on_eq[pred_set_conv]: "totalp_on A (λx y. (x, y)  r)  total_on A r"
  by (simp add: totalp_on_def total_on_def)

lemma total_onI [intro?]:
  "(x y. x  A  y  A  x  y  (x, y)  r  (y, x)  r)  total_on A r"
  unfolding total_on_def by blast

lemma totalI: "(x y. x  y  (x, y)  r  (y, x)  r)  total r"
  by (rule total_onI)

lemma totalp_onI: "(x y. x  A  y  A  x  y  R x y  R y x)  totalp_on A R"
  by (rule total_onI[to_pred])

lemma totalpI: "(x y. x  y  R x y  R y x)  totalp R"
  by (rule totalI[to_pred])

lemma totalp_onD:
  "totalp_on A R  x  A  y  A  x  y  R x y  R y x"
  by (simp add: totalp_on_def)

lemma totalpD: "totalp R  x  y  R x y  R y x"
  by (simp add: totalp_onD)

lemma total_on_subset: "total_on A r  B  A  total_on B r"
  by (auto simp: total_on_def)

lemma totalp_on_subset: "totalp_on A R  B  A  totalp_on B R"
  by (auto intro: totalp_onI dest: totalp_onD)

lemma totalp_on_image:
  assumes "inj_on f A"
  shows "totalp_on (f ` A) R  totalp_on A (λa b. R (f a) (f b))"
  using assms by (auto simp: totalp_on_def inj_on_def)

lemma total_on_empty [simp]: "total_on {} r"
  by (simp add: total_on_def)

lemma totalp_on_empty [simp]: "totalp_on {} R"
  by (simp add: totalp_on_def)

lemma total_on_singleton [simp]: "total_on {x} r"
  by (simp add: total_on_def)

lemma totalp_on_singleton [simp]: "totalp_on {x} R"
  by (simp add: totalp_on_def)

lemma (in linorder) totalp_on_less[simp]: "totalp_on A (<)"
  by (auto intro: totalp_onI)

lemma (in linorder) totalp_on_greater[simp]: "totalp_on A (>)"
  by (auto intro: totalp_onI)

lemma (in linorder) totalp_on_le[simp]: "totalp_on A (≤)"
  by (rule totalp_onI, rule linear)

lemma (in linorder) totalp_on_ge[simp]: "totalp_on A (≥)"
  by (rule totalp_onI, rule linear)


subsubsection ‹Single valued relations›

definition single_valued :: "('a × 'b) set  bool"
  where "single_valued r  (x y. (x, y)  r  (z. (x, z)  r  y = z))"

definition single_valuedp :: "('a  'b  bool)  bool"
  where "single_valuedp r  (x y. r x y  (z. r x z  y = z))"

lemma single_valuedp_single_valued_eq [pred_set_conv]:
  "single_valuedp (λx y. (x, y)  r)  single_valued r"
  by (simp add: single_valued_def single_valuedp_def)

lemma single_valuedp_iff_Uniq:
  "single_valuedp r  (x. 1y. r x y)"
  unfolding Uniq_def single_valuedp_def by auto

lemma single_valuedI:
  "(x y. (x, y)  r  (z. (x, z)  r  y = z))  single_valued r"
  unfolding single_valued_def by blast

lemma single_valuedpI:
  "(x y. r x y  (z. r x z  y = z))  single_valuedp r"
  by (fact single_valuedI [to_pred])

lemma single_valuedD:
  "single_valued r  (x, y)  r  (x, z)  r  y = z"
  by (simp add: single_valued_def)

lemma single_valuedpD:
  "single_valuedp r  r x y  r x z  y = z"
  by (fact single_valuedD [to_pred])

lemma single_valued_empty [simp]:
  "single_valued {}"
  by (simp add: single_valued_def)

lemma single_valuedp_bot [simp]:
  "single_valuedp "
  by (fact single_valued_empty [to_pred])

lemma single_valued_subset:
  "r  s  single_valued s  single_valued r"
  unfolding single_valued_def by blast

lemma single_valuedp_less_eq:
  "r  s  single_valuedp s  single_valuedp r"
  by (fact single_valued_subset [to_pred])


subsection ‹Relation operations›

subsubsection ‹The identity relation›

definition Id :: "'a rel"
  where "Id = {p. x. p = (x, x)}"

lemma IdI [intro]: "(a, a)  Id"
  by (simp add: Id_def)

lemma IdE [elim!]: "p  Id  (x. p = (x, x)  P)  P"
  unfolding Id_def by (iprover elim: CollectE)

lemma pair_in_Id_conv [iff]: "(a, b)  Id  a = b"
  unfolding Id_def by blast

lemma refl_Id: "refl Id"
  by (simp add: refl_on_def)

lemma antisym_Id: "antisym Id"
  ― ‹A strange result, since Id› is also symmetric.›
  by (simp add: antisym_def)

lemma sym_Id: "sym Id"
  by (simp add: sym_def)

lemma trans_Id: "trans Id"
  by (simp add: trans_def)

lemma single_valued_Id [simp]: "single_valued Id"
  by (unfold single_valued_def) blast

lemma irrefl_diff_Id [simp]: "irrefl (r - Id)"
  by (simp add: irrefl_def)

lemma trans_diff_Id: "trans r  antisym r  trans (r - Id)"
  unfolding antisym_def trans_def by blast

lemma total_on_diff_Id [simp]: "total_on A (r - Id) = total_on A r"
  by (simp add: total_on_def)

lemma Id_fstsnd_eq: "Id = {x. fst x = snd x}"
  by force


subsubsection ‹Diagonal: identity over a set›

definition Id_on :: "'a set  'a rel"
  where "Id_on A = (xA. {(x, x)})"

lemma Id_on_empty [simp]: "Id_on {} = {}"
  by (simp add: Id_on_def)

lemma Id_on_eqI: "a = b  a  A  (a, b)  Id_on A"
  by (simp add: Id_on_def)

lemma Id_onI [intro!]: "a  A  (a, a)  Id_on A"
  by (rule Id_on_eqI) (rule refl)

lemma Id_onE [elim!]: "c  Id_on A  (x. x  A  c = (x, x)  P)  P"
  ― ‹The general elimination rule.›
  unfolding Id_on_def by (iprover elim!: UN_E singletonE)

lemma Id_on_iff: "(x, y)  Id_on A  x = y  x  A"
  by blast

lemma Id_on_def' [nitpick_unfold]: "Id_on {x. A x} = Collect (λ(x, y). x = y  A x)"
  by auto

lemma Id_on_subset_Times: "Id_on A  A × A"
  by blast

lemma refl_on_Id_on: "refl_on A (Id_on A)"
  by (rule refl_onI [OF Id_on_subset_Times Id_onI])

lemma antisym_Id_on [simp]: "antisym (Id_on A)"
  unfolding antisym_def by blast

lemma sym_Id_on [simp]: "sym (Id_on A)"
  by (rule symI) clarify

lemma trans_Id_on [simp]: "trans (Id_on A)"
  by (fast intro: transI elim: transD)

lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
  unfolding single_valued_def by blast


subsubsection ‹Composition›

inductive_set relcomp  :: "('a × 'b) set  ('b × 'c) set  ('a × 'c) set"  (infixr "O" 75)
  for r :: "('a × 'b) set" and s :: "('b × 'c) set"
  where relcompI [intro]: "(a, b)  r  (b, c)  s  (a, c)  r O s"

notation relcompp (infixr "OO" 75)

lemmas relcomppI = relcompp.intros

text ‹
  For historic reasons, the elimination rules are not wholly corresponding.
  Feel free to consolidate this.
›

inductive_cases relcompEpair: "(a, c)  r O s"
inductive_cases relcomppE [elim!]: "(r OO s) a c"

lemma relcompE [elim!]: "xz  r O s 
  (x y z. xz = (x, z)  (x, y)  r  (y, z)  s   P)  P"
  apply (cases xz)
  apply simp
  apply (erule relcompEpair)
  apply iprover
  done

lemma R_O_Id [simp]: "R O Id = R"
  by fast

lemma Id_O_R [simp]: "Id O R = R"
  by fast

lemma relcomp_empty1 [simp]: "{} O R = {}"
  by blast

lemma relcompp_bot1 [simp]: " OO R = "
  by (fact relcomp_empty1 [to_pred])

lemma relcomp_empty2 [simp]: "R O {} = {}"
  by blast

lemma relcompp_bot2 [simp]: "R OO  = "
  by (fact relcomp_empty2 [to_pred])

lemma O_assoc: "(R O S) O T = R O (S O T)"
  by blast

lemma relcompp_assoc: "(r OO s) OO t = r OO (s OO t)"
  by (fact O_assoc [to_pred])

lemma trans_O_subset: "trans r  r O r  r"
  by (unfold trans_def) blast

lemma transp_relcompp_less_eq: "transp r  r OO r  r "
  by (fact trans_O_subset [to_pred])

lemma relcomp_mono: "r'  r  s'  s  r' O s'  r O s"
  by blast

lemma relcompp_mono: "r'  r  s'  s  r' OO s'  r OO s "
  by (fact relcomp_mono [to_pred])

lemma relcomp_subset_Sigma: "r  A × B  s  B × C  r O s  A × C"
  by blast

lemma relcomp_distrib [simp]: "R O (S  T) = (R O S)  (R O T)"
  by auto

lemma relcompp_distrib [simp]: "R OO (S  T) = R OO S  R OO T"
  by (fact relcomp_distrib [to_pred])

lemma relcomp_distrib2 [simp]: "(S  T) O R = (S O R)  (T O R)"
  by auto

lemma relcompp_distrib2 [simp]: "(S  T) OO R = S OO R  T OO R"
  by (fact relcomp_distrib2 [to_pred])

lemma relcomp_UNION_distrib: "s O (r ` I) = (iI. s O r i) "
  by auto

lemma relcompp_SUP_distrib: "s OO (r ` I) = (iI. s OO r i)"
  by (fact relcomp_UNION_distrib [to_pred])
    
lemma relcomp_UNION_distrib2: "(r ` I) O s = (iI. r i O s) "
  by auto

lemma relcompp_SUP_distrib2: "(r ` I) OO s = (iI. r i OO s)"
  by (fact relcomp_UNION_distrib2 [to_pred])
    
lemma single_valued_relcomp: "single_valued r  single_valued s  single_valued (r O s)"
  unfolding single_valued_def by blast

lemma relcomp_unfold: "r O s = {(x, z). y. (x, y)  r  (y, z)  s}"
  by (auto simp add: set_eq_iff)

lemma relcompp_apply: "(R OO S) a c  (b. R a b  S b c)"
  unfolding relcomp_unfold [to_pred] ..

lemma eq_OO: "(=) OO R = R"
  by blast

lemma OO_eq: "R OO (=) = R"
  by blast


subsubsection ‹Converse›

inductive_set converse :: "('a × 'b) set  ('b × 'a) set"  ("(_¯)" [1000] 999)
  for r :: "('a × 'b) set"
  where "(a, b)  r  (b, a)  r¯"

notation conversep  ("(_¯¯)" [1000] 1000)

notation (ASCII)
  converse  ("(_^-1)" [1000] 999) and
  conversep ("(_^--1)" [1000] 1000)

lemma converseI [sym]: "(a, b)  r  (b, a)  r¯"
  by (fact converse.intros)

lemma conversepI (* CANDIDATE [sym] *): "r a b  r¯¯ b a"
  by (fact conversep.intros)

lemma converseD [sym]: "(a, b)  r¯  (b, a)  r"
  by (erule converse.cases) iprover

lemma conversepD (* CANDIDATE [sym] *): "r¯¯ b a  r a b"
  by (fact converseD [to_pred])

lemma converseE [elim!]: "yx  r¯  (x y. yx = (y, x)  (x, y)  r  P)  P"
  ― ‹More general than converseD›, as it ``splits'' the member of the relation.›
  apply (cases yx)
  apply simp
  apply (erule converse.cases)
  apply iprover
  done

lemmas conversepE [elim!] = conversep.cases

lemma converse_iff [iff]: "(a, b)  r¯  (b, a)  r"
  by (auto intro: converseI)

lemma conversep_iff [iff]: "r¯¯ a b = r b a"
  by (fact converse_iff [to_pred])

lemma converse_converse [simp]: "(r¯)¯ = r"
  by (simp add: set_eq_iff)

lemma conversep_conversep [simp]: "(r¯¯)¯¯ = r"
  by (fact converse_converse [to_pred])

lemma converse_empty[simp]: "{}¯ = {}"
  by auto

lemma converse_UNIV[simp]: "UNIV¯ = UNIV"
  by auto

lemma converse_relcomp: "(r O s)¯ = s¯ O r¯"
  by blast

lemma converse_relcompp: "(r OO s)¯¯ = s¯¯ OO r¯¯"
  by (iprover intro: order_antisym conversepI relcomppI elim: relcomppE dest: conversepD)

lemma converse_Int: "(r  s)¯ = r¯  s¯"
  by blast

lemma converse_meet: "(r  s)¯¯ = r¯¯  s¯¯"
  by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)

lemma converse_Un: "(r  s)¯ = r¯  s¯"
  by blast

lemma converse_join: "(r  s)¯¯ = r¯¯  s¯¯"
  by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)

lemma converse_INTER: "((r ` S))¯ = (xS. (r x)¯)"
  by fast

lemma converse_UNION: "((r ` S))¯ = (xS. (r x)¯)"
  by blast

lemma converse_mono[simp]: "r¯  s ¯  r  s"
  by auto

lemma conversep_mono[simp]: "r¯¯  s ¯¯  r  s"
  by (fact converse_mono[to_pred])

lemma converse_inject[simp]: "r¯ = s ¯  r = s"
  by auto

lemma conversep_inject[simp]: "r¯¯ = s ¯¯  r = s"
  by (fact converse_inject[to_pred])

lemma converse_subset_swap: "r  s ¯  r ¯  s"
  by auto

lemma conversep_le_swap: "r  s ¯¯  r ¯¯  s"
  by (fact converse_subset_swap[to_pred])

lemma converse_Id [simp]: "Id¯ = Id"
  by blast

lemma converse_Id_on [simp]: "(Id_on A)¯ = Id_on A"
  by blast

lemma refl_on_converse [simp]: "refl_on A (r¯) = refl_on A r"
  by (auto simp: refl_on_def)

lemma reflp_on_conversp [simp]: "reflp_on A R¯¯  reflp_on A R"
  by (auto simp: reflp_on_def)

lemma irrefl_on_converse [simp]: "irrefl_on A (r¯) = irrefl_on A r"
  by (simp add: irrefl_on_def)

lemma irreflp_on_converse [simp]: "irreflp_on A (r¯¯) = irreflp_on A r"
  by (rule irrefl_on_converse[to_pred])

lemma sym_on_converse [simp]: "sym_on A (r¯) = sym_on A r"
  by (auto intro: sym_onI dest: sym_onD)

lemma symp_on_conversep [simp]: "symp_on A R¯¯ = symp_on A R"
  by (rule sym_on_converse[to_pred])

lemma asym_on_converse [simp]: "asym_on A (r¯) = asym_on A r"
  by (auto dest: asym_onD)

lemma asymp_on_conversep [simp]: "asymp_on A R¯¯ = asymp_on A R"
  by (rule asym_on_converse[to_pred])

lemma antisym_on_converse [simp]: "antisym_on A (r¯) = antisym_on A r"
  by (auto intro: antisym_onI dest: antisym_onD)

lemma antisymp_on_conversep [simp]: "antisymp_on A R¯¯ = antisymp_on A R"
  by (rule antisym_on_converse[to_pred])

lemma trans_on_converse [simp]: "trans_on A (r¯) = trans_on A r"
  by (auto intro: trans_onI dest: trans_onD)

lemma transp_on_conversep [simp]: "transp_on A R¯¯ = transp_on A R"
  by (rule trans_on_converse[to_pred])

lemma sym_conv_converse_eq: "sym r  r¯ = r"
  unfolding sym_def by fast

lemma sym_Un_converse: "sym (r  r¯)"
  unfolding sym_def by blast

lemma sym_Int_converse: "sym (r  r¯)"
  unfolding sym_def by blast

lemma total_on_converse [simp]: "total_on A (r¯) = total_on A r"
  by (auto simp: total_on_def)

lemma totalp_on_converse [simp]: "totalp_on A R¯¯ = totalp_on A R"
  by (rule total_on_converse[to_pred])

lemma conversep_noteq [simp]: "(≠)¯¯ = (≠)"
  by (auto simp add: fun_eq_iff)

lemma conversep_eq [simp]: "(=)¯¯ = (=)"
  by (auto simp add: fun_eq_iff)

lemma converse_unfold [code]: "r¯ = {(y, x). (x, y)  r}"
  by (simp add: set_eq_iff)


subsubsection ‹Domain, range and field›

inductive_set Domain :: "('a × 'b) set  'a set" for r :: "('a × 'b) set"
  where DomainI [intro]: "(a, b)  r  a  Domain r"

lemmas DomainPI = Domainp.DomainI

inductive_cases DomainE [elim!]: "a  Domain r"
inductive_cases DomainpE [elim!]: "Domainp r a"

inductive_set Range :: "('a × 'b) set  'b set" for r :: "('a × 'b) set"
  where RangeI [intro]: "(a, b)  r  b  Range r"

lemmas RangePI = Rangep.RangeI

inductive_cases RangeE [elim!]: "b  Range r"
inductive_cases RangepE [elim!]: "Rangep r b"

definition Field :: "'a rel  'a set"
  where "Field r = Domain r  Range r"

lemma Field_iff: "x  Field r  (y. (x,y)  r  (y,x)  r)"
  by (auto simp: Field_def)

lemma FieldI1: "(i, j)  R  i  Field R"
  unfolding Field_def by blast

lemma FieldI2: "(i, j)  R  j  Field R"
  unfolding Field_def by auto

lemma Domain_fst [code]: "Domain r = fst ` r"
  by force

lemma Range_snd [code]: "Range r = snd ` r"
  by force

lemma fst_eq_Domain: "fst ` R = Domain R"
  by force

lemma snd_eq_Range: "snd ` R = Range R"
  by force

lemma range_fst [simp]: "range fst = UNIV"
  by (auto simp: fst_eq_Domain)

lemma range_snd [simp]: "range snd = UNIV"
  by (auto simp: snd_eq_Range)

lemma Domain_empty [simp]: "Domain {} = {}"
  by auto

lemma Range_empty [simp]: "Range {} = {}"
  by auto

lemma Field_empty [simp]: "Field {} = {}"
  by (simp add: Field_def)

lemma Domain_empty_iff: "Domain r = {}  r = {}"
  by auto

lemma Range_empty_iff: "Range r = {}  r = {}"
  by auto

lemma Domain_insert [simp]: "Domain (insert (a, b) r) = insert a (Domain r)"
  by blast

lemma Range_insert [simp]: "Range (insert (a, b) r) = insert b (Range r)"
  by blast

lemma Field_insert [simp]: "Field (insert (a, b) r) = {a, b}  Field r"
  by (auto simp add: Field_def)

lemma Domain_iff: "a  Domain r  (y. (a, y)  r)"
  by blast

lemma Range_iff: "a  Range r  (y. (y, a)  r)"
  by blast

lemma Domain_Id [simp]: "Domain Id = UNIV"
  by blast

lemma Range_Id [simp]: "Range Id = UNIV"
  by blast

lemma Domain_Id_on [simp]: "Domain (Id_on A) = A"
  by blast

lemma Range_Id_on [simp]: "Range (Id_on A) = A"
  by blast

lemma Domain_Un_eq: "Domain (A  B) = Domain A  Domain B"
  by blast

lemma Range_Un_eq: "Range (A  B) = Range A  Range B"
  by blast

lemma Field_Un [simp]: "Field (r  s) = Field r  Field s"
  by (auto simp: Field_def)

lemma Domain_Int_subset: "Domain (A  B)  Domain A  Domain B"
  by blast

lemma Range_Int_subset: "Range (A  B)  Range A  Range B"
  by blast

lemma Domain_Diff_subset: "Domain A - Domain B  Domain (A - B)"
  by blast

lemma Range_Diff_subset: "Range A - Range B  Range (A - B)"
  by blast

lemma Domain_Union: "Domain (S) = (AS. Domain A)"
  by blast

lemma Range_Union: "Range (S) = (AS. Range A)"
  by blast

lemma Field_Union [simp]: "Field (R) = (Field ` R)"
  by (auto simp: Field_def)

lemma Domain_converse [simp]: "Domain (r¯) = Range r"
  by auto

lemma Range_converse [simp]: "Range (r¯) = Domain r"
  by blast

lemma Field_converse [simp]: "Field (r¯) = Field r"
  by (auto simp: Field_def)

lemma Domain_Collect_case_prod [simp]: "Domain {(x, y). P x y} = {x. y. P x y}"
  by auto

lemma Range_Collect_case_prod [simp]: "Range {(x, y). P x y} = {y. x. P x y}"
  by auto

lemma Domain_mono: "r  s  Domain r  Domain s"
  by blast

lemma Range_mono: "r  s  Range r  Range s"
  by blast

lemma mono_Field: "r  s  Field r  Field s"
  by (auto simp: Field_def Domain_def Range_def)

lemma Domain_unfold: "Domain r = {x. y. (x, y)  r}"
  by blast

lemma Field_square [simp]: "Field (x × x) = x"
  unfolding Field_def by blast


subsubsection ‹Image of a set under a relation›

definition Image :: "('a × 'b) set  'a set  'b set"  (infixr "``" 90)
  where "r `` s = {y. xs. (x, y)  r}"

lemma Image_iff: "b  r``A  (xA. (x, b)  r)"
  by (simp add: Image_def)

lemma Image_singleton: "r``{a} = {b. (a, b)  r}"
  by (simp add: Image_def)

lemma Image_singleton_iff [iff]: "b  r``{a}  (a, b)  r"
  by (rule Image_iff [THEN trans]) simp

lemma ImageI [intro]: "(a, b)  r  a  A  b  r``A"
  unfolding Image_def by blast

lemma ImageE [elim!]: "b  r `` A  (x. (x, b)  r  x  A  P)  P"
  unfolding Image_def by (iprover elim!: CollectE bexE)

lemma rev_ImageI: "a  A  (a, b)  r  b  r `` A"
  ― ‹This version's more effective when we already have the required a›
  by blast

lemma Image_empty1 [simp]: "{} `` X = {}"
by auto

lemma Image_empty2 [simp]: "R``{} = {}"
by blast

lemma Image_Id [simp]: "Id `` A = A"
  by blast

lemma Image_Id_on [simp]: "Id_on A `` B = A  B"
  by blast

lemma Image_Int_subset: "R `` (A  B)  R `` A  R `` B"
  by blast

lemma Image_Int_eq: "single_valued (converse R)  R `` (A  B) = R `` A  R `` B"
  by (auto simp: single_valued_def)

lemma Image_Un: "R `` (A  B) = R `` A  R `` B"
  by blast

lemma Un_Image: "(R  S) `` A = R `` A  S `` A"
  by blast

lemma Image_subset: "r  A × B  r``C  B"
  by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)

lemma Image_eq_UN: "r``B = (y B. r``{y})"
  ― ‹NOT suitable for rewriting›
  by blast

lemma Image_mono: "r'  r  A'  A  (r' `` A')  (r `` A)"
  by blast

lemma Image_UN: "r `` ((B ` A)) = (xA. r `` (B x))"
  by blast

lemma UN_Image: "(iI. X i) `` S = (iI. X i `` S)"
  by auto

lemma Image_INT_subset: "(r `` ((B ` A)))  (xA. r `` (B x))"
  by blast

text ‹Converse inclusion requires some assumptions›
lemma Image_INT_eq:
  assumes "single_valued (r¯)"
    and "A  {}"
  shows "r `` ((B ` A)) = (xA. r `` B x)"
proof(rule equalityI, rule Image_INT_subset)
  show "(xA. r `` B x)  r ``  (B ` A)"
  proof
    fix x
    assume "x  (xA. r `` B x)"
    then show "x  r ``  (B ` A)"
      using assms unfolding single_valued_def by simp blast
  qed
qed

lemma Image_subset_eq: "r``A  B  A  - ((r¯) `` (- B))"
  by blast

lemma Image_Collect_case_prod [simp]: "{(x, y). P x y} `` A = {y. xA. P x y}"
  by auto

lemma Sigma_Image: "(SIGMA x:A. B x) `` X = (xX  A. B x)"
  by auto

lemma relcomp_Image: "(X O Y) `` Z = Y `` (X `` Z)"
  by auto


subsubsection ‹Inverse image›

definition inv_image :: "'b rel  ('a  'b)  'a rel"
  where "inv_image r f = {(x, y). (f x, f y)  r}"

definition inv_imagep :: "('b  'b  bool)  ('a  'b)  'a  'a  bool"
  where "inv_imagep r f = (λx y. r (f x) (f y))"

lemma [pred_set_conv]: "inv_imagep (λx y. (x, y)  r) f = (λx y. (x, y)  inv_image r f)"
  by (simp add: inv_image_def inv_imagep_def)

lemma sym_inv_image: "sym r  sym (inv_image r f)"
  unfolding sym_def inv_image_def by blast

lemma trans_inv_image: "trans r  trans (inv_image r f)"
  unfolding trans_def inv_image_def
  by (simp (no_asm)) blast

lemma total_inv_image: "inj f; total r  total (inv_image r f)"
  unfolding inv_image_def total_on_def by (auto simp: inj_eq)

lemma asym_inv_image: "asym R  asym (inv_image R f)"
  by (simp add: inv_image_def asym_iff)

lemma in_inv_image[simp]: "(x, y)  inv_image r f  (f x, f y)  r"
  by (auto simp: inv_image_def)

lemma converse_inv_image[simp]: "(inv_image R f)¯ = inv_image (R¯) f"
  unfolding inv_image_def converse_unfold by auto

lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
  by (simp add: inv_imagep_def)


subsubsection ‹Powerset›

definition Powp :: "('a  bool)  'a set  bool"
  where "Powp A = (λB. x  B. A x)"

lemma Powp_Pow_eq [pred_set_conv]: "Powp (λx. x  A) = (λx. x  Pow A)"
  by (auto simp add: Powp_def fun_eq_iff)

lemmas Powp_mono [mono] = Pow_mono [to_pred]

end