| author | wenzelm | 
| Sun, 15 Jan 2006 19:58:57 +0100 | |
| changeset 18694 | 83f50ac6ddcb | 
| parent 16184 | 80617b8d33c5 | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/Lift.thy | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 1998 University of Cambridge | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 5 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 6 | The Lift-Control Example | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 7 | *) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 8 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 9 | theory Lift | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 10 | imports "../UNITY_Main" | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 11 | |
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 12 | begin | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 13 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 14 | record state = | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 15 |   floor :: "int"	    --{*current position of the lift*}
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 16 |   "open" :: "bool"	    --{*whether the door is opened at floor*}
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 17 |   stop  :: "bool"	    --{*whether the lift is stopped at floor*}
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 18 |   req   :: "int set"	    --{*for each floor, whether the lift is requested*}
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 19 |   up    :: "bool"	    --{*current direction of movement*}
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 20 |   move  :: "bool"	    --{*whether moving takes precedence over opening*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | consts | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 23 |   Min :: "int"       --{*least and greatest floors*}
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 24 |   Max :: "int"       --{*least and greatest floors*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 25 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 26 | axioms | 
| 13806 | 27 | Min_le_Max [iff]: "Min \<le> Max" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 28 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 29 | constdefs | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 30 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 31 |   --{*Abbreviations: the "always" part*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 32 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 33 | above :: "state set" | 
| 13806 | 34 |     "above == {s. \<exists>i. floor s < i & i \<le> Max & i \<in> req s}"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 35 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 36 | below :: "state set" | 
| 13806 | 37 |     "below == {s. \<exists>i. Min \<le> i & i < floor s & i \<in> req s}"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 38 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 39 | queueing :: "state set" | 
| 13806 | 40 | "queueing == above \<union> below" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 41 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 42 | goingup :: "state set" | 
| 13806 | 43 |     "goingup   == above \<inter> ({s. up s}  \<union> -below)"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 44 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 45 | goingdown :: "state set" | 
| 13806 | 46 |     "goingdown == below \<inter> ({s. ~ up s} \<union> -above)"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 47 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 48 | ready :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 49 |     "ready == {s. stop s & ~ open s & move s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 50 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 51 |   --{*Further abbreviations*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 52 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 53 | moving :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 54 |     "moving ==  {s. ~ stop s & ~ open s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 55 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 56 | stopped :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 57 |     "stopped == {s. stop s  & ~ open s & ~ move s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 58 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 59 | opened :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 60 |     "opened ==  {s. stop s  &  open s  &  move s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 61 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 62 |   closed :: "state set"  --{*but this is the same as ready!!*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 63 |     "closed ==  {s. stop s  & ~ open s &  move s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 64 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 65 | atFloor :: "int => state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 66 |     "atFloor n ==  {s. floor s = n}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 67 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 68 | Req :: "int => state set" | 
| 13806 | 69 |     "Req n ==  {s. n \<in> req s}"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 70 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 71 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 72 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 73 |   --{*The program*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 74 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 75 | request_act :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 76 |     "request_act == {(s,s'). s' = s (|stop:=True, move:=False|)
 | 
| 13806 | 77 | & ~ stop s & floor s \<in> req s}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 78 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 79 | open_act :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 80 | "open_act == | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 81 |          {(s,s'). s' = s (|open :=True,
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 82 | 			   req  := req s - {floor s},
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 83 | move := True|) | 
| 13806 | 84 | & stop s & ~ open s & floor s \<in> req s | 
| 85 | & ~(move s & s \<in> queueing)}" | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 86 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 87 | close_act :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 88 |     "close_act == {(s,s'). s' = s (|open := False|) & open s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 89 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 90 | req_up :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 91 | "req_up == | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 92 |          {(s,s'). s' = s (|stop  :=False,
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 93 | floor := floor s + 1, | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 94 | up := True|) | 
| 13806 | 95 | & s \<in> (ready \<inter> goingup)}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 96 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 97 | req_down :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 98 | "req_down == | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 99 |          {(s,s'). s' = s (|stop  :=False,
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 100 | floor := floor s - 1, | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 101 | up := False|) | 
| 13806 | 102 | & s \<in> (ready \<inter> goingdown)}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 103 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 104 | move_up :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 105 | "move_up == | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 106 |          {(s,s'). s' = s (|floor := floor s + 1|)
 | 
| 13806 | 107 | & ~ stop s & up s & floor s \<notin> req s}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 108 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 109 | move_down :: "(state*state) set" | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 110 | "move_down == | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 111 |          {(s,s'). s' = s (|floor := floor s - 1|)
 | 
| 13806 | 112 | & ~ stop s & ~ up s & floor s \<notin> req s}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 113 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 114 | button_press :: "(state*state) set" | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 115 |       --{*This action is omitted from prior treatments, which therefore are
 | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 116 | unrealistic: nobody asks the lift to do anything! But adding this | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 117 | action invalidates many of the existing progress arguments: various | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 118 | "ensures" properties fail. Maybe it should be constrained to only | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 119 | allow button presses in the current direction of travel, like in a | 
| 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 120 | real lift.*} | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 121 | "button_press == | 
| 13806 | 122 |          {(s,s'). \<exists>n. s' = s (|req := insert n (req s)|)
 | 
| 123 | & Min \<le> n & n \<le> Max}" | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 124 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 125 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 126 | Lift :: "state program" | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 127 |     --{*for the moment, we OMIT @{text button_press}*}
 | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 128 | "Lift == mk_total_program | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 129 |                 ({s. floor s = Min & ~ up s & move s & stop s &
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 130 | 		          ~ open s & req s = {}},
 | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 131 | 		 {request_act, open_act, close_act,
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 132 | req_up, req_down, move_up, move_down}, | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 133 | UNIV)" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 134 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 135 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 136 |   --{*Invariants*}
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 137 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 138 | bounded :: "state set" | 
| 13806 | 139 |     "bounded == {s. Min \<le> floor s & floor s \<le> Max}"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 140 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 141 | open_stop :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 142 |     "open_stop == {s. open s --> stop s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 143 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 144 | open_move :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 145 |     "open_move == {s. open s --> move s}"
 | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 146 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 147 | stop_floor :: "state set" | 
| 13806 | 148 |     "stop_floor == {s. stop s & ~ move s --> floor s \<in> req s}"
 | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 149 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 150 | moving_up :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 151 |     "moving_up == {s. ~ stop s & up s -->
 | 
| 13806 | 152 | (\<exists>f. floor s \<le> f & f \<le> Max & f \<in> req s)}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 153 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 154 | moving_down :: "state set" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 155 |     "moving_down == {s. ~ stop s & ~ up s -->
 | 
| 13806 | 156 | (\<exists>f. Min \<le> f & f \<le> floor s & f \<in> req s)}" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 157 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 158 | metric :: "[int,state] => int" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 159 | "metric == | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 160 | %n s. if floor s < n then (if up s then n - floor s | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 161 | else (floor s - Min) + (n-Min)) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 162 | else | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 163 | if n < floor s then (if up s then (Max - floor s) + (Max-n) | 
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 164 | else floor s - n) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 165 | else 0" | 
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 166 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 167 | locale Floor = | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 168 | fixes n | 
| 13806 | 169 | assumes Min_le_n [iff]: "Min \<le> n" | 
| 170 | and n_le_Max [iff]: "n \<le> Max" | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 171 | |
| 13806 | 172 | lemma not_mem_distinct: "[| x \<notin> A; y \<in> A |] ==> x \<noteq> y" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 173 | by blast | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 174 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 175 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 176 | declare Lift_def [THEN def_prg_Init, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 177 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 178 | declare request_act_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 179 | declare open_act_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 180 | declare close_act_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 181 | declare req_up_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 182 | declare req_down_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 183 | declare move_up_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 184 | declare move_down_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 185 | declare button_press_def [THEN def_act_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 186 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 187 | (*The ALWAYS properties*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 188 | declare above_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 189 | declare below_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 190 | declare queueing_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 191 | declare goingup_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 192 | declare goingdown_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 193 | declare ready_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 194 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 195 | (*Basic definitions*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 196 | declare bounded_def [simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 197 | open_stop_def [simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 198 | open_move_def [simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 199 | stop_floor_def [simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 200 | moving_up_def [simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 201 | moving_down_def [simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 202 | |
| 13806 | 203 | lemma open_stop: "Lift \<in> Always open_stop" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 204 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 205 | apply (unfold Lift_def, safety) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 206 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 207 | |
| 13806 | 208 | lemma stop_floor: "Lift \<in> Always stop_floor" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 209 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 210 | apply (unfold Lift_def, safety) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 211 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 212 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 213 | (*This one needs open_stop, which was proved above*) | 
| 13806 | 214 | lemma open_move: "Lift \<in> Always open_move" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 215 | apply (cut_tac open_stop) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 216 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 217 | apply (unfold Lift_def, safety) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 218 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 219 | |
| 13806 | 220 | lemma moving_up: "Lift \<in> Always moving_up" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 221 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 222 | apply (unfold Lift_def, safety) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
13812diff
changeset | 223 | apply (auto dest: order_le_imp_less_or_eq simp add: add1_zle_eq) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 224 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 225 | |
| 13806 | 226 | lemma moving_down: "Lift \<in> Always moving_down" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 227 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 228 | apply (unfold Lift_def, safety) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
13812diff
changeset | 229 | apply (blast dest: order_le_imp_less_or_eq) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 230 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 231 | |
| 13806 | 232 | lemma bounded: "Lift \<in> Always bounded" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 233 | apply (cut_tac moving_up moving_down) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 234 | apply (rule AlwaysI, force) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 235 | apply (unfold Lift_def, safety, auto) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 236 | apply (drule not_mem_distinct, assumption, arith)+ | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 237 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 238 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 239 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 240 | subsection{*Progress*}
 | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 241 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 242 | declare moving_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 243 | declare stopped_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 244 | declare opened_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 245 | declare closed_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 246 | declare atFloor_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 247 | declare Req_def [THEN def_set_simp, simp] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 248 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 249 | |
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14378diff
changeset | 250 | text{*The HUG'93 paper mistakenly omits the Req n from these!*}
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 251 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 252 | (** Lift_1 **) | 
| 13806 | 253 | lemma E_thm01: "Lift \<in> (stopped \<inter> atFloor n) LeadsTo (opened \<inter> atFloor n)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 254 | apply (cut_tac stop_floor) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 255 | apply (unfold Lift_def, ensures_tac "open_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 256 | done (*lem_lift_1_5*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 257 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 258 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 259 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 260 | |
| 13806 | 261 | lemma E_thm02: "Lift \<in> (Req n \<inter> stopped - atFloor n) LeadsTo | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 262 | (Req n \<inter> opened - atFloor n)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 263 | apply (cut_tac stop_floor) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 264 | apply (unfold Lift_def, ensures_tac "open_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 265 | done (*lem_lift_1_1*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 266 | |
| 13806 | 267 | lemma E_thm03: "Lift \<in> (Req n \<inter> opened - atFloor n) LeadsTo | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 268 | (Req n \<inter> closed - (atFloor n - queueing))" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 269 | apply (unfold Lift_def, ensures_tac "close_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 270 | done (*lem_lift_1_2*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 271 | |
| 13806 | 272 | lemma E_thm04: "Lift \<in> (Req n \<inter> closed \<inter> (atFloor n - queueing)) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13806diff
changeset | 273 | LeadsTo (opened \<inter> atFloor n)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 274 | apply (unfold Lift_def, ensures_tac "open_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 275 | done (*lem_lift_1_7*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 276 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 277 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 278 | (** Lift 2. Statements of thm05a and thm05b were wrong! **) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 279 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 280 | lemmas linorder_leI = linorder_not_less [THEN iffD1] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 281 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 282 | lemmas (in Floor) le_MinD = Min_le_n [THEN order_antisym] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 283 | and Max_leD = n_le_Max [THEN [2] order_antisym] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 284 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 285 | declare (in Floor) le_MinD [dest!] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 286 | and linorder_leI [THEN le_MinD, dest!] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 287 | and Max_leD [dest!] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 288 | and linorder_leI [THEN Max_leD, dest!] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 289 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 290 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 291 | (*lem_lift_2_0 | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 292 | NOT an ensures_tac property, but a mere inclusion | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 293 | don't know why script lift_2.uni says ENSURES*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 294 | lemma (in Floor) E_thm05c: | 
| 13806 | 295 | "Lift \<in> (Req n \<inter> closed - (atFloor n - queueing)) | 
| 296 | LeadsTo ((closed \<inter> goingup \<inter> Req n) \<union> | |
| 297 | (closed \<inter> goingdown \<inter> Req n))" | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
13812diff
changeset | 298 | by (auto intro!: subset_imp_LeadsTo simp add: linorder_neq_iff) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 299 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 300 | (*lift_2*) | 
| 13806 | 301 | lemma (in Floor) lift_2: "Lift \<in> (Req n \<inter> closed - (atFloor n - queueing)) | 
| 302 | LeadsTo (moving \<inter> Req n)" | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 303 | apply (rule LeadsTo_Trans [OF E_thm05c LeadsTo_Un]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 304 | apply (unfold Lift_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 305 | apply (ensures_tac [2] "req_down") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 306 | apply (ensures_tac "req_up", auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 307 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 308 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 309 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 310 | (** Towards lift_4 ***) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 311 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 312 | declare split_if_asm [split] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 313 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 314 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 315 | (*lem_lift_4_1 *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 316 | lemma (in Floor) E_thm12a: | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 317 | "0 < N ==> | 
| 13806 | 318 |       Lift \<in> (moving \<inter> Req n \<inter> {s. metric n s = N} \<inter> 
 | 
| 319 |               {s. floor s \<notin> req s} \<inter> {s. up s})    
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 320 | LeadsTo | 
| 13806 | 321 |                (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 322 | apply (cut_tac moving_up) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 323 | apply (unfold Lift_def, ensures_tac "move_up", safe) | 
| 13806 | 324 | (*this step consolidates two formulae to the goal metric n s' \<le> metric n s*) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 325 | apply (erule linorder_leI [THEN order_antisym, symmetric]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 326 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 327 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 328 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 329 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 330 | (*lem_lift_4_3 *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 331 | lemma (in Floor) E_thm12b: "0 < N ==> | 
| 13806 | 332 |       Lift \<in> (moving \<inter> Req n \<inter> {s. metric n s = N} \<inter> 
 | 
| 333 |               {s. floor s \<notin> req s} - {s. up s})    
 | |
| 334 |              LeadsTo (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 335 | apply (cut_tac moving_down) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 336 | apply (unfold Lift_def, ensures_tac "move_down", safe) | 
| 13806 | 337 | (*this step consolidates two formulae to the goal metric n s' \<le> metric n s*) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 338 | apply (erule linorder_leI [THEN order_antisym, symmetric]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 339 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 340 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 341 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 342 | (*lift_4*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 343 | lemma (in Floor) lift_4: | 
| 13806 | 344 |      "0<N ==> Lift \<in> (moving \<inter> Req n \<inter> {s. metric n s = N} \<inter> 
 | 
| 345 |                             {s. floor s \<notin> req s}) LeadsTo      
 | |
| 346 |                            (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 347 | apply (rule LeadsTo_Trans [OF subset_imp_LeadsTo | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 348 | LeadsTo_Un [OF E_thm12a E_thm12b]], auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 349 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 350 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 351 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 352 | (** towards lift_5 **) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 353 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 354 | (*lem_lift_5_3*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 355 | lemma (in Floor) E_thm16a: "0<N | 
| 13806 | 356 |   ==> Lift \<in> (closed \<inter> Req n \<inter> {s. metric n s = N} \<inter> goingup) LeadsTo  
 | 
| 357 |              (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 358 | apply (cut_tac bounded) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 359 | apply (unfold Lift_def, ensures_tac "req_up") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 360 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 361 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 362 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 363 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 364 | (*lem_lift_5_1 has ~goingup instead of goingdown*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 365 | lemma (in Floor) E_thm16b: "0<N ==> | 
| 13806 | 366 |       Lift \<in> (closed \<inter> Req n \<inter> {s. metric n s = N} \<inter> goingdown) LeadsTo  
 | 
| 367 |                    (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 368 | apply (cut_tac bounded) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 369 | apply (unfold Lift_def, ensures_tac "req_down") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 370 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 371 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 372 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 373 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 374 | (*lem_lift_5_0 proves an intersection involving ~goingup and goingup, | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 375 | i.e. the trivial disjunction, leading to an asymmetrical proof.*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 376 | lemma (in Floor) E_thm16c: | 
| 13806 | 377 |      "0<N ==> Req n \<inter> {s. metric n s = N} \<subseteq> goingup \<union> goingdown"
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 378 | by (force simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 379 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 380 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 381 | (*lift_5*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 382 | lemma (in Floor) lift_5: | 
| 13806 | 383 |      "0<N ==> Lift \<in> (closed \<inter> Req n \<inter> {s. metric n s = N}) LeadsTo    
 | 
| 384 |                      (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 385 | apply (rule LeadsTo_Trans [OF subset_imp_LeadsTo | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 386 | LeadsTo_Un [OF E_thm16a E_thm16b]]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 387 | apply (drule E_thm16c, auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 388 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 389 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 390 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 391 | (** towards lift_3 **) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 392 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 393 | (*lemma used to prove lem_lift_3_1*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 394 | lemma (in Floor) metric_eq_0D [dest]: | 
| 13806 | 395 | "[| metric n s = 0; Min \<le> floor s; floor s \<le> Max |] ==> floor s = n" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 396 | by (force simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 397 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 398 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 399 | (*lem_lift_3_1*) | 
| 13806 | 400 | lemma (in Floor) E_thm11: "Lift \<in> (moving \<inter> Req n \<inter> {s. metric n s = 0}) LeadsTo    
 | 
| 401 | (stopped \<inter> atFloor n)" | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 402 | apply (cut_tac bounded) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 403 | apply (unfold Lift_def, ensures_tac "request_act", auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 404 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 405 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 406 | (*lem_lift_3_5*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 407 | lemma (in Floor) E_thm13: | 
| 13806 | 408 |   "Lift \<in> (moving \<inter> Req n \<inter> {s. metric n s = N} \<inter> {s. floor s \<in> req s})  
 | 
| 409 |   LeadsTo (stopped \<inter> Req n \<inter> {s. metric n s = N} \<inter> {s. floor s \<in> req s})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 410 | apply (unfold Lift_def, ensures_tac "request_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 411 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 412 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 413 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 414 | (*lem_lift_3_6*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 415 | lemma (in Floor) E_thm14: "0 < N ==> | 
| 13806 | 416 | Lift \<in> | 
| 417 |         (stopped \<inter> Req n \<inter> {s. metric n s = N} \<inter> {s. floor s \<in> req s})  
 | |
| 418 |         LeadsTo (opened \<inter> Req n \<inter> {s. metric n s = N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 419 | apply (unfold Lift_def, ensures_tac "open_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 420 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 421 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 422 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 423 | (*lem_lift_3_7*) | 
| 13806 | 424 | lemma (in Floor) E_thm15: "Lift \<in> (opened \<inter> Req n \<inter> {s. metric n s = N})   
 | 
| 425 |              LeadsTo (closed \<inter> Req n \<inter> {s. metric n s = N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 426 | apply (unfold Lift_def, ensures_tac "close_act") | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 427 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 428 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 429 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 430 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 431 | (** the final steps **) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 432 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 433 | lemma (in Floor) lift_3_Req: "0 < N ==> | 
| 13806 | 434 | Lift \<in> | 
| 435 |         (moving \<inter> Req n \<inter> {s. metric n s = N} \<inter> {s. floor s \<in> req s})    
 | |
| 436 |         LeadsTo (moving \<inter> Req n \<inter> {s. metric n s < N})"
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 437 | apply (blast intro!: E_thm13 E_thm14 E_thm15 lift_5 intro: LeadsTo_Trans) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 438 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 439 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 440 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 441 | (*Now we observe that our integer metric is really a natural number*) | 
| 13806 | 442 | lemma (in Floor) Always_nonneg: "Lift \<in> Always {s. 0 \<le> metric n s}"
 | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 443 | apply (rule bounded [THEN Always_weaken]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 444 | apply (auto simp add: metric_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 445 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 446 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 447 | lemmas (in Floor) R_thm11 = Always_LeadsTo_weaken [OF Always_nonneg E_thm11] | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 448 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 449 | lemma (in Floor) lift_3: | 
| 13806 | 450 | "Lift \<in> (moving \<inter> Req n) LeadsTo (stopped \<inter> atFloor n)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 451 | apply (rule Always_nonneg [THEN integ_0_le_induct]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 452 | apply (case_tac "0 < z") | 
| 13806 | 453 | (*If z \<le> 0 then actually z = 0*) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 454 | prefer 2 apply (force intro: R_thm11 order_antisym simp add: linorder_not_less) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 455 | apply (rule LeadsTo_weaken_R [OF asm_rl Un_upper1]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 456 | apply (rule LeadsTo_Trans [OF subset_imp_LeadsTo | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 457 | LeadsTo_Un [OF lift_4 lift_3_Req]], auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 458 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 459 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 460 | |
| 13806 | 461 | lemma (in Floor) lift_1: "Lift \<in> (Req n) LeadsTo (opened \<inter> atFloor n)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 462 | apply (rule LeadsTo_Trans) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 463 | prefer 2 | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 464 | apply (rule LeadsTo_Un [OF E_thm04 LeadsTo_Un_post]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 465 | apply (rule E_thm01 [THEN [2] LeadsTo_Trans_Un]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 466 | apply (rule lift_3 [THEN [2] LeadsTo_Trans_Un]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 467 | apply (rule lift_2 [THEN [2] LeadsTo_Trans_Un]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 468 | apply (rule LeadsTo_Trans_Un [OF E_thm02 E_thm03]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 469 | apply (rule open_move [THEN Always_LeadsToI]) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 470 | apply (rule Always_LeadsToI [OF open_stop subset_imp_LeadsTo], clarify) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 471 | (*The case split is not essential but makes the proof much faster.*) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 472 | apply (case_tac "open x", auto) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 473 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
11868diff
changeset | 474 | |
| 11195 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 475 | |
| 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 476 | end |