src/HOLCF/IOA/meta_theory/Abstraction.thy
author mueller
Thu, 26 Nov 1998 16:37:56 +0100
changeset 5976 44290b71a85f
parent 4816 64f075872f69
child 10835 f4745d77e620
permissions -rw-r--r--
tuning to assimiliate it with PhD;
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$
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     3
    Author:     Olaf M"uller
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     4
    Copyright   1997  TU Muenchen
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     5
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     6
Abstraction Theory -- tailored for I/O automata
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     7
*)   
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     8
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     9
		       
4816
64f075872f69 simplification of explicit theory usage and merges
oheimb
parents: 4577
diff changeset
    10
Abstraction = LiveIOA + 
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    11
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    12
default term
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    13
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    15
consts
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
4577
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    17
  cex_abs      ::"('s1 => 's2) => ('a,'s1)execution => ('a,'s2)execution"
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    18
  cex_absSeq   ::"('s1 => 's2) => ('a option,'s1)transition Seq => ('a option,'s2)transition Seq"
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    19
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
  is_abstraction ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
  weakeningIOA       :: "('a,'s2)ioa => ('a,'s1)ioa => ('s1 => 's2) => bool" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    23
  temp_weakening     :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
  temp_strengthening :: "('a,'s2)ioa_temp => ('a,'s1)ioa_temp => ('s1 => 's2) => bool" 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
  state_weakening       :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
  state_strengthening   :: "('a,'s2)step_pred => ('a,'s1)step_pred => ('s1 => 's2) => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
  is_live_abstraction  :: "('s1 => 's2) => ('a,'s1)live_ioa => ('a,'s2)live_ioa => bool"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
defs
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
is_abstraction_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
  "is_abstraction f C A ==                          
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
   (!s:starts_of(C). f(s):starts_of(A)) &        
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
   (!s t a. reachable C s & s -a--C-> t     
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
            --> (f s) -a--A-> (f t))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    39
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    40
is_live_abstraction_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    41
  "is_live_abstraction h CL AM == 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
      is_abstraction h (fst CL) (fst AM) &
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    43
      temp_weakening (snd AM) (snd CL) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    45
cex_abs_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    46
  "cex_abs f ex == (f (fst ex), Map (%(a,t). (a,f t))`(snd ex))"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
4577
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    48
(* equals cex_abs on Sequneces -- after ex2seq application -- *)
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    49
cex_absSeq_def
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    50
  "cex_absSeq f == % s. Map (%(s,a,t). (f s,a,f t))`s"
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    51
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    52
weakeningIOA_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    53
   "weakeningIOA A C h == ! ex. ex : executions C --> cex_abs h ex : executions A"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    54
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    55
temp_weakening_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    56
   "temp_weakening Q P h == temp_strengthening (.~ Q) (.~ P) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    57
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    58
temp_strengthening_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    59
   "temp_strengthening Q P h == ! ex. (cex_abs h ex |== Q) --> (ex |== P)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    60
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    61
state_weakening_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    62
  "state_weakening Q P h == state_strengthening (.~Q) (.~P) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    63
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    64
state_strengthening_def
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    65
  "state_strengthening Q P h == ! s t a.  Q (h(s),a,h(t)) --> P (s,a,t)"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    66
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    67
rules
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    68
4577
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    69
(* thm about ex2seq which is not provable by induction as ex2seq is not continous *)
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    70
ex2seq_abs_cex
674b0b354feb added thms wrt weakening and strengthening in Abstraction;
mueller
parents: 4559
diff changeset
    71
  "ex2seq (cex_abs h ex) = cex_absSeq h (ex2seq ex)" 
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    72
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 4816
diff changeset
    73
(* analog to the proved thm strength_Box - proof skipped as trivial *)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    74
weak_Box
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    75
"temp_weakening P Q h 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    76
 ==> temp_weakening ([] P) ([] Q) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    77
5976
44290b71a85f tuning to assimiliate it with PhD;
mueller
parents: 4816
diff changeset
    78
(* analog to the proved thm strength_Next - proof skipped as trivial *)
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    79
weak_Next
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    80
"temp_weakening P Q h 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    81
 ==> temp_weakening (Next P) (Next Q) h"
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    82
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    83
end