summaryrefslogtreecommitdiff
path: root/backend/Registers.v
blob: 578e4b87d6af315cb1dc30004774b98cb5d51965 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
(** Pseudo-registers and register states.

  This library defines the type of pseudo-registers used in the RTL
  intermediate language, and of mappings from pseudo-registers to
  values as used in the dynamic semantics of RTL. *)

Require Import Coqlib.
Require Import AST.
Require Import Maps.
Require Import Ordered.
Require Import FSets.
Require FSetAVL.

Definition reg: Set := positive.

Module Reg.

Definition eq := peq.

Definition typenv := PMap.t typ.

End Reg.

(** Mappings from registers to some type. *)

Module Regmap := PMap.

Set Implicit Arguments.

Definition regmap_optget
    (A: Set) (or: option reg) (dfl: A) (rs: Regmap.t A) : A :=
  match or with
  | None => dfl
  | Some r => Regmap.get r rs
  end.

Definition regmap_optset
    (A: Set) (or: option reg) (v: A) (rs: Regmap.t A) : Regmap.t A :=
  match or with
  | None => rs
  | Some r => Regmap.set r v rs
  end.

Notation "a # b" := (Regmap.get b a) (at level 1).
Notation "a ## b" := (List.map (fun r => Regmap.get r a) b) (at level 1).
Notation "a # b <- c" := (Regmap.set b c a) (at level 1, b at next level).

(** Sets of registers *)

Module Regset := FSetAVL.Make(OrderedPositive).