Skip to content

Commit 461c421

Browse files
committed
Add support for more assembly
I want to be able to handle the output of gcc/clang on our C code. Right now, I have added support for parsing the assembly output. Equivalence checking is still to come.
1 parent 847d6a7 commit 461c421

32 files changed

+24195
-252
lines changed

src/Assembly/Equality.v

+30-2
Original file line numberDiff line numberDiff line change
@@ -32,11 +32,39 @@ Declare Scope REG_scope.
3232
Delimit Scope REG_scope with REG.
3333
Bind Scope REG_scope with REG.
3434

35+
36+
Definition REG_beq (r1 r2 : REG) : bool :=
37+
(prod_beq _ _ (prod_beq _ _ N.eqb N.eqb) N.eqb)
38+
(index_and_shift_and_bitcount_of_reg r1) (index_and_shift_and_bitcount_of_reg r2).
39+
40+
Lemma REG_dec_lb : forall r1 r2 : REG, r1 = r2 -> REG_beq r1 r2 = true.
41+
Proof.
42+
intros r1 r2 H.
43+
subst r2; destruct r1.
44+
all: reflexivity.
45+
Defined.
46+
47+
Lemma REG_dec_bl : forall r1 r2 : REG, REG_beq r1 r2 = true -> r1 = r2.
48+
Proof.
49+
cbv [REG_beq].
50+
intros r1 r2 H.
51+
rewrite <- (reg_of_index_and_shift_and_bitcount_of_index_and_shift_and_bitcount_of_reg r1), <- (reg_of_index_and_shift_and_bitcount_of_index_and_shift_and_bitcount_of_reg r2).
52+
reflect_hyps.
53+
rewrite H.
54+
reflexivity.
55+
Defined.
56+
57+
Definition REG_eq_dec (x y : REG) : {x = y} + {x <> y} :=
58+
(if REG_beq x y as b return (REG_beq x y = b -> _)
59+
then fun pf => left (REG_dec_bl x y pf)
60+
else fun pf => right (fun pf' => diff_false_true (eq_trans (eq_sym pf) (REG_dec_lb x y pf'))))
61+
eq_refl.
62+
3563
Infix "=?" := REG_beq : REG_scope.
3664

3765
Global Instance REG_beq_spec : reflect_rel (@eq REG) REG_beq | 10
38-
:= reflect_of_beq internal_REG_dec_bl internal_REG_dec_lb.
39-
Definition REG_beq_eq x y : (x =? y)%REG = true <-> x = y := conj (@internal_REG_dec_bl _ _) (@internal_REG_dec_lb _ _).
66+
:= reflect_of_beq REG_dec_bl REG_dec_lb.
67+
Definition REG_beq_eq x y : (x =? y)%REG = true <-> x = y := conj (@REG_dec_bl _ _) (@REG_dec_lb _ _).
4068
Lemma REG_beq_neq x y : (x =? y)%REG = false <-> x <> y.
4169
Proof. rewrite <- REG_beq_eq; destruct (x =? y)%REG; intuition congruence. Qed.
4270
Global Instance REG_beq_compat : Proper (eq ==> eq ==> eq) REG_beq | 10.

src/Assembly/Equivalence.v

+3-2
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ From Coq Require Import ZArith.
44
From Coq Require Import NArith.
55
Require Import Crypto.Assembly.Syntax.
66
Require Import Crypto.Assembly.Parse.
7+
Require Import Crypto.Assembly.Equality.
78
Require Import Crypto.Assembly.Symbolic.
89
Require Import Crypto.Util.Strings.Parse.Common.
910
Require Import Crypto.Util.ErrorT.
@@ -277,8 +278,8 @@ Definition show_annotated_Line : Show AnnotatedLine
277278
end)%string.
278279

279280
Global Instance show_lines_AnnotatedLines : ShowLines AnnotatedLines
280-
:= fun '(ls, ss)
281-
=> let d := dag.eager.force ss.(dag_state) in
281+
:= fun '(ls, sst)
282+
=> let d := dag.eager.force sst.(dag_state) in
282283
List.map (fun l => show_annotated_Line (l, d)) ls.
283284

284285
Fixpoint remove_common_indices {T} (eqb : T -> T -> bool) (xs ys : list T) (start_idx : nat) : list (nat * T) * list T

src/Assembly/EquivalenceProofs.v

+1-1
Original file line numberDiff line numberDiff line change
@@ -63,7 +63,7 @@ Local Open Scope list_scope.
6363

6464
(* TODO: move to global settings *)
6565
Local Set Keyed Unification.
66-
66+
Locate set_reg.
6767
Definition eval_idx_Z (G : symbol -> option Z) (dag : dag) (idx : idx) (v : Z) : Prop
6868
:= eval G dag (ExprRef idx) v.
6969
Definition eval_idx_or_list_idx (G : symbol -> option Z) (d : dag) (v1 : idx + list idx) (v2 : Z + list Z)

