src/HOLCF/IOA/meta_theory/TLS.thy
author mueller
Wed, 14 Jan 1998 16:38:04 +0100
changeset 4577 674b0b354feb
parent 4559 8e604d885b54
child 5677 4feffde494cf
permissions -rw-r--r--
added thms wrt weakening and strengthening in Abstraction;
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
Temporal Logic of Steps -- tailored for I/O automata
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
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
TLS = IOA + TL + 
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
default term
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
types
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
 ('a,'s)ioa_temp       = "('a option,'s)transition temporal" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
 ('a,'s)step_pred      = "('a option,'s)transition predicate"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
  's state_pred        = "'s predicate"
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
consts
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
option_lift :: "('a => 'b) => 'b => ('a option => 'b)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
plift       :: "('a => bool) => ('a option => bool)"
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
temp_sat   :: "('a,'s)execution => ('a,'s)ioa_temp => bool"    (infixr "|==" 22)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
xt1        :: "'s predicate => ('a,'s)step_pred"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
xt2        :: "'a option predicate => ('a,'s)step_pred"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
validTE    :: "('a,'s)ioa_temp => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
validIOA   :: "('a,'s)ioa => ('a,'s)ioa_temp => bool"
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
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
ex2seq     :: "('a,'s)execution => ('a option,'s)transition Seq"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
ex2seqC    :: "('a,'s)pairs -> ('s => ('a option,'s)transition Seq)" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
defs
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
(* should be in Option as option_lift1, and option_map should be renamed to 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    41
   option_lift2 *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
option_lift_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    43
  "option_lift f s y == case y of None => s | Some x => (f x)"
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
(* plift is used to determine that None action is always false in 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    46
   transition predicates *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
plift_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
  "plift P == option_lift P False"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    49
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    50
temp_sat_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
  "ex |== P == ((ex2seq ex) |= P)"
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
xt1_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    54
  "xt1 P tr == P (fst tr)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    55
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    56
xt2_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    57
  "xt2 P tr == P (fst (snd tr))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
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
(* FIX: Clarify type and Herkunft of a !! *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    61
ex2seq_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    62
  "ex2seq ex == ((ex2seqC `(snd ex)) (fst ex))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    63
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    64
ex2seqC_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    65
  "ex2seqC == (fix`(LAM h ex. (%s. case ex of 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    66
      nil =>  (s,None,s)>>nil
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    67
    | x##xs => (flift1 (%pr.
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    68
                (s,Some (fst pr), snd pr)>> (h`xs) (snd pr)) 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    69
                `x)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    70
      )))"
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
validTE_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    73
  "validTE P == ! ex. (ex |== P)"
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
validIOA_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    76
  "validIOA A P == ! ex : executions A . (ex |== P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    77
4577
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    78
rules
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    79
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    80
ex2seq_UUAlt
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    81
  "ex2seq (s,UU) = (s,None,s)>>UU"
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    82
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    83
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    84
end
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    85