src/HOLCF/IOA/meta_theory/Pred.thy
changeset 12338 de0f4a63baa5
parent 12218 6597093b77e7
child 14565 c6dc17aab88a
equal deleted inserted replaced
12337:7c6a970f0808 12338:de0f4a63baa5
     6 Logical Connectives lifted to predicates.
     6 Logical Connectives lifted to predicates.
     7 *)   
     7 *)   
     8 	       
     8 	       
     9 Pred = Main +
     9 Pred = Main +
    10 
    10 
    11 default term
    11 default type
    12 
    12 
    13 types
    13 types
    14 
    14 
    15 'a predicate      = "'a => bool"
    15 'a predicate      = "'a => bool"
    16 
    16