aboutsummaryrefslogtreecommitdiff
path: root/src/Assembly/Medial.v
blob: 5f082aa12e8781f245fb1be2be359c9ffe9761ab (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
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
Require Import QhasmEvalCommon Pseudo.
Require Import Language.
Require Import List.

Module Medial (Arch: PseudoMachine) <: Language.
  Import MedState EvalUtil Arch.

  Definition W := word width.

  (* Program Types *)
  Definition State := MedState width.
  Transparent State.

  Inductive MAssignment : Type :=
    | MAVar: nat -> nat -> MAssignment 
    | MAMem: forall m, nat -> nat -> Index m -> MAssignment 
    | MAConst: nat -> W -> MAssignment.

  Inductive MOperation :=
    | MIOpConst: IntOp -> nat -> W -> MOperation
    | MIOpReg: IntOp -> nat -> nat -> MOperation
    | MCOpReg: CarryOp -> nat -> nat -> MOperation
    | MDOpReg: DualOp -> nat -> nat -> option nat -> MOperation
    | MOpRot: RotOp -> nat -> Index width -> MOperation.

  Inductive MConditional :=
    | MZ: nat -> MConditional
    | MCReg: TestOp -> nat -> nat -> MConditional
    | MCConst: TestOp -> nat -> W -> MConditional.

  Inductive Medial: Set :=
    | MSkip: Medial
    | MSeq: Medial -> Medial -> Medial
    | MAssign: MAssignment -> Medial
    | MOp: MOperation -> Medial
    | MCond: MConditional -> Medial -> Medial -> Medial
    | MFunexp: nat -> Medial -> Medial
    | MDef: nat -> Medial -> Medial -> Medial
    | MCall: nat -> Medial.

  Definition Program := Medial.

  Fixpoint inline (l: nat) (f p: Medial) :=
    match p with
    | MSeq a b => MSeq (inline l f a) (inline l f b)
    | MCond c a b => MCond c (inline l f a) (inline l f b)
    | MFunexp e a => MFunexp e (inline l f a)
    | MDef l' f' p' =>
      if (Nat.eq_dec l l')
      then p
      else MDef l' (inline l f f') (inline l f p')
    | MCall l' =>
      if (Nat.eq_dec l l')
      then f
      else p
    | _ => p
    end.

  Fixpoint inlineAll (p: Medial) :=
    match p with
    | MSeq a b => MSeq (inlineAll a) (inlineAll b)
    | MCond c a b => MCond c (inlineAll a) (inlineAll b)
    | MFunexp e a => MFunexp e (inlineAll a)
    | MDef l' f' p' => inline l' f' (inlineAll p')
    | _ => p
    end.

  Fixpoint meval (m: Medial) (st: State): option State :=
    let omap := fun {A B} (x: option A) (f: A -> option B) =>
      match x with | Some y => f y | _ => None end in

    match m with
    | MSkip => Some st

    | MSeq a b => omap (meval a st) (fun s => meval b s)

    | MAssign a =>
      match a with
      | MAVar x y =>
        match (getVar y st) with
        | Some v => Some (setVar x v st)
        | _ => None
        end

      | MAMem _ x y i =>
        match (getMem y i st) with
        | Some v => Some (setVar x v st)
        | _ => None
        end

      | MAConst x c => Some (setVar x c st)
      end

    | MOp o =>
      match o with
      | MIOpConst io a c =>
        omap (getVar a st) (fun v =>
          let (res, c') := evalIntOp io v c in
          Some (setVar a res (setCarryOpt c' st)))

      | MIOpReg io a b =>
        omap (getVar a st) (fun va =>
          omap (getVar b st) (fun vb =>
            let (res, c') := evalIntOp io va vb in
            Some (setVar a res (setCarryOpt c' st))))

      | MCOpReg o a b =>
        omap (getVar a st) (fun va =>
          omap (getVar b st) (fun vb =>
            let c := match (snd st) with | Some b => b | _ => false end in
            let (res, c') := evalCarryOp o va vb c in
            Some (setVar a res (setCarry c' st))))

      | MDOpReg duo a b (Some x) =>
        omap (getVar a st) (fun va =>
          omap (getVar b st) (fun vb =>
            let (low, high) := evalDualOp duo va vb in
            Some (setVar a low (setVar x high st))))

      | MDOpReg duo a b None =>
        omap (getVar a st) (fun va =>
          omap (getVar b st) (fun vb =>
            let (low, high) := evalDualOp duo va vb in
            Some (setVar a low st)))

      | MOpRot ro a x =>
        omap (getVar a st) (fun v =>
          let res := evalRotOp ro v (proj1_sig x) in
          Some (setVar a res st))
      end

    | MCond c a b =>
      match c with
      | MCConst t x y =>
        omap (getVar x st) (fun vx =>
          if (evalTest t vx y)
          then meval a st
          else meval b st)

      | MCReg t x y =>
        omap (getVar x st) (fun vx =>

          (* TODO: why can't we infer width here? *)
          omap (@getVar width y st) (fun vy =>
            if (evalTest t vx vy)
            then meval a st
            else meval b st))

      | MZ x =>
        omap (getVar x st) (fun vx =>
          if (evalTest TEq vx (wzero width))
          then meval a st
          else meval b st)

      end

    | MFunexp n a =>
      (fix fexp (m: nat) (f: State -> option State) (s: State) :=
        match m with
        | O => Some s
        | S m' =>
          match f s with
          | None => None
          | Some s' => fexp m' f s'
          end
        end
      ) n (meval a) st

    | _ => None
    end.

  Definition evaluatesTo := fun m st0 st1 => meval (inlineAll m) st0 = Some st1.

  (* world peace *)
End Medial.