(* Title: HOL/Predicate.thy ID: $Id: Predicate.thy,v 1.9 2008/05/07 08:56:39 berghofe Exp $ Author: Stefan Berghofer, TU Muenchen *) header {* Predicates *} theory Predicate imports Inductive Relation begin subsection {* Equality and Subsets *} lemma pred_equals_eq: "((λx. x ∈ R) = (λx. x ∈ S)) = (R = S)" by (simp add: mem_def) lemma pred_equals_eq2 [pred_set_conv]: "((λx y. (x, y) ∈ R) = (λx y. (x, y) ∈ S)) = (R = S)" by (simp add: expand_fun_eq mem_def) lemma pred_subset_eq: "((λx. x ∈ R) <= (λx. x ∈ S)) = (R <= S)" by (simp add: mem_def) lemma pred_subset_eq2 [pred_set_conv]: "((λx y. (x, y) ∈ R) <= (λx y. (x, y) ∈ S)) = (R <= S)" by fast subsection {* Top and bottom elements *} lemma top1I [intro!]: "top x" by (simp add: top_fun_eq top_bool_eq) lemma top2I [intro!]: "top x y" by (simp add: top_fun_eq top_bool_eq) lemma bot1E [elim!]: "bot x ==> P" by (simp add: bot_fun_eq bot_bool_eq) lemma bot2E [elim!]: "bot x y ==> P" by (simp add: bot_fun_eq bot_bool_eq) subsection {* The empty set *} lemma bot_empty_eq: "bot = (λx. x ∈ {})" by (auto simp add: expand_fun_eq) lemma bot_empty_eq2: "bot = (λx y. (x, y) ∈ {})" by (auto simp add: expand_fun_eq) subsection {* Binary union *} lemma sup1_iff [simp]: "sup A B x <-> A x | B x" by (simp add: sup_fun_eq sup_bool_eq) lemma sup2_iff [simp]: "sup A B x y <-> A x y | B x y" by (simp add: sup_fun_eq sup_bool_eq) lemma sup_Un_eq [pred_set_conv]: "sup (λx. x ∈ R) (λx. x ∈ S) = (λx. x ∈ R ∪ S)" by (simp add: expand_fun_eq) lemma sup_Un_eq2 [pred_set_conv]: "sup (λx y. (x, y) ∈ R) (λx y. (x, y) ∈ S) = (λx y. (x, y) ∈ R ∪ S)" by (simp add: expand_fun_eq) lemma sup1I1 [elim?]: "A x ==> sup A B x" by simp lemma sup2I1 [elim?]: "A x y ==> sup A B x y" by simp lemma sup1I2 [elim?]: "B x ==> sup A B x" by simp lemma sup2I2 [elim?]: "B x y ==> sup A B x y" by simp text {* \medskip Classical introduction rule: no commitment to @{text A} vs @{text B}. *} lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x" by auto lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y" by auto lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" by simp iprover lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" by simp iprover subsection {* Binary intersection *} lemma inf1_iff [simp]: "inf A B x <-> A x ∧ B x" by (simp add: inf_fun_eq inf_bool_eq) lemma inf2_iff [simp]: "inf A B x y <-> A x y ∧ B x y" by (simp add: inf_fun_eq inf_bool_eq) lemma inf_Int_eq [pred_set_conv]: "inf (λx. x ∈ R) (λx. x ∈ S) = (λx. x ∈ R ∩ S)" by (simp add: expand_fun_eq) lemma inf_Int_eq2 [pred_set_conv]: "inf (λx y. (x, y) ∈ R) (λx y. (x, y) ∈ S) = (λx y. (x, y) ∈ R ∩ S)" by (simp add: expand_fun_eq) lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" by simp lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" by simp lemma inf1D1: "inf A B x ==> A x" by simp lemma inf2D1: "inf A B x y ==> A x y" by simp lemma inf1D2: "inf A B x ==> B x" by simp lemma inf2D2: "inf A B x y ==> B x y" by simp lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" by simp lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" by simp subsection {* Unions of families *} lemma SUP1_iff [simp]: "(SUP x:A. B x) b = (EX x:A. B x b)" by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast lemma SUP2_iff [simp]: "(SUP x:A. B x) b c = (EX x:A. B x b c)" by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b" by auto lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c" by auto lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R" by auto lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R" by auto lemma SUP_UN_eq: "(SUP i. (λx. x ∈ r i)) = (λx. x ∈ (UN i. r i))" by (simp add: expand_fun_eq) lemma SUP_UN_eq2: "(SUP i. (λx y. (x, y) ∈ r i)) = (λx y. (x, y) ∈ (UN i. r i))" by (simp add: expand_fun_eq) subsection {* Intersections of families *} lemma INF1_iff [simp]: "(INF x:A. B x) b = (ALL x:A. B x b)" by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast lemma INF2_iff [simp]: "(INF x:A. B x) b c = (ALL x:A. B x b c)" by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b" by auto lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" by auto lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" by auto lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" by auto lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" by auto lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" by auto lemma INF_INT_eq: "(INF i. (λx. x ∈ r i)) = (λx. x ∈ (INT i. r i))" by (simp add: expand_fun_eq) lemma INF_INT_eq2: "(INF i. (λx y. (x, y) ∈ r i)) = (λx y. (x, y) ∈ (INT i. r i))" by (simp add: expand_fun_eq) subsection {* Composition of two relations *} inductive pred_comp :: "['b => 'c => bool, 'a => 'b => bool] => 'a => 'c => bool" (infixr "OO" 75) for r :: "'b => 'c => bool" and s :: "'a => 'b => bool" where pred_compI [intro]: "s a b ==> r b c ==> (r OO s) a c" inductive_cases pred_compE [elim!]: "(r OO s) a c" lemma pred_comp_rel_comp_eq [pred_set_conv]: "((λx y. (x, y) ∈ r) OO (λx y. (x, y) ∈ s)) = (λx y. (x, y) ∈ r O s)" by (auto simp add: expand_fun_eq elim: pred_compE) subsection {* Converse *} inductive conversep :: "('a => 'b => bool) => 'b => 'a => bool" ("(_^--1)" [1000] 1000) for r :: "'a => 'b => bool" where conversepI: "r a b ==> r^--1 b a" notation (xsymbols) conversep ("(_¯¯)" [1000] 1000) lemma conversepD: assumes ab: "r^--1 a b" shows "r b a" using ab by cases simp lemma conversep_iff [iff]: "r^--1 a b = r b a" by (iprover intro: conversepI dest: conversepD) lemma conversep_converse_eq [pred_set_conv]: "(λx y. (x, y) ∈ r)^--1 = (λx y. (x, y) ∈ r^-1)" by (auto simp add: expand_fun_eq) lemma conversep_conversep [simp]: "(r^--1)^--1 = r" by (iprover intro: order_antisym conversepI dest: conversepD) lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1" by (iprover intro: order_antisym conversepI pred_compI elim: pred_compE dest: conversepD) lemma converse_meet: "(inf r s)^--1 = inf r^--1 s^--1" by (simp add: inf_fun_eq inf_bool_eq) (iprover intro: conversepI ext dest: conversepD) lemma converse_join: "(sup r s)^--1 = sup r^--1 s^--1" by (simp add: sup_fun_eq sup_bool_eq) (iprover intro: conversepI ext dest: conversepD) lemma conversep_noteq [simp]: "(op ~=)^--1 = op ~=" by (auto simp add: expand_fun_eq) lemma conversep_eq [simp]: "(op =)^--1 = op =" by (auto simp add: expand_fun_eq) subsection {* Domain *} inductive DomainP :: "('a => 'b => bool) => 'a => bool" for r :: "'a => 'b => bool" where DomainPI [intro]: "r a b ==> DomainP r a" inductive_cases DomainPE [elim!]: "DomainP r a" lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (λx y. (x, y) ∈ r) = (λx. x ∈ Domain r)" by (blast intro!: Orderings.order_antisym predicate1I) subsection {* Range *} inductive RangeP :: "('a => 'b => bool) => 'b => bool" for r :: "'a => 'b => bool" where RangePI [intro]: "r a b ==> RangeP r b" inductive_cases RangePE [elim!]: "RangeP r b" lemma RangeP_Range_eq [pred_set_conv]: "RangeP (λx y. (x, y) ∈ r) = (λx. x ∈ Range r)" by (blast intro!: Orderings.order_antisym predicate1I) subsection {* Inverse image *} 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 in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" by (simp add: inv_imagep_def) subsection {* The Powerset operator *} 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 expand_fun_eq) lemmas Powp_mono [mono] = Pow_mono [to_pred pred_subset_eq] subsection {* Properties of relations - predicate versions *} abbreviation antisymP :: "('a => 'a => bool) => bool" where "antisymP r == antisym {(x, y). r x y}" abbreviation transP :: "('a => 'a => bool) => bool" where "transP r == trans {(x, y). r x y}" abbreviation single_valuedP :: "('a => 'b => bool) => bool" where "single_valuedP r == single_valued {(x, y). r x y}" end
lemma pred_equals_eq:
((λx. x ∈ R) = (λx. x ∈ S)) = (R = S)
lemma pred_equals_eq2:
((λx y. (x, y) ∈ R) = (λx y. (x, y) ∈ S)) = (R = S)
lemma pred_subset_eq:
((λx. x ∈ R) ⊆ (λx. x ∈ S)) = (R ⊆ S)
lemma pred_subset_eq2:
((λx y. (x, y) ∈ R) ≤ (λx y. (x, y) ∈ S)) = (R ⊆ S)
lemma top1I:
top x
lemma top2I:
top x y
lemma bot1E:
bot x ==> P
lemma bot2E:
bot x y ==> P
lemma bot_empty_eq:
bot = (λx. x ∈ {})
lemma bot_empty_eq2:
bot = (λx y. (x, y) ∈ {})
lemma sup1_iff:
sup A B x = (A x ∨ B x)
lemma sup2_iff:
sup A B x y = (A x y ∨ B x y)
lemma sup_Un_eq:
sup (λx. x ∈ R) (λx. x ∈ S) = (λx. x ∈ R ∪ S)
lemma sup_Un_eq2:
sup (λx y. (x, y) ∈ R) (λx y. (x, y) ∈ S) = (λx y. (x, y) ∈ R ∪ S)
lemma sup1I1:
A x ==> sup A B x
lemma sup2I1:
A x y ==> sup A B x y
lemma sup1I2:
B x ==> sup A B x
lemma sup2I2:
B x y ==> sup A B x y
lemma sup1CI:
(¬ B x ==> A x) ==> sup A B x
lemma sup2CI:
(¬ B x y ==> A x y) ==> sup A B x y
lemma sup1E:
[| sup A B x; A x ==> P; B x ==> P |] ==> P
lemma sup2E:
[| sup A B x y; A x y ==> P; B x y ==> P |] ==> P
lemma inf1_iff:
inf A B x = (A x ∧ B x)
lemma inf2_iff:
inf A B x y = (A x y ∧ B x y)
lemma inf_Int_eq:
inf (λx. x ∈ R) (λx. x ∈ S) = (λx. x ∈ R ∩ S)
lemma inf_Int_eq2:
inf (λx y. (x, y) ∈ R) (λx y. (x, y) ∈ S) = (λx y. (x, y) ∈ R ∩ S)
lemma inf1I:
[| A x; B x |] ==> inf A B x
lemma inf2I:
[| A x y; B x y |] ==> inf A B x y
lemma inf1D1:
inf A B x ==> A x
lemma inf2D1:
inf A B x y ==> A x y
lemma inf1D2:
inf A B x ==> B x
lemma inf2D2:
inf A B x y ==> B x y
lemma inf1E:
[| inf A B x; [| A x; B x |] ==> P |] ==> P
lemma inf2E:
[| inf A B x y; [| A x y; B x y |] ==> P |] ==> P
lemma SUP1_iff:
(SUP x:A. B x) b = (∃x∈A. B x b)
lemma SUP2_iff:
(SUP x:A. B x) b c = (∃x∈A. B x b c)
lemma SUP1_I:
[| a ∈ A; B a b |] ==> (SUP x:A. B x) b
lemma SUP2_I:
[| a ∈ A; B a b c |] ==> (SUP x:A. B x) b c
lemma SUP1_E:
[| (SUP x:A. B x) b; !!x. [| x ∈ A; B x b |] ==> R |] ==> R
lemma SUP2_E:
[| (SUP x:A. B x) b c; !!x. [| x ∈ A; B x b c |] ==> R |] ==> R
lemma SUP_UN_eq:
(SUP i. (λx. x ∈ r i)) = (λx. x ∈ (UN i. r i))
lemma SUP_UN_eq2:
(SUP i. (λx y. (x, y) ∈ r i)) = (λx y. (x, y) ∈ (UN i. r i))
lemma INF1_iff:
(INF x:A. B x) b = (∀x∈A. B x b)
lemma INF2_iff:
(INF x:A. B x) b c = (∀x∈A. B x b c)
lemma INF1_I:
(!!x. x ∈ A ==> B x b) ==> (INF x:A. B x) b
lemma INF2_I:
(!!x. x ∈ A ==> B x b c) ==> (INF x:A. B x) b c
lemma INF1_D:
[| (INF x:A. B x) b; a ∈ A |] ==> B a b
lemma INF2_D:
[| (INF x:A. B x) b c; a ∈ A |] ==> B a b c
lemma INF1_E:
[| (INF x:A. B x) b; B a b ==> R; a ∉ A ==> R |] ==> R
lemma INF2_E:
[| (INF x:A. B x) b c; B a b c ==> R; a ∉ A ==> R |] ==> R
lemma INF_INT_eq:
(INF i. (λx. x ∈ r i)) = (λx. x ∈ (INT i. r i))
lemma INF_INT_eq2:
(INF i. (λx y. (x, y) ∈ r i)) = (λx y. (x, y) ∈ (INT i. r i))
lemma pred_comp_rel_comp_eq:
(λx y. (x, y) ∈ r) OO (λx y. (x, y) ∈ s) = (λx y. (x, y) ∈ r O s)
lemma conversepD:
r^--1 a b ==> r b a
lemma conversep_iff:
r^--1 a b = r b a
lemma conversep_converse_eq:
(λx y. (x, y) ∈ r)^--1 = (λx y. (x, y) ∈ r^-1)
lemma conversep_conversep:
r^--1^--1 = r
lemma converse_pred_comp:
(r OO s)^--1 = s^--1 OO r^--1
lemma converse_meet:
(inf r s)^--1 = inf r^--1 s^--1
lemma converse_join:
(sup r s)^--1 = sup r^--1 s^--1
lemma conversep_noteq:
op ≠^--1 = op ≠
lemma conversep_eq:
op =^--1 = op =
lemma DomainP_Domain_eq:
DomainP (λx y. (x, y) ∈ r) = (λx. x ∈ Domain r)
lemma RangeP_Range_eq:
RangeP (λx y. (x, y) ∈ r) = (λx. x ∈ Range r)
lemma
inv_imagep (λx y. (x, y) ∈ r) f = (λx y. (x, y) ∈ inv_image r f)
lemma in_inv_imagep:
inv_imagep r f x y = r (f x) (f y)
lemma Powp_Pow_eq:
Powp (λx. x ∈ A) = (λx. x ∈ Pow A)
lemma Powp_mono:
A ⊆ B ==> Powp A ⊆ Powp B