src/HOL/UNITY/Lift.thy
author oheimb
Mon, 21 Sep 1998 23:25:27 +0200
changeset 5526 e7617b57a3e6
parent 5426 566f47250bd0
child 5563 228b92552d1f
permissions -rw-r--r--
*** empty log message ***
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
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     7
*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     8
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
     9
Lift = SubstAx +
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
record state =
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    12
  floor :: nat		(*current position of the lift*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    13
  open  :: bool		(*whether the door is open at floor*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    14
  stop  :: bool		(*whether the lift is stopped at floor*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    15
  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
    16
  up    :: bool		(*current direction of movement*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    17
  move  :: bool		(*whether moving takes precedence over opening*)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    18
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    19
consts
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    20
  Min, Max :: nat       (*least and greatest floors*)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    21
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    22
rules
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    23
  Min_le_Max  "Min <= Max"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    24
  
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    25
  (** Linear arithmetic: justified by a separate call to arith_oracle_tac **)
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    26
  arith1      "m < fl ==> n - Suc fl < fl - m + (n - m)"
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    27
  arith2      "[| m<n;  m <= fl |] ==> n - Suc fl < fl - m + (n - m)"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    28
  arith3      "[| Suc ix < m |] ==> ix - n < m - Suc ix + (m - n)"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    29
  arith4      "[| ix < m;  n < ix |] ==> ix - n < m - Suc (ix) + (m - n)"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    30
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    31
constdefs
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    32
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    33
  (** Abbreviations: the "always" part **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    34
  
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    35
  above :: state set
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    36
    "above == {s. EX i. floor s < i & i <= Max & i : req s}"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    37
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    38
  below :: state set
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    39
    "below == {s. EX i. Min <= i & i < floor s & i : req s}"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    40
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    41
  queueing :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    42
    "queueing == above Un below"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    43
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    44
  goingup :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    45
    "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
    46
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    47
  goingdown :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    48
    "goingdown == below Int ({s. ~ up s} Un Compl above)"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    49
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    50
  ready :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    51
    "ready == {s. stop s & ~ open s & move s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    52
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    53
 
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    54
  (** Further abbreviations **)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    55
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    56
  moving :: state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    57
    "moving ==  {s. ~ stop s & ~ open s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    58
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    59
  stopped :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    60
    "stopped == {s. stop s  & ~ open s & ~ move s}"
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    61
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    62
  opened :: state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    63
    "opened ==  {s. stop s  &  open s  &  move s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    64
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    65
  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
    66
    "closed ==  {s. stop s  & ~ open s &  move s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    67
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    68
  atFloor :: nat => state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    69
    "atFloor n ==  {s. floor s = n}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    70
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    71
  Req :: nat => state set
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    72
    "Req n ==  {s. n : req s}"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    73
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    74
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    75
  
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    76
  (** The program **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    77
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    78
  request_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    79
    "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
    80
		                  & ~ stop s & floor s : req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    81
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    82
  open_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    83
    "open_act ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    84
         {(s,s'). s' = s (|open :=True,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    85
			   req  := req s - {floor s},
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    86
			   move := True|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    87
		       & stop s & ~ open s & floor s : req s
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    88
	               & ~(move s & s: queueing)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    89
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    90
  close_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    91
    "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
    92
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    93
  req_up :: "(state*state) set"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    94
    "req_up ==
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    95
         {(s,s'). s' = s (|stop  :=False,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    96
			   floor := Suc (floor s),
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    97
			   up    := True|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    98
		       & s : (ready Int goingup)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    99
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   100
  req_down :: "(state*state) set"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   101
    "req_down ==
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   102
         {(s,s'). s' = s (|stop  :=False,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   103
			   floor := floor s - 1,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   104
			   up    := False|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   105
		       & s : (ready Int goingdown)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   106
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   107
  move_up :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   108
    "move_up ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   109
         {(s,s'). s' = s (|floor := Suc (floor s)|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   110
		       & ~ stop s & up s & floor s ~: req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   111
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   112
  move_down :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   113
    "move_down ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   114
         {(s,s'). s' = s (|floor := floor s - 1|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   115
		       & ~ stop s & ~ up s & floor s ~: req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   116
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   117
  button_press  :: "(state*state) set"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   118
    "button_press ==
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   119
         {(s,s'). EX n. s' = s (|req := insert n (req s)|)
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   120
		        & Min <= n & n <= Max}"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   121
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   122
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   123
  Lprg :: state program
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   124
    (*for the moment, we OMIT button_press*)
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   125
    "Lprg == (|Init = {s. floor s = Min & ~ up s & move s & stop s &
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   126
		          ~ open s & req s = {}},
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   127
	       Acts = {id, request_act, open_act, close_act,
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   128
		       req_up, req_down, move_up, move_down}|)"
5277
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
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   131
  (** Invariants **)
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
  bounded :: state set
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   134
    "bounded == {s. Min <= floor s & floor s <= Max}"
5277
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
  open_stop :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   137
    "open_stop == {s. open s --> stop s}"
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
  open_move :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   140
    "open_move == {s. open s --> move s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   141
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   142
  stop_floor :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   143
    "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
   144
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   145
  moving_up :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   146
    "moving_up == {s. ~ stop s & up s -->
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   147
                   (EX f. floor s <= f & f <= Max & f : req s)}"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   148
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   149
  moving_down :: state set
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   150
    "moving_down == {s. ~ stop s & ~ up s -->
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   151
                     (EX f. Min <= f & f <= floor s & f : req s)}"
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   152
  
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   153
  metric :: [nat,state] => nat
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   154
    "metric ==
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   155
       %n s. if up s & floor s < n then n - floor s
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   156
	     else if ~ up s & n < floor s then floor s - n
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   157
	     else if up s & n < floor s then (Max - floor s) + (Max-n)
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   158
	     else if ~ up s & floor s < n then (floor s - Min) + (n-Min)
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   159
	     else 0"
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   160
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   161
locale floor =
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   162
  fixes 
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   163
    n	:: nat
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   164
  assumes
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   165
    Min_le_n    "Min <= n"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   166
    n_le_Max    "n <= Max"
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   167
  defines
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   168
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   169
end