theory Ex
imports Main
begin
section {* Sledgehammer *}
lemma "inj f ==> inj g ==> inj (f o g)"
sledgehammer e
apply (metis comp_inj_on inj_on_inverseI inv_f_eq top_set_eq) done
section {* Threads in ML *}
text {* some helper functions *}
ML {* fun sleep time = OS.Process.sleep (Time.fromSeconds time) *}
ML {* sleep 5 *}
ML {* val count = ref 0 *}
ML {*
fun slow_inc name time =
let
val a = !count
val _ = sleep time
val result = a + 1
val _ = count := result
in
priority (name ^ ": " ^ string_of_int result)
end
*}
ML {* slow_inc "incremented" 10 *}
ML {* SimpleThread.fork *}
ML {*
SimpleThread.fork true (fn () =>
slow_inc "parallel counter" 10)
*}
section {* Synchronization *}
text{* Example: lost update *}
ML {*
count := 0;
SimpleThread.fork true (fn () =>
slow_inc "counter1" 10);
SimpleThread.fork true (fn () =>
slow_inc "counter2" 10)
*}
subsection {* pthreads-primitives *}
text {* functions for mutexes *}
ML {* Mutex.mutex *}
ML {* Mutex.lock *}
ML {* Mutex.unlock *}
ML {*
(* protecting count *)
val mutex = Mutex.mutex()
*}
ML {*
fun fork_inc name time =
SimpleThread.fork true (fn () =>
let
val _ = Mutex.lock mutex
val _ = slow_inc name time
val _ = Mutex.unlock mutex
in () end);
*}
ML {*
fork_inc "A" 5;
fork_inc "B" 5;
*}
text {* CAUTION: mutex stays locked after interrupt! *}
ML {* fork_inc "C" 10000 *}
ML {* SimpleThread.interrupt it *}
ML {* fork_inc "D" 1 *}
ML {* Mutex.unlock mutex (* actually undefined behaviour!! *) *}
text {* Condition Variables *}
ML {*
val mutex = Mutex.mutex()
val cond = ConditionVar.conditionVar()
val state = ref 0
*}
text {* double state when signalled *}
ML {*
SimpleThread.fork true (fn () =>
let
val _ = Mutex.lock mutex
val _ = ConditionVar.wait (cond, mutex)
val _ = state := (!state) * 2
val _ = sleep 1
val _ = priority ("Changed to: " ^ string_of_int (!state))
val _ = Mutex.unlock mutex
in () end)
*}
text {* change state and signal condition variable*}
ML {* Mutex.lock mutex *}
ML {* state := 21 *}
ML {* ConditionVar.signal cond *}
ML {* Mutex.unlock mutex*}
subsection{* Isabelle/ML combinators*}
ML {* SimpleThread.synchronized *}
text {* protect count *}
ML {*
fun fork_inc_safe name time =
SimpleThread.fork true (fn () =>
SimpleThread.synchronized name mutex (fn () =>
slow_inc name time))
*}
ML {* val e = fork_inc_safe "E" 10000 *}
ML {* SimpleThread.interrupt e *}
ML {* fork_inc_safe "F" 1 *}
subsection {* State Variables with synchronized access *}
ML {* Synchronized.var *}
ML {* val state = Synchronized.var "goods" 0 *}
ML {* Synchronized.value state *}
text{* modifying state *}
ML {* Synchronized.change *}
ML {* Synchronized.change state (fn i => i + 1) *}
ML {* Synchronized.value state *}
text{* change with result *}
ML {* Synchronized.change;
Synchronized.change_result *}
ML {*
val previous_state_string =
Synchronized.change_result state (fn i =>
(string_of_int i, i + 1))
*}
text{* change (with result) when state meets some property *}
ML {* Synchronized.change_result;
Synchronized.guarded_access *}
text{* example: producer consumer *}
ML {* Synchronized.change state (K 0) *}
ML {* fun produce k = Synchronized.change state (fn i => i + k) *}
ML {*
fun consume () = Synchronized.guarded_access state (fn i =>
if i <= 0
then NONE
else SOME(i - 1, i - 1));
*}
ML {*
fun consumer name = SimpleThread.fork true (fn () =>
while true do
let
val amount = consume ()
val _ = sleep 1
in priority (name ^ " consumed; " ^ string_of_int amount ^ " left") end)
*}
ML {* val a = consumer "A" *}
ML {* produce 3 *}
ML {* val b = consumer "B" *}
ML {* produce 4 *}
ML {* map SimpleThread.interrupt [a,b] *}
text {* timeout while waiting on signal *}
ML {* Synchronized.guarded_access;
Synchronized.timed_access*}
ML {*
fun consume_t timeout =
Synchronized.timed_access
state
(* absolute time! *)
(fn _ => SOME (Time.+ (Time.now (),Time.fromSeconds (timeout))))
(fn i => if i <= 0 then NONE else SOME(i - 1, i - 1))
*}
ML {* produce 2 *}
ML {* consume_t 3 *}
ML {* consume_t 3 *}
ML {* consume_t 3 *}
section{* Thread Attributes *}
ML {*
fun start_work () = SimpleThread.fork true (fn () =>
uninterruptible (fn restore_attributes => fn () =>
let
val _ = sleep 5
val _ = priority "... initialized"
val _ = restore_attributes (fn () =>
(sleep 5; priority "done work")) ()
in () end) ())
*}
ML {* val thread = start_work () *}
ML {* SimpleThread.interrupt thread *}
section{* External processes *}
ML {* system_out *}
ML {* val bash_script = "sleep 10; echo woke up" *}
ML {* system_out bash_script *}
ML {* val bash_script = "while [ 0 ]; do x=$x; done" *}
ML {*
val thread = SimpleThread.fork true (fn () =>
let val _ = system_out bash_script
in () end)
*}
ML {*SimpleThread.interrupt thread*}
ML {* val path_to_isabelle = Path.explode "~~/bin/isabelle" *}
text {* File in File.shell_path (spaces!) *}
ML {* val command = File.shell_path path_to_isabelle ^ " version" *}
ML {* system_out command *}
end