| author | wenzelm | 
| Tue, 31 Oct 2023 16:49:03 +0100 | |
| changeset 78869 | f464f6bc5809 | 
| parent 78720 | 909dc00766a0 | 
| child 82720 | 956ecf2c07a0 | 
| 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
 | 
| 77998 | 18 | type 'a weak_ref = 'a ref option ref | 
| 19 | val weak_ref: 'a -> 'a weak_ref | |
| 78001 | 20 | val weak_peek: 'a weak_ref -> 'a option | 
| 78019 | 21 | val weak_active: 'a weak_ref -> bool | 
| 62935 | 22 | end; | 
| 23 | ||
| 24 | structure Unsynchronized: UNSYNCHRONIZED = | |
| 32737 | 25 | struct | 
| 26 | ||
| 78001 | 27 | (* regular references *) | 
| 28 | ||
| 32737 | 29 | datatype ref = datatype ref; | 
| 30 | ||
| 31 | val op := = op :=; | |
| 32 | val ! = !; | |
| 33 | ||
| 34 | fun change r f = r := f (! r); | |
| 35 | fun change_result r f = let val (x, y) = f (! r) in r := y; x end; | |
| 36 | ||
| 37 | fun inc i = (i := ! i + (1: int); ! i); | |
| 38 | fun dec i = (i := ! i - (1: int); ! i); | |
| 73696 
03e134d5f867
reactive "sledgehammer_filter": statically correct, but untested (no proof_file);
 wenzelm parents: 
62935diff
changeset | 39 | fun add i n = (i := ! i + (n: int); ! i); | 
| 32737 | 40 | |
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 41 | fun setmp flag value f x = | 
| 78720 | 42 | Thread_Attributes.uninterruptible_body (fn run => | 
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 43 | let | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 44 | val orig_value = ! flag; | 
| 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 45 | val _ = flag := value; | 
| 78716 | 46 | val result = Exn.capture0 (run f) x; | 
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 47 | val _ = flag := orig_value; | 
| 78720 | 48 | in Exn.release result end); | 
| 39616 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 wenzelm parents: 
38799diff
changeset | 49 | |
| 78001 | 50 | |
| 51 | (* weak references *) | |
| 52 | ||
| 78002 | 53 | (*see also $ML_SOURCES/basis/Weak.sml*) | 
| 54 | ||
| 77998 | 55 | type 'a weak_ref = 'a ref option ref; | 
| 78001 | 56 | |
| 77998 | 57 | fun weak_ref a = Weak.weak (SOME (ref a)); | 
| 58 | ||
| 78001 | 59 | fun weak_peek (ref (SOME (ref a))) = SOME a | 
| 60 | | weak_peek _ = NONE; | |
| 61 | ||
| 78019 | 62 | fun weak_active (ref (SOME (ref _))) = true | 
| 63 | | weak_active _ = false; | |
| 64 | ||
| 32737 | 65 | end; | 
| 62818 | 66 | |
| 62935 | 67 | ML_Name_Space.forget_val "ref"; | 
| 68 | ML_Name_Space.forget_type "ref"; |