aboutsummaryrefslogtreecommitdiffstats
path: root/aarch64/Asmblockgen.v
blob: fc07bbb9367b9572592e6bd4d4973f4df01a5503 (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
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
(* *************************************************************)
(*                                                             *)
(*             The Compcert verified compiler                  *)
(*                                                             *)
(*           Sylvain Boulmé     Grenoble-INP, VERIMAG          *)
(*           Léo Gourdin        UGA, VERIMAG                   *)
(*           Xavier Leroy       INRIA Paris-Rocquencourt       *)
(*           David Monniaux     CNRS, VERIMAG                  *)
(*           Cyril Six          Kalray                         *)
(*                                                             *)
(*  Copyright Kalray. Copyright VERIMAG. All rights reserved.  *)
(*  This file is distributed under the terms of the INRIA      *)
(*  Non-Commercial License Agreement.                          *)
(*                                                             *)
(* *************************************************************)

(** * Translation from Machblock to AArch64 assembly block language (Asmblock) 
    Inspired from the Mach->Asm pass of original Leroy's backend, but adapted to the block structure like the KVX backend. *)

Require Import Recdef Coqlib Zwf Zbits.
Require Import Errors AST Integers Floats Op.
Require Import Locations Machblock Asm Asmblock.

Local Open Scope string_scope.
Local Open Scope list_scope.
Local Open Scope error_monad_scope.

(** Extracting integer or float registers. *)

Definition ireg_of (r: mreg) : res ireg :=
  match preg_of r with
  |  IR' irs => match irs with
                | RR1 mr => OK mr
                | _ => Error(msg "Asmgenblock.ireg_of")
                end
  | _ => Error(msg "Asmgenblock.iregsp_of")
  end.

Definition freg_of (r: mreg) : res freg :=
  match preg_of r with FR' mr => OK mr | _ => Error(msg "Asmgenblock.freg_of") end.

(** Recognition of immediate arguments for logical integer operations.*)

(** Valid immediate arguments are repetitions of a bit pattern [B]
  of length [e] = 2, 4, 8, 16, 32 or 64.
  The bit pattern [B] must be of the form [0*1*0*] or [1*0*1*]
  but must not be all zeros or all ones. *)

(** The following automaton recognizes [0*1*0*|1*0*1*].
<<
               0          1          0
              / \        / \        / \
              \ /        \ /        \ /
        -0--> [B] --1--> [D] --0--> [F]
       /
     [A]
       \
        -1--> [C] --0--> [E] --1--> [G]
              / \        / \        / \
              \ /        \ /        \ /
               1          0          1
>>
*)

Module Automaton.

Inductive state : Type := SA | SB | SC | SD | SE | SF | SG | Sbad.

Definition start := SA.

Definition next (s: state) (b: bool) :=
  match s, b with
    | SA,false => SB      | SA,true => SC
    | SB,false => SB      | SB,true => SD
    | SC,false => SE      | SC,true => SC
    | SD,false => SF      | SD,true => SD
    | SE,false => SE      | SE,true => SG
    | SF,false => SF      | SF,true => Sbad
    | SG,false => Sbad    | SG,true => SG
    | Sbad,_ => Sbad
  end.

Definition accepting (s: state) :=
  match s with
  | SA | SB | SC | SD | SE | SF | SG => true
  | Sbad => false
  end.

Fixpoint run (len: nat) (s: state) (x: Z) : bool :=
  match len with
  | Datatypes.O => accepting s
  | Datatypes.S len => run len (next s (Z.odd x)) (Z.div2 x)
  end.

End Automaton.

(** The following function determines the candidate length [e],
    ensuring that [x] is a repetition [BB...B] 
    of a bit pattern [B] of length [e]. *)

Definition logical_imm_length (x: Z) (sixtyfour: bool) : nat :=
  (** [test n] checks that the low [2n] bits of [x] are of the
      form [BB], that is, two occurrences of the same [n] bits *)
  let test (n: Z) : bool :=
    Z.eqb (Zzero_ext n x) (Zzero_ext n (Z.shiftr x n)) in
  (** If [test n] fails, we know that the candidate length [e] is
      at least [2n].  Hence we test with decreasing values of [n]:
      32, 16, 8, 4, 2. *)
  if sixtyfour && negb (test 32) then 64%nat
  else if negb (test 16) then 32%nat
  else if negb (test 8) then 16%nat
  else if negb (test 4) then 8%nat
  else if negb (test 2) then 4%nat
  else 2%nat.

(** A valid logical immediate is 
- neither [0] nor [-1];
- composed of a repetition [BBBBB] of a bit-pattern [B] of length [e]
- the low [e] bits of the number, that is, [B], match [0*1*0*] or [1*0*1*].
*)

Definition is_logical_imm32 (x: int) : bool :=
  negb (Int.eq x Int.zero) && negb (Int.eq x Int.mone) &&
  Automaton.run (logical_imm_length (Int.unsigned x) false)
                Automaton.start (Int.unsigned x).

Definition is_logical_imm64 (x: int64) : bool :=
  negb (Int64.eq x Int64.zero) && negb (Int64.eq x Int64.mone) &&
  Automaton.run (logical_imm_length (Int64.unsigned x) true)
                Automaton.start (Int64.unsigned x).

Program Definition single_basic (bi: basic): bblock :=
  {| header := nil; body:= bi::nil; exit := None |}.

(* insert [bi] at the head of [k] *)
Program Definition insert_basic (bi: basic) (k:bblocks): bblocks :=
  match k with
  | bb::k' =>
    match bb.(header) with
    | nil => {| header := nil; body := bi :: (body bb); exit := exit bb |}::k'
    | _ => (single_basic bi)::k
    end
  | _ => (single_basic bi)::k
  end.

Notation "bi ::b k" := (insert_basic bi k) (at level 49, right associativity).

(* insert [ctl] at the head of [k] *)
Program Definition insert_ctl (ctl: control) (k:bblocks): bblocks :=
  {| header := nil; body := nil; exit := Some ctl |}::k.

Notation "ctl ::c k" := (insert_ctl ctl k) (at level 49, right associativity).


(** Alignment check for symbols *)

Parameter symbol_is_aligned : ident -> Z -> bool.
(** [symbol_is_aligned id sz] checks whether the symbol [id] is [sz] aligned *)

(***************************************************************************************)
(* STUB inspired from kvx/Asmblockgen.v and the reference aarch64/Asmgen.v (see below) *)

Definition indexed_memory_access (insn: addressing -> basic)
                                 (sz: Z) (base: iregsp) (ofs: ptrofs) (k: bblocks) : bblocks :=
  let ofs := Ptrofs.to_int64 ofs in
  insn (ADimm base ofs) ::b k. (* STUB: change me ! See Asmgen below *)


Definition loadptr (base: iregsp) (ofs: ptrofs) (dst: ireg) (k: bblocks): bblocks :=
  indexed_memory_access (PLoad Pldrx dst) 8 base ofs k.

Definition storeptr (src: ireg) (base: iregsp) (ofs: ptrofs) (k: bblocks): bblocks :=
  indexed_memory_access (PStore Pstrx src) 8 base ofs k.

(** Function epilogue *)

Definition make_epilogue (f: Machblock.function) (k: bblocks) : bblocks :=
  (* FIXME
     Cannot be used because memcpy destroys X30;
     issue being discussed with X. Leroy *)
  (* if is_leaf_function f
  then Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k
  else*) loadptr XSP f.(fn_retaddr_ofs) RA
         (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) ::b k).

(** Decompose integer literals into 16-bit fragments *)

Fixpoint decompose_int (N: nat) (n p: Z) {struct N} : list (Z * Z) :=
  match N with
  | Datatypes.O => nil
  | Datatypes.S N =>
    let frag := Zzero_ext 16 (Z.shiftr n p) in
    if Z.eqb frag 0 then
      decompose_int N n (p + 16)
    else
      (frag, p) :: decompose_int N (Z.ldiff n (Z.shiftl 65535 p)) (p + 16)
  end.

Definition negate_decomposition (l: list (Z * Z)) :=
  List.map (fun np => (Z.lxor (fst np) 65535, snd np)) l.

(* TODO: Coercions to move in Asmblock ? *)
Coercion PArithP: arith_p >-> Funclass.
Coercion PArithPP: arith_pp >-> Funclass.
Coercion PArithRR0: arith_rr0 >-> Funclass.
Coercion PArith: ar_instruction >-> basic.

Definition bcode := list basic.

Notation "i ::bi k" := (cons (A:=basic) i k) (at level 49, right associativity).
(* NB: this notation helps the Coq typechecker to infer coercion [PArith] in [bcode] expressions *)


Definition loadimm_k (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: bcode) : bcode :=
    List.fold_right (fun np k => Pmovk sz (fst np) (snd np) rd rd ::bi k) k l.

Definition loadimm_z (sz: isize) (rd: ireg) (l: list (Z * Z))  (k: bcode) : bcode :=
  match l with
  | nil => Pmovz sz 0 0 rd ::bi k
  | (n1, p1) :: l => (Pmovz sz n1 p1 rd) ::bi loadimm_k sz rd l k
  end.

Definition loadimm_n (sz: isize) (rd: ireg) (l: list (Z * Z)) (k: bcode) : bcode :=
  match l with
  | nil => Pmovn sz 0 0 rd ::bi k
  | (n1, p1) :: l => Pmovn sz n1 p1 rd ::bi loadimm_k sz rd (negate_decomposition l) k
  end.

Definition loadimm (sz: isize) (rd: ireg) (n: Z) (k: bcode) : bcode :=
  let N := match sz with W => 2%nat | X => 4%nat end in
  let dz := decompose_int N n 0 in
  let dn := decompose_int N (Z.lnot n) 0 in
  if Nat.leb (List.length dz) (List.length dn)
  then loadimm_z sz rd dz k
  else loadimm_n sz rd dn k.

Definition loadimm32 (rd: ireg) (n: int) (k: bcode) : bcode :=
  if is_logical_imm32 n
  then Porrimm W (Int.unsigned n) rd XZR ::bi k
  else loadimm W rd (Int.unsigned n) k.

Definition loadimm64 (rd: ireg) (n: int64) (k: bcode) : bcode :=
  if is_logical_imm64 n
  then Porrimm X (Int64.unsigned n) rd XZR ::bi k
  else loadimm X rd (Int64.unsigned n) k.

(** Translation of the arithmetic operation [r <- op(args)].
  The corresponding instructions are prepended to [k]. *)

Definition transl_op
              (op: operation) (args: list mreg) (res: mreg) (k: bcode) :=
  match op, args with
  | Ointconst n, nil =>
      do rd <- ireg_of res;
      OK (loadimm32 rd n k)
  | _, _ => Error(msg "Not implemented yet")
  end.

(** Translation of a Machblock instruction. *)

Definition transl_instr_basic (f: Machblock.function) (i: Machblock.basic_inst)
                              (ep: bool) (k: bcode) :=
  match i with
  | MBop op args res =>
      transl_op op args res k
  | _ => Error(msg "Not implemented yet")
  (*| MBgetstack ofs ty dst =>*)
      (*loadind SP ofs ty dst k*)
  (*| MBsetstack src ofs ty =>*)
      (*storeind src SP ofs ty k*)
  (*| MBgetparam ofs ty dst =>*)
      (*[> load via the frame pointer if it is valid <]*)
      (*do c <- loadind FP ofs ty dst k;*)
      (*OK (if ep then c*)
                (*else (loadind_ptr SP f.(fn_link_ofs) FP) ::i c)*)
  (*| MBop op args res =>*)
      (*transl_op op args res k*)
  (*| MBload trap chunk addr args dst =>*)
      (*transl_load trap chunk addr args dst k*)
  (*| MBstore chunk addr args src =>*)
      (*transl_store chunk addr args src k*)
  end.

(** Translation of a code sequence *)

(* TODO to remove ? *)
(*Definition fp_is_parent (before: bool) (i: Machblock.basic_inst) : bool :=*)
  (*match i with*)
  (*| MBgetstack ofs ty dst => before && negb (mreg_eq dst MFP)*)
  (*| MBsetstack src ofs ty => before*)
  (*| MBgetparam ofs ty dst => negb (mreg_eq dst MFP)*)
  (*| MBop op args res => before && negb (mreg_eq res MFP)*)
  (*| MBload trapping_mode chunk addr args dst => before && negb (mreg_eq dst MFP)*)
  (*| MBstore chunk addr args res => before*)
  (*end.*)

Definition it1_is_parent (before: bool) (i: Machblock.basic_inst) : bool :=
  match i with
  (*| Msetstack src ofs ty => before*)
  (*| Mgetparam ofs ty dst => negb (mreg_eq dst R29)*)
  | MBop op args res => before && negb (mreg_eq res R29)
  | _ => false
  end.

(** This is an equivalent definition in continuation-passing style
  that runs in constant stack space. *)
(*
Fixpoint transl_code_rec (f: Machblock.function) (il: list Machblock.basic_inst)
                         (it1p: bool) (k: bcode -> res bcode) :=
  match il with
  | nil => k nil
  | i1 :: il' =>
      transl_code_rec f il' (it1_is_parent it1p i1)
        (fun c1 => do c2 <- transl_instr_basic f i1 it1p c1; k c2)
  end.

Definition transl_code' (f: Machblock.function) (il: list Machblock.basic_inst) (it1p: bool) :=
  transl_code_rec f il it1p (fun c => OK c).
*)

Fixpoint transl_basic_code (f: Machblock.function) (il: list Machblock.basic_inst) (it1p: bool) (k: bcode):=
  match il with
  | nil => OK nil
  | i1 :: il' =>
      do k1 <- transl_basic_code f il' (it1_is_parent it1p i1) k;
      transl_instr_basic f i1 it1p k1
  end.

(** Translation of a whole function.  Note that we must check
  that the generated code contains less than [2^64] instructions,
  otherwise the offset part of the [PC] code pointer could wrap
  around, leading to incorrect executions. *)

(* NB Sylvain: this issue with PExpand seems specific to kvx -- and not necessary here *)
(* gen_bblocks can generate two bblocks if the ctl is a PExpand (since the PExpand must be alone in its block) *)
(*
Program Definition gen_bblocks (hd: list label) (c: bcode) (ctl: list instruction) :=
  match (extract_ctl ctl) with
  | None => 
      match c with
      | nil => {| header := hd; body := Pnop::nil; exit := None |} :: nil
      | i :: c => {| header := hd; body := ((i :: c) ++ extract_basic ctl); exit := None |} :: nil
      end
  | Some (PExpand (Pbuiltin ef args res)) =>
      match c with
      | nil => {| header := hd; body := nil; exit := Some (PExpand (Pbuiltin ef args res)) |} :: nil
      | _ => {| header := hd; body := c; exit := None |} 
              :: {| header := nil; body := nil; exit := Some (PExpand (Pbuiltin ef args res)) |} :: nil
      end
  | Some ex => {| header := hd; body := (c ++ extract_basic ctl); exit := Some ex |} :: nil
  end
.
*)

(* XXX: the simulation proof may be complex with this pattern. We may fix this later *)
Program Definition cons_bblocks (ll: list label) (bdy: list basic) (ex: option control) (k:bblocks): bblocks :=
  match non_empty_bblockb bdy ex with
  | true => {| header := ll; body:= bdy; exit := ex |}::k
  | false => k
  end.
Obligation 1.
  rewrite <- Heq_anonymous. constructor.
Qed.

Definition transl_control (f: Machblock.function) (ctl: control_flow_inst) (ep: bool) : res (bcode*control) :=
  Error (msg ("not yet implemented")).

Definition transl_exit (f: Machblock.function) (ext: option control_flow_inst) (ep: bool) : res (bcode*option control) :=
  match ext with
    Some ctl => do (b,c) <- transl_control f ctl ep; OK (b, Some c)
  | None => OK (nil, None)
  end.

Definition transl_block (f: Machblock.function) (fb: Machblock.bblock) (ep: bool) (k:bblocks): res (list bblock) :=
  let stub := false in (* TODO: FIXME *) 
  do (bdy, ex) <- transl_exit f fb.(Machblock.exit) stub;
  do bdy' <- transl_basic_code f fb.(Machblock.body) ep bdy;
  OK (cons_bblocks fb.(Machblock.header) bdy' ex k) 
  .

Fixpoint transl_blocks (f: Machblock.function) (lmb: Machblock.code) (ep: bool): res bblocks :=
  match lmb with
  | nil => OK nil
  | mb :: lmb => 
      do lb <- transl_blocks f lmb false; (* TODO: check [false] here *)
      transl_block f mb (if Machblock.header mb then ep else false) lb  (* TODO: check this *)
  end.
  (* OK (make_epilogue f ((Pret X0)::c nil)). (* STUB: TODO CHANGE ME ! *)*)

(* Currently, we assume to be after the PseudoAsmblockproof.transf_program pass... *)
Program Definition make_prologue (f:  Machblock.function) (k:bblocks) :=
   Pallocframe f.(fn_stacksize) f.(fn_link_ofs) ::b
   storeptr RA XSP f.(fn_retaddr_ofs) k.

Definition transl_function (f: Machblock.function) : res Asmblock.function :=
  do lb <- transl_blocks f f.(Machblock.fn_code) true;
  OK (mkfunction f.(Machblock.fn_sig)
        (make_prologue f lb)).

Definition transf_fundef (f: Machblock.fundef) : res Asmblock.fundef :=
  transf_partial_fundef transl_function f. (* TODO: do we need to check the size here ? (issue only in proofs) *)

Definition transf_program (p: Machblock.program) : res Asmblock.program :=
  transform_partial_program transf_fundef p.