src/HOLCF/IOA/meta_theory/LiveIOA.thy
author wenzelm
Fri, 17 Nov 2006 02:20:03 +0100
changeset 21404 eb85850d3eb7
parent 19741 f65265d71426
child 25135 4f8176c940cf
permissions -rw-r--r--
more robust syntax for definition/abbreviation/notation;
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/LiveIOA.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 {* Live I/O automata -- specified by temproal formulas *}
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 LiveIOA
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
imports TLS
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
begin
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    11
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) live_ioa = "('a,'s)ioa * ('a,'s)ioa_temp"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
consts
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
validLIOA   :: "('a,'s)live_ioa => ('a,'s)ioa_temp  => bool"
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
WF         :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
SF         :: "('a,'s)ioa => 'a set => ('a,'s)ioa_temp"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
liveexecutions    :: "('a,'s)live_ioa => ('a,'s)execution set"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
livetraces        :: "('a,'s)live_ioa => 'a trace set"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
live_implements   :: "('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
is_live_ref_map   :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    29
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    32
validLIOA_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
  "validLIOA AL P == validIOA (fst AL) ((snd AL) .--> P)"
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
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    36
WF_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
  "WF A acts ==  <> [] <%(s,a,t) . Enabled A acts s> .--> [] <> <xt2 (plift (%a. a : acts))>"
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
SF_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    40
  "SF A acts ==  [] <> <%(s,a,t) . Enabled A acts s> .--> [] <> <xt2 (plift (%a. a : acts))>"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    41
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    43
liveexecutions_def:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
   "liveexecutions AP == {exec. exec : executions (fst AP) & (exec |== (snd AP))}"
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
livetraces_def:
10835
nipkow
parents: 4816
diff changeset
    47
  "livetraces AP == {mk_trace (fst AP)$(snd ex) | ex. ex:liveexecutions AP}"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    49
live_implements_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    50
  "live_implements CL AM == (inp (fst CL) = inp (fst AM)) &
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
                            (out (fst CL) = out (fst AM)) &
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
                            livetraces CL <= livetraces AM"
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
is_live_ref_map_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    55
   "is_live_ref_map f CL AM ==
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    56
            is_ref_map f (fst CL ) (fst AM) &
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    57
            (! exec : executions (fst CL). (exec |== (snd CL)) -->
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
                                           ((corresp_ex (fst AM) f exec) |== (snd AM)))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    59
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    60
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    61
declare split_paired_Ex [simp del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    62
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    63
lemma live_implements_trans: 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    64
"!!LC. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |]  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    65
      ==> live_implements (A,LA) (C,LC)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    66
apply (unfold live_implements_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    67
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    68
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    69
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    70
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    71
subsection "Correctness of live refmap"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    72
	
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    73
lemma live_implements: "[| inp(C)=inp(A); out(C)=out(A);  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    74
                   is_live_ref_map f (C,M) (A,L) |]  
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    75
                ==> live_implements (C,M) (A,L)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    76
apply (simp add: is_live_ref_map_def live_implements_def livetraces_def liveexecutions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    77
apply (tactic "safe_tac set_cs")
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    78
apply (rule_tac x = "corresp_ex A f ex" in exI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    79
apply (tactic "safe_tac set_cs")
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    80
  (* Traces coincide, Lemma 1 *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    81
  apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    82
  apply (erule lemma_1 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    83
  apply (simp (no_asm) add: externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    84
  apply (auto)[1]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    85
  apply (simp add: executions_def reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    86
 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    87
  (* corresp_ex is execution, Lemma 2 *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    88
  apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    89
  apply (simp add: executions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    90
  (* start state *) 
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    91
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
  apply (simp add: is_ref_map_def corresp_ex_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    93
  (* is-execution-fragment *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    94
  apply (erule lemma_2 [THEN spec, THEN mp])
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    95
  apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    96
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
 (* Liveness *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    98
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    99
done
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
   100
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   101
end