src/HOL/UNITY/Lift.thy
author paulson
Thu, 26 Aug 1999 11:33:24 +0200
changeset 7359 98a2afab3f86
parent 6718 e869ff059252
child 10064 1a77667b21ef
permissions -rw-r--r--
extra syntax for JN, making it more like UN
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 =
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    12
  floor :: int		(*current position of the lift*)
5277
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*)
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    15
  req   :: int set	(*for each floor, whether the lift is requested*)
5277
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
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    20
  Min, Max :: int       (*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
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    25
constdefs
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    26
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    27
  (** Abbreviations: the "always" part **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    28
  
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    29
  above :: state set
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    30
    "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
    31
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    32
  below :: state set
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    33
    "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
    34
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    35
  queueing :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    36
    "queueing == above Un below"
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
  goingup :: state set
5931
325300576da7 Finally removing "Compl" from HOL
paulson
parents: 5596
diff changeset
    39
    "goingup   == above Int  ({s. up s}  Un -below)"
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
  goingdown :: state set
5931
325300576da7 Finally removing "Compl" from HOL
paulson
parents: 5596
diff changeset
    42
    "goingdown == below Int ({s. ~ up s} Un -above)"
5277
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
  ready :: state set
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    45
    "ready == {s. stop s & ~ open s & move s}"
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
 
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    48
  (** Further abbreviations **)
5277
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
  moving :: state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    51
    "moving ==  {s. ~ stop s & ~ open s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    52
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    53
  stopped :: state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    54
    "stopped == {s. stop s  & ~ open s & ~ move s}"
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    55
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    56
  opened :: state set
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    57
    "opened ==  {s. stop s  &  open s  &  move s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    58
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    59
  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
    60
    "closed ==  {s. stop s  & ~ open s &  move s}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    61
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    62
  atFloor :: int => state set
5320
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    63
    "atFloor n ==  {s. floor s = n}"
79b326bceafb now trans_tac is part of the claset...
paulson
parents: 5277
diff changeset
    64
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    65
  Req :: int => state set
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    66
    "Req n ==  {s. n : req s}"
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
    67
5320
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
  
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    70
  (** The program **)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    71
  
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    72
  request_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    73
    "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
    74
		                  & ~ stop s & floor s : req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    75
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    76
  open_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    77
    "open_act ==
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    78
         {(s,s'). s' = s (|open :=True,
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    79
			   req  := req s - {floor s},
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    80
			   move := True|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    81
		       & stop s & ~ open s & floor s : req s
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    82
	               & ~(move s & s: queueing)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    83
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    84
  close_act :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    85
    "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
    86
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    87
  req_up :: "(state*state) set"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    88
    "req_up ==
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    89
         {(s,s'). s' = s (|stop  :=False,
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    90
			   floor := floor s + #1,
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    91
			   up    := True|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    92
		       & s : (ready Int goingup)}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    93
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    94
  req_down :: "(state*state) set"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
    95
    "req_down ==
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    96
         {(s,s'). s' = s (|stop  :=False,
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
    97
			   floor := floor s - #1,
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    98
			   up    := False|)
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
    99
		       & s : (ready Int goingdown)}"
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
  move_up :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   102
    "move_up ==
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   103
         {(s,s'). s' = s (|floor := floor s + #1|)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   104
		       & ~ stop s & up s & floor s ~: req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   105
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   106
  move_down :: "(state*state) set"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   107
    "move_down ==
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   108
         {(s,s'). s' = s (|floor := floor s - #1|)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   109
		       & ~ stop s & ~ up s & floor s ~: req s}"
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   110
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   111
  button_press  :: "(state*state) set"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   112
    "button_press ==
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   113
         {(s,s'). EX n. s' = s (|req := insert n (req s)|)
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   114
		        & Min <= n & n <= Max}"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   115
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   116
6718
e869ff059252 renamed Lprg to Lift; simplified proof of Always_nonneg
paulson
parents: 6295
diff changeset
   117
  Lift :: state program
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   118
    (*for the moment, we OMIT button_press*)
6718
e869ff059252 renamed Lprg to Lift; simplified proof of Always_nonneg
paulson
parents: 6295
diff changeset
   119
    "Lift == mk_program ({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
   120
		          ~ open s & req s = {}},
5596
b29d18d8c4d2 abstype of programs
paulson
parents: 5584
diff changeset
   121
			 {request_act, open_act, close_act,
b29d18d8c4d2 abstype of programs
paulson
parents: 5584
diff changeset
   122
			  req_up, req_down, move_up, move_down})"
5277
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
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   128
    "bounded == {s. Min <= floor s & floor s <= Max}"
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
  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 -->
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   141
                   (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
   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 -->
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   145
                     (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
   146
  
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   147
  metric :: [int,state] => int
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   148
    "metric ==
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   149
       %n s. if floor s < n then (if up s then n - floor s
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   150
			          else (floor s - Min) + (n-Min))
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   151
             else
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   152
             if n < floor s then (if up s then (Max - floor s) + (Max-n)
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   153
		                  else floor s - n)
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   154
             else #0"
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   155
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   156
locale floor =
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   157
  fixes 
5563
228b92552d1f Now uses integers instead of naturals
paulson
parents: 5426
diff changeset
   158
    n	:: int
5340
d75c03cf77b5 Misc changes
paulson
parents: 5320
diff changeset
   159
  assumes
5357
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   160
    Min_le_n    "Min <= n"
6efb2b87610c New theory Lift
paulson
parents: 5340
diff changeset
   161
    n_le_Max    "n <= Max"
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6015
diff changeset
   162
  defines
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   163
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
diff changeset
   164
end