add quantifier syntax for eventually
authorhoelzl
Sun, 12 Apr 2015 11:33:30 +0200
changeset 60037 071a99649dde
parent 60036 218fcc645d22
child 60038 ca431cbce2a3
add quantifier syntax for eventually
src/HOL/Filter.thy
--- a/src/HOL/Filter.thy	Sun Apr 12 11:33:19 2015 +0200
+++ b/src/HOL/Filter.thy	Sun Apr 12 11:33:30 2015 +0200
@@ -39,6 +39,15 @@
 definition eventually :: "('a \<Rightarrow> bool) \<Rightarrow> 'a filter \<Rightarrow> bool"
   where "eventually P F \<longleftrightarrow> Rep_filter F P"
 
+syntax
+  "_eventually"  :: "pttrn => 'a filter => bool => bool"      ("(3ALL _ in _./ _)" [0, 0, 10] 10)
+
+syntax (xsymbols)
+  "_eventually"  :: "pttrn => 'a filter => bool => bool"      ("(3\<forall>_ in _./ _)" [0, 0, 10] 10)
+
+translations
+  "\<forall>x in F. P" == "CONST eventually (%x. P) F"
+
 lemma eventually_Abs_filter:
   assumes "is_filter F" shows "eventually P (Abs_filter F) = F P"
   unfolding eventually_def using assms by (simp add: Abs_filter_inverse)