src/HOLCF/IOA/meta_theory/TL.thy
author oheimb
Thu, 24 Sep 1998 17:17:56 +0200
changeset 5554 3cae5d6510c2
parent 4577 674b0b354feb
child 5976 44290b71a85f
permissions -rw-r--r--
removed addcongs2 and delcongs2 simplified CLASIMP_DATA
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     1
(*  Title:      HOLCF/IOA/meta_theory/TLS.thy
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     2
    ID:         $Id$
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     3
    Author:     Olaf M"uller
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     4
    Copyright   1997  TU Muenchen
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     5
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     6
A General Temporal Logic
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     7
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     8
Version 2: Interface directly after Sequeces, i.e. predicates and predicate transformers are in HOL
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     9
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    10
*)   
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    11
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    12
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    13
		       
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
TL = Pred + Sequence +  
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    15
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
default term
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
types
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    19
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
'a temporal      = 'a Seq predicate
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
consts
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
suffix     :: "'a Seq => 'a Seq => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
tsuffix    :: "'a Seq => 'a Seq => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
validT     :: "'a Seq predicate => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
unlift     ::  "'a lift => 'a"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
Init         ::"'a predicate => 'a temporal"          ("<_>" [0] 1000)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
Box          ::"'a temporal => 'a temporal"   ("[] (_)" [80] 80)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
Diamond      ::"'a temporal => 'a temporal"   ("<> (_)" [80] 80)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
Next         ::"'a temporal => 'a temporal"   
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
Leadsto      ::"'a temporal => 'a temporal => 'a temporal"  (infixr "~>" 22)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    39
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    40
syntax (symbols)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    41
   "Box"        ::"'a temporal => 'a temporal"   ("\\<box> (_)" [80] 80)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
   "Diamond"    ::"'a temporal => 'a temporal"   ("\\<diamond> (_)" [80] 80)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    43
   "Leadsto"    ::"'a temporal => 'a temporal => 'a temporal"  (infixr "\\<leadsto>" 22)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    45
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    46
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
unlift_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    49
  "unlift x == (case x of 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    50
                 Undef   => arbitrary
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
               | Def y   => y)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    53
(* this means that for nil and UU the effect is unpredictable *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    54
Init_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    55
  "Init P s ==  (P (unlift (HD`s)))" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    56
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    57
suffix_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
  "suffix s2 s == ? s1. (Finite s1  & s = s1 @@ s2)" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    59
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    60
tsuffix_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    61
  "tsuffix s2 s == s2 ~= nil & s2 ~= UU & suffix s2 s"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    62
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    63
Box_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    64
  "([] P) s == ! s2. tsuffix s2 s --> P s2"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    65
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    66
Next_def
4577
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    67
  "(Next P) s == if (TL`s=UU | TL`s=nil) then (P s) else P (TL`s)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    68
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    69
Diamond_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    70
  "<> P == .~ ([] (.~ P))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    71
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    72
Leadsto_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    73
   "P ~> Q == ([] (P .--> (<> Q)))"  
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    74
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    75
validT_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    76
  "validT P == ! s. s~=UU & s~=nil --> (s |= P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    77
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    78
end