src/HOL/HOLCF/IOA/Pred.thy
changeset 62192 bdaa0eb0fc74
parent 62116 bc178c0fe1a1
child 62194 0aabc5931361
equal deleted inserted replaced
62191:eb9f5ee249f9 62192:bdaa0eb0fc74
    13 type_synonym 'a predicate = "'a \<Rightarrow> bool"
    13 type_synonym 'a predicate = "'a \<Rightarrow> bool"
    14 
    14 
    15 definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool"  ("_ \<Turnstile> _" [100, 9] 8)
    15 definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool"  ("_ \<Turnstile> _" [100, 9] 8)
    16   where "(s \<Turnstile> P) \<longleftrightarrow> P s"
    16   where "(s \<Turnstile> P) \<longleftrightarrow> P s"
    17 
    17 
    18 definition valid :: "'a predicate \<Rightarrow> bool"  (*  ("|-") *)
    18 definition valid :: "'a predicate \<Rightarrow> bool"  (*  FIXME ("|-") *)
    19   where "valid P \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
    19   where "valid P \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
    20 
    20 
    21 definition NOT :: "'a predicate \<Rightarrow> 'a predicate"  ("\<^bold>\<not> _" [40] 40)
    21 definition NOT :: "'a predicate \<Rightarrow> 'a predicate"  ("\<^bold>\<not> _" [40] 40)
    22   where "NOT P s \<longleftrightarrow> \<not> P s"
    22   where "NOT P s \<longleftrightarrow> \<not> P s"
    23 
    23