diff -r c23bdb4ed2f6 -r c5316f843f72 src/HOL/BNF_Def.thy --- a/src/HOL/BNF_Def.thy Mon Sep 01 13:23:05 2014 +0200 +++ b/src/HOL/BNF_Def.thy Mon Sep 01 13:23:39 2014 +0200 @@ -15,6 +15,9 @@ "bnf" :: thy_goal begin +lemma Collect_splitD: "x \ Collect (split A) \ A (fst x) (snd x)" + by auto + definition rel_fun :: "('a \ 'c \ bool) \ ('b \ 'd \ bool) \ ('a \ 'b) \ ('c \ 'd) \ bool" where @@ -30,6 +33,20 @@ shows "B (f x) (g y)" using assms by (simp add: rel_fun_def) +definition rel_set :: "('a \ 'b \ bool) \ 'a set \ 'b set \ bool" + where "rel_set R = (\A B. (\x\A. \y\B. R x y) \ (\y\B. \x\A. R x y))" + +lemma rel_setI: + assumes "\x. x \ A \ \y\B. R x y" + assumes "\y. y \ B \ \x\A. R x y" + shows "rel_set R A B" + using assms unfolding rel_set_def by simp + +lemma predicate2_transferD: + "\rel_fun R1 (rel_fun R2 (op =)) P Q; a \ A; b \ B; A \ {(x, y). R1 x y}; B \ {(x, y). R2 x y}\ \ + P (fst a) (fst b) \ Q (snd a) (snd b)" + unfolding rel_fun_def by (blast dest!: Collect_splitD) + definition collect where "collect F x = (\f \ F. f x)"