Skip to content

Commit e57e990

Browse files
committed
simulations: add simulation of the add instruction
1 parent e77b5ff commit e57e990

File tree

4 files changed

+51
-75
lines changed

4 files changed

+51
-75
lines changed

CoqOfRust/links/M.v

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1038,14 +1038,14 @@ Module LowM.
10381038
Inductive t (R Output : Set) : Set :=
10391039
| Pure (value : Output.t R Output)
10401040
| CallPrimitive {A : Set} (primitive : Primitive.t A) (k : A -> t R Output)
1041-
| Call {A : Set} (e : t A A) (k : SuccessOrPanic.t A -> t R Output)
1041+
| Call {A : Set} `{Link A} {e : M} (run : {{ e 🔽 A, A }}) (k : SuccessOrPanic.t A -> t R Output)
10421042
| LetAlloc {A : Set} `{Link A}
10431043
(e : t R A)
10441044
(k : Output.t R (Ref.t Pointer.Kind.Raw A) -> t R Output)
10451045
| Loop {A : Set} (body : t R A) (k : Output.t R A -> t R Output).
10461046
Arguments Pure {_ _}.
10471047
Arguments CallPrimitive {_ _ _}.
1048-
Arguments Call {_ _ _}.
1048+
Arguments Call {_ _ _ _ _}.
10491049
Arguments LetAlloc {_ _ _ _}.
10501050
Arguments Loop {_ _ _}.
10511051
End LowM.
@@ -1144,7 +1144,7 @@ Proof.
11441144
}
11451145
{ (* CallClosure *)
11461146
eapply (LowM.Call (A := Output')). {
1147-
exact (evaluate _ _ _ _ _ run).
1147+
exact run.
11481148
}
11491149
intros output'; eapply evaluate.
11501150
match goal with

CoqOfRust/revm/revm_interpreter/instructions/simulations/arithmetic.v

Lines changed: 4 additions & 63 deletions
Original file line numberDiff line numberDiff line change
@@ -6,8 +6,10 @@ Require Import revm.revm_interpreter.links.interpreter_types.
66
Require Import revm.revm_interpreter.simulations.gas.
77
Require Import revm.revm_interpreter.simulations.interpreter_types.
88
Require Import revm.revm_interpreter.instructions.links.arithmetic.
9+
Require Import ruint.simulations.add.
910

1011
Import Impl_Gas.
12+
Import add.Impl_Uint.
1113

1214
Instance run_add Stack
1315
{WIRE H : Set} `{Link WIRE} `{Link H}
@@ -20,69 +22,8 @@ Proof.
2022
constructor.
2123
destruct run_InterpreterTypes_for_WIRE.
2224
destruct run_StackTrait_for_Stack, run_LoopControl_for_Control.
23-
(* destruct popn_top as [? []].
24-
destruct gas as [? []].
25-
destruct set_instruction_result as [? []]. *)
26-
Time simulate.
27-
{ Time simulate_destruct; simulate.
28-
match goal with
29-
| H : _ |- Run.Trait ?Stack (_ ?x1 ?x2) =>
30-
apply (H Stack x1 x2)
31-
end.
32-
}
33-
progress cbn in *.
34-
idtac.
35-
(* destruct link; cbn in *.
36-
destruct run_StackTrait_for_Stack, run_LoopControl_for_Control; cbn in *. *)
37-
generalize set_instruction_result; clear; intros.
38-
match goal with
39-
| |- Run.Trait _ (?f _ _) =>
40-
set (f' := f) in |- *
41-
end.
42-
destruct WIRE_types; cbn in *.
43-
destruct link, run_LoopControl_for_Control, set_instruction_result0; cbn in *.
44-
progress fold f' in set_instruction_result.
45-
unfold InterpreterTypes.Types.IsLinkControl in *; cbn in *.
46-
with_strategy transparent [Φ φ] apply set_instruction_result.
47-
progress fold f' in set_instruction_result.
48-
apply set_instruction_result.
49-
match goal with
50-
| H : forall _ _ _, ?f _ _ _ = _ |- _ =>
51-
specialize (H _ _ _ eq_refl)
52-
end.
53-
Check set_instruction_result.
54-
refine (set_instruction_result _ _ _).
55-
best.
56-
match goal with.
57-
dependent apply set_instruction_result.
58-
epose proof (set_instruction_result Stack
59-
(* ((Ref.cast_to Pointer.Kind.MutRef
60-
{|
61-
Ref.core :=
62-
SubPointer.Runner.apply interpreter.(Ref.core)
63-
Interpreter.get_control
64-
|})) *)
65-
(* instruction_result.InstructionResult.OutOfGas *)
66-
).
67-
apply H5.
68-
apply set_instruction_result.
69-
typeclasses eauto.
70-
match goal with
71-
| |- context[if ?condition then _ else _] =>
72-
let condition' := eval cbv in condition in
73-
change condition with condition'
74-
end.
75-
Time match goal with
76-
| |- context[if ?condition then _ else _] =>
77-
destruct condition
78-
end.
79-
{ Show.
80-
simulate_reduce.
81-
82-
}
83-
4: { intros. apply Impl_Gas.run_record_cost.
84-
cbn in output.
85-
Show.
25+
simulate.
26+
Defined.
8627

8728
(*
8829
Require Import revm.interpreter.links.interpreter.

CoqOfRust/ruint/simulations/add.v

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,18 @@
1+
Require Import CoqOfRust.CoqOfRust.
2+
Require Import CoqOfRust.links.M.
3+
Require Import CoqOfRust.simulations.M.
4+
Require Import ruint.links.add.
5+
Require Import ruint.links.lib.
6+
7+
Module Impl_Uint.
8+
(* Uint<BITS, LIMBS> *)
9+
Definition Self (BITS LIMBS : Usize.t) : Set :=
10+
Uint.t BITS LIMBS.
11+
12+
Instance run_wrapping_add Stack
13+
(BITS LIMBS : Usize.t)
14+
(x1 x2 : Self BITS LIMBS) :
15+
Run.Trait Stack
16+
(Impl_Uint.run_wrapping_add BITS LIMBS x1 x2).
17+
Admitted.
18+
End Impl_Uint.

CoqOfRust/simulations/M.v

Lines changed: 26 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -613,19 +613,25 @@ Module Stack.
613613
end
614614
end.
615615

616+
Axiom admit_access_for_now :
617+
forall {A : Set} `{Link A} (Stack : Stack.t) (ref_core : Ref.Core.t A),
618+
t Stack ref_core.
619+
616620
Ltac infer :=
617621
cbn ||
618622
match goal with
619623
| |- t _ (SubPointer.Runner.apply _ ?r) =>
620624
apply (runner r)
621625
end ||
622-
assumption.
626+
assumption ||
627+
apply Stack.CanAccess.Immediate ||
628+
apply admit_access_for_now.
623629
End CanAccess.
624630
End Stack.
625631

626632
(** Here we define an execution mode where we keep dynamic cast to retrieve data from the stack. In
627633
practice, these casts should always be correct as the original Rust code was well typed. *)
628-
Module StackM.
634+
(* Module StackM.
629635
Inductive t (A : Set) : Set :=
630636
| Pure (value : A)
631637
| GetCanAccess {B : Set} `{Link B}
@@ -747,9 +753,8 @@ Module Run.
747753
{{ StackM.GetCanAccess Stack ref_core k 🌲 result }}
748754
749755
where "{{ e 🌲 result }}" := (t result e).
750-
End Run.
756+
End Run. *)
751757

752-
(*
753758
Module Run.
754759
Reserved Notation "{{ StackIn 🌲 e }}".
755760

@@ -927,17 +932,29 @@ Ltac simulate_one_step :=
927932
(apply Run.StateRead; [repeat Stack.CanAccess.infer | intros]) ||
928933
(apply Run.StateWrite; [repeat Stack.CanAccess.infer |]) ||
929934
simulate_get_sub_pointer ||
930-
(apply Run.Call; [unshelve eapply Run.simulation | intros []]) ||
935+
(apply Run.Call; [
936+
(
937+
unshelve eapply Run.simulation;
938+
try match goal with
939+
| H : _ |- Run.Trait ?Stack (_ ?x1) =>
940+
apply (H Stack x1)
941+
| H : _ |- Run.Trait ?Stack (_ ?x1 ?x2) =>
942+
apply (H Stack x1 x2)
943+
| H : _ |- Run.Trait ?Stack (_ ?x1 ?x2 ?x3) =>
944+
apply (H Stack x1 x2 x3)
945+
end
946+
) |
947+
intros []
948+
]) ||
931949
(apply Run.LetAlloc; [|intros []]).
932950

933-
Ltac simulate :=
934-
progress repeat simulate_one_step.
935-
936951
Ltac simulate_destruct :=
937952
match goal with
938953
| |- {{ _ 🌲 ?e }} =>
939954
match e with
940955
| context [match ?e with _ => _ end] => destruct e
941956
end
942957
end.
943-
*)
958+
959+
Ltac simulate :=
960+
progress repeat (simulate_one_step || simulate_destruct).

0 commit comments

Comments
 (0)