# HG changeset patch # User haftmann # Date 1325165020 -3600 # Node ID 6a86cc88b02f0add6ea9b612ea327b079de5b9b2 # Parent 313be214e40a1e4d7866e5b6c59b6745e63e2bd2 fundamental theorems on Set.bind diff -r 313be214e40a -r 6a86cc88b02f src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Thu Dec 29 14:44:44 2011 +0100 +++ b/src/HOL/Complete_Lattices.thy Thu Dec 29 14:23:40 2011 +0100 @@ -928,11 +928,15 @@ "Set.bind A f = UNION A f" by (simp add: bind_def UNION_eq) +lemma member_bind [simp]: + "x \ Set.bind P f \ x \ UNION P f " + by (simp add: bind_UNION) + lemma Union_image_eq [simp]: "\(B ` A) = (\x\A. B x)" by (rule sym) (fact SUP_def) -lemma UN_iff [simp]: "(b \ (\x\A. B x)) = (\x\A. b \ B x)" +lemma UN_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" by (auto simp add: SUP_def image_def) lemma UN_I [intro]: "a \ A \ b \ B a \ b \ (\x\A. B x)" diff -r 313be214e40a -r 6a86cc88b02f src/HOL/Set.thy --- a/src/HOL/Set.thy Thu Dec 29 14:44:44 2011 +0100 +++ b/src/HOL/Set.thy Thu Dec 29 14:23:40 2011 +0100 @@ -1790,6 +1790,22 @@ hide_const (open) bind +lemma bind_bind: + fixes A :: "'a set" + shows "Set.bind (Set.bind A B) C = Set.bind A (\x. Set.bind (B x) C)" + by (auto simp add: bind_def) + +lemma empty_bind [simp]: + "Set.bind {} B = {}" + by (simp add: bind_def) + +lemma nonempty_bind_const: + "A \ {} \ Set.bind A (\_. B) = B" + by (auto simp add: bind_def) + +lemma bind_const: "Set.bind A (\_. B) = (if A = {} then {} else B)" + by (auto simp add: bind_def) + subsubsection {* Operations for execution *}