src/HOL/UNITY/Traces.thy
author paulson
Tue, 29 Sep 1998 15:58:47 +0200
changeset 5584 aad639e56d4e
parent 5313 1861a564d7e2
child 5596 b29d18d8c4d2
permissions -rw-r--r--
Now id:(Acts prg) is implicit
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
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
Traces = UNITY +
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
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    27
record 'a program =
5584
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    28
  Init  :: 'a set
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    29
  Acts0 :: "('a * 'a)set set"
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    30
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    31
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    32
constdefs
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    33
    Acts :: "'a program => ('a * 'a)set set"
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5313
diff changeset
    34
    "Acts prg == insert id (Acts0 prg)"
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    35
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    36
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    37
consts reachable :: "'a program => 'a set"
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    38
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    39
inductive "reachable prg"
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    40
  intrs 
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    41
    Init  "s: Init prg ==> s : reachable prg"
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    42
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    43
    Acts  "[| act: Acts prg;  s : reachable prg;  (s,s'): act |]
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    44
	   ==> s' : reachable prg"
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    45
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    46
end