src/HOL/BNF/BNF_Def.thy
changeset 54961 e60428f432bc
parent 54841 af71b753c459
--- a/src/HOL/BNF/BNF_Def.thy	Fri Jan 10 12:09:11 2014 +0100
+++ b/src/HOL/BNF/BNF_Def.thy	Fri Jan 10 11:47:10 2014 +0100
@@ -150,6 +150,9 @@
 lemma convol_image_vimage2p: "<f o fst, g o snd> ` Collect (split (vimage2p f g R)) \<subseteq> Collect (split R)"
   unfolding vimage2p_def convol_def by auto
 
+lemma vimage2p_Grp: "vimage2p f g P = Grp UNIV f OO P OO (Grp UNIV g)\<inverse>\<inverse>"
+  unfolding vimage2p_def Grp_def by auto
+
 (*FIXME: duplicates lemma from Record.thy*)
 lemma o_eq_dest_lhs: "a o b = c \<Longrightarrow> a (b v) = c v"
   by clarsimp