diff options
author | Rutger Broekhoff | 2025-07-07 21:52:08 +0200 |
---|---|---|
committer | Rutger Broekhoff | 2025-07-07 21:52:08 +0200 |
commit | ba61dfd69504ec6263a9dee9931d93adeb6f3142 (patch) | |
tree | d6c9b78e50eeab24e0c1c09ab45909a6ae3fd5db /theories/evallang/operational_props.v | |
download | verified-dyn-lang-interp-ba61dfd69504ec6263a9dee9931d93adeb6f3142.tar.gz verified-dyn-lang-interp-ba61dfd69504ec6263a9dee9931d93adeb6f3142.zip |
Initialize repository
Diffstat (limited to 'theories/evallang/operational_props.v')
-rw-r--r-- | theories/evallang/operational_props.v | 33 |
1 files changed, 33 insertions, 0 deletions
diff --git a/theories/evallang/operational_props.v b/theories/evallang/operational_props.v new file mode 100644 index 0000000..31724c0 --- /dev/null +++ b/theories/evallang/operational_props.v | |||
@@ -0,0 +1,33 @@ | |||
1 | From mininix Require Export evallang.operational. | ||
2 | From stdpp Require Import options. | ||
3 | |||
4 | Module Import evallang. | ||
5 | Export evallang. | ||
6 | |||
7 | (** Properties of operational semantics *) | ||
8 | Lemma step_not_final e1 e2 : e1 --> e2 → ¬final e1. | ||
9 | Proof. induction 1; simpl; repeat case_match; naive_solver. Qed. | ||
10 | Lemma final_nf e : final e → nf step e. | ||
11 | Proof. by intros ? [??%step_not_final]. Qed. | ||
12 | |||
13 | Lemma SAbsL_rtc ex1 ex1' e : ex1 -->* ex1' → EAbs ex1 e -->* EAbs ex1' e. | ||
14 | Proof. induction 1; econstructor; eauto using step. Qed. | ||
15 | Lemma SAppL_rtc e1 e1' e2 : e1 -->* e1' → EApp e1 e2 -->* EApp e1' e2. | ||
16 | Proof. induction 1; econstructor; eauto using step. Qed. | ||
17 | Lemma SEval_rtc ds e1 e1' : e1 -->* e1' → EEval ds e1 -->* EEval ds e1'. | ||
18 | Proof. induction 1; econstructor; eauto using step. Qed. | ||
19 | |||
20 | Ltac inv_step := repeat | ||
21 | match goal with H : ?e --> _ |- _ => assert_fails (is_var e); inv H end. | ||
22 | |||
23 | Lemma step_det e d1 d2 : | ||
24 | e --> d1 → | ||
25 | e --> d2 → | ||
26 | d1 = d2. | ||
27 | Proof. | ||
28 | intros Hred1. revert d2. | ||
29 | induction Hred1; intros d2 Hred2; inv Hred2; try by inv_step; | ||
30 | f_equal; by apply IHHred1. | ||
31 | Qed. | ||
32 | |||
33 | End evallang. | ||