src/HOL/UNITY/Reach.thy
author paulson
Wed, 28 Apr 1999 13:36:31 +0200
changeset 6535 880f31a62784
parent 6295 351b3c2b0d83
child 10064 1a77667b21ef
permissions -rw-r--r--
eliminated theory UNITY/Traces
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Reach.thy
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
Reachability in Directed Graphs.  From Chandy and Misra, section 6.4.
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
6535
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 6295
diff changeset
     9
Reach = FP + SubstAx +
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
types   vertex
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
        state = "vertex=>bool"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
arities vertex :: term
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
consts
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
  init ::  "vertex"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
  edges :: "(vertex*vertex) set"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
constdefs
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
  asgt  :: "[vertex,vertex] => (state*state) set"
5071
548f398d770b Consequences of the change from [ := ] to ( := ) in theory Update.
nipkow
parents: 4896
diff changeset
    24
    "asgt u v == {(s,s'). s' = s(v:= s u | s v)}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
5253
82a5ca6290aa New record type of programs
paulson
parents: 5240
diff changeset
    26
  Rprg :: state program
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
    27
    "Rprg == mk_program ({%v. v=init}, UN (u,v): edges. {asgt u v})"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
5240
bbcd79ef7cf2 Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents: 5232
diff changeset
    29
  reach_invariant :: state set
bbcd79ef7cf2 Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents: 5232
diff changeset
    30
    "reach_invariant == {s. (ALL v. s v --> (init, v) : edges^*) & s init}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
  fixedpoint :: state set
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
    "fixedpoint == {s. ALL (u,v): edges. s u --> s v}"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
  metric :: state => nat
5426
566f47250bd0 A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents: 5253
diff changeset
    36
    "metric s == card {v. ~ s v}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    37
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    38
rules
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    39
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    40
  (*We assume that the set of vertices is finite*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
  finite_graph "finite (UNIV :: vertex set)"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
  
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
end