src/HOLCF/IOA/meta_theory/Abstraction.thy
author wenzelm
Wed, 16 Apr 2008 21:52:58 +0200
changeset 26690 e30b8d500c7d
parent 26359 6d437bde2f1d
child 27208 5fe899199f85
permissions -rw-r--r--
Facts.dest_static;
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/Abstraction.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 {* Abstraction Theory -- 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 Abstraction
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
imports LiveIOA
23560
e43686246de4 proper use of ioa_package.ML;
wenzelm
parents: 19741
diff changeset
    10
uses ("ioa_package.ML")
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    11
begin
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    12
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    13
defaultsort type
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    15
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    16
  cex_abs :: "('s1 => 's2) => ('a,'s1)execution => ('a,'s2)execution" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    17
  "cex_abs f ex = (f (fst ex), Map (%(a,t). (a,f t))$(snd ex))"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    18
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    19
  -- {* equals cex_abs on Sequences -- after ex2seq application *}
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    20
  cex_absSeq :: "('s1 => 's2) => ('a option,'s1)transition Seq => ('a option,'s2)transition Seq" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    21
  "cex_absSeq f = (%s. Map (%(s,a,t). (f s,a,f t))$s)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    23
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    24
  is_abstraction ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    25
  "is_abstraction f C A =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    26
   ((!s:starts_of(C). f(s):starts_of(A)) &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    27
   (!s t a. reachable C s & s -a--C-> t
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    28
            --> (f s) -a--A-> (f t)))"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    30
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    31
  weakeningIOA :: "('a,'s2)ioa => ('a,'s1)ioa => ('s1 => 's2) => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    32
  "weakeningIOA A C h = (!ex. ex : executions C --> cex_abs h ex : executions A)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    33
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    34
  temp_strengthening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    35
  "temp_strengthening Q P h = (!ex. (cex_abs h ex |== Q) --> (ex |== P))"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    36
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    37
  temp_weakening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    38
  "temp_weakening Q P h = temp_strengthening (.~ Q) (.~ P) h"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    39
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    40
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    41
  state_strengthening :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    42
  "state_strengthening Q P h = (!s t a.  Q (h(s),a,h(t)) --> P (s,a,t))"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    43
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    44
  state_weakening :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    45
  "state_weakening Q P h = state_strengthening (.~Q) (.~P) h"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    46
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    47
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    48
  is_live_abstraction :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    49
  "is_live_abstraction h CL AM =
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    50
     (is_abstraction h (fst CL) (fst AM) &
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    51
      temp_weakening (snd AM) (snd CL) h)"
4559
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
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    54
axiomatization where
4577
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    55
(* thm about ex2seq which is not provable by induction as ex2seq is not continous *)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    56
ex2seq_abs_cex:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    57
  "ex2seq (cex_abs h ex) = cex_absSeq h (ex2seq ex)"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    59
axiomatization where
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 4816
diff changeset
    60
(* analog to the proved thm strength_Box - proof skipped as trivial *)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    61
weak_Box:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    62
"temp_weakening P Q h
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    63
 ==> temp_weakening ([] P) ([] Q) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    64
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    65
axiomatization where
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 4816
diff changeset
    66
(* analog to the proved thm strength_Next - proof skipped as trivial *)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    67
weak_Next:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    68
"temp_weakening P Q h
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    69
 ==> temp_weakening (Next P) (Next Q) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    70
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    71
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    72
subsection "cex_abs"
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    73
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    74
lemma cex_abs_UU: "cex_abs f (s,UU) = (f s, UU)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    75
  by (simp add: cex_abs_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    76
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    77
lemma cex_abs_nil: "cex_abs f (s,nil) = (f s, nil)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    78
  by (simp add: cex_abs_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    79
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    80
lemma cex_abs_cons: "cex_abs f (s,(a,t)>>ex) = (f s, (a,f t) >> (snd (cex_abs f (t,ex))))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    81
  by (simp add: cex_abs_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    82
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    83
declare cex_abs_UU [simp] cex_abs_nil [simp] cex_abs_cons [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    84
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
    85
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    86
subsection "lemmas"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    87
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    88
lemma temp_weakening_def2: "temp_weakening Q P h = (! ex. (ex |== P) --> (cex_abs h ex |== Q))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    89
  apply (simp add: temp_weakening_def temp_strengthening_def NOT_def temp_sat_def satisfies_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    90
  apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    91
  done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    92
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    93
lemma state_weakening_def2: "state_weakening Q P h = (! s t a. P (s,a,t) --> Q (h(s),a,h(t)))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    94
  apply (simp add: state_weakening_def state_strengthening_def NOT_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    95
  apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    96
  done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    97
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    98
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    99
subsection "Abstraction Rules for Properties"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   100
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   101
lemma exec_frag_abstraction [rule_format]:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   102
 "[| is_abstraction h C A |] ==>
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   103
  !s. reachable C s & is_exec_frag C (s,xs)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   104
  --> is_exec_frag A (cex_abs h (s,xs))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   105
apply (unfold cex_abs_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   106
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   107
apply (tactic {* pair_induct_tac "xs" [thm "is_exec_frag_def"] 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   108
txt {* main case *}
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   109
apply (auto dest: reachable.reachable_n simp add: is_abstraction_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   110
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   111
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   112
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   113
lemma abs_is_weakening: "is_abstraction h C A ==> weakeningIOA A C h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   114
apply (simp add: weakeningIOA_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   115
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   116
apply (simp add: executions_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   117
txt {* start state *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   118
apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   119
apply (simp add: is_abstraction_def cex_abs_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   120
txt {* is-execution-fragment *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   121
apply (erule exec_frag_abstraction)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   122
apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   123
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   124
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   125
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   126
lemma AbsRuleT1: "[|is_abstraction h C A; validIOA A Q; temp_strengthening Q P h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   127
          ==> validIOA C P"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   128
apply (drule abs_is_weakening)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   129
apply (simp add: weakeningIOA_def validIOA_def temp_strengthening_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   130
apply (auto simp add: split_paired_all)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   131
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   132
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   133
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   134
(* FIX: Nach TLS.ML *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   135
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   136
lemma IMPLIES_temp_sat: "(ex |== P .--> Q) = ((ex |== P) --> (ex |== Q))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   137
  by (simp add: IMPLIES_def temp_sat_def satisfies_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   138
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   139
lemma AND_temp_sat: "(ex |== P .& Q) = ((ex |== P) & (ex |== Q))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   140
  by (simp add: AND_def temp_sat_def satisfies_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   141
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   142
lemma OR_temp_sat: "(ex |== P .| Q) = ((ex |== P) | (ex |== Q))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   143
  by (simp add: OR_def temp_sat_def satisfies_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   144
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   145
lemma NOT_temp_sat: "(ex |== .~ P) = (~ (ex |== P))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   146
  by (simp add: NOT_def temp_sat_def satisfies_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   147
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   148
declare IMPLIES_temp_sat [simp] AND_temp_sat [simp] OR_temp_sat [simp] NOT_temp_sat [simp]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   149
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   150
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   151
lemma AbsRuleT2:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   152
   "[|is_live_abstraction h (C,L) (A,M);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   153
          validLIOA (A,M) Q;  temp_strengthening Q P h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   154
          ==> validLIOA (C,L) P"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   155
apply (unfold is_live_abstraction_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   156
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   157
apply (drule abs_is_weakening)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   158
apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   159
apply (auto simp add: split_paired_all)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   160
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   161
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   162
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   163
lemma AbsRuleTImprove:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   164
   "[|is_live_abstraction h (C,L) (A,M);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   165
          validLIOA (A,M) (H1 .--> Q);  temp_strengthening Q P h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   166
          temp_weakening H1 H2 h; validLIOA (C,L) H2 |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   167
          ==> validLIOA (C,L) P"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   168
apply (unfold is_live_abstraction_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   169
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   170
apply (drule abs_is_weakening)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   171
apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   172
apply (auto simp add: split_paired_all)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   173
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   174
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   175
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   176
subsection "Correctness of safe abstraction"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   177
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   178
lemma abstraction_is_ref_map:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   179
"is_abstraction h C A ==> is_ref_map h C A"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   180
apply (unfold is_abstraction_def is_ref_map_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   181
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   182
apply (rule_tac x = "(a,h t) >>nil" in exI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   183
apply (simp add: move_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   184
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   185
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   186
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   187
lemma abs_safety: "[| inp(C)=inp(A); out(C)=out(A);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   188
                   is_abstraction h C A |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   189
                ==> C =<| A"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   190
apply (simp add: ioa_implements_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   191
apply (rule trace_inclusion)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   192
apply (simp (no_asm) add: externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   193
apply (auto)[1]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   194
apply (erule abstraction_is_ref_map)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   195
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   196
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   197
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   198
subsection "Correctness of life abstraction"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   199
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   200
(* Reduces to Filter (Map fst x) = Filter (Map fst (Map (%(a,t). (a,x)) x),
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   201
   that is to special Map Lemma *)
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   202
lemma traces_coincide_abs:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   203
  "ext C = ext A
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   204
         ==> mk_trace C$xs = mk_trace A$(snd (cex_abs f (s,xs)))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   205
apply (unfold cex_abs_def mk_trace_def filter_act_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   206
apply simp
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   207
apply (tactic {* pair_induct_tac "xs" [] 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   208
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   209
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   210
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   211
(* Does not work with abstraction_is_ref_map as proof of abs_safety, because
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   212
   is_live_abstraction includes temp_strengthening which is necessarily based
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   213
   on cex_abs and not on corresp_ex. Thus, the proof is redoone in a more specific
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   214
   way for cex_abs *)
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   215
lemma abs_liveness: "[| inp(C)=inp(A); out(C)=out(A);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   216
                   is_live_abstraction h (C,M) (A,L) |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   217
                ==> live_implements (C,M) (A,L)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   218
apply (simp add: is_live_abstraction_def live_implements_def livetraces_def liveexecutions_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   219
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   220
apply (rule_tac x = "cex_abs h ex" in exI)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   221
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   222
  (* Traces coincide *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   223
  apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   224
  apply (rule traces_coincide_abs)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   225
  apply (simp (no_asm) add: externals_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   226
  apply (auto)[1]
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   227
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   228
  (* cex_abs is execution *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   229
  apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   230
  apply (simp add: executions_def)
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   231
  (* start state *)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   232
  apply (rule conjI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   233
  apply (simp add: is_abstraction_def cex_abs_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   234
  (* is-execution-fragment *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   235
  apply (erule exec_frag_abstraction)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   236
  apply (simp add: reachable.reachable_0)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   237
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   238
 (* Liveness *)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   239
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   240
 apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   241
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   242
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   243
(* FIX: NAch Traces.ML bringen *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   244
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   245
lemma implements_trans:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   246
"[| A =<| B; B =<| C|] ==> A =<| C"
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   247
by (auto simp add: ioa_implements_def)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   248
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   249
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   250
subsection "Abstraction Rules for Automata"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   251
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   252
lemma AbsRuleA1: "[| inp(C)=inp(A); out(C)=out(A);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   253
                   inp(Q)=inp(P); out(Q)=out(P);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   254
                   is_abstraction h1 C A;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   255
                   A =<| Q ;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   256
                   is_abstraction h2 Q P |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   257
                ==> C =<| P"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   258
apply (drule abs_safety)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   259
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   260
apply (drule abs_safety)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   261
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   262
apply (erule implements_trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   263
apply (erule implements_trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   264
apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   265
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   266
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   267
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   268
lemma AbsRuleA2: "!!LC. [| inp(C)=inp(A); out(C)=out(A);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   269
                   inp(Q)=inp(P); out(Q)=out(P);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   270
                   is_live_abstraction h1 (C,LC) (A,LA);
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   271
                   live_implements (A,LA) (Q,LQ) ;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   272
                   is_live_abstraction h2 (Q,LQ) (P,LP) |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   273
                ==> live_implements (C,LC) (P,LP)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   274
apply (drule abs_liveness)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   275
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   276
apply (drule abs_liveness)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   277
apply assumption+
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   278
apply (erule live_implements_trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   279
apply (erule live_implements_trans)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   280
apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   281
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   282
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   283
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   284
declare split_paired_All [simp del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   285
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   286
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   287
subsection "Localizing Temporal Strengthenings and Weakenings"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   288
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   289
lemma strength_AND:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   290
"[| temp_strengthening P1 Q1 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   291
          temp_strengthening P2 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   292
       ==> temp_strengthening (P1 .& P2) (Q1 .& Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   293
apply (unfold temp_strengthening_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   294
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   295
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   296
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   297
lemma strength_OR:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   298
"[| temp_strengthening P1 Q1 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   299
          temp_strengthening P2 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   300
       ==> temp_strengthening (P1 .| P2) (Q1 .| Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   301
apply (unfold temp_strengthening_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   302
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   303
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   304
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   305
lemma strength_NOT:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   306
"[| temp_weakening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   307
       ==> temp_strengthening (.~ P) (.~ Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   308
apply (unfold temp_strengthening_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   309
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   310
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   311
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   312
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   313
lemma strength_IMPLIES:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   314
"[| temp_weakening P1 Q1 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   315
          temp_strengthening P2 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   316
       ==> temp_strengthening (P1 .--> P2) (Q1 .--> Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   317
apply (unfold temp_strengthening_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   318
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   319
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   320
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   321
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   322
lemma weak_AND:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   323
"[| temp_weakening P1 Q1 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   324
          temp_weakening P2 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   325
       ==> temp_weakening (P1 .& P2) (Q1 .& Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   326
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   327
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   328
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   329
lemma weak_OR:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   330
"[| temp_weakening P1 Q1 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   331
          temp_weakening P2 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   332
       ==> temp_weakening (P1 .| P2) (Q1 .| Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   333
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   334
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   335
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   336
lemma weak_NOT:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   337
"[| temp_strengthening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   338
       ==> temp_weakening (.~ P) (.~ Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   339
apply (unfold temp_strengthening_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   340
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   341
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   342
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   343
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   344
lemma weak_IMPLIES:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   345
"[| temp_strengthening P1 Q1 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   346
          temp_weakening P2 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   347
       ==> temp_weakening (P1 .--> P2) (Q1 .--> Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   348
apply (unfold temp_strengthening_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   349
apply (simp add: temp_weakening_def2)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   350
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   351
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   352
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   353
subsubsection {* Box *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   354
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   355
(* FIX: should be same as nil_is_Conc2 when all nils are turned to right side !! *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   356
lemma UU_is_Conc: "(UU = x @@ y) = (((x::'a Seq)= UU) | (x=nil & y=UU))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   357
apply (tactic {* Seq_case_simp_tac "x" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   358
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   359
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   360
lemma ex2seqConc [rule_format]:
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   361
"Finite s1 -->
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   362
  (! ex. (s~=nil & s~=UU & ex2seq ex = s1 @@ s) --> (? ex'. s = ex2seq ex'))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   363
apply (rule impI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   364
apply (tactic {* Seq_Finite_induct_tac 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   365
apply blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   366
(* main case *)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   367
apply clarify
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   368
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   369
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   370
(* UU case *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   371
apply (simp add: nil_is_Conc)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   372
(* nil case *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   373
apply (simp add: nil_is_Conc)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   374
(* cons case *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   375
apply (tactic {* pair_tac "aa" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   376
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   377
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   378
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   379
(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   380
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   381
lemma ex2seq_tsuffix:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   382
"tsuffix s (ex2seq ex) ==> ? ex'. s = (ex2seq ex')"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   383
apply (unfold tsuffix_def suffix_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   384
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   385
apply (drule ex2seqConc)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   386
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   387
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   388
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   389
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   390
(* FIX: NAch Sequence.ML bringen *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   391
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   392
lemma Mapnil: "(Map f$s = nil) = (s=nil)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   393
apply (tactic {* Seq_case_simp_tac "s" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   394
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   395
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   396
lemma MapUU: "(Map f$s = UU) = (s=UU)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   397
apply (tactic {* Seq_case_simp_tac "s" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   398
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   399
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   400
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   401
(* important property of cex_absSeq: As it is a 1to1 correspondence,
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   402
  properties carry over *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   403
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   404
lemma cex_absSeq_tsuffix:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   405
"tsuffix s t ==> tsuffix (cex_absSeq h s) (cex_absSeq h t)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   406
apply (unfold tsuffix_def suffix_def cex_absSeq_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   407
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   408
apply (simp add: Mapnil)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   409
apply (simp add: MapUU)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   410
apply (rule_tac x = "Map (% (s,a,t) . (h s,a, h t))$s1" in exI)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   411
apply (simp add: Map2Finite MapConc)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   412
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   413
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   414
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   415
lemma strength_Box:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   416
"[| temp_strengthening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   417
       ==> temp_strengthening ([] P) ([] Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   418
apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Box_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   419
apply clarify
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   420
apply (frule ex2seq_tsuffix)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   421
apply clarify
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   422
apply (drule_tac h = "h" in cex_absSeq_tsuffix)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   423
apply (simp add: ex2seq_abs_cex)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   424
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   425
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   426
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   427
subsubsection {* Init *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   428
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   429
lemma strength_Init:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   430
"[| state_strengthening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   431
       ==> temp_strengthening (Init P) (Init Q) h"
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   432
apply (unfold temp_strengthening_def state_strengthening_def
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   433
  temp_sat_def satisfies_def Init_def unlift_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   434
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   435
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   436
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   437
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   438
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   439
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   440
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   441
subsubsection {* Next *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   442
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   443
lemma TL_ex2seq_UU:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   444
"(TL$(ex2seq (cex_abs h ex))=UU) = (TL$(ex2seq ex)=UU)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   445
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   446
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   447
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   448
apply (tactic {* Seq_case_simp_tac "s" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   449
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   450
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   451
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   452
lemma TL_ex2seq_nil:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   453
"(TL$(ex2seq (cex_abs h ex))=nil) = (TL$(ex2seq ex)=nil)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   454
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   455
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   456
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   457
apply (tactic {* Seq_case_simp_tac "s" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   458
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   459
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   460
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   461
(* FIX: put to Sequence Lemmas *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   462
lemma MapTL: "Map f$(TL$s) = TL$(Map f$s)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   463
apply (tactic {* Seq_induct_tac "s" [] 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   464
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   465
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   466
(* important property of cex_absSeq: As it is a 1to1 correspondence,
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   467
  properties carry over *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   468
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   469
lemma cex_absSeq_TL:
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   470
"cex_absSeq h (TL$s) = (TL$(cex_absSeq h s))"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   471
apply (unfold cex_absSeq_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   472
apply (simp add: MapTL)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   473
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   474
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   475
(* important property of ex2seq: can be shiftet, as defined "pointwise" *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   476
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   477
lemma TLex2seq: "[| (snd ex)~=UU ; (snd ex)~=nil |] ==> (? ex'. TL$(ex2seq ex) = ex2seq ex')"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   478
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   479
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   480
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   481
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   482
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   483
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   484
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   485
lemma ex2seqnilTL: "(TL$(ex2seq ex)~=nil) = ((snd ex)~=nil & (snd ex)~=UU)"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   486
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   487
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   488
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   489
apply (tactic {* Seq_case_simp_tac "s" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   490
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   491
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   492
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   493
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   494
lemma strength_Next:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   495
"[| temp_strengthening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   496
       ==> temp_strengthening (Next P) (Next Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   497
apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Next_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   498
apply simp
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   499
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   500
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   501
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   502
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   503
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   504
(* cons case *)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   505
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU ex2seq_abs_cex cex_absSeq_TL [symmetric] ex2seqnilTL)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   506
apply (erule conjE)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   507
apply (drule TLex2seq)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   508
apply assumption
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   509
apply auto
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   510
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   511
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   512
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   513
text {* Localizing Temporal Weakenings     - 2 *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   514
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   515
lemma weak_Init:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   516
"[| state_weakening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   517
       ==> temp_weakening (Init P) (Init Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   518
apply (simp add: temp_weakening_def2 state_weakening_def2
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   519
  temp_sat_def satisfies_def Init_def unlift_def)
26359
6d437bde2f1d more antiquotations
haftmann
parents: 25135
diff changeset
   520
apply auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   521
apply (tactic {* pair_tac "ex" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   522
apply (tactic {* Seq_case_simp_tac "y" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   523
apply (tactic {* pair_tac "a" 1 *})
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   524
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   525
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   526
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   527
text {* Localizing Temproal Strengthenings - 3 *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   528
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   529
lemma strength_Diamond:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   530
"[| temp_strengthening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   531
       ==> temp_strengthening (<> P) (<> Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   532
apply (unfold Diamond_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   533
apply (rule strength_NOT)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   534
apply (rule weak_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   535
apply (erule weak_NOT)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   536
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   537
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   538
lemma strength_Leadsto:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   539
"[| temp_weakening P1 P2 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   540
          temp_strengthening Q1 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   541
       ==> temp_strengthening (P1 ~> Q1) (P2 ~> Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   542
apply (unfold Leadsto_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   543
apply (rule strength_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   544
apply (erule strength_IMPLIES)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   545
apply (erule strength_Diamond)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   546
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   547
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   548
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   549
text {* Localizing Temporal Weakenings - 3 *}
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   550
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   551
lemma weak_Diamond:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   552
"[| temp_weakening P Q h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   553
       ==> temp_weakening (<> P) (<> Q) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   554
apply (unfold Diamond_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   555
apply (rule weak_NOT)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   556
apply (rule strength_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   557
apply (erule strength_NOT)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   558
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   559
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   560
lemma weak_Leadsto:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   561
"[| temp_strengthening P1 P2 h;
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   562
          temp_weakening Q1 Q2 h |]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   563
       ==> temp_weakening (P1 ~> Q1) (P2 ~> Q2) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   564
apply (unfold Leadsto_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   565
apply (rule weak_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   566
apply (erule weak_IMPLIES)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   567
apply (erule weak_Diamond)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   568
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   569
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   570
lemma weak_WF:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   571
  " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   572
    ==> temp_weakening (WF A acts) (WF C acts) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   573
apply (unfold WF_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   574
apply (rule weak_IMPLIES)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   575
apply (rule strength_Diamond)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   576
apply (rule strength_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   577
apply (rule strength_Init)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   578
apply (rule_tac [2] weak_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   579
apply (rule_tac [2] weak_Diamond)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   580
apply (rule_tac [2] weak_Init)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   581
apply (auto simp add: state_weakening_def state_strengthening_def
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   582
  xt2_def plift_def option_lift_def NOT_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   583
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   584
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   585
lemma weak_SF:
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   586
  " !!A. [| !! s. Enabled A acts (h s) ==> Enabled C acts s|]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   587
    ==> temp_weakening (SF A acts) (SF C acts) h"
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   588
apply (unfold SF_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   589
apply (rule weak_IMPLIES)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   590
apply (rule strength_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   591
apply (rule strength_Diamond)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   592
apply (rule strength_Init)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   593
apply (rule_tac [2] weak_Box)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   594
apply (rule_tac [2] weak_Diamond)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   595
apply (rule_tac [2] weak_Init)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   596
apply (auto simp add: state_weakening_def state_strengthening_def
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   597
  xt2_def plift_def option_lift_def NOT_def)
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   598
done
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   599
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   600
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   601
lemmas weak_strength_lemmas =
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   602
  weak_OR weak_AND weak_NOT weak_IMPLIES weak_Box weak_Next weak_Init
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   603
  weak_Diamond weak_Leadsto strength_OR strength_AND strength_NOT
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   604
  strength_IMPLIES strength_Box strength_Next strength_Init
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   605
  strength_Diamond strength_Leadsto weak_WF weak_SF
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   606
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   607
ML {*
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   608
fun abstraction_tac i =
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   609
    SELECT_GOAL (CLASIMPSET (fn (cs, ss) =>
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   610
      auto_tac (cs addSIs @{thms weak_strength_lemmas},
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 23560
diff changeset
   611
        ss addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}]))) i
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   612
*}
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   613
23560
e43686246de4 proper use of ioa_package.ML;
wenzelm
parents: 19741
diff changeset
   614
use "ioa_package.ML"
e43686246de4 proper use of ioa_package.ML;
wenzelm
parents: 19741
diff changeset
   615
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
   616
end