| 62008 |      1 | (*  Title:      HOL/HOLCF/IOA/Pred.thy
 | 
| 40945 |      2 |     Author:     Olaf Müller
 | 
| 17233 |      3 | *)
 | 
| 4559 |      4 | 
 | 
| 62002 |      5 | section \<open>Logical Connectives lifted to predicates\<close>
 | 
| 4559 |      6 | 
 | 
| 17233 |      7 | theory Pred
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
| 36452 |     11 | default_sort type
 | 
| 4559 |     12 | 
 | 
| 62005 |     13 | type_synonym 'a predicate = "'a \<Rightarrow> bool"
 | 
| 4559 |     14 | 
 | 
| 62116 |     15 | definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool"  ("_ \<Turnstile> _" [100, 9] 8)
 | 
| 62005 |     16 |   where "(s \<Turnstile> P) \<longleftrightarrow> P s"
 | 
| 4559 |     17 | 
 | 
| 62194 |     18 | definition valid :: "'a predicate \<Rightarrow> bool"  ("\<TTurnstile> _" [9] 8)
 | 
|  |     19 |   where "(\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
 | 
| 4559 |     20 | 
 | 
| 62005 |     21 | definition NOT :: "'a predicate \<Rightarrow> 'a predicate"  ("\<^bold>\<not> _" [40] 40)
 | 
| 62116 |     22 |   where "NOT P s \<longleftrightarrow> \<not> P s"
 | 
| 4559 |     23 | 
 | 
| 62005 |     24 | definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<and>" 35)
 | 
|  |     25 |   where "(P \<^bold>\<and> Q) s \<longleftrightarrow> P s \<and> Q s"
 | 
| 4559 |     26 | 
 | 
| 62005 |     27 | definition OR :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<or>" 30)
 | 
|  |     28 |   where "(P \<^bold>\<or> Q) s \<longleftrightarrow> P s \<or> Q s"
 | 
| 4559 |     29 | 
 | 
| 62005 |     30 | definition IMPLIES :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<longrightarrow>" 25)
 | 
|  |     31 |   where "(P \<^bold>\<longrightarrow> Q) s \<longleftrightarrow> P s \<longrightarrow> Q s"
 | 
| 4559 |     32 | 
 | 
| 7661 |     33 | end
 |