diff -r c059d550afec -r 23904fa13e03 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Tue Nov 23 23:10:13 2010 +0100 +++ b/src/HOL/Predicate.thy Tue Nov 23 23:11:06 2010 +0100 @@ -476,15 +476,15 @@ by (simp add: single_def) definition bind :: "'a pred \ ('a \ 'b pred) \ 'b pred" (infixl "\=" 70) where - "P \= f = (SUPR (Predicate.eval P) f)" + "P \= f = (SUPR {x. Predicate.eval P x} f)" lemma eval_bind [simp]: - "eval (P \= f) = Predicate.eval (SUPR (Predicate.eval P) f)" + "eval (P \= f) = Predicate.eval (SUPR {x. Predicate.eval P x} f)" by (simp add: bind_def) lemma bind_bind: "(P \= Q) \= R = P \= (\x. Q x \= R)" - by (rule pred_eqI) simp + by (rule pred_eqI) auto lemma bind_single: "P \= single = P" @@ -496,15 +496,15 @@ lemma bottom_bind: "\ \= P = \" - by (rule pred_eqI) simp + by (rule pred_eqI) auto lemma sup_bind: "(P \ Q) \= R = P \= R \ Q \= R" - by (rule pred_eqI) simp + by (rule pred_eqI) auto lemma Sup_bind: "(\A \= f) = \((\x. x \= f) ` A)" - by (rule pred_eqI) simp + by (rule pred_eqI) auto lemma pred_iffI: assumes "\x. eval A x \ eval B x" @@ -933,8 +933,12 @@ where "the A = (THE x. eval A x)" -lemma the_eq[code]: "the A = singleton (\x. not_unique A) A" -by (auto simp add: the_def singleton_def not_unique_def) +lemma the_eqI: + "(THE x. Predicate.eval P x) = x \ Predicate.the P = x" + by (simp add: the_def) + +lemma the_eq [code]: "the A = singleton (\x. not_unique A) A" + by (rule the_eqI) (simp add: singleton_def not_unique_def) code_abort not_unique