aboutsummaryrefslogtreecommitdiffstats
path: root/binop.v
diff options
context:
space:
mode:
authorLibravatar Rutger Broekhoff2024-06-26 20:50:18 +0200
committerLibravatar Rutger Broekhoff2024-06-26 20:50:18 +0200
commit73df1945b31c0beee88cf4476df4ccd09d31403b (patch)
treeed00db26b711442e643f38b66888a3df56e33ebd /binop.v
downloadmininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.tar.gz
mininix-formalization-73df1945b31c0beee88cf4476df4ccd09d31403b.zip
Import Coq project
Diffstat (limited to 'binop.v')
-rw-r--r--binop.v87
1 files changed, 87 insertions, 0 deletions
diff --git a/binop.v b/binop.v
new file mode 100644
index 0000000..c8f87fd
--- /dev/null
+++ b/binop.v
@@ -0,0 +1,87 @@
1Require Import Coq.NArith.BinNat.
2Require Import Coq.Strings.Ascii.
3Require Import Coq.Strings.String.
4From stdpp Require Import fin_sets gmap option.
5From mininix Require Import expr maptools shared.
6
7Open Scope string_scope.
8Open Scope N_scope.
9Open Scope Z_scope.
10Open Scope nat_scope.
11
12Definition right_union {A} (bs1 bs2 : gmap string A) : gmap string A :=
13 bs2 ∪ bs1.
14
15Variant binop_r : op → value → value → value → Prop :=
16 | Bo_PlusInt n1 n2 : binop_r O_Plus (V_Int n1) (V_Int n2) (V_Int (n1 + n2))
17 | Bo_PlusStr s1 s2 : binop_r O_Plus (V_Str s1) (V_Str s2) (V_Str (s1 ++ s2))
18 | Bo_MinInt n1 n2 : binop_r O_Min (V_Int n1) (V_Int n2) (V_Int (n1 - n2))
19 | Bo_UpdAttrset bs1 bs2 :
20 binop_r O_Upd (V_Attrset bs1) (V_Attrset bs2)
21 (V_Attrset (right_union bs1 bs2))
22 | Bo_Eq v1 v2 : binop_r O_Eq v1 v2 (expr_eq v1 v2)
23 | Bo_DivInt (n1 n2 : Z) : n2 ≠ 0%Z → binop_r O_Div n1 n2 (n1 / n2)%Z
24 | Bo_LtInt (n1 n2 : Z) : binop_r O_Lt n1 n2 (bool_decide (n1 < n2)%Z)
25 | Bo_LtStr s1 s2 : binop_r O_Lt (V_Str s1) (V_Str s2) (string_lt s1 s2).
26
27Notation "u1 '⟦' op '⟧' u2 '-⊚->' v" := (binop_r op u1 u2 v) (at level 55).
28
29Definition binop_eval (op : op) (v1 v2 : value) : option value :=
30 match op with
31 | O_Plus =>
32 match v1, v2 with
33 | V_Int n1, V_Int n2 => Some (V_Int (n1 + n2))
34 | V_Str s1, V_Str s2 => Some (V_Str (s1 ++ s2))
35 | _, _ => None
36 end
37 | O_Min =>
38 match v1, v2 with
39 | V_Int n1, V_Int n2 => Some (V_Int (n1 - n2))
40 | _, _ => None
41 end
42 | O_Upd =>
43 match v1, v2 with
44 | V_Attrset bs1, V_Attrset bs2 =>
45 Some (V_Attrset (right_union bs1 bs2))
46 | _, _ => None
47 end
48 | O_Eq => Some (V_Bool (expr_eq (X_V v1) (X_V v2)))
49 | O_Div =>
50 match v1, v2 with
51 | V_Int n1, V_Int n2 =>
52 if decide (n2 = 0)%Z
53 then None
54 else Some (V_Int (n1 / n2)%Z)
55 | _, _ => None
56 end
57 | O_Lt =>
58 match v1, v2 with
59 | V_Int n1, V_Int n2 => Some (V_Bool (bool_decide (n1 < n2)%Z))
60 | V_Str s1, V_Str s2 => Some (V_Bool (string_lt s1 s2))
61 | _, _ => None
62 end
63 end.
64
65Theorem binop_eval_sound op u1 u2 v :
66 binop_eval op u1 u2 = Some v → binop_r op u1 u2 v.
67Proof.
68 intros Heval.
69 destruct op, u1, u2; try discriminate; simplify_eq/=; try constructor.
70 destruct (decide (n0 = 0)%Z); try discriminate.
71 simplify_eq/=. by constructor.
72Qed.
73
74Theorem binop_eval_complete op u1 u2 v :
75 binop_r op u1 u2 v → binop_eval op u1 u2 = Some v.
76Proof.
77 intros Heval. inv Heval; try done.
78 unfold binop_eval. by apply decide_False.
79Qed.
80
81Theorem binop_deterministic op u1 u2 v1 v2 :
82 u1 ⟦op⟧ u2 -⊚-> v1 → u1 ⟦op⟧ u2 -⊚-> v2 → v1 = v2.
83Proof.
84 intros Heval1 Heval2.
85 apply binop_eval_complete in Heval1, Heval2.
86 congruence.
87Qed.