Library Ltac2.Ref
Type of a reference cell, similar to OCaml's 'a ref type.
Ltac2 Type 'a ref := 'a Init.ref.
Ltac2 ref (v : 'a) : 'a ref := { contents := v}.
Ltac2 get (r : 'a ref) : 'a := r.(contents).
Ltac2 set (r : 'a ref) (v : 'a) : unit := r.(contents) := v.
Ltac2 incr (r : int ref) : unit := r.(contents) := add (r.(contents)) 1.
Ltac2 decr (r : int ref) : unit := r.(contents) := sub (r.(contents)) 1.
Ltac2 update (r : 'a ref) (f : 'a -> 'a) : unit :=
r.(contents) := f (r.(contents)).
Ltac2 ref (v : 'a) : 'a ref := { contents := v}.
Ltac2 get (r : 'a ref) : 'a := r.(contents).
Ltac2 set (r : 'a ref) (v : 'a) : unit := r.(contents) := v.
Ltac2 incr (r : int ref) : unit := r.(contents) := add (r.(contents)) 1.
Ltac2 decr (r : int ref) : unit := r.(contents) := sub (r.(contents)) 1.
Ltac2 update (r : 'a ref) (f : 'a -> 'a) : unit :=
r.(contents) := f (r.(contents)).