| author | wenzelm | 
| Sun, 26 Feb 2023 21:17:10 +0100 | |
| changeset 77386 | cae3d891adff | 
| parent 73696 | 03e134d5f867 | 
| child 77996 | afa6117bace4 | 
| permissions | -rw-r--r-- | 
| 62508 
d0b68218ea55
discontinued RAW session: bootstrap directly from isabelle_process RAW_ML_SYSTEM;
 wenzelm parents: 
61925diff
changeset | 1 | (* Title: Pure/Concurrent/unsynchronized.ML | 
| 32737 | 2 | Author: Makarius | 
| 3 | ||
| 4 | Raw ML references as unsynchronized state variables. | |
| 5 | *) | |
| 6 | ||
| 62935 | 7 | signature UNSYNCHRONIZED = | 
| 8 | sig | |
| 9 | datatype ref = datatype ref | |
| 10 | val := : 'a ref * 'a -> unit | |
| 11 | val ! : 'a ref -> 'a | |
| 12 |   val change: 'a ref -> ('a -> 'a) -> unit
 | |
| 13 |   val change_result: 'a ref -> ('a -> 'b * 'a) -> 'b
 | |
| 14 | val inc: int ref -> int | |
| 15 | val dec: int ref -> int | |
| 73696 
03e134d5f867
reactive "sledgehammer_filter": statically correct, but untested (no proof_file);
 wenzelm parents: 
62935diff
changeset | 16 | val add: int ref -> int -> int | 
| 62935 | 17 |   val setmp: 'a ref -> 'a -> ('b -> 'c) -> 'b -> 'c
 | 
| 18 | end; | |
| 19 | ||
| 20 | structure Unsynchronized: UNSYNCHRONIZED = | |
| 32737 | 21 | struct | 
| 22 | ||
| 23 | datatype ref = datatype ref; | |
| 24 | ||
| 25 | val op := = op :=; | |
| 26 | val ! = !; | |
| 27 | ||
| 28 | fun change r f = r := f (! r); | |
| 29 | fun change_result r f = let val (x, y) = f (! r) in r := y; x end; | |
| 30 | ||
| 31 | fun inc i = (i := ! i + (1: int); ! i); | |
| 32 | fun dec i = (i := ! i - (1: int); ! i); | |
| 73696 
03e134d5f867
reactive "sledgehammer_filter": statically correct, but untested (no proof_file);
 wenzelm parents: 
62935diff
changeset | 33 | fun add i n = (i := ! i + (n: int); ! i); | 
| 32737 | 34 | |
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 35 | fun setmp flag value f x = | 
| 62923 | 36 | Thread_Attributes.uninterruptible (fn restore_attributes => fn () => | 
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 37 | let | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 38 | val orig_value = ! flag; | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 39 | val _ = flag := value; | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 40 | val result = Exn.capture (restore_attributes f) x; | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 41 | val _ = flag := orig_value; | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 42 | in Exn.release result end) (); | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 43 | |
| 32737 | 44 | end; | 
| 62818 | 45 | |
| 62935 | 46 | ML_Name_Space.forget_val "ref"; | 
| 47 | ML_Name_Space.forget_type "ref"; |