Skip to content

Commit

Permalink
Complete wasm runtime EJson compare
Browse files Browse the repository at this point in the history
  • Loading branch information
pkel committed Jun 1, 2021
1 parent ad0786e commit cd8c746
Show file tree
Hide file tree
Showing 3 changed files with 83 additions and 23 deletions.
30 changes: 27 additions & 3 deletions runtimes/assemblyscript/assembly/index.ts
Original file line number Diff line number Diff line change
Expand Up @@ -620,13 +620,37 @@ function compare(a: EjValue, b: EjValue): i32 {
return compare_base<i64>(aa.value, bb.value);
}
if (a instanceof EjString && b instanceof EjString) {
throw new Error("compare: does not support String");
let aa : EjString = changetype<EjString>(a) ;
let bb : EjString = changetype<EjString>(b) ;
return compare_base<string>(aa.value, bb.value);
}
if (a instanceof EjArray && b instanceof EjArray) {
throw new Error("compare: does not support Array");
let va = changetype<EjArray>(a).values.slice() ;
let vb = changetype<EjArray>(b).values.slice() ;
va.sort(compare);
vb.sort(compare);
let i = 0;
while (i < va.length && i < vb.length) {
let c = compare(va[i], vb[i]);
if (c != 0) { return c; };
i++;
}
return compare_base<i32>(va.length, vb.length);
}
if (a instanceof EjObject && b instanceof EjObject) {
throw new Error("compare: does not support Object");
let oa = changetype<EjObject>(a).values ;
let ob = changetype<EjObject>(b).values ;
let ka = oa.keys().sort();
let kb = ob.keys().sort();
let i = 0
while (i < ka.length && i < kb.length) {
let c = compare_base<string>(ka[i], kb[i]);
if (c != 0) { return c; };
c = compare(oa[ka[i]], ob[kb[i]]);
if (c != 0) { return c; };
i++;
}
return compare_base<i32>(ka.length, kb.length);
}
return compare_base<u32>(type_id(a), type_id(b));
}
Expand Down
Binary file modified runtimes/assemblyscript/runtime.wasm
Binary file not shown.
76 changes: 56 additions & 20 deletions tests/operators/main.ml
Original file line number Diff line number Diff line change
Expand Up @@ -131,6 +131,11 @@ module Arg = struct
Coq_ejobject (List.map (fun (k, v) -> Util.char_list_of_string k, v) x)
end

let sort_array (x : _ ejson) =
match x with
| Coq_ejarray z -> Coq_ejarray (List.sort compare z)
| z -> z

let _ =
let open Arg in
test_op
Expand Down Expand Up @@ -241,16 +246,52 @@ let _ =
EJsonOpMathTrunc
[ [ num 1.1 ]
; [ num 10000.7 ]
; [ num (-1.)]
; [ num (-3.14)]
; [ num (-3.94)]
; [ num (-1000000.94)]
];
test_rtop
EJsonRuntimeEqual
[ [ bool false; bool true ]
; [ bool true; bool true]
; [ num (-1.) ]
; [ num (-3.14) ]
; [ num (-3.94) ]
; [ num (-1000000.94) ]
];
let compare_values =
[ null; bool true; bool false
; int 0; int 1; num 0.; num 1.
; str "hello world"
; str ""
; arr [ null ]
; arr [bool false]
; arr []
; obj [ "a" , null ]
; obj [ "a" , bool true ]
]
in
let compare_pairs =
List.map (fun a ->
List.map (fun b ->
[ a; b ]
) compare_values
) compare_values
|> List.concat
in
test_rtop EJsonRuntimeEqual compare_pairs;
test_rtop EJsonRuntimeEqual
[ [ arr [int 0; int 1]; arr [int 1; int 0] ]
];
test_rtop
EJsonRuntimeCompare (* Wasm and JS compare support arguments of different
type. This is used to speed up set operations.
Coq/Imp implements set operations naively and only
compares integers and numbers. *)
[ [ int 0; int 1]
; [ int 1; int 0]
; [ int 1; int 1]
; [ num 0.; num 1.]
; [ num 1.; num 0.]
; [ num 1.; num 1.]
];
test_rtop ~fixup:sort_array
EJsonRuntimeMax (* ... we trigger the wasm compare on heterogeneous types
using set union. *)
(List.map (fun x -> [arr x; arr (List.rev x)]) compare_pairs)
;
test_rtop
EJsonRuntimeToString
[ [ null ]
Expand Down Expand Up @@ -316,7 +357,7 @@ let _ =
[ [ obj ["a", null]; str "a" ]
; [ obj ["a", int 0; "b", int 1; "c", int 2]; str "a" ]
; [ obj ["b", int 0; "a", int 1; "c", int 2]; str "a" ]
(* ; [ obj []; str "a"] invalid / missing key *)
(* ; [ obj []; str "a"] invalid / missing key *)
];
test_rtop
EJsonRuntimeArray
Expand All @@ -341,7 +382,7 @@ let _ =
; [ arr [int 0; int 1; int 2]; int 0 ]
; [ arr [int 0; int 1; int 2]; int 1 ]
; [ arr [int 0; int 1; int 2]; int 2 ]
(* ; [ arr [null]; int 1 ] (* invalid / out of bounds *) *)
(* ; [ arr [null]; int 1 ] (* invalid / out of bounds *) *)
];
test_rtop
EJsonRuntimeUnbrand
Expand Down Expand Up @@ -400,11 +441,6 @@ let _ =
; [ arr [null]; arr [] ]
; [ arr []; arr [] ]
];
let sort_array (x : _ ejson) =
match x with
| Coq_ejarray z -> Coq_ejarray (List.sort compare z)
| z -> z
in
(* ImpEJson eval does not sort the arrays.
* JS runtime and WASM runtime do sort the arrays. *)
test_rtop ~fixup:sort_array
Expand Down Expand Up @@ -459,7 +495,7 @@ let _ =
[ [ str "" ]
; [ str "a" ]
; [ str "abcc" ]
(* ; [ str "abccᵤ" ] *) (* Coq uses char list, JS/Wasm use UTF16. *)
(* ; [ str "abccᵤ" ] *) (* Coq uses char list, JS/Wasm use UTF16. *)
];
test_rtop
EJsonRuntimeSubstring (* string, start, target length *)
Expand Down Expand Up @@ -549,9 +585,9 @@ let _ =
; [int 43]
; [int (1 lsl 20)]
; [int (1 lsl 61)]
(* ; [int (-1)] (* Imp returns 0 *) *)
(* ; [int max_int] (* Wasm off by one *) *)
(* ; [int min_int] (* Imp returns 0 *) *)
(* ; [int (-1)] (* Imp returns 0 *) *)
(* ; [int max_int] (* Wasm off by one *) *)
(* ; [int min_int] (* Imp returns 0 *) *)
];
test_rtop
EJsonRuntimeNatSqrt
Expand Down

0 comments on commit cd8c746

Please sign in to comment.