wenzelm@10358: (* Title: HOL/Relation.thy haftmann@46664: Author: Lawrence C Paulson, Cambridge University Computer Laboratory; Stefan Berghofer, TU Muenchen nipkow@1128: *) nipkow@1128: haftmann@46664: header {* Relations – as sets of pairs, and binary predicates *} berghofe@12905: nipkow@15131: theory Relation haftmann@32850: imports Datatype Finite_Set nipkow@15131: begin paulson@5978: haftmann@46664: notation haftmann@46664: bot ("\") and haftmann@46664: top ("\") and haftmann@46664: inf (infixl "\" 70) and haftmann@46664: sup (infixl "\" 65) and haftmann@46664: Inf ("\_" [900] 900) and haftmann@46664: Sup ("\_" [900] 900) haftmann@46664: haftmann@46664: syntax (xsymbols) haftmann@46664: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@46664: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@46664: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@46664: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@46664: haftmann@46664: haftmann@46664: subsection {* Classical rules for reasoning on predicates *} haftmann@46664: haftmann@46689: (* CANDIDATE declare predicate1I [Pure.intro!, intro!] *) haftmann@46664: declare predicate1D [Pure.dest?, dest?] haftmann@46689: (* CANDIDATE declare predicate1D [Pure.dest, dest] *) haftmann@46664: declare predicate2I [Pure.intro!, intro!] haftmann@46664: declare predicate2D [Pure.dest, dest] haftmann@46664: declare bot1E [elim!] haftmann@46664: declare bot2E [elim!] haftmann@46664: declare top1I [intro!] haftmann@46664: declare top2I [intro!] haftmann@46664: declare inf1I [intro!] haftmann@46664: declare inf2I [intro!] haftmann@46664: declare inf1E [elim!] haftmann@46664: declare inf2E [elim!] haftmann@46664: declare sup1I1 [intro?] haftmann@46664: declare sup2I1 [intro?] haftmann@46664: declare sup1I2 [intro?] haftmann@46664: declare sup2I2 [intro?] haftmann@46664: declare sup1E [elim!] haftmann@46664: declare sup2E [elim!] haftmann@46664: declare sup1CI [intro!] haftmann@46664: declare sup2CI [intro!] haftmann@46664: declare INF1_I [intro!] haftmann@46664: declare INF2_I [intro!] haftmann@46664: declare INF1_D [elim] haftmann@46664: declare INF2_D [elim] haftmann@46664: declare INF1_E [elim] haftmann@46664: declare INF2_E [elim] haftmann@46664: declare SUP1_I [intro] haftmann@46664: declare SUP2_I [intro] haftmann@46664: declare SUP1_E [elim!] haftmann@46664: declare SUP2_E [elim!] haftmann@46664: haftmann@46664: haftmann@46664: subsection {* Conversions between set and predicate relations *} haftmann@46664: haftmann@46664: lemma pred_equals_eq [pred_set_conv]: "((\x. x \ R) = (\x. x \ S)) \ (R = S)" haftmann@46664: by (simp add: set_eq_iff fun_eq_iff) haftmann@46664: haftmann@46664: lemma pred_equals_eq2 [pred_set_conv]: "((\x y. (x, y) \ R) = (\x y. (x, y) \ S)) \ (R = S)" haftmann@46664: by (simp add: set_eq_iff fun_eq_iff) haftmann@46664: haftmann@46664: lemma pred_subset_eq [pred_set_conv]: "((\x. x \ R) \ (\x. x \ S)) \ (R \ S)" haftmann@46664: by (simp add: subset_iff le_fun_def) haftmann@46664: haftmann@46664: lemma pred_subset_eq2 [pred_set_conv]: "((\x y. (x, y) \ R) \ (\x y. (x, y) \ S)) \ (R \ S)" haftmann@46664: by (simp add: subset_iff le_fun_def) haftmann@46664: haftmann@46689: lemma bot_empty_eq (* CANDIDATE [pred_set_conv] *): "\ = (\x. x \ {})" haftmann@46689: by (auto simp add: fun_eq_iff) haftmann@46689: haftmann@46689: lemma bot_empty_eq2 (* CANDIDATE [pred_set_conv] *): "\ = (\x y. (x, y) \ {})" haftmann@46664: by (auto simp add: fun_eq_iff) haftmann@46664: haftmann@46689: (* CANDIDATE lemma top_empty_eq [pred_set_conv]: "\ = (\x. x \ UNIV)" haftmann@46689: by (auto simp add: fun_eq_iff) *) haftmann@46689: haftmann@46689: (* CANDIDATE lemma top_empty_eq2 [pred_set_conv]: "\ = (\x y. (x, y) \ UNIV)" haftmann@46689: by (auto simp add: fun_eq_iff) *) haftmann@46664: haftmann@46664: lemma inf_Int_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" haftmann@46664: by (simp add: inf_fun_def) haftmann@46664: haftmann@46664: lemma inf_Int_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" haftmann@46664: by (simp add: inf_fun_def) haftmann@46664: haftmann@46664: lemma sup_Un_eq [pred_set_conv]: "(\x. x \ R) \ (\x. x \ S) = (\x. x \ R \ S)" haftmann@46664: by (simp add: sup_fun_def) haftmann@46664: haftmann@46664: lemma sup_Un_eq2 [pred_set_conv]: "(\x y. (x, y) \ R) \ (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" haftmann@46664: by (simp add: sup_fun_def) haftmann@46664: haftmann@46689: lemma INF_INT_eq (* CANDIDATE [pred_set_conv] *): "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" haftmann@46664: by (simp add: INF_apply fun_eq_iff) haftmann@46664: haftmann@46689: lemma INF_INT_eq2 (* CANDIDATE [pred_set_conv] *): "(\i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i. r i))" haftmann@46664: by (simp add: INF_apply fun_eq_iff) haftmann@46664: haftmann@46664: lemma SUP_UN_eq [pred_set_conv]: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" haftmann@46664: by (simp add: SUP_apply fun_eq_iff) haftmann@46664: haftmann@46664: lemma SUP_UN_eq2 [pred_set_conv]: "(\i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i. r i))" haftmann@46664: by (simp add: SUP_apply fun_eq_iff) haftmann@46664: haftmann@46664: haftmann@46664: subsection {* Relations as sets of pairs *} wenzelm@12913: nipkow@46372: type_synonym 'a rel = "('a * 'a) set" nipkow@46372: wenzelm@19656: definition wenzelm@21404: converse :: "('a * 'b) set => ('b * 'a) set" wenzelm@21404: ("(_^-1)" [1000] 999) where haftmann@45137: "r^-1 = {(y, x). (x, y) : r}" paulson@7912: wenzelm@21210: notation (xsymbols) wenzelm@19656: converse ("(_\)" [1000] 999) wenzelm@19656: wenzelm@19656: definition krauss@32235: rel_comp :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set" wenzelm@21404: (infixr "O" 75) where haftmann@45137: "r O s = {(x,z). EX y. (x, y) : r & (y, z) : s}" wenzelm@12913: wenzelm@21404: definition wenzelm@21404: Image :: "[('a * 'b) set, 'a set] => 'b set" wenzelm@21404: (infixl "``" 90) where haftmann@45137: "r `` s = {y. EX x:s. (x,y):r}" paulson@7912: wenzelm@21404: definition wenzelm@21404: Id :: "('a * 'a) set" where -- {* the identity relation *} haftmann@45137: "Id = {p. EX x. p = (x,x)}" paulson@7912: wenzelm@21404: definition nipkow@30198: Id_on :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *} haftmann@45137: "Id_on A = (\x\A. {(x,x)})" wenzelm@12913: wenzelm@21404: definition wenzelm@21404: Domain :: "('a * 'b) set => 'a set" where haftmann@45137: "Domain r = {x. EX y. (x,y):r}" paulson@5978: wenzelm@21404: definition wenzelm@21404: Range :: "('a * 'b) set => 'b set" where haftmann@45137: "Range r = Domain(r^-1)" paulson@5978: wenzelm@21404: definition wenzelm@21404: Field :: "('a * 'a) set => 'a set" where haftmann@45137: "Field r = Domain r \ Range r" paulson@10786: wenzelm@21404: definition nipkow@30198: refl_on :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *} haftmann@45137: "refl_on A r \ r \ A \ A & (ALL x: A. (x,x) : r)" paulson@6806: nipkow@26297: abbreviation nipkow@30198: refl :: "('a * 'a) set => bool" where -- {* reflexivity over a type *} haftmann@45137: "refl \ refl_on UNIV" nipkow@26297: wenzelm@21404: definition wenzelm@21404: sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *} haftmann@45137: "sym r \ (ALL x y. (x,y): r --> (y,x): r)" paulson@6806: wenzelm@21404: definition wenzelm@21404: antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *} haftmann@45137: "antisym r \ (ALL x y. (x,y):r --> (y,x):r --> x=y)" paulson@6806: wenzelm@21404: definition wenzelm@21404: trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *} haftmann@45137: "trans r \ (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)" paulson@5978: wenzelm@21404: definition haftmann@45137: irrefl :: "('a * 'a) set => bool" where haftmann@45137: "irrefl r \ (\x. (x,x) \ r)" nipkow@29859: nipkow@29859: definition haftmann@45137: total_on :: "'a set => ('a * 'a) set => bool" where haftmann@45137: "total_on A r \ (\x\A.\y\A. x\y \ (x,y)\r \ (y,x)\r)" nipkow@29859: nipkow@29859: abbreviation "total \ total_on UNIV" nipkow@29859: nipkow@29859: definition wenzelm@21404: single_valued :: "('a * 'b) set => bool" where haftmann@45137: "single_valued r \ (ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z))" berghofe@7014: wenzelm@21404: definition wenzelm@21404: inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where haftmann@45137: "inv_image r f = {(x, y). (f x, f y) : r}" oheimb@11136: berghofe@12905: haftmann@46664: subsubsection {* The identity relation *} berghofe@12905: berghofe@12905: lemma IdI [intro]: "(a, a) : Id" nipkow@26271: by (simp add: Id_def) berghofe@12905: berghofe@12905: lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P" nipkow@26271: by (unfold Id_def) (iprover elim: CollectE) berghofe@12905: berghofe@12905: lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)" nipkow@26271: by (unfold Id_def) blast berghofe@12905: nipkow@30198: lemma refl_Id: "refl Id" nipkow@30198: by (simp add: refl_on_def) berghofe@12905: berghofe@12905: lemma antisym_Id: "antisym Id" berghofe@12905: -- {* A strange result, since @{text Id} is also symmetric. *} nipkow@26271: by (simp add: antisym_def) berghofe@12905: huffman@19228: lemma sym_Id: "sym Id" nipkow@26271: by (simp add: sym_def) huffman@19228: berghofe@12905: lemma trans_Id: "trans Id" nipkow@26271: by (simp add: trans_def) berghofe@12905: berghofe@12905: haftmann@46664: subsubsection {* Diagonal: identity over a set *} berghofe@12905: nipkow@30198: lemma Id_on_empty [simp]: "Id_on {} = {}" nipkow@30198: by (simp add: Id_on_def) paulson@13812: nipkow@30198: lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A" nipkow@30198: by (simp add: Id_on_def) berghofe@12905: blanchet@35828: lemma Id_onI [intro!,no_atp]: "a : A ==> (a, a) : Id_on A" nipkow@30198: by (rule Id_on_eqI) (rule refl) berghofe@12905: nipkow@30198: lemma Id_onE [elim!]: nipkow@30198: "c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P" wenzelm@12913: -- {* The general elimination rule. *} nipkow@30198: by (unfold Id_on_def) (iprover elim!: UN_E singletonE) berghofe@12905: nipkow@30198: lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)" nipkow@26271: by blast berghofe@12905: haftmann@45967: lemma Id_on_def' [nitpick_unfold]: haftmann@44278: "Id_on {x. A x} = Collect (\(x, y). x = y \ A x)" haftmann@44278: by auto bulwahn@40923: nipkow@30198: lemma Id_on_subset_Times: "Id_on A \ A \ A" nipkow@26271: by blast berghofe@12905: berghofe@12905: haftmann@46664: subsubsection {* Composition of two relations *} berghofe@12905: wenzelm@12913: lemma rel_compI [intro]: krauss@32235: "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s" nipkow@26271: by (unfold rel_comp_def) blast berghofe@12905: wenzelm@12913: lemma rel_compE [elim!]: "xz : r O s ==> krauss@32235: (!!x y z. xz = (x, z) ==> (x, y) : r ==> (y, z) : s ==> P) ==> P" nipkow@26271: by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE) berghofe@12905: berghofe@12905: lemma rel_compEpair: krauss@32235: "(a, c) : r O s ==> (!!y. (a, y) : r ==> (y, c) : s ==> P) ==> P" nipkow@26271: by (iprover elim: rel_compE Pair_inject ssubst) berghofe@12905: berghofe@12905: lemma R_O_Id [simp]: "R O Id = R" nipkow@26271: by fast berghofe@12905: berghofe@12905: lemma Id_O_R [simp]: "Id O R = R" nipkow@26271: by fast berghofe@12905: krauss@23185: lemma rel_comp_empty1[simp]: "{} O R = {}" nipkow@26271: by blast krauss@23185: krauss@23185: lemma rel_comp_empty2[simp]: "R O {} = {}" nipkow@26271: by blast krauss@23185: berghofe@12905: lemma O_assoc: "(R O S) O T = R O (S O T)" nipkow@26271: by blast berghofe@12905: wenzelm@12913: lemma trans_O_subset: "trans r ==> r O r \ r" nipkow@26271: by (unfold trans_def) blast berghofe@12905: wenzelm@12913: lemma rel_comp_mono: "r' \ r ==> s' \ s ==> (r' O s') \ (r O s)" nipkow@26271: by blast berghofe@12905: berghofe@12905: lemma rel_comp_subset_Sigma: krauss@32235: "r \ A \ B ==> s \ B \ C ==> (r O s) \ A \ C" nipkow@26271: by blast berghofe@12905: krauss@28008: lemma rel_comp_distrib[simp]: "R O (S \ T) = (R O S) \ (R O T)" krauss@28008: by auto krauss@28008: krauss@28008: lemma rel_comp_distrib2[simp]: "(S \ T) O R = (S O R) \ (T O R)" krauss@28008: by auto krauss@28008: krauss@36772: lemma rel_comp_UNION_distrib: "s O UNION I r = UNION I (%i. s O r i)" krauss@36772: by auto krauss@36772: krauss@36772: lemma rel_comp_UNION_distrib2: "UNION I r O s = UNION I (%i. r i O s)" krauss@36772: by auto krauss@36772: wenzelm@12913: haftmann@46664: subsubsection {* Reflexivity *} wenzelm@12913: nipkow@30198: lemma refl_onI: "r \ A \ A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r" nipkow@30198: by (unfold refl_on_def) (iprover intro!: ballI) berghofe@12905: nipkow@30198: lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r" nipkow@30198: by (unfold refl_on_def) blast berghofe@12905: nipkow@30198: lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A" nipkow@30198: by (unfold refl_on_def) blast huffman@19228: nipkow@30198: lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A" nipkow@30198: by (unfold refl_on_def) blast huffman@19228: nipkow@30198: lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \ B) (r \ s)" nipkow@30198: by (unfold refl_on_def) blast huffman@19228: nipkow@30198: lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \ B) (r \ s)" nipkow@30198: by (unfold refl_on_def) blast huffman@19228: nipkow@30198: lemma refl_on_INTER: nipkow@30198: "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)" nipkow@30198: by (unfold refl_on_def) fast huffman@19228: nipkow@30198: lemma refl_on_UNION: nipkow@30198: "ALL x:S. refl_on (A x) (r x) \ refl_on (UNION S A) (UNION S r)" nipkow@30198: by (unfold refl_on_def) blast huffman@19228: nipkow@30198: lemma refl_on_empty[simp]: "refl_on {} {}" nipkow@30198: by(simp add:refl_on_def) nipkow@26297: nipkow@30198: lemma refl_on_Id_on: "refl_on A (Id_on A)" nipkow@30198: by (rule refl_onI [OF Id_on_subset_Times Id_onI]) huffman@19228: blanchet@41792: lemma refl_on_def' [nitpick_unfold, code]: bulwahn@41056: "refl_on A r = ((\(x, y) \ r. x : A \ y : A) \ (\x \ A. (x, x) : r))" bulwahn@41056: by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2) wenzelm@12913: haftmann@46664: haftmann@46664: subsubsection {* Antisymmetry *} berghofe@12905: berghofe@12905: lemma antisymI: berghofe@12905: "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r" nipkow@26271: by (unfold antisym_def) iprover berghofe@12905: berghofe@12905: lemma antisymD: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b" nipkow@26271: by (unfold antisym_def) iprover berghofe@12905: huffman@19228: lemma antisym_subset: "r \ s ==> antisym s ==> antisym r" nipkow@26271: by (unfold antisym_def) blast wenzelm@12913: huffman@19228: lemma antisym_empty [simp]: "antisym {}" nipkow@26271: by (unfold antisym_def) blast huffman@19228: nipkow@30198: lemma antisym_Id_on [simp]: "antisym (Id_on A)" nipkow@26271: by (unfold antisym_def) blast huffman@19228: huffman@19228: haftmann@46664: subsubsection {* Symmetry *} huffman@19228: huffman@19228: lemma symI: "(!!a b. (a, b) : r ==> (b, a) : r) ==> sym r" nipkow@26271: by (unfold sym_def) iprover paulson@15177: paulson@15177: lemma symD: "sym r ==> (a, b) : r ==> (b, a) : r" nipkow@26271: by (unfold sym_def, blast) berghofe@12905: huffman@19228: lemma sym_Int: "sym r ==> sym s ==> sym (r \ s)" nipkow@26271: by (fast intro: symI dest: symD) huffman@19228: huffman@19228: lemma sym_Un: "sym r ==> sym s ==> sym (r \ s)" nipkow@26271: by (fast intro: symI dest: symD) huffman@19228: huffman@19228: lemma sym_INTER: "ALL x:S. sym (r x) ==> sym (INTER S r)" nipkow@26271: by (fast intro: symI dest: symD) huffman@19228: huffman@19228: lemma sym_UNION: "ALL x:S. sym (r x) ==> sym (UNION S r)" nipkow@26271: by (fast intro: symI dest: symD) huffman@19228: nipkow@30198: lemma sym_Id_on [simp]: "sym (Id_on A)" nipkow@26271: by (rule symI) clarify huffman@19228: huffman@19228: haftmann@46664: subsubsection {* Transitivity *} huffman@19228: haftmann@46127: lemma trans_join [code]: haftmann@45012: "trans r \ (\(x, y1) \ r. \(y2, z) \ r. y1 = y2 \ (x, z) \ r)" haftmann@45012: by (auto simp add: trans_def) haftmann@45012: berghofe@12905: lemma transI: berghofe@12905: "(!!x y z. (x, y) : r ==> (y, z) : r ==> (x, z) : r) ==> trans r" nipkow@26271: by (unfold trans_def) iprover berghofe@12905: berghofe@12905: lemma transD: "trans r ==> (a, b) : r ==> (b, c) : r ==> (a, c) : r" nipkow@26271: by (unfold trans_def) iprover berghofe@12905: huffman@19228: lemma trans_Int: "trans r ==> trans s ==> trans (r \ s)" nipkow@26271: by (fast intro: transI elim: transD) huffman@19228: huffman@19228: lemma trans_INTER: "ALL x:S. trans (r x) ==> trans (INTER S r)" nipkow@26271: by (fast intro: transI elim: transD) huffman@19228: nipkow@30198: lemma trans_Id_on [simp]: "trans (Id_on A)" nipkow@26271: by (fast intro: transI elim: transD) huffman@19228: nipkow@29859: lemma trans_diff_Id: " trans r \ antisym r \ trans (r-Id)" nipkow@29859: unfolding antisym_def trans_def by blast nipkow@29859: haftmann@46664: haftmann@46664: subsubsection {* Irreflexivity *} nipkow@29859: haftmann@46127: lemma irrefl_distinct [code]: haftmann@45012: "irrefl r \ (\(x, y) \ r. x \ y)" haftmann@45012: by (auto simp add: irrefl_def) haftmann@45012: nipkow@29859: lemma irrefl_diff_Id[simp]: "irrefl(r-Id)" nipkow@29859: by(simp add:irrefl_def) nipkow@29859: haftmann@45139: haftmann@46664: subsubsection {* Totality *} nipkow@29859: nipkow@29859: lemma total_on_empty[simp]: "total_on {} r" nipkow@29859: by(simp add:total_on_def) nipkow@29859: nipkow@29859: lemma total_on_diff_Id[simp]: "total_on A (r-Id) = total_on A r" nipkow@29859: by(simp add: total_on_def) berghofe@12905: haftmann@46664: haftmann@46664: subsubsection {* Converse *} wenzelm@12913: wenzelm@12913: lemma converse_iff [iff]: "((a,b): r^-1) = ((b,a) : r)" nipkow@26271: by (simp add: converse_def) berghofe@12905: nipkow@13343: lemma converseI[sym]: "(a, b) : r ==> (b, a) : r^-1" nipkow@26271: by (simp add: converse_def) berghofe@12905: nipkow@13343: lemma converseD[sym]: "(a,b) : r^-1 ==> (b, a) : r" nipkow@26271: by (simp add: converse_def) berghofe@12905: berghofe@12905: lemma converseE [elim!]: berghofe@12905: "yx : r^-1 ==> (!!x y. yx = (y, x) ==> (x, y) : r ==> P) ==> P" wenzelm@12913: -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *} nipkow@26271: by (unfold converse_def) (iprover elim!: CollectE splitE bexE) berghofe@12905: berghofe@12905: lemma converse_converse [simp]: "(r^-1)^-1 = r" nipkow@26271: by (unfold converse_def) blast berghofe@12905: berghofe@12905: lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1" nipkow@26271: by blast berghofe@12905: huffman@19228: lemma converse_Int: "(r \ s)^-1 = r^-1 \ s^-1" nipkow@26271: by blast huffman@19228: huffman@19228: lemma converse_Un: "(r \ s)^-1 = r^-1 \ s^-1" nipkow@26271: by blast huffman@19228: huffman@19228: lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)" nipkow@26271: by fast huffman@19228: huffman@19228: lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)" nipkow@26271: by blast huffman@19228: berghofe@12905: lemma converse_Id [simp]: "Id^-1 = Id" nipkow@26271: by blast berghofe@12905: nipkow@30198: lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A" nipkow@26271: by blast berghofe@12905: nipkow@30198: lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r" nipkow@30198: by (unfold refl_on_def) auto berghofe@12905: huffman@19228: lemma sym_converse [simp]: "sym (converse r) = sym r" nipkow@26271: by (unfold sym_def) blast huffman@19228: huffman@19228: lemma antisym_converse [simp]: "antisym (converse r) = antisym r" nipkow@26271: by (unfold antisym_def) blast berghofe@12905: huffman@19228: lemma trans_converse [simp]: "trans (converse r) = trans r" nipkow@26271: by (unfold trans_def) blast berghofe@12905: huffman@19228: lemma sym_conv_converse_eq: "sym r = (r^-1 = r)" nipkow@26271: by (unfold sym_def) fast huffman@19228: huffman@19228: lemma sym_Un_converse: "sym (r \ r^-1)" nipkow@26271: by (unfold sym_def) blast huffman@19228: huffman@19228: lemma sym_Int_converse: "sym (r \ r^-1)" nipkow@26271: by (unfold sym_def) blast huffman@19228: nipkow@29859: lemma total_on_converse[simp]: "total_on A (r^-1) = total_on A r" nipkow@29859: by (auto simp: total_on_def) nipkow@29859: wenzelm@12913: haftmann@46664: subsubsection {* Domain *} berghofe@12905: blanchet@35828: declare Domain_def [no_atp] paulson@24286: berghofe@12905: lemma Domain_iff: "(a : Domain r) = (EX y. (a, y) : r)" nipkow@26271: by (unfold Domain_def) blast berghofe@12905: berghofe@12905: lemma DomainI [intro]: "(a, b) : r ==> a : Domain r" nipkow@26271: by (iprover intro!: iffD2 [OF Domain_iff]) berghofe@12905: berghofe@12905: lemma DomainE [elim!]: berghofe@12905: "a : Domain r ==> (!!y. (a, y) : r ==> P) ==> P" nipkow@26271: by (iprover dest!: iffD1 [OF Domain_iff]) berghofe@12905: haftmann@46127: lemma Domain_fst [code]: haftmann@45012: "Domain r = fst ` r" haftmann@45012: by (auto simp add: image_def Bex_def) haftmann@45012: berghofe@12905: lemma Domain_empty [simp]: "Domain {} = {}" nipkow@26271: by blast berghofe@12905: paulson@32876: lemma Domain_empty_iff: "Domain r = {} \ r = {}" paulson@32876: by auto paulson@32876: berghofe@12905: lemma Domain_insert: "Domain (insert (a, b) r) = insert a (Domain r)" nipkow@26271: by blast berghofe@12905: berghofe@12905: lemma Domain_Id [simp]: "Domain Id = UNIV" nipkow@26271: by blast berghofe@12905: nipkow@30198: lemma Domain_Id_on [simp]: "Domain (Id_on A) = A" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Domain_Un_eq: "Domain(A \ B) = Domain(A) \ Domain(B)" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Domain_Int_subset: "Domain(A \ B) \ Domain(A) \ Domain(B)" nipkow@26271: by blast berghofe@12905: wenzelm@12913: lemma Domain_Diff_subset: "Domain(A) - Domain(B) \ Domain(A - B)" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Domain_Union: "Domain (Union S) = (\A\S. Domain A)" nipkow@26271: by blast nipkow@26271: nipkow@26271: lemma Domain_converse[simp]: "Domain(r^-1) = Range r" nipkow@26271: by(auto simp:Range_def) berghofe@12905: wenzelm@12913: lemma Domain_mono: "r \ s ==> Domain r \ Domain s" nipkow@26271: by blast berghofe@12905: krauss@36729: lemma fst_eq_Domain: "fst ` R = Domain R" huffman@44921: by force paulson@22172: haftmann@29609: lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)" haftmann@29609: by auto haftmann@29609: haftmann@29609: lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)" haftmann@29609: by auto haftmann@29609: berghofe@12905: haftmann@46664: subsubsection {* Range *} berghofe@12905: berghofe@12905: lemma Range_iff: "(a : Range r) = (EX y. (y, a) : r)" nipkow@26271: by (simp add: Domain_def Range_def) berghofe@12905: berghofe@12905: lemma RangeI [intro]: "(a, b) : r ==> b : Range r" nipkow@26271: by (unfold Range_def) (iprover intro!: converseI DomainI) berghofe@12905: berghofe@12905: lemma RangeE [elim!]: "b : Range r ==> (!!x. (x, b) : r ==> P) ==> P" nipkow@26271: by (unfold Range_def) (iprover elim!: DomainE dest!: converseD) berghofe@12905: haftmann@46127: lemma Range_snd [code]: haftmann@45012: "Range r = snd ` r" haftmann@45012: by (auto simp add: image_def Bex_def) haftmann@45012: berghofe@12905: lemma Range_empty [simp]: "Range {} = {}" nipkow@26271: by blast berghofe@12905: paulson@32876: lemma Range_empty_iff: "Range r = {} \ r = {}" paulson@32876: by auto paulson@32876: berghofe@12905: lemma Range_insert: "Range (insert (a, b) r) = insert b (Range r)" nipkow@26271: by blast berghofe@12905: berghofe@12905: lemma Range_Id [simp]: "Range Id = UNIV" nipkow@26271: by blast berghofe@12905: nipkow@30198: lemma Range_Id_on [simp]: "Range (Id_on A) = A" nipkow@26271: by auto berghofe@12905: paulson@13830: lemma Range_Un_eq: "Range(A \ B) = Range(A) \ Range(B)" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Range_Int_subset: "Range(A \ B) \ Range(A) \ Range(B)" nipkow@26271: by blast berghofe@12905: wenzelm@12913: lemma Range_Diff_subset: "Range(A) - Range(B) \ Range(A - B)" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Range_Union: "Range (Union S) = (\A\S. Range A)" nipkow@26271: by blast nipkow@26271: nipkow@26271: lemma Range_converse[simp]: "Range(r^-1) = Domain r" nipkow@26271: by blast berghofe@12905: krauss@36729: lemma snd_eq_Range: "snd ` R = Range R" huffman@44921: by force nipkow@26271: nipkow@26271: haftmann@46664: subsubsection {* Field *} nipkow@26271: nipkow@26271: lemma mono_Field: "r \ s \ Field r \ Field s" nipkow@26271: by(auto simp:Field_def Domain_def Range_def) nipkow@26271: nipkow@26271: lemma Field_empty[simp]: "Field {} = {}" nipkow@26271: by(auto simp:Field_def) nipkow@26271: nipkow@26271: lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \ Field r" nipkow@26271: by(auto simp:Field_def) nipkow@26271: nipkow@26271: lemma Field_Un[simp]: "Field (r \ s) = Field r \ Field s" nipkow@26271: by(auto simp:Field_def) nipkow@26271: nipkow@26271: lemma Field_Union[simp]: "Field (\R) = \(Field ` R)" nipkow@26271: by(auto simp:Field_def) nipkow@26271: nipkow@26271: lemma Field_converse[simp]: "Field(r^-1) = Field r" nipkow@26271: by(auto simp:Field_def) paulson@22172: berghofe@12905: haftmann@46664: subsubsection {* Image of a set under a relation *} berghofe@12905: blanchet@35828: declare Image_def [no_atp] paulson@24286: wenzelm@12913: lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)" nipkow@26271: by (simp add: Image_def) berghofe@12905: wenzelm@12913: lemma Image_singleton: "r``{a} = {b. (a, b) : r}" nipkow@26271: by (simp add: Image_def) berghofe@12905: wenzelm@12913: lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)" nipkow@26271: by (rule Image_iff [THEN trans]) simp berghofe@12905: blanchet@35828: lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A" nipkow@26271: by (unfold Image_def) blast berghofe@12905: berghofe@12905: lemma ImageE [elim!]: wenzelm@12913: "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P" nipkow@26271: by (unfold Image_def) (iprover elim!: CollectE bexE) berghofe@12905: berghofe@12905: lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A" berghofe@12905: -- {* This version's more effective when we already have the required @{text a} *} nipkow@26271: by blast berghofe@12905: berghofe@12905: lemma Image_empty [simp]: "R``{} = {}" nipkow@26271: by blast berghofe@12905: berghofe@12905: lemma Image_Id [simp]: "Id `` A = A" nipkow@26271: by blast berghofe@12905: nipkow@30198: lemma Image_Id_on [simp]: "Id_on A `` B = A \ B" nipkow@26271: by blast paulson@13830: paulson@13830: lemma Image_Int_subset: "R `` (A \ B) \ R `` A \ R `` B" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Image_Int_eq: paulson@13830: "single_valued (converse R) ==> R `` (A \ B) = R `` A \ R `` B" nipkow@26271: by (simp add: single_valued_def, blast) berghofe@12905: paulson@13830: lemma Image_Un: "R `` (A \ B) = R `` A \ R `` B" nipkow@26271: by blast berghofe@12905: paulson@13812: lemma Un_Image: "(R \ S) `` A = R `` A \ S `` A" nipkow@26271: by blast paulson@13812: wenzelm@12913: lemma Image_subset: "r \ A \ B ==> r``C \ B" nipkow@26271: by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2) berghofe@12905: paulson@13830: lemma Image_eq_UN: "r``B = (\y\ B. r``{y})" berghofe@12905: -- {* NOT suitable for rewriting *} nipkow@26271: by blast berghofe@12905: wenzelm@12913: lemma Image_mono: "r' \ r ==> A' \ A ==> (r' `` A') \ (r `` A)" nipkow@26271: by blast berghofe@12905: paulson@13830: lemma Image_UN: "(r `` (UNION A B)) = (\x\A. r `` (B x))" nipkow@26271: by blast paulson@13830: paulson@13830: lemma Image_INT_subset: "(r `` INTER A B) \ (\x\A. r `` (B x))" nipkow@26271: by blast berghofe@12905: paulson@13830: text{*Converse inclusion requires some assumptions*} paulson@13830: lemma Image_INT_eq: paulson@13830: "[|single_valued (r\); A\{}|] ==> r `` INTER A B = (\x\A. r `` B x)" paulson@13830: apply (rule equalityI) paulson@13830: apply (rule Image_INT_subset) paulson@13830: apply (simp add: single_valued_def, blast) paulson@13830: done berghofe@12905: wenzelm@12913: lemma Image_subset_eq: "(r``A \ B) = (A \ - ((r^-1) `` (-B)))" nipkow@26271: by blast berghofe@12905: berghofe@12905: haftmann@46664: subsubsection {* Single valued relations *} wenzelm@12913: wenzelm@12913: lemma single_valuedI: berghofe@12905: "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r" nipkow@26271: by (unfold single_valued_def) berghofe@12905: berghofe@12905: lemma single_valuedD: berghofe@12905: "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z" nipkow@26271: by (simp add: single_valued_def) berghofe@12905: huffman@19228: lemma single_valued_rel_comp: huffman@19228: "single_valued r ==> single_valued s ==> single_valued (r O s)" nipkow@26271: by (unfold single_valued_def) blast huffman@19228: huffman@19228: lemma single_valued_subset: huffman@19228: "r \ s ==> single_valued s ==> single_valued r" nipkow@26271: by (unfold single_valued_def) blast huffman@19228: huffman@19228: lemma single_valued_Id [simp]: "single_valued Id" nipkow@26271: by (unfold single_valued_def) blast huffman@19228: nipkow@30198: lemma single_valued_Id_on [simp]: "single_valued (Id_on A)" nipkow@26271: by (unfold single_valued_def) blast huffman@19228: berghofe@12905: haftmann@46664: subsubsection {* Graphs given by @{text Collect} *} berghofe@12905: berghofe@12905: lemma Domain_Collect_split [simp]: "Domain{(x,y). P x y} = {x. EX y. P x y}" nipkow@26271: by auto berghofe@12905: berghofe@12905: lemma Range_Collect_split [simp]: "Range{(x,y). P x y} = {y. EX x. P x y}" nipkow@26271: by auto berghofe@12905: berghofe@12905: lemma Image_Collect_split [simp]: "{(x,y). P x y} `` A = {y. EX x:A. P x y}" nipkow@26271: by auto berghofe@12905: berghofe@12905: haftmann@46664: subsubsection {* Inverse image *} berghofe@12905: huffman@19228: lemma sym_inv_image: "sym r ==> sym (inv_image r f)" nipkow@26271: by (unfold sym_def inv_image_def) blast huffman@19228: wenzelm@12913: lemma trans_inv_image: "trans r ==> trans (inv_image r f)" berghofe@12905: apply (unfold trans_def inv_image_def) berghofe@12905: apply (simp (no_asm)) berghofe@12905: apply blast berghofe@12905: done berghofe@12905: krauss@32463: lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)" krauss@32463: by (auto simp:inv_image_def) krauss@32463: krauss@33218: lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f" krauss@33218: unfolding inv_image_def converse_def by auto krauss@33218: haftmann@23709: haftmann@46664: subsubsection {* Finiteness *} haftmann@29609: haftmann@29609: lemma finite_converse [iff]: "finite (r^-1) = finite r" haftmann@29609: apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r") haftmann@29609: apply simp haftmann@29609: apply (rule iffI) haftmann@29609: apply (erule finite_imageD [unfolded inj_on_def]) haftmann@29609: apply (simp split add: split_split) haftmann@29609: apply (erule finite_imageI) haftmann@29609: apply (simp add: converse_def image_def, auto) haftmann@29609: apply (rule bexI) haftmann@29609: prefer 2 apply assumption haftmann@29609: apply simp haftmann@29609: done haftmann@29609: paulson@32876: lemma finite_Domain: "finite r ==> finite (Domain r)" paulson@32876: by (induct set: finite) (auto simp add: Domain_insert) paulson@32876: paulson@32876: lemma finite_Range: "finite r ==> finite (Range r)" paulson@32876: by (induct set: finite) (auto simp add: Range_insert) haftmann@29609: haftmann@29609: lemma finite_Field: "finite r ==> finite (Field r)" haftmann@29609: -- {* A finite relation has a finite field (@{text "= domain \ range"}. *} haftmann@29609: apply (induct set: finite) haftmann@29609: apply (auto simp add: Field_def Domain_insert Range_insert) haftmann@29609: done haftmann@29609: haftmann@29609: haftmann@46664: subsubsection {* Miscellaneous *} krauss@36728: krauss@36728: text {* Version of @{thm[source] lfp_induct} for binary relations *} haftmann@23709: haftmann@23709: lemmas lfp_induct2 = haftmann@23709: lfp_induct_set [of "(a, b)", split_format (complete)] haftmann@23709: krauss@36728: text {* Version of @{thm[source] subsetI} for binary relations *} krauss@36728: krauss@36728: lemma subrelI: "(\x y. (x, y) \ r \ (x, y) \ s) \ r \ s" krauss@36728: by auto krauss@36728: haftmann@46664: haftmann@46664: subsection {* Relations as binary predicates *} haftmann@46664: haftmann@46664: subsubsection {* Composition *} haftmann@46664: haftmann@46664: inductive pred_comp :: "['a \ 'b \ bool, 'b \ 'c \ bool] \ 'a \ 'c \ bool" (infixr "OO" 75) haftmann@46664: for r :: "'a \ 'b \ bool" and s :: "'b \ 'c \ bool" where haftmann@46664: pred_compI [intro]: "r a b \ s b c \ (r OO s) a c" haftmann@46664: haftmann@46664: inductive_cases pred_compE [elim!]: "(r OO s) a c" haftmann@46664: haftmann@46664: lemma pred_comp_rel_comp_eq [pred_set_conv]: haftmann@46664: "((\x y. (x, y) \ r) OO (\x y. (x, y) \ s)) = (\x y. (x, y) \ r O s)" haftmann@46664: by (auto simp add: fun_eq_iff) haftmann@46664: haftmann@46664: haftmann@46664: subsubsection {* Converse *} haftmann@46664: haftmann@46664: inductive conversep :: "('a \ 'b \ bool) \ 'b \ 'a \ bool" ("(_^--1)" [1000] 1000) haftmann@46664: for r :: "'a \ 'b \ bool" where haftmann@46664: conversepI: "r a b \ r^--1 b a" haftmann@46664: haftmann@46664: notation (xsymbols) haftmann@46664: conversep ("(_\\)" [1000] 1000) haftmann@46664: haftmann@46664: lemma conversepD: haftmann@46664: assumes ab: "r^--1 a b" haftmann@46664: shows "r b a" using ab haftmann@46664: by cases simp haftmann@46664: haftmann@46664: lemma conversep_iff [iff]: "r^--1 a b = r b a" haftmann@46664: by (iprover intro: conversepI dest: conversepD) haftmann@46664: haftmann@46664: lemma conversep_converse_eq [pred_set_conv]: haftmann@46664: "(\x y. (x, y) \ r)^--1 = (\x y. (x, y) \ r^-1)" haftmann@46664: by (auto simp add: fun_eq_iff) haftmann@46664: haftmann@46664: lemma conversep_conversep [simp]: "(r^--1)^--1 = r" haftmann@46664: by (iprover intro: order_antisym conversepI dest: conversepD) haftmann@46664: haftmann@46664: lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1" haftmann@46664: by (iprover intro: order_antisym conversepI pred_compI haftmann@46664: elim: pred_compE dest: conversepD) haftmann@46664: haftmann@46664: lemma converse_meet: "(r \ s)^--1 = r^--1 \ s^--1" haftmann@46664: by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD) haftmann@46664: haftmann@46664: lemma converse_join: "(r \ s)^--1 = r^--1 \ s^--1" haftmann@46664: by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD) haftmann@46664: haftmann@46664: lemma conversep_noteq [simp]: "(op \)^--1 = op \" haftmann@46664: by (auto simp add: fun_eq_iff) haftmann@46664: haftmann@46664: lemma conversep_eq [simp]: "(op =)^--1 = op =" haftmann@46664: by (auto simp add: fun_eq_iff) haftmann@46664: haftmann@46664: haftmann@46664: subsubsection {* Domain *} haftmann@46664: haftmann@46664: inductive DomainP :: "('a \ 'b \ bool) \ 'a \ bool" haftmann@46664: for r :: "'a \ 'b \ bool" where haftmann@46664: DomainPI [intro]: "r a b \ DomainP r a" haftmann@46664: haftmann@46664: inductive_cases DomainPE [elim!]: "DomainP r a" haftmann@46664: haftmann@46664: lemma DomainP_Domain_eq [pred_set_conv]: "DomainP (\x y. (x, y) \ r) = (\x. x \ Domain r)" haftmann@46664: by (blast intro!: Orderings.order_antisym predicate1I) haftmann@46664: haftmann@46664: haftmann@46664: subsubsection {* Range *} haftmann@46664: haftmann@46664: inductive RangeP :: "('a \ 'b \ bool) \ 'b \ bool" haftmann@46664: for r :: "'a \ 'b \ bool" where haftmann@46664: RangePI [intro]: "r a b \ RangeP r b" haftmann@46664: haftmann@46664: inductive_cases RangePE [elim!]: "RangeP r b" haftmann@46664: haftmann@46664: lemma RangeP_Range_eq [pred_set_conv]: "RangeP (\x y. (x, y) \ r) = (\x. x \ Range r)" haftmann@46664: by (blast intro!: Orderings.order_antisym predicate1I) haftmann@46664: haftmann@46664: haftmann@46664: subsubsection {* Inverse image *} haftmann@46664: haftmann@46664: definition inv_imagep :: "('b \ 'b \ bool) \ ('a \ 'b) \ 'a \ 'a \ bool" where haftmann@46664: "inv_imagep r f = (\x y. r (f x) (f y))" haftmann@46664: haftmann@46664: lemma [pred_set_conv]: "inv_imagep (\x y. (x, y) \ r) f = (\x y. (x, y) \ inv_image r f)" haftmann@46664: by (simp add: inv_image_def inv_imagep_def) haftmann@46664: haftmann@46664: lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)" haftmann@46664: by (simp add: inv_imagep_def) haftmann@46664: haftmann@46664: haftmann@46664: subsubsection {* Powerset *} haftmann@46664: haftmann@46664: definition Powp :: "('a \ bool) \ 'a set \ bool" where haftmann@46664: "Powp A = (\B. \x \ B. A x)" haftmann@46664: haftmann@46664: lemma Powp_Pow_eq [pred_set_conv]: "Powp (\x. x \ A) = (\x. x \ Pow A)" haftmann@46664: by (auto simp add: Powp_def fun_eq_iff) haftmann@46664: haftmann@46664: lemmas Powp_mono [mono] = Pow_mono [to_pred] haftmann@46664: haftmann@46664: haftmann@46664: subsubsection {* Properties of predicate relations *} haftmann@46664: haftmann@46664: abbreviation antisymP :: "('a \ 'a \ bool) \ bool" where haftmann@46664: "antisymP r \ antisym {(x, y). r x y}" haftmann@46664: haftmann@46664: abbreviation transP :: "('a \ 'a \ bool) \ bool" where haftmann@46664: "transP r \ trans {(x, y). r x y}" haftmann@46664: haftmann@46664: abbreviation single_valuedP :: "('a \ 'b \ bool) \ bool" where haftmann@46664: "single_valuedP r \ single_valued {(x, y). r x y}" haftmann@46664: haftmann@46664: (*FIXME inconsistencies: abbreviations vs. definitions, suffix `P` vs. suffix `p`*) haftmann@46664: haftmann@46664: definition reflp :: "('a \ 'a \ bool) \ bool" where haftmann@46664: "reflp r \ refl {(x, y). r x y}" haftmann@46664: haftmann@46664: definition symp :: "('a \ 'a \ bool) \ bool" where haftmann@46664: "symp r \ sym {(x, y). r x y}" haftmann@46664: haftmann@46664: definition transp :: "('a \ 'a \ bool) \ bool" where haftmann@46664: "transp r \ trans {(x, y). r x y}" haftmann@46664: haftmann@46664: lemma reflpI: haftmann@46664: "(\x. r x x) \ reflp r" haftmann@46664: by (auto intro: refl_onI simp add: reflp_def) haftmann@46664: haftmann@46664: lemma reflpE: haftmann@46664: assumes "reflp r" haftmann@46664: obtains "r x x" haftmann@46664: using assms by (auto dest: refl_onD simp add: reflp_def) haftmann@46664: haftmann@46664: lemma sympI: haftmann@46664: "(\x y. r x y \ r y x) \ symp r" haftmann@46664: by (auto intro: symI simp add: symp_def) haftmann@46664: haftmann@46664: lemma sympE: haftmann@46664: assumes "symp r" and "r x y" haftmann@46664: obtains "r y x" haftmann@46664: using assms by (auto dest: symD simp add: symp_def) haftmann@46664: haftmann@46664: lemma transpI: haftmann@46664: "(\x y z. r x y \ r y z \ r x z) \ transp r" haftmann@46664: by (auto intro: transI simp add: transp_def) haftmann@46664: haftmann@46664: lemma transpE: haftmann@46664: assumes "transp r" and "r x y" and "r y z" haftmann@46664: obtains "r x z" haftmann@46664: using assms by (auto dest: transD simp add: transp_def) haftmann@46664: haftmann@46664: no_notation haftmann@46664: bot ("\") and haftmann@46664: top ("\") and haftmann@46664: inf (infixl "\" 70) and haftmann@46664: sup (infixl "\" 65) and haftmann@46664: Inf ("\_" [900] 900) and haftmann@46664: Sup ("\_" [900] 900) haftmann@46664: haftmann@46664: no_syntax (xsymbols) haftmann@46664: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@46664: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@46664: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@46664: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@46664: nipkow@1128: end haftmann@46689: