|
1 theory Ex |
|
2 imports Main |
|
3 begin |
|
4 |
|
5 section {* Sledgehammer *} |
|
6 lemma "inj f ==> inj g ==> inj (f o g)" |
|
7 sledgehammer e |
|
8 apply (metis comp_inj_on inj_on_inverseI inv_f_eq top_set_eq) done |
|
9 |
|
10 |
|
11 |
|
12 |
|
13 section {* Threads in ML *} |
|
14 |
|
15 text {* some helper functions *} |
|
16 |
|
17 ML {* fun sleep time = OS.Process.sleep (Time.fromSeconds time) *} |
|
18 ML {* sleep 5 *} |
|
19 |
|
20 ML {* val count = ref 0 *} |
|
21 ML {* |
|
22 fun slow_inc name time = |
|
23 let |
|
24 val a = !count |
|
25 val _ = sleep time |
|
26 val result = a + 1 |
|
27 val _ = count := result |
|
28 in |
|
29 priority (name ^ ": " ^ string_of_int result) |
|
30 end |
|
31 *} |
|
32 ML {* slow_inc "incremented" 10 *} |
|
33 |
|
34 ML {* SimpleThread.fork *} |
|
35 ML {* |
|
36 SimpleThread.fork true (fn () => |
|
37 slow_inc "parallel counter" 10) |
|
38 *} |
|
39 |
|
40 |
|
41 |
|
42 |
|
43 section {* Synchronization *} |
|
44 |
|
45 text{* Example: lost update *} |
|
46 ML {* |
|
47 count := 0; |
|
48 |
|
49 SimpleThread.fork true (fn () => |
|
50 slow_inc "counter1" 10); |
|
51 |
|
52 SimpleThread.fork true (fn () => |
|
53 slow_inc "counter2" 10) |
|
54 *} |
|
55 |
|
56 |
|
57 |
|
58 |
|
59 subsection {* pthreads-primitives *} |
|
60 |
|
61 text {* functions for mutexes *} |
|
62 ML {* Mutex.mutex *} |
|
63 ML {* Mutex.lock *} |
|
64 ML {* Mutex.unlock *} |
|
65 |
|
66 ML {* |
|
67 (* protecting count *) |
|
68 val mutex = Mutex.mutex() |
|
69 *} |
|
70 |
|
71 ML {* |
|
72 fun fork_inc name time = |
|
73 SimpleThread.fork true (fn () => |
|
74 let |
|
75 val _ = Mutex.lock mutex |
|
76 val _ = slow_inc name time |
|
77 val _ = Mutex.unlock mutex |
|
78 in () end); |
|
79 *} |
|
80 |
|
81 ML {* |
|
82 fork_inc "A" 5; |
|
83 fork_inc "B" 5; |
|
84 *} |
|
85 |
|
86 text {* CAUTION: mutex stays locked after interrupt! *} |
|
87 ML {* fork_inc "C" 10000 *} |
|
88 ML {* SimpleThread.interrupt it *} |
|
89 ML {* fork_inc "D" 1 *} |
|
90 ML {* Mutex.unlock mutex (* actually undefined behaviour!! *) *} |
|
91 |
|
92 |
|
93 |
|
94 |
|
95 text {* Condition Variables *} |
|
96 ML {* |
|
97 val mutex = Mutex.mutex() |
|
98 val cond = ConditionVar.conditionVar() |
|
99 val state = ref 0 |
|
100 *} |
|
101 |
|
102 text {* double state when signalled *} |
|
103 ML {* |
|
104 SimpleThread.fork true (fn () => |
|
105 let |
|
106 val _ = Mutex.lock mutex |
|
107 val _ = ConditionVar.wait (cond, mutex) |
|
108 val _ = state := (!state) * 2 |
|
109 val _ = sleep 1 |
|
110 val _ = priority ("Changed to: " ^ string_of_int (!state)) |
|
111 val _ = Mutex.unlock mutex |
|
112 in () end) |
|
113 *} |
|
114 |
|
115 text {* change state and signal condition variable*} |
|
116 ML {* Mutex.lock mutex *} |
|
117 ML {* state := 21 *} |
|
118 ML {* ConditionVar.signal cond *} |
|
119 ML {* Mutex.unlock mutex*} |
|
120 |
|
121 |
|
122 |
|
123 |
|
124 subsection{* Isabelle/ML combinators*} |
|
125 |
|
126 ML {* SimpleThread.synchronized *} |
|
127 |
|
128 text {* protect count *} |
|
129 ML {* |
|
130 fun fork_inc_safe name time = |
|
131 SimpleThread.fork true (fn () => |
|
132 SimpleThread.synchronized name mutex (fn () => |
|
133 slow_inc name time)) |
|
134 *} |
|
135 |
|
136 ML {* val e = fork_inc_safe "E" 10000 *} |
|
137 ML {* SimpleThread.interrupt e *} |
|
138 ML {* fork_inc_safe "F" 1 *} |
|
139 |
|
140 |
|
141 |
|
142 |
|
143 subsection {* State Variables with synchronized access *} |
|
144 |
|
145 ML {* Synchronized.var *} |
|
146 |
|
147 ML {* val state = Synchronized.var "goods" 0 *} |
|
148 ML {* Synchronized.value state *} |
|
149 |
|
150 text{* modifying state *} |
|
151 ML {* Synchronized.change *} |
|
152 ML {* Synchronized.change state (fn i => i + 1) *} |
|
153 ML {* Synchronized.value state *} |
|
154 |
|
155 text{* change with result *} |
|
156 ML {* Synchronized.change; |
|
157 Synchronized.change_result *} |
|
158 ML {* |
|
159 val previous_state_string = |
|
160 Synchronized.change_result state (fn i => |
|
161 (string_of_int i, i + 1)) |
|
162 *} |
|
163 |
|
164 |
|
165 text{* change (with result) when state meets some property *} |
|
166 |
|
167 ML {* Synchronized.change_result; |
|
168 Synchronized.guarded_access *} |
|
169 |
|
170 |
|
171 text{* example: producer consumer *} |
|
172 |
|
173 ML {* Synchronized.change state (K 0) *} |
|
174 |
|
175 ML {* fun produce k = Synchronized.change state (fn i => i + k) *} |
|
176 |
|
177 ML {* |
|
178 fun consume () = Synchronized.guarded_access state (fn i => |
|
179 if i <= 0 |
|
180 then NONE |
|
181 else SOME(i - 1, i - 1)); |
|
182 *} |
|
183 |
|
184 ML {* |
|
185 fun consumer name = SimpleThread.fork true (fn () => |
|
186 while true do |
|
187 let |
|
188 val amount = consume () |
|
189 val _ = sleep 1 |
|
190 in priority (name ^ " consumed; " ^ string_of_int amount ^ " left") end) |
|
191 *} |
|
192 |
|
193 ML {* val a = consumer "A" *} |
|
194 ML {* produce 3 *} |
|
195 ML {* val b = consumer "B" *} |
|
196 ML {* produce 4 *} |
|
197 |
|
198 ML {* map SimpleThread.interrupt [a,b] *} |
|
199 |
|
200 |
|
201 text {* timeout while waiting on signal *} |
|
202 ML {* Synchronized.guarded_access; |
|
203 Synchronized.timed_access*} |
|
204 |
|
205 ML {* |
|
206 fun consume_t timeout = |
|
207 Synchronized.timed_access |
|
208 state |
|
209 (* absolute time! *) |
|
210 (fn _ => SOME (Time.+ (Time.now (),Time.fromSeconds (timeout)))) |
|
211 (fn i => if i <= 0 then NONE else SOME(i - 1, i - 1)) |
|
212 *} |
|
213 |
|
214 ML {* produce 2 *} |
|
215 ML {* consume_t 3 *} |
|
216 ML {* consume_t 3 *} |
|
217 ML {* consume_t 3 *} |
|
218 |
|
219 |
|
220 |
|
221 |
|
222 section{* Thread Attributes *} |
|
223 |
|
224 ML {* |
|
225 fun start_work () = SimpleThread.fork true (fn () => |
|
226 uninterruptible (fn restore_attributes => fn () => |
|
227 let |
|
228 val _ = sleep 5 |
|
229 val _ = priority "... initialized" |
|
230 val _ = restore_attributes (fn () => |
|
231 (sleep 5; priority "done work")) () |
|
232 in () end) ()) |
|
233 *} |
|
234 |
|
235 ML {* val thread = start_work () *} |
|
236 ML {* SimpleThread.interrupt thread *} |
|
237 |
|
238 |
|
239 |
|
240 |
|
241 section{* External processes *} |
|
242 |
|
243 ML {* system_out *} |
|
244 |
|
245 ML {* val bash_script = "sleep 10; echo woke up" *} |
|
246 ML {* system_out bash_script *} |
|
247 |
|
248 ML {* val bash_script = "while [ 0 ]; do x=$x; done" *} |
|
249 ML {* |
|
250 val thread = SimpleThread.fork true (fn () => |
|
251 let val _ = system_out bash_script |
|
252 in () end) |
|
253 *} |
|
254 ML {*SimpleThread.interrupt thread*} |
|
255 |
|
256 |
|
257 ML {* val path_to_isabelle = Path.explode "~~/bin/isabelle" *} |
|
258 |
|
259 text {* File in File.shell_path (spaces!) *} |
|
260 ML {* val command = File.shell_path path_to_isabelle ^ " version" *} |
|
261 ML {* system_out command *} |
|
262 |
|
263 end |
|
264 |