src/HOL/UNITY/Traces.thy
author wenzelm
Wed, 03 Feb 1999 16:42:40 +0100
changeset 6188 c40e5ac04e3e
parent 6012 1894bfc4aee9
child 6295 351b3c2b0d83
permissions -rw-r--r--
added is_draft; renamed sig to PRIVATE_THEORY;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     1
(*  Title:      HOL/UNITY/Traces
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     2
    ID:         $Id$
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     4
    Copyright   1998  University of Cambridge
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     5
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     6
Inductive definitions of
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     7
  * traces: the possible execution traces
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     8
  * reachable: the set of reachable states
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
     9
*)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    10
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    11
Traces = LessThan +
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    13
consts traces :: "['a set, ('a * 'a)set set] => ('a * 'a list) set"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    15
  (*Initial states and program => (final state, reversed trace to it)...
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    16
    Arguments MUST be curried in an inductive definition*)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    17
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    18
inductive "traces Init Acts"  
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
  intrs 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
         (*Initial trace is empty*)
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    21
    Init  "s: Init ==> (s,[]) : traces Init Acts"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    23
    Acts  "[| act: Acts;  (s,evs) : traces Init Acts;  (s,s'): act |]
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4837
diff changeset
    24
	   ==> (s', s#evs) : traces Init Acts"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    26
5596
b29d18d8c4d2 abstype of programs
paulson
parents: 5584
diff changeset
    27
typedef (Program)
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    28
  'a program = "{(states :: 'a set,
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    29
		  init   :: 'a set,
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    30
		  acts   :: ('a * 'a)set set).
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    31
		 init <= states &
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    32
		 acts <= Pow(states Times states) &
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    33
		 diag states : acts}"
5584
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    34
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    35
constdefs
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    36
  restrict_rel :: "['a set, ('a * 'a) set] => ('a * 'a) set"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    37
    "restrict_rel A r == (A Times A) Int r"
5596
b29d18d8c4d2 abstype of programs
paulson
parents: 5584
diff changeset
    38
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    39
  mk_program :: "('a set * 'a set * ('a * 'a)set set) => 'a program"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    40
    "mk_program ==
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    41
       %(states, init, acts).
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    42
	Abs_Program (states,
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    43
		     states Int init,
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    44
		     restrict_rel states `` (insert Id acts))"
5596
b29d18d8c4d2 abstype of programs
paulson
parents: 5584
diff changeset
    45
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    46
  States :: "'a program => 'a set"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    47
    "States F == (%(states, init, acts). states) (Rep_Program F)"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    48
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    49
  Init :: "'a program => 'a set"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    50
    "Init F == (%(states, init, acts). init) (Rep_Program F)"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    51
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    52
  Acts :: "'a program => ('a * 'a)set set"
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5648
diff changeset
    53
    "Acts F == (%(states, init, acts). acts) (Rep_Program F)"
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    54
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    55
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    56
consts reachable :: "'a program => 'a set"
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    57
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5608
diff changeset
    58
inductive "reachable F"
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    59
  intrs 
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5608
diff changeset
    60
    Init  "s: Init F ==> s : reachable F"
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    61
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5608
diff changeset
    62
    Acts  "[| act: Acts F;  s : reachable F;  (s,s'): act |]
3ac11c4af76a tidying and renaming
paulson
parents: 5608
diff changeset
    63
	   ==> s' : reachable F"
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    64
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
end