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