src/HOL/HOLCF/IOA/Pred.thy
author wenzelm
Sat, 16 Jan 2016 23:35:55 +0100
changeset 62194 0aabc5931361
parent 62192 bdaa0eb0fc74
child 74101 d804e93ae9ff
permissions -rw-r--r--
tuned syntax;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62005
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/Pred.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     3
*)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62000
diff changeset
     5
section \<open>Logical Connectives lifted to predicates\<close>
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory Pred
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports Main
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
begin
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35355
diff changeset
    11
default_sort type
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    12
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    13
type_synonym 'a predicate = "'a \<Rightarrow> bool"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    15
definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool"  ("_ \<Turnstile> _" [100, 9] 8)
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    16
  where "(s \<Turnstile> P) \<longleftrightarrow> P s"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
62194
0aabc5931361 tuned syntax;
wenzelm
parents: 62192
diff changeset
    18
definition valid :: "'a predicate \<Rightarrow> bool"  ("\<TTurnstile> _" [9] 8)
0aabc5931361 tuned syntax;
wenzelm
parents: 62192
diff changeset
    19
  where "(\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    21
definition NOT :: "'a predicate \<Rightarrow> 'a predicate"  ("\<^bold>\<not> _" [40] 40)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    22
  where "NOT P s \<longleftrightarrow> \<not> P s"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    24
definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<and>" 35)
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    25
  where "(P \<^bold>\<and> Q) s \<longleftrightarrow> P s \<and> Q s"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    27
definition OR :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<or>" 30)
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    28
  where "(P \<^bold>\<or> Q) s \<longleftrightarrow> P s \<or> Q s"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
62005
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    30
definition IMPLIES :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate"  (infixr "\<^bold>\<longrightarrow>" 25)
68db98c2cd97 modernized defs;
wenzelm
parents: 62002
diff changeset
    31
  where "(P \<^bold>\<longrightarrow> Q) s \<longleftrightarrow> P s \<longrightarrow> Q s"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
7661
8c3190b173aa depend on Main;
wenzelm
parents: 6340
diff changeset
    33
end