# HG changeset patch # User wenzelm # Date 1331981823 -3600 # Node ID 216a839841bc007269388542cec4ba2861892e07 # Parent 144d94446378df605226d7291e9bd2e36ee24732# Parent ef4b0d6b2fb61ec7900bb5d227e86f0576d161f8 merged; diff -r ef4b0d6b2fb6 -r 216a839841bc NEWS --- a/NEWS Sat Mar 17 11:23:14 2012 +0100 +++ b/NEWS Sat Mar 17 11:57:03 2012 +0100 @@ -118,6 +118,8 @@ Domain_def ~> Domain_unfold Range_def ~> Domain_converse [symmetric] +Generalized theorems INF_INT_eq, INF_INT_eq2, SUP_UN_eq, SUP_UN_eq2. + INCOMPATIBILITY. * Consolidated various theorem names relating to Finite_Set.fold diff -r ef4b0d6b2fb6 -r 216a839841bc src/HOL/Library/Zorn.thy --- a/src/HOL/Library/Zorn.thy Sat Mar 17 11:23:14 2012 +0100 +++ b/src/HOL/Library/Zorn.thy Sat Mar 17 11:57:03 2012 +0100 @@ -12,38 +12,36 @@ begin (* Define globally? In Set.thy? *) -definition chain_subset :: "'a set set \ bool" ("chain\<^bsub>\\<^esub>") where -"chain\<^bsub>\\<^esub> C \ \A\C.\B\C. A \ B \ B \ A" +definition chain_subset :: "'a set set \ bool" ("chain\<^bsub>\\<^esub>") +where + "chain\<^bsub>\\<^esub> C \ \A\C.\B\C. A \ B \ B \ A" text{* The lemma and section numbers refer to an unpublished article \cite{Abrial-Laffitte}. *} -definition - chain :: "'a set set => 'a set set set" where - "chain S = {F. F \ S & chain\<^bsub>\\<^esub> F}" +definition chain :: "'a set set \ 'a set set set" +where + "chain S = {F. F \ S \ chain\<^bsub>\\<^esub> F}" -definition - super :: "['a set set,'a set set] => 'a set set set" where - "super S c = {d. d \ chain S & c \ d}" - -definition - maxchain :: "'a set set => 'a set set set" where - "maxchain S = {c. c \ chain S & super S c = {}}" +definition super :: "'a set set \ 'a set set \ 'a set set set" +where + "super S c = {d. d \ chain S \ c \ d}" -definition - succ :: "['a set set,'a set set] => 'a set set" where - "succ S c = - (if c \ chain S | c \ maxchain S - then c else SOME c'. c' \ super S c)" +definition maxchain :: "'a set set \ 'a set set set" +where + "maxchain S = {c. c \ chain S \ super S c = {}}" -inductive_set - TFin :: "'a set set => 'a set set set" - for S :: "'a set set" - where - succI: "x \ TFin S ==> succ S x \ TFin S" - | Pow_UnionI: "Y \ Pow(TFin S) ==> Union(Y) \ TFin S" +definition succ :: "'a set set \ 'a set set \ 'a set set" +where + "succ S c = (if c \ chain S \ c \ maxchain S then c else SOME c'. c' \ super S c)" + +inductive_set TFin :: "'a set set \ 'a set set set" +for S :: "'a set set" +where + succI: "x \ TFin S \ succ S x \ TFin S" +| Pow_UnionI: "Y \ Pow (TFin S) \ \Y \ TFin S" subsection{*Mathematical Preamble*} diff -r ef4b0d6b2fb6 -r 216a839841bc src/HOL/Relation.thy --- a/src/HOL/Relation.thy Sat Mar 17 11:23:14 2012 +0100 +++ b/src/HOL/Relation.thy Sat Mar 17 11:57:03 2012 +0100 @@ -95,6 +95,18 @@ 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]: "(\i\S. (\x. x \ r i)) = (\x. x \ (\i\S. r i))" + by (simp add: fun_eq_iff) + +lemma INF_INT_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" + by (simp add: fun_eq_iff) + +lemma SUP_UN_eq [pred_set_conv]: "(\i\S. (\x. x \ r i)) = (\x. x \ (\i\S. r i))" + by (simp add: fun_eq_iff) + +lemma SUP_UN_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" + by (simp add: fun_eq_iff) + lemma Inf_INT_eq [pred_set_conv]: "\S = (\x. x \ INTER S Collect)" by (simp add: fun_eq_iff) @@ -119,19 +131,6 @@ lemma SUP_Sup_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ i)) = (\x y. (x, y) \ \S)" by (simp add: fun_eq_iff) -lemma INF_INT_eq [pred_set_conv]: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" - by (simp add: fun_eq_iff) - -lemma INF_INT_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" - by (simp add: fun_eq_iff) - -lemma SUP_UN_eq [pred_set_conv]: "(\i. (\x. x \ r i)) = (\x. x \ (\i. r i))" - by (simp add: fun_eq_iff) - -lemma SUP_UN_eq2 [pred_set_conv]: "(\i\S. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (\i\S. r i))" - by (simp add: fun_eq_iff) - - subsection {* Properties of relations *} @@ -277,13 +276,17 @@ "\x\S. sym (r x) \ sym (INTER S r)" by (fast intro: symI elim: symE) -(* FIXME thm sym_INTER [to_pred] *) +lemma symp_INF: + "\x\S. symp (r x) \ symp (INFI S r)" + by (fact sym_INTER [to_pred]) lemma sym_UNION: "\x\S. sym (r x) \ sym (UNION S r)" by (fast intro: symI elim: symE) -(* FIXME thm sym_UNION [to_pred] *) +lemma symp_SUP: + "\x\S. symp (r x) \ symp (SUPR S r)" + by (fact sym_UNION [to_pred]) subsubsection {* Antisymmetry *}