src/HOL/UNITY/Lift.thy
author paulson
Wed, 19 Aug 1998 10:34:31 +0200
changeset 5340 d75c03cf77b5
parent 5320 79b326bceafb
child 5357 6efb2b87610c
permissions -rw-r--r--
Misc changes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Lift.thy
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     2
    ID:         $Id$
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     5
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     6
The Lift-Control Example
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
     7
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
     8
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     9
*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    10
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    11
Lift = SubstAx +
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    12
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    13
record state =
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    14
  floor :: nat		(*current position of the lift*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    15
  open  :: bool		(*whether the door is open at floor*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    16
  stop  :: bool		(*whether the lift is stopped at floor*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    17
  req   :: nat set	(*for each floor, whether the lift is requested*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    18
  up    :: bool		(*current direction of movement*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    19
  move  :: bool		(*whether moving takes precedence over opening*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    20
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    21
consts
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    22
  min, max :: nat       (*least and greatest floors*)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    23
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    24
rules
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    25
  min_le_max  "min <= max"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    26
  
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    27
  (** Linear arithmetic: justified by a separate call to arith_oracle_tac **)
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    28
  arith1      "m < fl ==> n - Suc fl < fl - m + (n - m)"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    29
  arith2      "[| Suc fl < n; m < n |] ==> n - Suc fl < fl - m + (n - m)"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    30
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    31
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    32
constdefs
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    33
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    34
  (** Abbreviations: the "always" part **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    35
  
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    36
  above :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    37
    "above == {s. EX i. floor s < i & i <= max & i : req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    38
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    39
  below :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    40
    "below == {s. EX i. min <= i & i < floor s & i : req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    41
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    42
  queueing :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    43
    "queueing == above Un below"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    44
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    45
  goingup :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    46
    "goingup   == above Int  ({s. up s}  Un Compl below)"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    47
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    48
  goingdown :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    49
    "goingdown == below Int ({s. ~ up s} Un Compl above)"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    50
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    51
  ready :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    52
    "ready == {s. stop s & ~ open s & move s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    53
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    54
 
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    55
  (** Further abbreviations **)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    56
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    57
  moving :: state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    58
    "moving ==  {s. ~ stop s & ~ open s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    59
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    60
  stopped :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    61
    "stopped == {s. stop s  & ~ open s & ~ move s}"
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    62
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    63
  opened :: state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    64
    "opened ==  {s. stop s  &  open s  &  move s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    65
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    66
  closed :: state set  (*but this is the same as ready!!*)
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    67
    "closed ==  {s. stop s  & ~ open s &  move s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    68
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    69
  atFloor :: nat => state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    70
    "atFloor n ==  {s. floor s = n}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    71
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    72
  Req :: nat => state set
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    73
    "Req n ==  {s. n : req s}"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    74
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    75
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    76
  
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    77
  (** The program **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    78
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    79
  request_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    80
    "request_act == {(s,s'). s' = s (|stop:=True, move:=False|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    81
		                  & ~ stop s & floor s : req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    82
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    83
  open_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    84
    "open_act ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    85
         {(s,s'). s' = s (|open :=True,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    86
			   req  := req s - {floor s},
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    87
			   move := True|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    88
		       & stop s & ~ open s & floor s : req s
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    89
	               & ~(move s & s: queueing)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    90
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    91
  close_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    92
    "close_act == {(s,s'). s' = s (|open := False|) & open s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    93
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    94
  req_up_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    95
    "req_up_act ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    96
         {(s,s'). s' = s (|stop  :=False,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    97
			   floor := Suc (floor s),
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    98
			   up    := True|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    99
		       & s : (ready Int goingup)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   100
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   101
  req_down_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   102
    "req_down_act ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   103
         {(s,s'). s' = s (|stop  :=False,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   104
			   floor := floor s - 1,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   105
			   up    := False|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   106
		       & s : (ready Int goingdown)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   107
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   108
  move_up :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   109
    "move_up ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   110
         {(s,s'). s' = s (|floor := Suc (floor s)|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   111
		       & ~ stop s & up s & floor s ~: req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   112
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   113
  move_down :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   114
    "move_down ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   115
         {(s,s'). s' = s (|floor := floor s - 1|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   116
		       & ~ stop s & ~ up s & floor s ~: req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   117
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   118
  Lprg :: state program
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   119
    "Lprg == (|Init = {s. floor s = min & ~ up s & move s & stop s &
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   120
		          ~ open s & req s = {}},
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   121
	       Acts = {id, request_act, open_act, close_act,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   122
		       req_up_act, req_down_act, move_up, move_down}|)"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   123
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   124
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   125
  (** Invariants **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   126
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   127
  bounded :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   128
    "bounded == {s. min <= floor s & floor s <= max}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   129
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   130
  open_stop :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   131
    "open_stop == {s. open s --> stop s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   132
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   133
  open_move :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   134
    "open_move == {s. open s --> move s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   135
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   136
  stop_floor :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   137
    "stop_floor == {s. stop s & ~ move s --> floor s : req s}"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   138
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   139
  moving_up :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   140
    "moving_up == {s. ~ stop s & up s -->
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   141
                   (EX f. floor s <= f & f <= max & f : req s)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   142
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   143
  moving_down :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   144
    "moving_down == {s. ~ stop s & ~ up s -->
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   145
                     (EX f. min <= f & f <= floor s & f : req s)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   146
  
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   147
  metric :: [nat,state] => nat
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   148
    "metric n s == if up s & floor s < n then n - floor s
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   149
		   else if ~ up s & n < floor s then floor s - n
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   150
		   else if up s & n < floor s then (max - floor s) + (max-n)
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   151
		   else if ~ up s & floor s < n then (floor s - min) + (n-min)
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   152
		   else 0"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   153
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   154
locale floor =
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   155
  fixes 
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   156
    n	:: nat
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   157
  assumes
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   158
    min_le_n    "min <= n"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   159
    n_le_max    "n <= max"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   160
  defines
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   161
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   162
end