src/Pure/Concurrent/unsynchronized.ML
author wenzelm
Wed, 10 Feb 2021 22:30:51 +0100
changeset 73246 b9c480878663
parent 62935 3c7a35c12e03
child 73696 03e134d5f867
permissions -rw-r--r--
tuned comments;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62508
d0b68218ea55 discontinued RAW session: bootstrap directly from isabelle_process RAW_ML_SYSTEM;
wenzelm
parents: 61925
diff changeset
     1
(*  Title:      Pure/Concurrent/unsynchronized.ML
32737
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
     3
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
     4
Raw ML references as unsynchronized state variables.
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
     5
*)
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
     6
62935
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
     7
signature UNSYNCHRONIZED =
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
     8
sig
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
     9
  datatype ref = datatype ref
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    10
  val := : 'a ref * 'a -> unit
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    11
  val ! : 'a ref -> 'a
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    12
  val change: 'a ref -> ('a -> 'a) -> unit
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    13
  val change_result: 'a ref -> ('a -> 'b * 'a) -> 'b
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    14
  val inc: int ref -> int
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    15
  val dec: int ref -> int
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    16
  val setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    17
end;
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    18
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    19
structure Unsynchronized: UNSYNCHRONIZED =
32737
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    20
struct
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    21
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    22
datatype ref = datatype ref;
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    23
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    24
val op := = op :=;
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    25
val ! = !;
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    26
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    27
fun change r f = r := f (! r);
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    28
fun change_result r f = let val (x, y) = f (! r) in r := y; x end;
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    29
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    30
fun inc i = (i := ! i + (1: int); ! i);
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    31
fun dec i = (i := ! i - (1: int); ! i);
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    32
39616
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    33
fun setmp flag value f x =
62923
3a122e1e352a clarified bootstrap;
wenzelm
parents: 62891
diff changeset
    34
  Thread_Attributes.uninterruptible (fn restore_attributes => fn () =>
39616
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    35
    let
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    36
      val orig_value = ! flag;
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    37
      val _ = flag := value;
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    38
      val result = Exn.capture (restore_attributes f) x;
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    39
      val _ = flag := orig_value;
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    40
    in Exn.release result end) ();
8052101883c3 renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
wenzelm
parents: 38799
diff changeset
    41
32737
76fa673eee8b Raw ML references as unsynchronized state variables.
wenzelm
parents:
diff changeset
    42
end;
62818
2733b240bfea clarified modules;
wenzelm
parents: 62508
diff changeset
    43
62935
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    44
ML_Name_Space.forget_val "ref";
3c7a35c12e03 proper signature for structure;
wenzelm
parents: 62923
diff changeset
    45
ML_Name_Space.forget_type "ref";