src/Assembly/Parse.v

+42-14
Original file line numberDiff line numberDiff line change
@@ -100,16 +100,26 @@ Definition parse_label : ParserAction string
100100
(fun '(char, ls) => string_of_list_ascii (char :: ls))
101101
(([a-zA-Z] || parse_any_ascii "._?$") ;;
102102
(([a-zA-Z] || parse_any_ascii "0123456789_$#@~.?")* )).
103+
Definition parse_non_access_size_label : ParserAction string
104+
:= parse_lookahead_not parse_AccessSize ;;R parse_label.
103105

104106
Definition parse_MEM : ParserAction MEM
105-
:= parse_map
106-
(fun '(access_size, (br (*base reg*), sr (*scale reg, including z *), offset, base_label))
107-
=> {| mem_bits_access_size := access_size:option AccessSize
108-
; mem_base_reg := br:option REG
109-
; mem_base_label := base_label
110-
; mem_scale_reg := sr:option (Z * REG)
111-
; mem_offset := offset:option Z |})
107+
:= parse_option_list_map
108+
(fun '(access_size, (constant_location_label, (br (*base reg*), sr (*scale reg, including z *), offset, base_label)))
109+
=> match base_label, constant_location_label with
110+
| Some _, Some _ => (* invalid? *) None
111+
| Some _ as lbl, None
112+
| None, Some _ as lbl
113+
| None, None as lbl =>
114+
Some
115+
{| mem_bits_access_size := access_size:option AccessSize
116+
; mem_base_reg := br:option REG
117+
; mem_base_label := lbl
118+
; mem_scale_reg := sr:option (Z * REG)
119+
; mem_offset := offset:option Z |}
120+
end)
112121
(((strip_whitespace_after parse_AccessSize)?) ;;
122+
(parse_non_access_size_label?) ;;
113123
(parse_option_list_map
114124
(fun '(offset, vars)
115125
=> (vars <-- List.map (fun '(c, (v, e), vs) => match vs, e with [], 1%Z => Some (c, v) | _, _ => None end) vars;
@@ -160,7 +170,8 @@ Definition parse_OpCode_list : list (string * OpCode)
160170
:= Eval vm_compute in
161171
List.map
162172
(fun r => (show r, r))
163-
(list_all OpCode).
173+
(list_all OpCode)
174+
++ [(".quad", dq); (".word", dw); (".byte", db)].
164175

165176
Definition parse_OpCode : ParserAction OpCode
166177
:= parse_strs_case_insensitive parse_OpCode_list.
@@ -254,7 +265,14 @@ Global Instance show_lvl_MEM : ShowLevel MEM
254265
:= fun m
255266
=> (match m.(mem_bits_access_size) with
256267
| Some n
257-
=> show_lvl_app (fun 'tt => if n =? 8 then "byte" else if n =? 64 then "QWORD PTR" else "BAD SIZE")%N (* TODO: Fix casing and stuff *)
268+
=> show_lvl_app (fun 'tt => if n =? 8 then "byte"
269+
else if n =? 16 then "word"
270+
else if n =? 32 then "dword"
271+
else if n =? 64 then "QWORD PTR"
272+
else if n =? 128 then "XMMWORD PTR"
273+
else if n =? 256 then "YMMWORD PTR"
274+
else if n =? 512 then "ZMMWORD PTR"
275+
else "BAD SIZE")%N (* TODO: Fix casing and stuff *)
258276
| None => show_lvl
259277
end)
260278
(fun 'tt
@@ -275,11 +293,21 @@ Global Instance show_lvl_MEM : ShowLevel MEM
275293
then "0x08 * " ++ Decimal.show_Z (offset / 8)
276294
else Hex.show_Z offset)
277295
end%Z) in
278-
"[" ++ match m.(mem_base_label) with
279-
| None => reg_part ++ offset_part
280-
| Some l => "((" ++ l ++ offset_part ++ "))"
281-
end
282-
++ "]").
296+
match m.(mem_base_label), m.(mem_base_reg), m.(mem_offset), m.(mem_scale_reg) with
297+
| Some lbl, Some rip, None, None => lbl ++ "[" ++ reg_part ++ offset_part ++ "]"
298+
| Some lbl, _, _, _ => let l_offset := lbl ++ offset_part in
299+
"[" ++
300+
(if reg_part =? ""
301+
then "((" ++ l_offset ++ "))"
302+
else reg_part ++ " + " ++ l_offset)
303+
++ "]"
304+
| None, _, _, _ =>
305+
"[" ++
306+
(if reg_part =? ""
307+
then "((" ++ offset_part ++ "))"
308+
else reg_part ++ " + " ++ offset_part)
309+
++ "]"
310+
end).
283311
Global Instance show_MEM : Show MEM := show_lvl_MEM.
284312

285313
Global Instance show_lvl_JUMP_LABEL : ShowLevel JUMP_LABEL

0 commit comments

Comments
 (0)