rev_bexI [intro?];
authorwenzelm
Tue, 07 May 2002 19:54:04 +0200
changeset 13113 5eb9be7b72a5
parent 13112 7275750553b7
child 13114 f2b00262bdfc
rev_bexI [intro?];
src/HOL/Set.thy
--- a/src/HOL/Set.thy	Tue May 07 19:15:11 2002 +0200
+++ b/src/HOL/Set.thy	Tue May 07 19:54:04 2002 +0200
@@ -260,7 +260,7 @@
     choice of @{prop "x:A"}. *}
   by (unfold Bex_def) blast
 
-lemma rev_bexI: "x:A ==> P x ==> EX x:A. P x"
+lemma rev_bexI [intro?]: "x:A ==> P x ==> EX x:A. P x"
   -- {* The best argument order when there is only one @{prop "x:A"}. *}
   by (unfold Bex_def) blast