src/HOLCF/IOA/meta_theory/TLS.thy
author wenzelm
Fri, 02 Sep 2005 17:23:59 +0200
changeset 17233 41eee2e7b465
parent 14981 e73f8140af78
child 19741 f65265d71426
permissions -rw-r--r--
converted specifications to Isar theories;
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$
12218
wenzelm
parents: 10835
diff changeset
     3
    Author:     Olaf Müller
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     4
*)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     5
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     6
header {* Temporal Logic of Steps -- tailored for I/O automata *}
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     7
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
theory TLS
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
imports IOA TL
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
begin
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    11
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    12
defaultsort type
4559
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
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    15
  ('a, 's) ioa_temp  = "('a option,'s)transition temporal"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    16
  ('a, 's) step_pred = "('a option,'s)transition predicate"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    17
  's state_pred      = "'s predicate"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    19
consts
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
option_lift :: "('a => 'b) => 'b => ('a option => 'b)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
plift       :: "('a => bool) => ('a option => bool)"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    23
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
temp_sat   :: "('a,'s)execution => ('a,'s)ioa_temp => bool"    (infixr "|==" 22)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
xt1        :: "'s predicate => ('a,'s)step_pred"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
xt2        :: "'a option predicate => ('a,'s)step_pred"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
validTE    :: "('a,'s)ioa_temp => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
validIOA   :: "('a,'s)ioa => ('a,'s)ioa_temp => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    31
mkfin      :: "'a Seq => 'a Seq"
4559
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
ex2seq     :: "('a,'s)execution => ('a option,'s)transition Seq"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    34
ex2seqC    :: "('a,'s)pairs -> ('s => ('a option,'s)transition Seq)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
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
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    39
mkfin_def:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    40
  "mkfin s == if Partial s then @t. Finite t & s = t @@ UU
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    41
                           else s"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    42
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    43
option_lift_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
  "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
    45
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    46
(* plift is used to determine that None action is always false in
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
   transition predicates *)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    48
plift_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    49
  "plift P == option_lift P False"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    50
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    51
temp_sat_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
  "ex |== P == ((ex2seq ex) |= P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    53
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    54
xt1_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    55
  "xt1 P tr == P (fst tr)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    56
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    57
xt2_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
  "xt2 P tr == P (fst (snd tr))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    59
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    60
ex2seq_def:
10835
nipkow
parents: 5976
diff changeset
    61
  "ex2seq ex == ((ex2seqC $(mkfin (snd ex))) (fst ex))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    62
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    63
ex2seqC_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    64
  "ex2seqC == (fix$(LAM h ex. (%s. case ex of
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    65
      nil =>  (s,None,s)>>nil
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    66
    | x##xs => (flift1 (%pr.
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    67
                (s,Some (fst pr), snd pr)>> (h$xs) (snd pr))
10835
nipkow
parents: 5976
diff changeset
    68
                $x)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    69
      )))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    70
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    71
validTE_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    72
  "validTE P == ! ex. (ex |== P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    73
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    74
validIOA_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    75
  "validIOA A P == ! ex : executions A . (ex |== P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    76
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 5677
diff changeset
    77
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    78
axioms
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 5677
diff changeset
    79
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    80
mkfin_UU:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    81
  "mkfin UU = nil"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    82
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    83
mkfin_nil:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    84
  "mkfin nil =nil"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    85
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    86
mkfin_cons:
5677
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    87
  "(mkfin (a>>s)) = (a>>(mkfin s))"
4feffde494cf another little bug ;-) and minor changes in TLS.*;
mueller
parents: 4577
diff changeset
    88
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    89
ML {* use_legacy_bindings (the_context ()) *}
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    90
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    91
end