From 6e697cccab5824e4baff4531b71ad5228eb217bc Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Fri, 9 Dec 2016 13:35:35 +0100
Subject: [PATCH] Merge heap_lang/derived with other files.

---
 _CoqProject            |  1 -
 heap_lang/derived.v    | 77 ------------------------------------------
 heap_lang/lang.v       | 10 ++++++
 heap_lang/notation.v   |  4 +--
 heap_lang/rules.v      | 59 ++++++++++++++++++++++++++++++++
 heap_lang/wp_tactics.v |  2 +-
 6 files changed, 72 insertions(+), 81 deletions(-)
 delete mode 100644 heap_lang/derived.v

diff --git a/_CoqProject b/_CoqProject
index 6c16a3342..11ce1cd1a 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -98,7 +98,6 @@ heap_lang/lang.v
 heap_lang/tactics.v
 heap_lang/wp_tactics.v
 heap_lang/rules.v
-heap_lang/derived.v
 heap_lang/notation.v
 heap_lang/lib/spawn.v
 heap_lang/lib/par.v
diff --git a/heap_lang/derived.v b/heap_lang/derived.v
deleted file mode 100644
index f327ef1f7..000000000
--- a/heap_lang/derived.v
+++ /dev/null
@@ -1,77 +0,0 @@
-From iris.heap_lang Require Export rules.
-Import uPred.
-
-(** Define some derived forms, and derived lemmas about them. *)
-Notation Lam x e := (Rec BAnon x e).
-Notation Let x e1 e2 := (App (Lam x e2) e1).
-Notation Seq e1 e2 := (Let BAnon e1 e2).
-Notation LamV x e := (RecV BAnon x e).
-Notation LetCtx x e2 := (AppRCtx (LamV x e2)).
-Notation SeqCtx e2 := (LetCtx BAnon e2).
-Notation Skip := (Seq (Lit LitUnit) (Lit LitUnit)).
-Notation Match e0 x1 e1 x2 e2 := (Case e0 (Lam x1 e1) (Lam x2 e2)).
-
-Section derived.
-Context `{heapG Σ}.
-Implicit Types P Q : iProp Σ.
-Implicit Types Φ : val → iProp Σ.
-
-(** Proof rules for the sugar *)
-Lemma wp_lam E x elam e1 e2 Φ :
-  e1 = Lam x elam →
-  is_Some (to_val e2) →
-  Closed (x :b: []) elam →
-  ▷ WP subst' x e2 elam @ E {{ Φ }} ⊢ WP App e1 e2 @ E {{ Φ }}.
-Proof. intros. by rewrite -(wp_rec _ BAnon) //. Qed.
-
-Lemma wp_let E x e1 e2 Φ :
-  is_Some (to_val e1) → Closed (x :b: []) e2 →
-  ▷ WP subst' x e1 e2 @ E {{ Φ }} ⊢ WP Let x e1 e2 @ E {{ Φ }}.
-Proof. by apply wp_lam. Qed.
-
-Lemma wp_seq E e1 e2 Φ :
-  is_Some (to_val e1) → Closed [] e2 →
-  ▷ WP e2 @ E {{ Φ }} ⊢ WP Seq e1 e2 @ E {{ Φ }}.
-Proof. intros ??. by rewrite -wp_let. Qed.
-
-Lemma wp_skip E Φ : ▷ Φ (LitV LitUnit) ⊢ WP Skip @ E {{ Φ }}.
-Proof. rewrite -wp_seq; last eauto. by rewrite -wp_value. Qed.
-
-Lemma wp_match_inl E e0 x1 e1 x2 e2 Φ :
-  is_Some (to_val e0) → Closed (x1 :b: []) e1 →
-  ▷ WP subst' x1 e0 e1 @ E {{ Φ }} ⊢ WP Match (InjL e0) x1 e1 x2 e2 @ E {{ Φ }}.
-Proof. intros. by rewrite -wp_case_inl // -[X in _ ⊢ X]later_intro -wp_let. Qed.
-
-Lemma wp_match_inr E e0 x1 e1 x2 e2 Φ :
-  is_Some (to_val e0) → Closed (x2 :b: []) e2 →
-  ▷ WP subst' x2 e0 e2 @ E {{ Φ }} ⊢ WP Match (InjR e0) x1 e1 x2 e2 @ E {{ Φ }}.
-Proof. intros. by rewrite -wp_case_inr // -[X in _ ⊢ X]later_intro -wp_let. Qed.
-
-Lemma wp_le E (n1 n2 : Z) P Φ :
-  (n1 ≤ n2 → P ⊢ ▷ Φ (LitV (LitBool true))) →
-  (n2 < n1 → P ⊢ ▷ Φ (LitV (LitBool false))) →
-  P ⊢ WP BinOp LeOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
-Proof.
-  intros. rewrite -wp_bin_op //; [].
-  destruct (bool_decide_reflect (n1 ≤ n2)); by eauto with omega.
-Qed.
-
-Lemma wp_lt E (n1 n2 : Z) P Φ :
-  (n1 < n2 → P ⊢ ▷ Φ (LitV (LitBool true))) →
-  (n2 ≤ n1 → P ⊢ ▷ Φ (LitV (LitBool false))) →
-  P ⊢ WP BinOp LtOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
-Proof.
-  intros. rewrite -wp_bin_op //; [].
-  destruct (bool_decide_reflect (n1 < n2)); by eauto with omega.
-Qed.
-
-Lemma wp_eq E e1 e2 v1 v2 P Φ :
-  to_val e1 = Some v1 → to_val e2 = Some v2 →
-  (v1 = v2 → P ⊢ ▷ Φ (LitV (LitBool true))) →
-  (v1 ≠ v2 → P ⊢ ▷ Φ (LitV (LitBool false))) →
-  P ⊢ WP BinOp EqOp e1 e2 @ E {{ Φ }}.
-Proof.
-  intros. rewrite -wp_bin_op //; [].
-  destruct (bool_decide_reflect (v1 = v2)); by eauto.
-Qed.
-End derived.
diff --git a/heap_lang/lang.v b/heap_lang/lang.v
index 3ea46c969..4ef9b4c8d 100644
--- a/heap_lang/lang.v
+++ b/heap_lang/lang.v
@@ -402,3 +402,13 @@ Canonical Structure heap_lang := ectx_lang (heap_lang.expr).
 
 (* Prefer heap_lang names over ectx_language names. *)
 Export heap_lang.
+
+(** Define some derived forms *)
+Notation Lam x e := (Rec BAnon x e).
+Notation Let x e1 e2 := (App (Lam x e2) e1).
+Notation Seq e1 e2 := (Let BAnon e1 e2).
+Notation LamV x e := (RecV BAnon x e).
+Notation LetCtx x e2 := (AppRCtx (LamV x e2)).
+Notation SeqCtx e2 := (LetCtx BAnon e2).
+Notation Skip := (Seq (Lit LitUnit) (Lit LitUnit)).
+Notation Match e0 x1 e1 x2 e2 := (Case e0 (Lam x1 e1) (Lam x2 e2)).
diff --git a/heap_lang/notation.v b/heap_lang/notation.v
index bd2bb6ff4..9406c5c0c 100644
--- a/heap_lang/notation.v
+++ b/heap_lang/notation.v
@@ -1,5 +1,5 @@
-From iris.heap_lang Require Export derived.
-Export heap_lang.
+From iris.program_logic Require Import language.
+From iris.heap_lang Require Export lang tactics.
 
 Coercion LitInt : Z >-> base_lit.
 Coercion LitBool : bool >-> base_lit.
diff --git a/heap_lang/rules.v b/heap_lang/rules.v
index efeb15ab3..4aac38ac2 100644
--- a/heap_lang/rules.v
+++ b/heap_lang/rules.v
@@ -214,4 +214,63 @@ Proof.
   iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
   iModIntro. iSplit=>//. by iApply "HΦ".
 Qed.
+
+(** Proof rules for derived constructs *)
+Lemma wp_lam E x elam e1 e2 Φ :
+  e1 = Lam x elam →
+  is_Some (to_val e2) →
+  Closed (x :b: []) elam →
+  ▷ WP subst' x e2 elam @ E {{ Φ }} ⊢ WP App e1 e2 @ E {{ Φ }}.
+Proof. intros. by rewrite -(wp_rec _ BAnon) //. Qed.
+
+Lemma wp_let E x e1 e2 Φ :
+  is_Some (to_val e1) → Closed (x :b: []) e2 →
+  ▷ WP subst' x e1 e2 @ E {{ Φ }} ⊢ WP Let x e1 e2 @ E {{ Φ }}.
+Proof. by apply wp_lam. Qed.
+
+Lemma wp_seq E e1 e2 Φ :
+  is_Some (to_val e1) → Closed [] e2 →
+  ▷ WP e2 @ E {{ Φ }} ⊢ WP Seq e1 e2 @ E {{ Φ }}.
+Proof. intros ??. by rewrite -wp_let. Qed.
+
+Lemma wp_skip E Φ : ▷ Φ (LitV LitUnit) ⊢ WP Skip @ E {{ Φ }}.
+Proof. rewrite -wp_seq; last eauto. by rewrite -wp_value. Qed.
+
+Lemma wp_match_inl E e0 x1 e1 x2 e2 Φ :
+  is_Some (to_val e0) → Closed (x1 :b: []) e1 →
+  ▷ WP subst' x1 e0 e1 @ E {{ Φ }} ⊢ WP Match (InjL e0) x1 e1 x2 e2 @ E {{ Φ }}.
+Proof. intros. by rewrite -wp_case_inl // -[X in _ ⊢ X]later_intro -wp_let. Qed.
+
+Lemma wp_match_inr E e0 x1 e1 x2 e2 Φ :
+  is_Some (to_val e0) → Closed (x2 :b: []) e2 →
+  ▷ WP subst' x2 e0 e2 @ E {{ Φ }} ⊢ WP Match (InjR e0) x1 e1 x2 e2 @ E {{ Φ }}.
+Proof. intros. by rewrite -wp_case_inr // -[X in _ ⊢ X]later_intro -wp_let. Qed.
+
+Lemma wp_le E (n1 n2 : Z) P Φ :
+  (n1 ≤ n2 → P ⊢ ▷ Φ (LitV (LitBool true))) →
+  (n2 < n1 → P ⊢ ▷ Φ (LitV (LitBool false))) →
+  P ⊢ WP BinOp LeOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
+Proof.
+  intros. rewrite -wp_bin_op //; [].
+  destruct (bool_decide_reflect (n1 ≤ n2)); by eauto with omega.
+Qed.
+
+Lemma wp_lt E (n1 n2 : Z) P Φ :
+  (n1 < n2 → P ⊢ ▷ Φ (LitV (LitBool true))) →
+  (n2 ≤ n1 → P ⊢ ▷ Φ (LitV (LitBool false))) →
+  P ⊢ WP BinOp LtOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
+Proof.
+  intros. rewrite -wp_bin_op //; [].
+  destruct (bool_decide_reflect (n1 < n2)); by eauto with omega.
+Qed.
+
+Lemma wp_eq E e1 e2 v1 v2 P Φ :
+  to_val e1 = Some v1 → to_val e2 = Some v2 →
+  (v1 = v2 → P ⊢ ▷ Φ (LitV (LitBool true))) →
+  (v1 ≠ v2 → P ⊢ ▷ Φ (LitV (LitBool false))) →
+  P ⊢ WP BinOp EqOp e1 e2 @ E {{ Φ }}.
+Proof.
+  intros. rewrite -wp_bin_op //; [].
+  destruct (bool_decide_reflect (v1 = v2)); by eauto.
+Qed.
 End rules.
diff --git a/heap_lang/wp_tactics.v b/heap_lang/wp_tactics.v
index 65fff20b5..85c1e30b9 100644
--- a/heap_lang/wp_tactics.v
+++ b/heap_lang/wp_tactics.v
@@ -1,4 +1,4 @@
-From iris.heap_lang Require Export tactics derived.
+From iris.heap_lang Require Export tactics rules.
 Import uPred.
 
 (** wp-specific helper tactics *)
-- 
GitLab