├── tutorial ├── slides │ ├── simp_lang.003.png │ ├── simp_lang.005.png │ ├── simp_lang.008.png │ ├── simp_lang.009.png │ ├── simp_lang.010.png │ ├── simp_lang.015.png │ ├── simp_lang.017.png │ ├── simp_lang.018.png │ └── simp_lang.019.png └── README.md ├── .gitignore ├── src ├── simp.v ├── examples │ ├── swap.v │ ├── par.v │ ├── spawn.v │ └── parallel_add.v ├── tactics.v ├── adequacy.v ├── heap_lib.v ├── notation.v ├── class_instances.v ├── primitive_laws.v ├── proofmode.v ├── heap_ra.v └── lang.v ├── _CoqProject ├── Makefile ├── iris-simp-lang.opam ├── LICENSE ├── .github └── workflows │ └── build.yml └── README.md /tutorial/slides/simp_lang.003.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.003.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.005.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.005.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.008.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.008.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.009.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.009.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.010.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.010.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.015.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.015.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.017.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.017.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.018.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.018.png -------------------------------------------------------------------------------- /tutorial/slides/simp_lang.019.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/tchajed/iris-simp-lang/HEAD/tutorial/slides/simp_lang.019.png -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | *.vo 2 | *.vos 3 | *.vok 4 | *.glob 5 | .*.aux 6 | # auto-generated 7 | .coqdeps.d 8 | .Makefile.coq.d 9 | Makefile.coq 10 | Makefile.coq.conf 11 | .lia.cache 12 | -------------------------------------------------------------------------------- /src/simp.v: -------------------------------------------------------------------------------- 1 | (** this file just gather up all the imports needed to work with simp_lang *) 2 | From iris.proofmode Require Export tactics. 3 | From iris_simp_lang Require Export 4 | notation class_instances primitive_laws proofmode. 5 | From iris Require Import options. 6 | Export lang. 7 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -Q src iris_simp_lang 2 | 3 | -arg -w -arg -redundant-canonical-projection 4 | -arg -w -arg -notation-overriden 5 | -arg -w -arg -deprecated-since-8.19 6 | 7 | src/lang.v 8 | src/notation.v 9 | src/tactics.v 10 | src/class_instances.v 11 | src/heap_ra.v 12 | src/heap_lib.v 13 | src/primitive_laws.v 14 | src/proofmode.v 15 | src/simp.v 16 | src/adequacy.v 17 | 18 | src/examples/swap.v 19 | src/examples/spawn.v 20 | src/examples/par.v 21 | src/examples/parallel_add.v 22 | -------------------------------------------------------------------------------- /src/examples/swap.v: -------------------------------------------------------------------------------- 1 | From iris_simp_lang Require Import simp. 2 | From iris Require Import options. 3 | 4 | Section proof. 5 | Context `{!simpGS Σ}. 6 | 7 | Definition swap: val := λ: "x" "y", 8 | let: "tmp" := !"x" in 9 | "x" <- !"y";; 10 | "y" <- "tmp". 11 | 12 | Lemma swap_spec (x y: loc) v1 v2 : 13 | {{{ x ↦ v1 ∗ y ↦ v2 }}} swap #x #y {{{ RET #(); x ↦ v2 ∗ y ↦ v1 }}}. 14 | Proof. 15 | iIntros (Φ) "[Hx Hy] Post". 16 | wp_rec. wp_pures. 17 | wp_load. wp_load. 18 | wp_store. wp_store. 19 | iApply "Post"; iFrame. 20 | Qed. 21 | 22 | End proof. 23 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | ## this Makefile, as well as the test setup in Makefile.coq.local, is copied 2 | ## from std++ (https://gitlab.mpi-sws.org/iris/stdpp) 3 | 4 | # Forward most targets to Coq makefile (with some trick to make this phony) 5 | %: Makefile.coq phony 6 | +@make -f Makefile.coq $@ 7 | 8 | all: Makefile.coq 9 | +@make -f Makefile.coq all 10 | .PHONY: all 11 | 12 | clean: Makefile.coq 13 | +@make -f Makefile.coq clean 14 | find src \( -name "*.d" -o -name "*.vo" -o -name "*.vo[sk]" -o -name "*.aux" -o -name "*.cache" -o -name "*.glob" -o -name "*.vio" \) -print -delete || true 15 | rm -f Makefile.coq .lia.cache 16 | .PHONY: clean 17 | 18 | # Create Coq Makefile. 19 | Makefile.coq: _CoqProject Makefile 20 | "$(COQBIN)coq_makefile" -f _CoqProject -o Makefile.coq 21 | 22 | # Some files that do *not* need to be forwarded to Makefile.coq 23 | Makefile: ; 24 | _CoqProject: ; 25 | 26 | # Phony wildcard targets 27 | phony: ; 28 | .PHONY: phony 29 | -------------------------------------------------------------------------------- /iris-simp-lang.opam: -------------------------------------------------------------------------------- 1 | opam-version: "2.0" 2 | maintainer: "tchajed@gmail.com" 3 | version: "dev" 4 | 5 | homepage: "https://github.com/tchajed/iris-simp-lang" 6 | dev-repo: "git+https://github.com/tchajed/iris-simp-lang.git" 7 | bug-reports: "https://github.com/tchajed/iris-simp-lang/issues" 8 | license: "MIT" 9 | 10 | synopsis: "Simple language to demo instantiating Iris" 11 | description: """ 12 | Iris allows supports a generic language interface, often instantiated using 13 | HeapLang. Here we instantiate that interface with a bare minimum of features, 14 | simplifying compared to HeapLang. 15 | """ 16 | 17 | depends: [ 18 | "coq" {>= "8.20"} 19 | "coq-iris" {>= "dev.2025-02-27.1.c773500a" | (>= "4.3.0" & < "5.0") | = "dev"} 20 | ] 21 | 22 | build: [make "-j%{jobs}%"] 23 | install: [make "install"] 24 | 25 | tags: [ 26 | "category:Miscellaneous" 27 | "keyword:iris" 28 | "logpath:iris_simp_lang" 29 | ] 30 | 31 | authors: [ 32 | "Tej Chajed" 33 | ] 34 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2021 Tej Chajed 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /src/examples/par.v: -------------------------------------------------------------------------------- 1 | From iris.base_logic.lib Require Import invariants. 2 | From iris_simp_lang Require Export simp examples.spawn. 3 | From iris.prelude Require Import options. 4 | Import uPred. 5 | 6 | Definition parN : namespace := nroot .@ "par". 7 | 8 | Definition par : val := 9 | λ: "e1" "e2", 10 | let: "handle" := spawn "e1" in 11 | let: "v2" := "e2" #() in 12 | let: "v1" := join "handle" in 13 | ("v1", "v2"). 14 | Notation "e1 ||| e2" := (par (λ: <>, e1)%E (λ: <>, e2)%E) : expr_scope. 15 | Notation "e1 ||| e2" := (par (λ: <>, e1)%V (λ: <>, e2)%V) : val_scope. 16 | 17 | Section proof. 18 | Context `{!simpGS Σ, !spawnG Σ}. 19 | 20 | Lemma wp_par (Ψ1 Ψ2 : val → iProp Σ) (e1 e2 : expr) (Φ : val → iProp Σ) : 21 | WP e1 {{ Ψ1 }} -∗ WP e2 {{ Ψ2 }} -∗ 22 | (∀ v1 v2, Ψ1 v1 ∗ Ψ2 v2 -∗ ▷ Φ (v1,v2)%V) -∗ 23 | WP (e1 ||| e2)%V {{ Φ }}. 24 | Proof using All. 25 | iIntros "Hf1 Hf2 HΦ". wp_lam. wp_pures. 26 | wp_apply (spawn_spec parN with "[Hf1]"). 27 | { by wp_lam. } 28 | iIntros (l) "Hl". wp_let. wp_pures. wp_bind e2. 29 | wp_apply (wp_wand with "Hf2"); iIntros (v) "H2". wp_let. 30 | wp_apply (join_spec with "[$Hl]"). iIntros (w) "H1". 31 | iSpecialize ("HΦ" with "[$H1 $H2]"). by wp_pures. 32 | Qed. 33 | 34 | End proof. 35 | -------------------------------------------------------------------------------- /.github/workflows/build.yml: -------------------------------------------------------------------------------- 1 | name: CI 2 | 3 | on: 4 | push: 5 | branches: 6 | - main 7 | pull_request: 8 | branches: 9 | - "**" 10 | schedule: 11 | # every Tuesday at 10am UTC (5am EST) 12 | - cron: "0 10 * * tue" 13 | 14 | jobs: 15 | build: 16 | runs-on: ubuntu-latest 17 | strategy: 18 | matrix: 19 | coq: 20 | - "dev" 21 | - "9.0" 22 | - "8.20" 23 | fail-fast: false 24 | steps: 25 | - uses: actions/checkout@v4 26 | - uses: coq-community/docker-coq-action@v1 27 | with: 28 | opam_file: "iris-simp-lang.opam" 29 | coq_version: "${{ matrix.coq }}" 30 | before_install: | 31 | startGroup "Add Iris repo" 32 | opam repo add iris-dev https://gitlab.mpi-sws.org/iris/opam.git 33 | endGroup 34 | startGroup "Print opam config" 35 | opam config list; opam repo list; opam list 36 | endGroup 37 | build-release: 38 | name: build (released iris, latest coq) 39 | runs-on: ubuntu-latest 40 | steps: 41 | - uses: actions/checkout@v4 42 | - uses: coq-community/docker-coq-action@v1 43 | with: 44 | opam_file: "iris-simp-lang.opam" 45 | coq_version: "latest" 46 | 47 | # See also: 48 | # https://github.com/coq-community/docker-coq-action#readme 49 | # https://github.com/erikmd/docker-coq-github-action-demo 50 | -------------------------------------------------------------------------------- /src/tactics.v: -------------------------------------------------------------------------------- 1 | From stdpp Require Import fin_maps. 2 | From iris_simp_lang Require Import lang. 3 | From iris Require Import options. 4 | 5 | (*| 6 | This file implements some low-level tactics used to implement simp_lang. 7 | `reshape_expr` is used to implement the proofmode support (especially tactics 8 | like `wp_bind` and `wp_pure`) while `inv_base_step` is convenient automation for 9 | proving typeclass instances that describe simp_lang's reduction rules. 10 | |*) 11 | 12 | (** The tactic [reshape_expr e tac] decomposes the expression [e] into an 13 | evaluation context [K] and a subexpression [e']. It calls the tactic [tac K e'] 14 | for each possible decomposition until [tac] succeeds. *) 15 | Ltac reshape_expr e tac := 16 | (* Note that the current context is spread into a list of fully-constructed 17 | items [K] A fully-constructed item is inserted into [K] by calling 18 | [add_item]. *) 19 | let rec go K e := 20 | match e with 21 | | _ => tac K e 22 | | App ?e (Val ?v) => add_item (AppLCtx v) K e 23 | | App ?e1 ?e2 => add_item (AppRCtx e1) K e2 24 | | UnOp ?op ?e => add_item (UnOpCtx op) K e 25 | | BinOp ?op ?e (Val ?v) => add_item (BinOpLCtx op v) K e 26 | | BinOp ?op ?e1 ?e2 => add_item (BinOpRCtx op e1) K e2 27 | | HeapOp ?op ?e (Val ?v) => add_item (HeapOpLCtx op v) K e 28 | | HeapOp ?op ?e1 ?e2 => add_item (HeapOpRCtx op e1) K e2 29 | | If ?e0 ?e1 ?e2 => add_item (IfCtx e1 e2) K e0 30 | end 31 | with add_item Ki K e := go (Ki :: K) e 32 | in go (@nil ectx_item) e. 33 | 34 | (** The tactic [inv_base_step] performs inversion on hypotheses of the shape 35 | [base_step]. The tactic will discharge head-reductions starting from values, and 36 | simplifies hypothesis related to conversions from and to values, and finite map 37 | operations. This tactic is slightly ad-hoc and tuned for proving our lifting 38 | lemmas. *) 39 | Ltac inv_base_step := 40 | repeat match goal with 41 | | _ => progress simplify_map_eq/= (* simplify memory stuff *) 42 | | H : to_val _ = Some _ |- _ => apply of_to_val in H 43 | | H : base_step ?e _ _ _ _ _ |- _ => 44 | try (is_var e; fail 1); (* inversion yields many goals if [e] is a variable 45 | and should thus better be avoided. *) 46 | inversion H; subst; clear H 47 | end. 48 | 49 | Create HintDb base_step. 50 | Global Hint Extern 0 (base_reducible _ _) => eexists _, _, _, _; simpl : base_step. 51 | 52 | (* [simpl apply] is too stupid, so we need extern hints here. *) 53 | Global Hint Extern 1 (base_step _ _ _ _ _ _) => econstructor : base_step. 54 | Global Hint Extern 0 (base_step (HeapOp AllocOp _ _) _ _ _ _ _) => apply alloc_fresh : base_step. 55 | -------------------------------------------------------------------------------- /src/adequacy.v: -------------------------------------------------------------------------------- 1 | From iris.program_logic Require Export adequacy. 2 | From iris_simp_lang Require Import simp heap_lib. 3 | From iris.prelude Require Import options. 4 | 5 | (*| 6 | =========== 7 | Adequacy 8 | =========== 9 | 10 | This is a really important part of setting up the language. The infrastructure 11 | we've set up will let us prove specifications in Iris for simp_lang, but what do 12 | these theorems mean? This file proves **adequacy** of the weakest preconditions, 13 | which lifts a weakest precondition from within separation logic to a safety 14 | theorem about the semantics that's independent of Iris. 15 | 16 | Most of this is proven already for the generic weakest precondition definition 17 | we're using. Only one thing is missing: we need to initialize the state 18 | interpretation for the initial state. This gets to execute a ghost update, which 19 | we use to create the initial auth element for the heap_ra ghost state. 20 | 21 | The Coq implementation mostly consists of an orthogonal problem related to these 22 | Σ and related assumptions we make all over the place; if you want the details 23 | you should read 24 | https://gitlab.mpi-sws.org/iris/iris/-/blob/master/docs/resource_algebras.md, 25 | but here is a brief explanation. This argument is a list of RA functors and 26 | determine which ghost state is available in an Iris proof (this is needed to 27 | support impredicative ghost state, that is ghost state that refers to other 28 | ghost state). The simpGS assumption over Σ not only assumes that some RAs are 29 | available but also bundles a ghost name for the heap. Here, we allocate that 30 | ghost name and associated state. 31 | |*) 32 | 33 | (** These assumptions are just functors in Σ, unlike simpGS which also has a 34 | ghost name. *) 35 | Class simpGpreS Σ := SimpPreG { 36 | simp_preG_iris :: invGpreS Σ; 37 | simp_preG_heap :: heap_mapGpreS loc val Σ; 38 | }. 39 | 40 | Definition simpΣ : gFunctors := 41 | #[invΣ; heap_mapΣ loc val]. 42 | 43 | Global Instance subG_heapGpreS {Σ} : subG simpΣ Σ → simpGpreS Σ. 44 | Proof. solve_inG. Qed. 45 | 46 | Lemma simp_adequacy Σ `{!simpGpreS Σ} 47 | (s: stuckness) (e: expr) (σ: state) (φ: val → Prop) : 48 | (∀ (simpGS0: simpGS Σ), ⊢ WP e @ s; ⊤ {{ v, ⌜φ v⌝ }}) → 49 | adequate s e σ (λ (v: val) _, φ v). 50 | Proof. 51 | intros Hwp; eapply (wp_adequacy _ _); iIntros (??) "". 52 | iMod (heap_map_init σ.(heap)) as (?) "Hh". 53 | iModIntro. 54 | iExists 55 | (λ σ κs, (heap_map_interp σ.(heap))%I), 56 | (λ _, True%I). 57 | iFrame. iApply (Hwp (SimpGS _ _ _)). 58 | Qed. 59 | 60 | (*| 61 | The thing to observe in the adequacy theorem's statement is that we assume 62 | `simpGpreS Σ` (these are just ordinary functors, which we'll get by including 63 | `simpΣ` in our definition of Σ) and then pass a `simpGS Σ` to a WP proof (this is 64 | higher-order, so you have to carefully follow the positive and negative 65 | occurrences). This is possible because `wp_adequacy` permits us to execute any 66 | initial ghost updates to create the first state interpretation. 67 | |*) 68 | -------------------------------------------------------------------------------- /src/examples/spawn.v: -------------------------------------------------------------------------------- 1 | From iris.algebra Require Import excl. 2 | From iris.base_logic.lib Require Import invariants. 3 | From iris_simp_lang Require Export simp. 4 | From iris.prelude Require Import options. 5 | 6 | (** we don't have sums in the language but we can emulate them with tagged 7 | pairs *) 8 | Definition NONE: expr := (#false, #()). 9 | Definition NONEV: val := (#false, #()). 10 | 11 | Definition SOME: expr := λ: "v", (#true, "v"). 12 | Definition SOMEV (v:val): val := (#true, v). 13 | 14 | Definition spawn : val := 15 | λ: "f", 16 | let: "c" := ref NONE in 17 | Fork ("c" <- SOME ("f" #())) ;; "c". 18 | Definition join : val := 19 | rec: "join" "c" := 20 | let: "r" := !"c" in 21 | if: Fst "r" then Snd "r" 22 | else "join" "c". 23 | 24 | (** The CMRA & functor we need. *) 25 | (* Not bundling simpGS, as it may be shared with other users. *) 26 | Class spawnG Σ := SpawnG { spawn_tokG :: inG Σ (exclR unitO) }. 27 | Definition spawnΣ : gFunctors := #[GFunctor (exclR unitO)]. 28 | 29 | Global Instance subG_spawnΣ {Σ} : subG spawnΣ Σ → spawnG Σ. 30 | Proof. solve_inG. Qed. 31 | 32 | (** Now we come to the Iris part of the proof. *) 33 | Section proof. 34 | Context `{!simpGS Σ, !spawnG Σ} (N : namespace). 35 | 36 | Definition spawn_inv (γ : gname) (l : loc) (Ψ : val → iProp Σ) : iProp Σ := 37 | ∃ lv, l ↦ lv ∗ (⌜lv = NONEV⌝ ∨ 38 | ∃ w, ⌜lv = SOMEV w⌝ ∗ (Ψ w ∨ own γ (Excl ()))). 39 | 40 | Definition join_handle (l : loc) (Ψ : val → iProp Σ) : iProp Σ := 41 | ∃ γ, own γ (Excl ()) ∗ inv N (spawn_inv γ l Ψ). 42 | 43 | Global Instance spawn_inv_ne n γ l : 44 | Proper (pointwise_relation val (dist n) ==> dist n) (spawn_inv γ l). 45 | Proof. solve_proper. Qed. 46 | Global Instance join_handle_ne n l : 47 | Proper (pointwise_relation val (dist n) ==> dist n) (join_handle l). 48 | Proof. solve_proper. Qed. 49 | 50 | (** The main proofs. *) 51 | Lemma spawn_spec (Ψ : val → iProp Σ) (f : val) : 52 | {{{ WP f #() {{ Ψ }} }}} spawn f {{{ l, RET #l; join_handle l Ψ }}}. 53 | Proof. 54 | iIntros (Φ) "Hf HΦ". rewrite /spawn /=. wp_lam. 55 | wp_alloc l as "Hl". 56 | iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done. 57 | iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?". 58 | { iNext. iExists NONEV. iFrame; eauto. } 59 | wp_apply (wp_fork with "[Hf]"). 60 | - iNext. wp_bind (f _). iApply (wp_wand with "Hf"); iIntros (v) "Hv". 61 | wp_pures. iInv N as (v') "[>Hl _]". 62 | wp_store. iSplitL; last done. iIntros "!> !>". iExists (SOMEV v). iFrame. eauto. 63 | - wp_pures. iApply "HΦ". rewrite /join_handle. eauto. 64 | Qed. 65 | 66 | Lemma join_spec (Ψ : val → iProp Σ) l : 67 | {{{ join_handle l Ψ }}} join #l {{{ v, RET v; Ψ v }}}. 68 | Proof. 69 | iIntros (Φ) "H HΦ". iDestruct "H" as (γ) "[Hγ #?]". 70 | iLöb as "IH". wp_rec. wp_bind (! _)%E. iInv N as (v) "[Hl Hinv]". 71 | wp_load. iDestruct "Hinv" as "[%|Hinv]"; subst. 72 | - iModIntro. iSplitL "Hl"; [iNext; iExists _; iFrame; eauto|]. 73 | wp_apply ("IH" with "Hγ [HΦ]"). auto. 74 | - iDestruct "Hinv" as (v' ->) "[HΨ|Hγ']". 75 | + iModIntro. iSplitL "Hl Hγ"; [iNext; iExists _; iFrame; eauto|]. 76 | wp_pures. by iApply "HΦ". 77 | + iDestruct (own_valid_2 with "Hγ Hγ'") as %[]. 78 | Qed. 79 | End proof. 80 | 81 | #[export] Typeclasses Opaque join_handle. 82 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # simp_lang 2 | 3 | [![CI](https://github.com/tchajed/iris-simp-lang/actions/workflows/build.yml/badge.svg)](https://github.com/tchajed/iris-simp-lang/actions/workflows/build.yml) 4 | 5 | simp_lang is a very simple programming language that we instantiate Iris with. 6 | On top of that Iris gives a program logic based on weakest preconditions. It is 7 | heavily inspired by heap_lang (and mostly copied from it) but aims to simplify 8 | things down as much as possible while still supporting verifying concurrent 9 | programs. 10 | 11 | You might want to start with a [high-level conceptual 12 | overview](https://youtu.be/HndwyM04KEU) (links to a YouTube video; if you'd 13 | prefer there's a [static version](tutorial/)): 14 | 15 | [![](tutorial/slides/simp_lang.019.png)](https://youtu.be/HndwyM04KEU) 16 | 17 | This overview might be useful before diving into this code, which works out all 18 | the details and goes a step beyond by also doing some program verification in our 19 | new language. 20 | 21 | The recommended reading order for this tutorial is the following: 22 | 23 | 1. [lang.v](src/lang.v) defines the syntax and semantics of simp_lang (instantiating `ectxi_language`) 24 | 2. [primitive_laws.v](src/primitive_laws.v) defines a _state interpretation_ 25 | for simp_lang (instantiating 26 | `irisG simp_lang`). This is the connection between the state of simp_lang (a 27 | heap from locations to values) and the Iris logic. 28 | 3. [heap_ra.v](src/heap_ra.v) and [heap_lib.v](src/heap_lib.v) are the mechanism for the state interpretation, 29 | which will make more sense after seeing them used. 30 | 4. [adequacy.v](src/adequacy.v) sets up the generic language adequacy theorem 31 | with an initialization of the state interpretation for simp_lang. 32 | 33 | Next, you can check out some examples from the [Iris POPL 2021 tutorial](https://gitlab.mpi-sws.org/iris/tutorial-popl21/) that are 34 | re-implemented and verified in simp_lang: 35 | 36 | 1. [examples/swap.v](src/examples/swap.v) verifies a version of swap. 37 | 2. [examples/parallel_add.v](src/examples/parallel_add.v) verifies the parallel 38 | increment example. It also demonstrates applying the adequacy theorem to 39 | derive a theorem about `parallel_add` whose statement is independent of Iris. 40 | 41 | There are a few files that are optional reading which make the tutorial work: 42 | 43 | - [tactics.v](src/tactics.v) and 44 | [class_instances.v](src/class_instances.v) are necessary parts of the 45 | implementation but aren't directly related to instantiating Iris. 46 | - [notation.v](src/notation.v) makes it possible to write programs in simp_lang 47 | - [proofmode.v](src/proofmode.v) gives enough proofmode support to actually 48 | verify programs written in simp_lang. 49 | - [examples/spawn.v](src/examples/spawn.v) and 50 | [examples/par.v](src/examples/par.v) implement and verify the par combinator 51 | (`e1 ||| e2`) used in the tutorial example. 52 | 53 | ## Compiling 54 | 55 | This development relies on a development version of Iris and Coq 8.20 or later. 56 | We test Coq 8.20 and master with Iris dev in CI, as well as the released 57 | version of Iris. 58 | 59 | You'll need to install Iris, which is easiest done through opam. There are 60 | installation instructions at https://gitlab.mpi-sws.org/iris/iris. 61 | -------------------------------------------------------------------------------- /src/examples/parallel_add.v: -------------------------------------------------------------------------------- 1 | From iris_simp_lang Require Import simp adequacy examples.par. 2 | From iris.bi.lib Require Import fractional. 3 | From iris.base_logic.lib Require Import invariants ghost_var. 4 | From iris Require Import options. 5 | 6 | Open Scope Z. 7 | 8 | Definition parallel_add: expr := 9 | let: "r" := ref #0 in 10 | (FAA "r" #2) ||| (FAA "r" #2);; 11 | !"r". 12 | 13 | Section proof. 14 | Context `{!simpGS Σ, !spawnG Σ, !ghost_varG Σ Z}. 15 | 16 | Definition parallel_add_inv (r : loc) (γ1 γ2 : gname) : iProp Σ := 17 | ∃ n1 n2 : Z, r ↦ #(n1 + n2) ∗ 18 | ghost_var γ1 (1/2) n1 ∗ ghost_var γ2 (1/2) n2. 19 | 20 | Lemma parallel_add_spec : 21 | {{{ True }}} parallel_add {{{ RET #4; True }}}. 22 | Proof using All. 23 | iIntros (Φ) "_ Post". 24 | unfold parallel_add. wp_alloc r as "Hr". wp_let. 25 | iMod (ghost_var_alloc 0) as (γ1) "[Hγ1 Hγ1●]". 26 | iMod (ghost_var_alloc 0) as (γ2) "[Hγ2 Hγ2●]". 27 | iMod (inv_alloc nroot _ (parallel_add_inv r γ1 γ2) with "[Hr Hγ1● Hγ2●]") as "#Hinv". 28 | { iExists 0, 0. iFrame. } 29 | wp_apply (wp_par (λ _, ghost_var γ1 (1/2) 2) (λ _, ghost_var γ2 (1/2) 2) 30 | with "[Hγ1] [Hγ2]"). 31 | - iInv "Hinv" as (n1 n2) ">(Hr & Hγ1● & Hγ2●)". 32 | wp_apply (wp_faa with "Hr"); iIntros "Hr". 33 | iDestruct (ghost_var_agree with "Hγ1● Hγ1") as %->. 34 | iMod (ghost_var_update_halves 2 with "Hγ1● Hγ1") as "[Hγ1● Hγ1]". 35 | iModIntro. 36 | iSplitR "Hγ1"; [ | by iFrame ]. 37 | iNext. iExists _, _; iFrame. 38 | replace (0 + n2 + 2) with (2 + n2) by lia; done. 39 | - iInv "Hinv" as (n1 n2) ">(Hr & Hγ1● & Hγ2●)". 40 | wp_apply (wp_faa with "Hr"); iIntros "Hr". 41 | iDestruct (ghost_var_agree with "Hγ2● Hγ2") as %->. 42 | iMod (ghost_var_update_halves 2 with "Hγ2● Hγ2") as "[Hγ2● Hγ2]". 43 | iModIntro. 44 | iSplitR "Hγ2"; [ | by iFrame ]. 45 | iNext. iExists _, _; iFrame. 46 | replace (n1 + 0) with n1 by lia; done. 47 | - iIntros (??) "[Hγ1 Hγ2] !>". wp_seq. 48 | iInv "Hinv" as (n1 n2) ">(Hr & Hγ1● & Hγ2●)". 49 | wp_load. 50 | iDestruct (ghost_var_agree with "Hγ1● Hγ1") as %->. 51 | iDestruct (ghost_var_agree with "Hγ2● Hγ2") as %->. 52 | iModIntro. 53 | iSplitL "Hr Hγ1● Hγ2●". 54 | { iNext. iExists _, _; iFrame. } 55 | by iApply "Post". 56 | Qed. 57 | 58 | End proof. 59 | 60 | (*| 61 | =============== 62 | Using adequacy 63 | =============== 64 | 65 | We proved an Iris specification, but what does it really mean about the program? 66 | We can use the simp_lang adequacy theorem to "extract" the pure postcondition 67 | into a statement about the execution semantics. 68 | 69 | The Iris triple says two things, which are bundled into the `adequate` 70 | definition used by `simp_adequacy`: `parallel_add` will not get stuck, and if it 71 | terminates in a value (with some additional forked threads [es]), that value 72 | will be 4. We'll only prove the second part here. 73 | 74 | To use adequacy and this triple, we'll finally gather up all the functors needed 75 | for Iris ghost state and produce one of the Σ assumptions used in all of our 76 | proofs; it has all the ghost state needed for the features used in this proof. 77 | 78 | Note that we aren't in a section and this theorem statement only refers to the 79 | language semantics, so we've eliminated Iris from our trusted computing base! 80 | The theorem refers to the semantics through erased_step, which ignores 81 | observations; this doesn't matter for this language because we aren't using 82 | prophecy variables. 83 | |*) 84 | 85 | Definition parallel_addΣ: gFunctors := 86 | #[simpΣ; spawnΣ; ghost_varΣ Z]. 87 | 88 | Lemma parallel_add_returns_4 σ σ' es v : 89 | rtc erased_step ([parallel_add], σ) (Val v :: es, σ') → 90 | v = LitV (LitInt 4). 91 | Proof. 92 | intros Hstep. 93 | cut (adequate NotStuck parallel_add σ (λ v _, v = LitV (LitInt 4))). 94 | { intros H. eapply adequate_alt in H as [Hval _]; eauto. } 95 | apply (simp_adequacy parallel_addΣ). 96 | intros. 97 | by iApply (parallel_add_spec with "[//]"). 98 | Qed. 99 | 100 | Print Assumptions parallel_add_returns_4. 101 | -------------------------------------------------------------------------------- /src/heap_lib.v: -------------------------------------------------------------------------------- 1 | From stdpp Require Import gmap. 2 | From iris.base_logic.lib Require Export own. 3 | From iris.proofmode Require Import tactics. 4 | From iris_simp_lang Require Import heap_ra. 5 | From iris.prelude Require Import options. 6 | 7 | (*| 8 | 9 | This file is a "ghost state" library on top of heap_ra. 10 | 11 | It defines Iris propositions for owning the ghost state that will track the 12 | simp_lang heap and the points-to facts that will give the right to read and write 13 | to it. These are built using the general Iris mechanisms for user-defined ghost 14 | state. Wrapping them in a library like this makes the API to the ghost state 15 | easier to use, since it is now stated in terms of updates in the Iris logic 16 | rather than theorems about the RA (so for example these theorems can be used 17 | directly in the IPM). 18 | 19 | |*) 20 | 21 | Class heap_mapGpreS (L V : Type) (Σ : gFunctors) `{Countable L} := { 22 | heap_mapGpreS_inG :: inG Σ (heap_mapR L V); 23 | }. 24 | 25 | Class heap_mapGS (L V : Type) (Σ : gFunctors) `{Countable L} := GenHeapGS { 26 | heap_map_inG :: heap_mapGpreS L V Σ; 27 | heap_map_name : gname; 28 | }. 29 | Global Arguments GenHeapGS L V Σ {_ _ _} _ : assert. 30 | Global Arguments heap_map_name {L V Σ _ _} _ : assert. 31 | 32 | Definition heap_mapΣ (L V : Type) `{Countable L} : gFunctors := #[ 33 | GFunctor (heap_mapR L V) 34 | ]. 35 | 36 | Global Instance subG_heap_mapGpreS {Σ L V} `{Countable L} : 37 | subG (heap_mapΣ L V) Σ → heap_mapGpreS L V Σ. 38 | Proof. solve_inG. Qed. 39 | 40 | Section definitions. 41 | Context `{Countable L, hG : !heap_mapGS L V Σ}. 42 | 43 | (*| 44 | These two definitions are the key idea behind the state interpretation. 45 | `heap_map_interp` is the authoritative element of this RA, which will be the 46 | state interpretation of `σ`, while `pointsto_def` has fragments that live outside 47 | the state interpretation and are owned by threads. `l ↦ v` will be notation for 48 | `pointsto`, with a full 1 fraction. 49 | |*) 50 | 51 | Definition heap_map_interp (σ : gmap L V) : iProp Σ := 52 | own (heap_map_name hG) (Auth (σ : gmap L V)). 53 | 54 | Definition pointsto_def (l : L) (v: V) : iProp Σ := 55 | own (heap_map_name hG) (Frag {[l := v]}). 56 | Definition pointsto_aux : seal (@pointsto_def). Proof. by eexists. Qed. 57 | Definition pointsto := pointsto_aux.(unseal). 58 | Definition pointsto_eq : @pointsto = @pointsto_def := pointsto_aux.(seal_eq). 59 | End definitions. 60 | 61 | Notation "l ↦ v" := (pointsto l v) 62 | (at level 20, format "l ↦ v") : bi_scope. 63 | 64 | Section heap_map. 65 | Context {L V} `{Countable L, !heap_mapGS L V Σ}. 66 | Implicit Types (P Q : iProp Σ). 67 | Implicit Types (Φ : V → iProp Σ). 68 | Implicit Types (σ : gmap L V) (l : L) (v : V). 69 | 70 | (** General properties of pointsto *) 71 | Global Instance pointsto_timeless l v : Timeless (l ↦ v). 72 | Proof. rewrite pointsto_eq. apply _. Qed. 73 | 74 | Lemma pointsto_conflict l v1 v2 : l ↦ v1 -∗ l ↦ v2 -∗ False. 75 | Proof. 76 | rewrite pointsto_eq. iIntros "H1 H2". 77 | iDestruct (own_valid_2 with "H1 H2") as %Hdisj%heap_map_frag_frag_valid. 78 | apply map_disjoint_dom in Hdisj. 79 | set_solver. 80 | Qed. 81 | 82 | (** Update lemmas *) 83 | Lemma heap_map_alloc σ l v : 84 | σ !! l = None → 85 | heap_map_interp σ ==∗ heap_map_interp (<[l:=v]>σ) ∗ l ↦ v. 86 | Proof. 87 | iIntros (Hσl). rewrite /heap_map_interp pointsto_eq /pointsto_def /=. 88 | iIntros "Hσ". 89 | iMod (own_update with "Hσ") as "[Hσ Hl]". 90 | { eapply (heap_map_alloc_update _ l); done. } 91 | iModIntro. iFrame. 92 | Qed. 93 | 94 | Lemma heap_map_valid σ l v : heap_map_interp σ -∗ l ↦ v -∗ ⌜σ !! l = Some v⌝. 95 | Proof. 96 | iIntros "Hσ Hl". 97 | rewrite /heap_map_interp pointsto_eq /pointsto_def. 98 | by iDestruct (own_valid_2 with "Hσ Hl") as %Hsub%heap_map_singleton_valid. 99 | Qed. 100 | 101 | Lemma heap_map_update σ l v1 v2 : 102 | heap_map_interp σ -∗ l ↦ v1 ==∗ heap_map_interp (<[l:=v2]>σ) ∗ l ↦ v2. 103 | Proof. 104 | iIntros "Hσ Hl". 105 | rewrite /heap_map_interp pointsto_eq /pointsto_def. 106 | iMod (own_update_2 with "Hσ Hl") as "[Hσ Hl]". 107 | { eapply heap_map_modify_update. } 108 | iModIntro. iFrame. 109 | Qed. 110 | End heap_map. 111 | 112 | Lemma heap_map_init `{Countable L, !heap_mapGpreS L V Σ} σ : 113 | ⊢ |==> ∃ _ : heap_mapGS L V Σ, heap_map_interp σ. 114 | Proof. 115 | iMod (own_alloc (Auth (σ : gmap L V))) as (γ) "Hσ". 116 | { exact: heap_map_auth_valid. } 117 | iExists (GenHeapGS _ _ _ γ). 118 | done. 119 | Qed. 120 | -------------------------------------------------------------------------------- /src/notation.v: -------------------------------------------------------------------------------- 1 | From stdpp Require Export binders strings. 2 | From iris_simp_lang Require Export lang. 3 | From iris Require Import options. 4 | 5 | (*| 6 | These notations are magic that make simp_lang, like heap_lang, easy to type and 7 | read. 8 | |*) 9 | 10 | 11 | (** Coercions to make programs easier to type. *) 12 | Coercion LitInt : Z >-> base_lit. 13 | Coercion LitBool : bool >-> base_lit. 14 | 15 | Coercion App : expr >-> Funclass. 16 | 17 | Coercion Val : val >-> expr. 18 | Coercion Var : string >-> expr. 19 | 20 | (** Define some derived forms. *) 21 | Notation Lam x e := (Rec BAnon x e) (only parsing). 22 | Notation Let x e1 e2 := (App (Lam x e2) e1) (only parsing). 23 | Notation Seq e1 e2 := (Let BAnon e1 e2) (only parsing). 24 | Notation LamV x e := (RecV BAnon x e) (only parsing). 25 | 26 | (** expressions use single constructors with a code to specify the specific 27 | operation; here we wrap those in notations to make them easy to read and 28 | write *) 29 | Notation Pair := (BinOp PairOp). 30 | Notation Fst := (UnOp FstOp). 31 | Notation Snd := (UnOp SndOp). 32 | Notation Alloc e := (HeapOp AllocOp e (Val (LitV LitUnit))). 33 | Notation Load e := (HeapOp LoadOp e (Val (LitV LitUnit))). 34 | Notation Store e1 e2 := (HeapOp StoreOp e1 e2). 35 | Notation FAA := (HeapOp FaaOp). 36 | 37 | (* Skip should be atomic, we sometimes open invariants around 38 | it. Hence, we need to explicitly use LamV instead of e.g., Seq. *) 39 | Notation Skip := (App (Val $ LamV BAnon (Val $ LitV LitUnit)) (Val $ LitV LitUnit)). 40 | 41 | (* No scope for the values, does not conflict and scope is often not inferred 42 | properly. *) 43 | Notation "# l" := (LitV l%Z%V%stdpp) (at level 8, format "# l"). 44 | 45 | (** Syntax inspired by Coq/Ocaml. Constructions with higher precedence come 46 | first. *) 47 | Notation "( e1 , e2 , .. , en )" := (Pair .. (Pair e1 e2) .. en) : expr_scope. 48 | Notation "( e1 , e2 , .. , en )" := (PairV .. (PairV e1 e2) .. en) : val_scope. 49 | 50 | Notation "()" := LitUnit : val_scope. 51 | Notation "! e" := (Load e%E) (at level 9, right associativity) : expr_scope. 52 | Notation "'ref' e" := (Alloc e%E) (at level 10) : expr_scope. 53 | (* The unicode ← is already part of the notation "_ ← _; _" for bind. *) 54 | Notation "e1 <- e2" := (Store e1%E e2%E) (at level 80) : expr_scope. 55 | 56 | Notation "e1 + e2" := (BinOp PlusOp e1%E e2%E) : expr_scope. 57 | 58 | (* The breaking point '/ ' makes sure that the body of the rec is indented 59 | by two spaces in case the whole rec does not fit on a single line. *) 60 | Notation "'rec:' f x := e" := (Rec f%binder x%binder e%E) 61 | (at level 200, f at level 1, x at level 1, e at level 200, 62 | format "'[' 'rec:' f x := '/ ' e ']'") : expr_scope. 63 | Notation "'rec:' f x := e" := (RecV f%binder x%binder e%E) 64 | (at level 200, f at level 1, x at level 1, e at level 200, 65 | format "'[' 'rec:' f x := '/ ' e ']'") : val_scope. 66 | Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E) 67 | (at level 200, e1, e2, e3 at level 200) : expr_scope. 68 | 69 | (** Derived notions, in order of declaration. The notations for let and seq 70 | are stated explicitly instead of relying on the Notations Let and Seq as 71 | defined above. This is needed because App is now a coercion, and these 72 | notations are otherwise not pretty printed back accordingly. *) 73 | Notation "'rec:' f x y .. z := e" := (Rec f%binder x%binder (Lam y%binder .. (Lam z%binder e%E) ..)) 74 | (at level 200, f, x, y, z at level 1, e at level 200, 75 | format "'[' 'rec:' f x y .. z := '/ ' e ']'") : expr_scope. 76 | Notation "'rec:' f x y .. z := e" := (RecV f%binder x%binder (Lam y%binder .. (Lam z%binder e%E) ..)) 77 | (at level 200, f, x, y, z at level 1, e at level 200, 78 | format "'[' 'rec:' f x y .. z := '/ ' e ']'") : val_scope. 79 | 80 | (* The breaking point '/ ' makes sure that the body of the λ: is indented 81 | by two spaces in case the whole λ: does not fit on a single line. *) 82 | Notation "λ: x , e" := (Lam x%binder e%E) 83 | (at level 200, x at level 1, e at level 200, 84 | format "'[' 'λ:' x , '/ ' e ']'") : expr_scope. 85 | Notation "λ: x y .. z , e" := (Lam x%binder (Lam y%binder .. (Lam z%binder e%E) ..)) 86 | (at level 200, x, y, z at level 1, e at level 200, 87 | format "'[' 'λ:' x y .. z , '/ ' e ']'") : expr_scope. 88 | 89 | Notation "λ: x , e" := (LamV x%binder e%E) 90 | (at level 200, x at level 1, e at level 200, 91 | format "'[' 'λ:' x , '/ ' e ']'") : val_scope. 92 | Notation "λ: x y .. z , e" := (LamV x%binder (Lam y%binder .. (Lam z%binder e%E) .. )) 93 | (at level 200, x, y, z at level 1, e at level 200, 94 | format "'[' 'λ:' x y .. z , '/ ' e ']'") : val_scope. 95 | 96 | Notation "'let:' x := e1 'in' e2" := (Lam x%binder e2%E e1%E) 97 | (at level 200, x at level 1, e1, e2 at level 200, 98 | format "'[' 'let:' x := '[' e1 ']' 'in' '/' e2 ']'") : expr_scope. 99 | Notation "e1 ;; e2" := (Lam BAnon e2%E e1%E) 100 | (at level 100, e2 at level 200, 101 | format "'[' '[hv' '[' e1 ']' ;; ']' '/' e2 ']'") : expr_scope. 102 | -------------------------------------------------------------------------------- /src/class_instances.v: -------------------------------------------------------------------------------- 1 | From iris_simp_lang Require Import notation tactics. 2 | From iris.prelude Require Import options. 3 | 4 | (*| 5 | These instances prove that various expressions are atomic or pure. 6 | 7 | `Atomic e` is defined generically for languages by saying `e` reduces to a value 8 | (recall: this is defined by `to_val e = Some _`) in a single step. 9 | 10 | `PureExec φ n e1 e2` shows that if φ holds (a pure Coq proposition), `e1` 11 | executes to `e2` in `n` steps. This is eventually needed to define a tactic 12 | `wp_pure _` that finds and reasons about pure reductions (this subsumes 13 | `wp_let`, `wp_seq`, `wp_app` and the like, which are just restrictions of 14 | `wp_pure`). 15 | |*) 16 | 17 | 18 | Global Instance into_val_val v : IntoVal (Val v) v. 19 | Proof. done. Qed. 20 | Global Instance as_val_val v : AsVal (Val v). 21 | Proof. by eexists. Qed. 22 | 23 | (** * Instances of the [Atomic] class *) 24 | Section atomic. 25 | Local Ltac solve_atomic := 26 | apply strongly_atomic_atomic, ectx_language_atomic; 27 | [inversion 1; naive_solver 28 | |apply ectxi_language_sub_redexes_are_values; intros [] **; naive_solver]. 29 | 30 | Global Instance rec_atomic s f x e : Atomic s (Rec f x e). 31 | Proof. solve_atomic. Qed. 32 | (** The instance below is a more general version of [Skip] *) 33 | Global Instance beta_atomic s f x v1 v2 : Atomic s (App (RecV f x (Val v1)) (Val v2)). 34 | Proof. destruct f, x; solve_atomic. Qed. 35 | Global Instance unop_atomic s op v : Atomic s (UnOp op (Val v)). 36 | Proof. solve_atomic. Qed. 37 | Global Instance binop_atomic s op v1 v2 : Atomic s (BinOp op (Val v1) (Val v2)). 38 | Proof. solve_atomic. Qed. 39 | Global Instance if_true_atomic s v1 e2 : 40 | Atomic s (If (Val $ LitV $ LitBool true) (Val v1) e2). 41 | Proof. solve_atomic. Qed. 42 | Global Instance if_false_atomic s e1 v2 : 43 | Atomic s (If (Val $ LitV $ LitBool false) e1 (Val v2)). 44 | Proof. solve_atomic. Qed. 45 | 46 | Global Instance fork_atomic s e : Atomic s (Fork e). 47 | Proof. solve_atomic. Qed. 48 | 49 | Global Instance heap_op_atomic op s v1 v2 : Atomic s (HeapOp op (Val v1) (Val v2)). 50 | Proof. solve_atomic. Qed. 51 | End atomic. 52 | 53 | (** * Instances of the [PureExec] class *) 54 | (** The behavior of the various [wp_] tactics with regard to lambda differs in 55 | the following way: 56 | 57 | - [wp_pures] does *not* reduce lambdas/recs that are hidden behind a definition. 58 | - [wp_rec] and [wp_lam] reduce lambdas/recs that are hidden behind a definition. 59 | 60 | To realize this behavior, we define the class [AsRecV v f x erec], which takes a 61 | value [v] as its input, and turns it into a [RecV f x erec] via the instance 62 | [AsRecV_recv : AsRecV (RecV f x e) f x e]. We register this instance via 63 | [Hint Extern] so that it is only used if [v] is syntactically a lambda/rec, and 64 | not if [v] contains a lambda/rec that is hidden behind a definition. 65 | 66 | To make sure that [wp_rec] and [wp_lam] do reduce lambdas/recs that are hidden 67 | behind a definition, we activate [AsRecV_recv] by hand in these tactics. *) 68 | Class AsRecV (v : val) (f x : binder) (erec : expr) := 69 | as_recv : v = RecV f x erec. 70 | Global Hint Mode AsRecV ! - - - : typeclass_instances. 71 | Definition AsRecV_recv f x e : AsRecV (RecV f x e) f x e := eq_refl. 72 | Global Hint Extern 0 (AsRecV (RecV _ _ _) _ _ _) => 73 | apply AsRecV_recv : typeclass_instances. 74 | 75 | Section pure_exec. 76 | Local Ltac solve_exec_safe := intros; subst; do 3 eexists; econstructor; eauto. 77 | Local Ltac solve_exec_puredet := simpl; intros; by inv_base_step. 78 | Local Ltac solve_pure_exec := 79 | subst; intros ?; apply nsteps_once, pure_base_step_pure_step; 80 | constructor; [solve_exec_safe | solve_exec_puredet]. 81 | 82 | Global Instance pure_recc f x (erec : expr) : 83 | PureExec True 1 (Rec f x erec) (Val $ RecV f x erec). 84 | Proof. solve_pure_exec. Qed. 85 | Global Instance pure_beta f x (erec : expr) (v1 v2 : val) `{!AsRecV v1 f x erec} : 86 | PureExec True 1 (App (Val v1) (Val v2)) (subst' x v2 (subst' f v1 erec)). 87 | Proof. unfold AsRecV in *. solve_pure_exec. Qed. 88 | 89 | Global Instance pure_unop op v v' : 90 | PureExec (un_op_eval op v = Some v') 1 (UnOp op (Val v)) (Val v'). 91 | Proof. solve_pure_exec. Qed. 92 | 93 | Global Instance pure_binop op v1 v2 v' : 94 | PureExec (bin_op_eval op v1 v2 = Some v') 1 (BinOp op (Val v1) (Val v2)) (Val v') | 10. 95 | Proof. solve_pure_exec. Qed. 96 | (* Higher-priority instance for [EqOp]. *) 97 | Global Instance pure_eqop v1 v2 : 98 | PureExec True 1 99 | (BinOp EqOp (Val v1) (Val v2)) 100 | (Val $ LitV $ LitBool $ bool_decide (v1 = v2)) | 1. 101 | Proof. solve_pure_exec. Qed. 102 | 103 | Global Instance pure_if_true e1 e2 : 104 | PureExec True 1 (If (Val $ LitV $ LitBool true) e1 e2) e1. 105 | Proof. solve_pure_exec. Qed. 106 | Global Instance pure_if_false e1 e2 : 107 | PureExec True 1 (If (Val $ LitV $ LitBool false) e1 e2) e2. 108 | Proof. solve_pure_exec. Qed. 109 | End pure_exec. 110 | -------------------------------------------------------------------------------- /tutorial/README.md: -------------------------------------------------------------------------------- 1 | # Instantiating Iris - a conceptual overview 2 | 3 | (this is a static version of a 12-minute overview I recorded which you can 4 | watch on [YouTube](https://youtu.be/HndwyM04KEU)) 5 | 6 | We're going to instantiate Iris with a new language. What that means is we'll 7 | define a language - a syntax and semantics - and then use Iris's general 8 | definition of weakest preconditions to get a program logic for our language. 9 | 10 | First, we need a syntax for the language. This works best with a lambda calculus 11 | because it is expected to be defined in terms of a type of expressions and a 12 | type of values, with functions `to_val` and `of_val` to go back and forth. 13 | `to_val : expr -> option val` is partial and when it returns Some we think of 14 | that expression as a value. This will be relevant later when thinking about when 15 | programs are done executing vs stuck. 16 | 17 | The `simp_lang` language has the standard lambda calculus constructs for 18 | variables and lambdas, some pure and heap operations, and `Fork` for 19 | concurrency. 20 | 21 | ![](slides/simp_lang.003.png) 22 | 23 | Next we need a semantics. This is a stateful language so we define a type of 24 | state; `simp_lang` is simple enough that the only state is a heap from locations 25 | to values. The semantics is expressed in standard lambda calculus fashion as a 26 | transition relation between a starting expression and state `(e, σ)` and a next 27 | expression and state `(e', σ')`. We can also spawn threads along the way with 28 | `Fork` which the semantics can interleave non-deterministically. 29 | 30 | Writing down this complete relation is a bit of work so Iris comes up with a 31 | simpler interface built on top of the general language interface. This 32 | interface, `ectxi_language`, allows you to describe the semantics using 33 | _evaluation context items_ that define which subexpression to evaluate, combined 34 | with _head reduction rules_ that describe how to finally reduce each construct 35 | in the language. As a simple example, `e □` and `□ v` are contexts that say we 36 | can reduce the argument to an application at any time and the function in an 37 | application once its argument is a value. This defines a deterministic 38 | right-to-left evaluation order. 39 | 40 | ![](slides/simp_lang.005.png) 41 | 42 | Now we want to take our language and plug it into Iris. What we need to finish 43 | the definition of weakest preconditions is something called a state 44 | interpretation, which maps states of the language to some sort of ownership 45 | within the Iris logic. To illustrate what the state interpretation does, I'll 46 | walk through how it appears in the definition of weakest preconditions. 47 | 48 | ![](slides/simp_lang.008.png) 49 | 50 | Now the complete definition of `wp` is a bit hairy and abstract, so I'll walk 51 | through what it boils down to for a simple Hoare spec about a language 52 | primitive, `Store(l, v)` (which we typically write with notation as `l <- v`). 53 | We can prove this triple in a low-level way by proving this Iris entailment: 54 | 55 | ![](slides/simp_lang.009.png) 56 | 57 | If we look at the beginning and end of this entailment, we can see that it 58 | _assumes_ the precondition and state interpretation (at `σ`) hold initially and 59 | has to _prove_ the state interpretation (at `σ'`) and postcondition hold at the 60 | end. This works out only because `Store` is atomic - the definition of weakest 61 | preconditions sort of threads the state interpretation through as an invariant, 62 | albeit one which is tied to some evolving state of the program. 63 | 64 | ![](slides/simp_lang.010.png) 65 | 66 | In between we prove that `Store(l, v)` is reducible - this is what guarantees 67 | that any weakest precondition implies the expression doesn't get stuck - and we 68 | get to assume that `Store` takes a step and prove the state interpretation of 69 | the new state. Finally, the `|==>` is a _fancy update_ modality that allows us 70 | to update ghost state before proving the new state interpretation. We'll need 71 | that because the state interpretation asserts ownership of some ghost state, and 72 | we need to update it to match the new state `σ'`. 73 | 74 | ![](slides/simp_lang.009.png) 75 | 76 | So what is the state interpretation? It's basically ownership of a particular 77 | RA, the auth RA for maps. If you have seen this RA, the state intepretation for 78 | a heap-based language has the auth element of this RA while the definition of 79 | points-to `l ↦ v` holds fragments. 80 | 81 | If you haven't seen it, here's a quick overview of the simplest RA that would 82 | work for the heap. We have two types of elements in the carrier - authoritative 83 | elements `●m` that intuitively will hold the entire heap `m` and fragments `◯m` 84 | that will hold portions of the heap. Fragments compose by disjoint union (while 85 | fragments that overlap compose to some invalid element), and we can have a 86 | single authoritative element and any number of fragments. Crucially this RA 87 | maintains a connection between the auth and its fragments: each fragment is 88 | always a subset of the entire auth. 89 | 90 | Now going back to the state interpretation and points-to definition we can see 91 | why the definition of points-to has meaning: it asserts _knowledge_ of what's in 92 | the heap because of validity (when combined with the state interpretation), and 93 | it asserts _exclusive ownership_ because no other thread can have the same 94 | fragment since the composition is invalid. 95 | 96 | This is the most basic RA where threads always have all-or-nothing access to a 97 | location - we can build a more sophisticated RA that also supports fractional 98 | permissions, and indeed that's what the code does. 99 | 100 | ![](slides/simp_lang.015.png) 101 | 102 | Now we can go back and understand how to use the state interpretation to prove a 103 | primitive Hoare triple. These are needed to finish a language instantiation; 104 | after all the primitives have specs the user can use the program logic to prove 105 | larger programs, but we need to use lower-level tools to prove the primitives 106 | correct. 107 | 108 | If we simplify this example down to what `Store` specifically needs (due to its 109 | particular execution semantics), and take care of proving that `Store` is 110 | reducible if we have `l ↦ v0`, the task remaining is just a particular 111 | viewshift. We need to take the combination of state interpretation and 112 | precondition ownership to a new state interpretation and postcondition points-to fact. 113 | 114 | ![](slides/simp_lang.017.png) 115 | 116 | Turns out this is just a frame-preserving update in the heap RA. Intuitively 117 | this works for the same reason `l ↦ v` is exclusive: the frame cannot also have 118 | `l ↦ w` because then the combination of the frame with this thread's ownership 119 | would be invalid. 120 | 121 | ![](slides/simp_lang.018.png) 122 | 123 | In summary: 124 | 125 | ![](slides/simp_lang.019.png) 126 | -------------------------------------------------------------------------------- /src/primitive_laws.v: -------------------------------------------------------------------------------- 1 | From iris.program_logic Require Export weakestpre. 2 | From iris.proofmode Require Import tactics. 3 | From iris.program_logic Require Import ectx_lifting. 4 | From iris_simp_lang Require Import notation tactics class_instances. 5 | From iris_simp_lang Require Import heap_lib. 6 | From iris Require Import options. 7 | 8 | (*| 9 | This is one of the most interesting parts of the instantiation. Now that we have 10 | a syntax and semantics, we want a program logic. There's exactly one more thing 11 | Iris needs before we can define weakest preconditions: a **state 12 | interpretation**. This is a function from state (recall, just a heap for 13 | simp_lang) to iProp Σ. The way this plugs into Iris is by instantiating the 14 | `irisGS simp_lang Σ` typeclass, which is an assumption for `wp`, the generic 15 | definition of weakest preconditions (this is the definition you usually interact 16 | with through either the `WP` notation or "Texan" Hoare-triple notation). 17 | 18 | The state interpretation for simp_lang maps `gmap loc val` into an appropriate 19 | RA. Most of the definition is related to `heap_mapG`, which is defined in 20 | `heap_ra.v`. We can think of the state interpretation as being an invariant 21 | maintained by the weakest precondition, except that it is a function of the 22 | state and thus has meaning tied to the program execution. Therefore we pick an 23 | RA which is like an auth of a gmap for the state interpretation and map `σ : 24 | gmap loc val` to something like `own γ (●σ)`. Then we can use fragments to 25 | define the core points-to connective for this language, something like `l ↦ v := 26 | own γ (◯{|l := v|})`. 27 | 28 | When we prove the WP for an atomic language primitive, we'll prove it more 29 | directly than usual. The proof obligation will be to transform the state 30 | interpretation of `σ` to the state interpretation of some `σ'` that's a valid 31 | transition of the primitive. It's here that we'll **update the ghost state** to 32 | match transitions like allocation and use agreement between the auth and 33 | fragments to prove the WP for loads, for example. These are the most interesting 34 | proofs about the language because they don't just reason about pure reduction 35 | steps but actually have to make use of the logic and the ghost state we set up 36 | to reason about its state --- the purely functional part of the language clearly 37 | doesn't need separation logic! 38 | 39 | The code implements this with two differences. First, the RA we actually use in 40 | `heap_ra.` is an Iris library `gmap_viewR loc val` which uses a generalization 41 | of auth called views. The important point is that the auth component is the 42 | state's heap and the fragments are sub-heaps; the view RA keeps track of the 43 | fact that the composition of all the fragments is a sub-heap of the 44 | authoritative element. It also adds something called discardable fractions, a 45 | generalization of the fraction RA, in order to model fractional and persistent 46 | permissions to heap locations, but we can mostly ignore this complication. 47 | 48 | Second, there's a pesky ghost name `γ` in the informal definitions above. These 49 | are hidden away in the `simpGS` typeclass as part of `heap_mapG` that all proofs 50 | about this language will carry. It'll be fixed once before any execution by the 51 | adequacy theorem, as you'll see in adequacy.v. After that we get it through a 52 | typeclass to avoid mentioning it explicitly in any proofs. 53 | 54 | If you were writing your own language, you would probably start with `heap_mapG` 55 | from the Iris standard library to get all the nice features and lemmas for the 56 | heap part of the state interpretation. Then you could add other algebras and 57 | global ghost names to the equivalent of `simpGS`, as long as you also instantiate 58 | them in `adequacy.v`. 59 | |*) 60 | 61 | 62 | Class simpGS Σ := SimpGS { 63 | simp_invGS : invGS Σ; 64 | simp_heap_mapG :: heap_mapGS loc val Σ; 65 | }. 66 | 67 | (* Observe that this instance assumes [simpGS Σ], which already has a fixed ghost 68 | name for the heap ghost state. We'll see in adequacy.v how to obtain a [simpGS Σ] 69 | after allocating that ghost state. *) 70 | Global Instance simpGS_irisGS `{!simpGS Σ} : irisGS simp_lang Σ := { 71 | iris_invGS := simp_invGS; 72 | state_interp σ _ κs _ := (heap_map_interp σ.(heap))%I; 73 | fork_post _ := True%I; 74 | (* These two fields are for a new feature that makes the number of laters per 75 | physical step flexible; picking 0 here means we get just one later per 76 | physical step, as older versions of Iris had. This is the same way heap_lang 77 | works. *) 78 | num_laters_per_step _ := 0; 79 | state_interp_mono _ _ _ _ := fupd_intro _ _ 80 | }. 81 | 82 | Notation "l ↦ v" := (pointsto l v) 83 | (at level 20, format "l ↦ v") : bi_scope. 84 | 85 | Section lifting. 86 | Context `{!simpGS Σ}. 87 | Implicit Types P Q : iProp Σ. 88 | Implicit Types Φ Ψ : val → iProp Σ. 89 | Implicit Types efs : list expr. 90 | Implicit Types σ : state. 91 | Implicit Types v : val. 92 | Implicit Types l : loc. 93 | 94 | (** Fork: Not using Texan triples to avoid some unnecessary [True] *) 95 | Lemma wp_fork s E e Φ : 96 | ▷ WP e @ s; ⊤ {{ _, True }} -∗ ▷ Φ (LitV LitUnit) -∗ WP Fork e @ s; E {{ Φ }}. 97 | Proof. 98 | iIntros "He HΦ". iApply wp_lift_atomic_base_step; [done|]. 99 | iIntros (σ1 κ κs n nt) "Hσ !>"; iSplit; first by eauto with base_step. 100 | iIntros "!>" (v2 σ2 efs Hstep) "_Hcred"; inv_base_step. by iFrame. 101 | Qed. 102 | 103 | (** Heap *) 104 | 105 | Lemma wp_alloc s E v : 106 | {{{ True }}} Alloc (Val v) @ s; E 107 | {{{ l, RET LitV (LitInt l); l ↦ v }}}. 108 | Proof. 109 | iIntros (Φ) "_ HΦ". iApply wp_lift_atomic_base_step_no_fork; first done. 110 | iIntros (σ1 κ κs n nt) "Hσ !>"; iSplit; first by auto with base_step. 111 | iIntros "!>" (v2 σ2 efs Hstep) "_Hcred"; inv_base_step. 112 | iMod (heap_map_alloc σ1.(heap) _ _ with "Hσ") as "[Hσ Hl]"; first done. 113 | iModIntro; iSplit=> //. iFrame. by iApply "HΦ". 114 | Qed. 115 | 116 | Lemma wp_load s E l v : 117 | {{{ l ↦ v }}} Load (Val $ LitV $ LitInt l) @ s; E {{{ RET v; l ↦ v }}}. 118 | Proof. 119 | iIntros (Φ) "Hl HΦ". iApply wp_lift_atomic_base_step_no_fork; first done. 120 | iIntros (σ1 κ κs n nt) "Hσ !>". iDestruct (heap_map_valid with "Hσ Hl") as %?. 121 | iSplit; first by eauto with base_step. 122 | iNext. iIntros (v2 σ2 efs Hstep) "_Hcred"; inv_base_step. 123 | iModIntro; iSplit=> //. iFrame. by iApply "HΦ". 124 | Qed. 125 | 126 | Lemma wp_store s E l v w : 127 | {{{ l ↦ v }}} Store (Val $ LitV $ LitInt l) (Val $ w) @ s; E {{{ RET #(); l ↦ w }}}. 128 | Proof. 129 | iIntros (Φ) "Hl HΦ". iApply wp_lift_atomic_base_step_no_fork; first done. 130 | iIntros (σ1 κ κs n nt) "Hσ !>". iDestruct (heap_map_valid with "Hσ Hl") as %?. 131 | iSplit; first by eauto with base_step. 132 | iNext. iIntros (v2 σ2 efs Hstep) "_Hcred"; inv_base_step. 133 | iMod (heap_map_update _ _ _ w with "Hσ Hl") as "[Hσ Hl]". 134 | iModIntro; iSplit=> //. iFrame. by iApply "HΦ". 135 | Qed. 136 | 137 | Lemma wp_faa s E l (n1 n2: Z) : 138 | {{{ l ↦ #n1 }}} 139 | FAA (Val $ LitV $ LitInt l) (Val $ LitV $ LitInt $ n2) @ s; E 140 | {{{ RET #n1; l ↦ #(n1+n2) }}}. 141 | Proof. 142 | iIntros (Φ) "Hl HΦ". iApply wp_lift_atomic_base_step_no_fork; first done. 143 | iIntros (σ1 κ κs n nt) "Hσ !>". iDestruct (heap_map_valid with "Hσ Hl") as %?. 144 | iSplit; first by eauto with base_step. 145 | iNext. iIntros (v2 σ2 efs Hstep) "_Hcred"; inv_base_step. 146 | iMod (heap_map_update _ _ _ #(n1 + n2) with "Hσ Hl") as "[Hσ Hl]". 147 | iModIntro; iSplit=> //. iFrame. by iApply "HΦ". 148 | Qed. 149 | 150 | End lifting. 151 | -------------------------------------------------------------------------------- /src/proofmode.v: -------------------------------------------------------------------------------- 1 | From iris.proofmode Require Import coq_tactics reduction. 2 | From iris.proofmode Require Export tactics. 3 | From iris.program_logic Require Import atomic. 4 | From iris_simp_lang Require Import tactics class_instances primitive_laws notation. 5 | From iris.prelude Require Import options. 6 | 7 | (*| 8 | This is a heavily stripped-down version of HeapLang's proofmode support. To make 9 | any program proofs reasonable we do need to implement `wp_pure` and `wp_bind`, 10 | and as a demo of the implementation we also implement `wp_load` in the 11 | reflective style typical in the IPM. `wp_pure` is the basis for a number of 12 | tactics like `wp_rec` and `wp_let` and such, while `wp_bind` is what powers 13 | `wp_apply`. 14 | |*) 15 | 16 | Lemma tac_wp_expr_eval `{!simpGS Σ} Δ s E Φ e e' : 17 | (∀ (e'':=e'), e = e'') → 18 | envs_entails Δ (WP e' @ s; E {{ Φ }}) → envs_entails Δ (WP e @ s; E {{ Φ }}). 19 | Proof. by intros ->. Qed. 20 | 21 | Tactic Notation "wp_expr_eval" tactic3(t) := 22 | iStartProof; 23 | lazymatch goal with 24 | | |- envs_entails _ (wp ?s ?E ?e ?Q) => 25 | notypeclasses refine (tac_wp_expr_eval _ _ _ _ e _ _ _); 26 | [let x := fresh in intros x; t; unfold x; notypeclasses refine eq_refl|] 27 | | _ => fail "wp_expr_eval: not a 'wp'" 28 | end. 29 | Ltac wp_expr_simpl := wp_expr_eval simpl. 30 | 31 | Lemma tac_wp_pure `{!simpGS Σ} Δ Δ' s E K e1 e2 φ n Φ : 32 | PureExec φ n e1 e2 → 33 | φ → 34 | MaybeIntoLaterNEnvs n Δ Δ' → 35 | envs_entails Δ' (WP (fill K e2) @ s; E {{ Φ }}) → 36 | envs_entails Δ (WP (fill K e1) @ s; E {{ Φ }}). 37 | Proof. 38 | rewrite envs_entails_unseal=> ??? HΔ'. rewrite into_laterN_env_sound /=. 39 | (* We want [pure_exec_fill] to be available to TC search locally. *) 40 | pose proof @pure_exec_fill. 41 | rewrite HΔ' -lifting.wp_pure_step_later //. 42 | iIntros "Hwp !> _" => //. 43 | Qed. 44 | 45 | Lemma tac_wp_value_nofupd `{!simpGS Σ} Δ s E Φ v : 46 | envs_entails Δ (Φ v) → envs_entails Δ (WP (Val v) @ s; E {{ Φ }}). 47 | Proof. rewrite envs_entails_unseal=> ->. by apply wp_value. Qed. 48 | 49 | Lemma tac_wp_value `{!simpGS Σ} Δ s E (Φ : val → iPropI Σ) v : 50 | envs_entails Δ (|={E}=> Φ v) → envs_entails Δ (WP (Val v) @ s; E {{ Φ }}). 51 | Proof. rewrite envs_entails_unseal=> ->. by rewrite wp_value_fupd. Qed. 52 | 53 | (** Simplify the goal if it is [WP] of a value. 54 | If the postcondition already allows a fupd, do not add a second one. 55 | But otherwise, *do* add a fupd. This ensures that all the lemmas applied 56 | here are bidirectional, so we never will make a goal unprovable. *) 57 | Ltac wp_value_head := 58 | lazymatch goal with 59 | | |- envs_entails _ (wp ?s ?E (Val _) (λ _, fupd ?E _ _)) => 60 | eapply tac_wp_value_nofupd 61 | | |- envs_entails _ (wp ?s ?E (Val _) (λ _, wp _ ?E _ _)) => 62 | eapply tac_wp_value_nofupd 63 | | |- envs_entails _ (wp ?s ?E (Val _) _) => 64 | eapply tac_wp_value 65 | end. 66 | 67 | Ltac wp_finish := 68 | wp_expr_simpl; (* simplify occurences of subst/fill *) 69 | try wp_value_head; (* in case we have reached a value, get rid of the WP *) 70 | pm_prettify. (* prettify ▷s caused by [MaybeIntoLaterNEnvs] and 71 | λs caused by wp_value *) 72 | 73 | Ltac solve_vals_compare_safe := 74 | (* The first branch is for when we have [vals_compare_safe] in the context. 75 | The other two branches are for when either one of the branches reduces to 76 | [True] or we have it in the context. *) 77 | fast_done || (left; fast_done) || (right; fast_done). 78 | 79 | (** The argument [efoc] can be used to specify the construct that should be 80 | reduced. For example, you can write [wp_pure (EIf _ _ _)], which will search 81 | for an [EIf _ _ _] in the expression, and reduce it. 82 | 83 | The use of [open_constr] in this tactic is essential. It will convert all holes 84 | (i.e. [_]s) into evars, that later get unified when an occurences is found 85 | (see [unify e' efoc] in the code below). *) 86 | Tactic Notation "wp_pure" open_constr(efoc) := 87 | iStartProof; 88 | lazymatch goal with 89 | | |- envs_entails _ (wp ?s ?E ?e ?Q) => 90 | let e := eval simpl in e in 91 | reshape_expr e ltac:(fun K e' => 92 | unify e' efoc; 93 | eapply (tac_wp_pure _ _ _ _ K e'); 94 | [tc_solve (* PureExec *) 95 | |try solve_vals_compare_safe (* The pure condition for PureExec -- 96 | handles trivial goals, including [vals_compare_safe] *) 97 | |tc_solve (* IntoLaters *) 98 | |wp_finish (* new goal *) 99 | ]) 100 | || fail "wp_pure: cannot find" efoc "in" e "or" efoc "is not a redex" 101 | | _ => fail "wp_pure: not a 'wp'" 102 | end. 103 | 104 | Ltac wp_pures := 105 | iStartProof; 106 | first [ (* The `;[]` makes sure that no side-condition magically spawns. *) 107 | progress repeat (wp_pure _; []) 108 | | wp_finish (* In case wp_pure never ran, make sure we do the usual cleanup. *) 109 | ]. 110 | 111 | (** Unlike [wp_pures], the tactics [wp_rec] and [wp_lam] should also reduce 112 | lambdas/recs that are hidden behind a definition, i.e. they should use 113 | [AsRecV_recv] as a proper instance instead of a [Hint Extern]. 114 | 115 | We achieve this by putting [AsRecV_recv] in the current environment so that it 116 | can be used as an instance by the typeclass resolution system. We then perform 117 | the reduction, and finally we clear this new hypothesis. *) 118 | Tactic Notation "wp_rec" := 119 | let H := fresh in 120 | pose proof (H := AsRecV_recv); 121 | wp_pure (App _ _); 122 | clear H. 123 | 124 | Tactic Notation "wp_if" := wp_pure (If _ _ _). 125 | Tactic Notation "wp_if_true" := wp_pure (If (LitV (LitBool true)) _ _). 126 | Tactic Notation "wp_if_false" := wp_pure (If (LitV (LitBool false)) _ _). 127 | Tactic Notation "wp_unop" := wp_pure (UnOp _ _). 128 | Tactic Notation "wp_binop" := wp_pure (BinOp _ _ _). 129 | Tactic Notation "wp_op" := wp_unop || wp_binop. 130 | Tactic Notation "wp_lam" := wp_rec. 131 | Tactic Notation "wp_let" := wp_pure (Rec BAnon (BNamed _) _); wp_lam. 132 | Tactic Notation "wp_seq" := wp_pure (Rec BAnon BAnon _); wp_lam. 133 | Tactic Notation "wp_proj" := wp_pure (Fst _) || wp_pure (Snd _). 134 | Tactic Notation "wp_pair" := wp_pure (Pair _ _). 135 | Tactic Notation "wp_closure" := wp_pure (Rec _ _ _). 136 | 137 | Lemma tac_wp_bind `{!simpGS Σ} K Δ s E Φ e f : 138 | f = (λ e, fill K e) → (* as an eta expanded hypothesis so that we can `simpl` it *) 139 | envs_entails Δ (WP e @ s; E {{ v, WP f (Val v) @ s; E {{ Φ }} }})%I → 140 | envs_entails Δ (WP fill K e @ s; E {{ Φ }}). 141 | Proof. rewrite envs_entails_unseal=> -> ->. by apply: wp_bind. Qed. 142 | 143 | Ltac wp_bind_core K := 144 | lazymatch eval hnf in K with 145 | | [] => idtac 146 | | _ => eapply (tac_wp_bind K); [simpl; reflexivity|reduction.pm_prettify] 147 | end. 148 | 149 | Tactic Notation "wp_bind" open_constr(efoc) := 150 | iStartProof; 151 | lazymatch goal with 152 | | |- envs_entails _ (wp ?s ?E ?e ?Q) => 153 | first [ reshape_expr e ltac:(fun K e' => unify e' efoc; wp_bind_core K) 154 | | fail 1 "wp_bind: cannot find" efoc "in" e ] 155 | | _ => fail "wp_bind: not a 'wp'" 156 | end. 157 | 158 | (*| 159 | ===================== 160 | Convenience tactics 161 | ===================== 162 | 163 | `wp_load` is just a shorthand for using the `wp_load` lemma. 164 | |*) 165 | 166 | 167 | (** Heap tactics *) 168 | Section heap. 169 | Context `{!simpGS Σ}. 170 | Implicit Types P Q : iProp Σ. 171 | Implicit Types Φ : val → iProp Σ. 172 | Implicit Types Δ : envs (uPredI (iResUR Σ)). 173 | Implicit Types (l: loc) (v : val) (z : Z). 174 | 175 | Lemma tac_wp_load Δ Δ' s E i K b (l: loc) v Φ : 176 | MaybeIntoLaterNEnvs 1 Δ Δ' → 177 | envs_lookup i Δ' = Some (b, l ↦ v)%I → 178 | envs_entails Δ' (WP fill K (Val v) @ s; E {{ Φ }}) → 179 | envs_entails Δ (WP fill K (Load (LitV l)) @ s; E {{ Φ }}). 180 | Proof. 181 | rewrite envs_entails_unseal=> ?? Hi. 182 | rewrite into_laterN_env_sound /=. 183 | iIntros "Henv". 184 | iDestruct (envs_lookup_split with "Henv") as "[Hl Henv]"; first by eauto. 185 | iApply wp_bind. 186 | destruct b; simpl. 187 | - iDestruct "Hl" as ">#Hl". 188 | iApply (wp_load with "Hl"). iIntros "!> _". 189 | iApply Hi. iApply ("Henv" with "Hl"). 190 | - iDestruct "Hl" as ">Hl". 191 | iApply (wp_load with "Hl"). iIntros "!> Hl". 192 | iApply Hi. iApply ("Henv" with "Hl"). 193 | Qed. 194 | 195 | End heap. 196 | 197 | (** Evaluate [lem] to a hypothesis [H] that can be applied, and then run 198 | [wp_bind K; tac H] for every possible evaluation context. [tac] can do 199 | [iApplyHyp H] to actually apply the hypothesis. TC resolution of [lem] premises 200 | happens *after* [tac H] got executed. *) 201 | Tactic Notation "wp_apply_core" open_constr(lem) tactic3(tac) := 202 | wp_pures; 203 | iPoseProofCore lem as false (fun H => 204 | lazymatch goal with 205 | | |- envs_entails _ (wp ?s ?E ?e ?Q) => 206 | reshape_expr e ltac:(fun K e' => 207 | wp_bind_core K; tac H) || 208 | lazymatch iTypeOf H with 209 | | Some (_,?P) => fail "wp_apply: cannot apply" P 210 | end 211 | | _ => fail "wp_apply: not a 'wp'" 212 | end). 213 | Tactic Notation "wp_apply" open_constr(lem) := 214 | wp_apply_core lem (fun H => iApplyHyp H; try iNext; try wp_expr_simpl). 215 | 216 | Tactic Notation "wp_load" := 217 | let solve_pointsto _ := 218 | let l := match goal with |- _ = Some (_, (?l ↦ _)%I) => l end in 219 | iAssumptionCore || fail "wp_load: cannot find" l "↦ ?" in 220 | wp_pures; 221 | lazymatch goal with 222 | | |- envs_entails _ (wp ?s ?E ?e ?Q) => 223 | first 224 | [reshape_expr e ltac:(fun K e' => eapply (tac_wp_load _ _ _ _ _ K)) 225 | |fail 1 "wp_load: cannot find 'Load' in" e]; 226 | [tc_solve 227 | |solve_pointsto () 228 | |wp_finish] 229 | | _ => fail "wp_load: not a 'wp'" 230 | end. 231 | 232 | (* These tactics are not implemented reflectively out of laziness, but they do 233 | get the job done. Don't be afraid to implement helpers like this, though! They 234 | can be slightly buggier, slower, and give worse error messages but it's great 235 | for prototyping. *) 236 | 237 | Tactic Notation "wp_alloc" ident(l) "as" constr(H) := 238 | wp_apply (wp_alloc with "[//]"); iIntros (l) H. 239 | 240 | Tactic Notation "wp_store" := 241 | wp_pures; 242 | wp_bind (Store _ _); 243 | lazymatch goal with 244 | | |- envs_entails ?Δ (wp ?s ?E (Store (Val (LitV (LitInt ?l))) _) ?Q) => 245 | lazymatch Δ with 246 | | context[Esnoc _ ?i (l ↦ _)%I] => 247 | wp_apply (wp_store with i); iIntros i 248 | | _ => fail "wp_store: could not find" l "↦ v" 249 | end 250 | | _ => fail "wp_store: not a 'wp'" 251 | end. 252 | -------------------------------------------------------------------------------- /src/heap_ra.v: -------------------------------------------------------------------------------- 1 | From stdpp Require Import countable. 2 | From stdpp Require Import gmap. 3 | From iris.algebra Require Export cmra. 4 | From iris.algebra Require Import proofmode_classes updates frac stepindex_finite. 5 | From iris.algebra Require Import agree. 6 | From iris.prelude Require Import options. 7 | 8 | (* 9 | ========== 10 | Heap CMRA 11 | ========== 12 | 13 | Here we define the CMRA we'll use to construct the "points-to" ghost state 14 | that will be built-in to the program logic, in the sense that its meaning is 15 | tied to the physical state of the program by the program logic definition itself. 16 | 17 | Ordinarily you wouldn't build a CMRA "by hand" like this, but instead would use 18 | the various combinators provided by Iris. In fact, there's a library called 19 | gen_heap built on top of the auth RA (and a few others) that works well for the 20 | purpose in simp_lang. However, by building it up ourselves it's easier to show 21 | how all the pieces fit together and reduce the amount of magic. 22 | 23 | The downside to constructing the RA by hand is that this file has a subsantial 24 | amount of proof. 25 | 26 | This library is parameterized over arbitrary types of keys (`K`) and values 27 | (`V`), although we will only use it for `loc` and `val` of simp_lang. 28 | Mathematically, we can think of this RA as having elements of the form `Auth m` 29 | and `Frag m` where `m : gmap K V`, which stand for "authorative ownership" and 30 | "fragmentary ownership". The idea is that there is a single `Auth σ` that holds 31 | the true, physical heap, while `Frag m` will be owned by threads and gives the 32 | right to read/modify a part of the heap. We'll also have a special `⊥` (invalid) 33 | element to represent compositions that don't make sense. 34 | 35 | These have the following composition/validity properties: 36 | 37 | - `Auth m ⋅ Auth m' = ⊥` (there's only one auth) 38 | - `Frag m ⋅ Frag m' = m ∪ m' if m and m' are disjoint (##ₘ) 39 | - `✓ (Auth m ⋅ Frag m') ↔ m' ⊆ m (fragments always agree with the authoritative copy) 40 | 41 | An important special case is `Frag {[k := v]}`, fragmentary ownership of a 42 | singleton. This is like a points-to fact, and in fact `l ↦ v` will be *defined* as 43 | `own (Frag {[l := v]})`. 44 | 45 | This CMRA permits the following important *local updates* and validity rules, which are 46 | "frame-preserving updates". `x ~~> y` is defined to be `∀ z, ✓ (x ⋅ z) → ✓ (y ⋅ 47 | z)`, which means that we can replace x with y locally and preserve global 48 | validity. 49 | 50 | - `Auth m ~~> Auth (<[k := v]> m) ⋅ Frag {[k := v]` if `m !! k = None` 51 | - `Auth m ⋅ Frag {[k := v]} ~~> Auth (<[k := v']> m) ⋅ Frag {[k := v']}` 52 | 53 | |*) 54 | 55 | Section heap_map. 56 | Context {K: Type} 57 | `{EqDecision0: !EqDecision K} 58 | `{Countable0: !Countable K} 59 | {V: Type} `{EqDecision V}. 60 | 61 | Record auth_map_val := 62 | { auth: option (gmap K V); 63 | frag: gmap K V; 64 | }. 65 | 66 | (** The carrier for the CMRA we're defining. We will define something that 67 | looks like a PCM (partial commutative monoid), with a partial composition 68 | operation that returns `Invalid` when composition doesn't make sense. We need 69 | to do this to fit into Iris's algebraic structure for resource, called a CMRA, 70 | which isntead has a notion of composing elements and a predicate for valid 71 | elements. The difference is important for higher-order ghost state, but this 72 | is not a higher-order CMRA (that is, it doesn't depend on `iProp`). *) 73 | Inductive heap_map := 74 | | Ok (v: auth_map_val) 75 | | Invalid. 76 | 77 | Canonical Structure heap_mapO : ofe := leibnizO heap_map. 78 | 79 | Global Instance heap_map_inhabited : Inhabited heap_map := populate Invalid. 80 | Local Lemma gmap_eq_dec : EqDecision (gmap K V). 81 | Proof using All. apply _. Qed. 82 | Global Instance auth_map_eq_dec : EqDecision auth_map_val. 83 | Proof using All. solve_decision. Qed. 84 | Global Instance heap_map_eq_dec : EqDecision heap_map. 85 | Proof using All. solve_decision. Qed. 86 | 87 | Global Instance heap_map_Ok_inj : Inj (=) (=) Ok. 88 | Proof. by injection 1. Qed. 89 | 90 | (** An element is valid if it's Ok v and the fragment is contained in the 91 | auth. This ensures that fragments really agree with the global auth state. *) 92 | Local Instance heap_map_valid_instance : Valid heap_map := λ hm, 93 | match hm with 94 | | Ok {| auth := Some m; frag := m' |} => m' ⊆ m 95 | | Ok {| auth := None; frag := m' |} => True 96 | | Invalid => False 97 | end. 98 | 99 | (** In a PCM there would be a unit for composition ε and `∀ x, x ⋅ ε = x`. 100 | Instead of a global unit, a CMRA requires one more piece: a partial core or 101 | `pcore`, that (might) give a unit for each element - the might is the 102 | "partial" in the name. The `pcore` is used to model the persistently modality 103 | (□P). It won't be relevant to our example, but this CMRA has a unit and we can 104 | use it to give a total pcore, which turns out to simplify some proofs later. 105 | *) 106 | Instance heap_map_unit : Unit heap_map := Ok {| auth := None; frag := ∅ |}. 107 | Local Instance heap_map_pcore_instance : PCore heap_map := λ _, Some ε. 108 | 109 | (** Composition has two facets: the auth part of a heap does not compose 110 | (there's only one auth), while fragments compose via disjoint union. Other 111 | combinations result in `Invalid`. *) 112 | Local Instance heap_map_op_instance : Op heap_map := λ hm1 hm2, 113 | match hm1, hm2 with 114 | | Ok {| auth := Some m1; frag := m1' |}, 115 | Ok {| auth := Some m2; frag := m2' |} => 116 | Invalid 117 | | Ok {| auth := Some m1; frag := m1' |}, 118 | Ok {| auth := None; frag := m2' |} => 119 | if decide (m1' ##ₘ m2') then 120 | Ok {| auth := Some m1; frag := m1' ∪ m2' |} 121 | else Invalid 122 | | Ok {| auth := None; frag := m1' |}, 123 | Ok {| auth := mm2; frag := m2' |} => 124 | if decide (m1' ##ₘ m2') then 125 | Ok {| auth := mm2; frag := m1' ∪ m2' |} 126 | else Invalid 127 | | Invalid, _ => Invalid 128 | | _, Invalid => Invalid 129 | end. 130 | 131 | Lemma heap_map_unit_ok (hm: heap_map) : ε ⋅ hm = hm. 132 | Proof. 133 | rewrite /ε /heap_map_unit. 134 | rewrite /op /heap_map_op_instance. 135 | destruct hm as [[[m|] m'] | ]; auto. 136 | - rewrite -> decide_True by apply map_disjoint_empty_l. 137 | rewrite map_empty_union //. 138 | - rewrite -> decide_True by apply map_disjoint_empty_l. 139 | rewrite map_empty_union //. 140 | Qed. 141 | 142 | Definition Auth (m: gmap K V) : heap_map := Ok {| auth := Some m; frag := ∅ |}. 143 | Definition Frag (m: gmap K V) : heap_map := Ok {| auth := None; frag := m |}. 144 | 145 | Lemma heap_map_frag_op (m1 m2: gmap K V) : 146 | m1 ##ₘ m2 → 147 | Frag m1 ⋅ Frag m2 = Frag (m1 ∪ m2). 148 | Proof. 149 | intros Hdisj. 150 | rewrite /op /heap_map_op_instance /=. 151 | rewrite decide_True //. 152 | Qed. 153 | 154 | Lemma heap_map_auth_op (m1 m2: gmap K V) : 155 | Auth m1 ⋅ Auth m2 = Invalid. 156 | Proof. 157 | rewrite /op /heap_map_op_instance //=. 158 | Qed. 159 | 160 | Local Lemma heap_map_op_auth_frag (m m': gmap K V) : 161 | Auth m ⋅ Frag m' = Ok {| auth := Some m; frag := m'; |}. 162 | Proof. 163 | rewrite /op /heap_map_op_instance /=. 164 | rewrite decide_True. 165 | - rewrite map_empty_union //. 166 | - apply map_disjoint_empty_l. 167 | Qed. 168 | 169 | Ltac map_solver := 170 | repeat match goal with 171 | | H: context[_ ##ₘ _] |- _ => rewrite map_disjoint_dom in H 172 | | |- context[_ ##ₘ _] => rewrite map_disjoint_dom 173 | | H: context[dom (_ ∪ _)] |- _ => rewrite dom_union_L in H 174 | | |- context[dom (_ ∪ _)] => rewrite dom_union_L 175 | | _ => rewrite map_union_assoc 176 | end; 177 | set_solver. 178 | 179 | (*| 180 | Below we need to prove that the composition operation and validity predicate 181 | respect all the algebraic laws of a CMRA. 182 | |*) 183 | 184 | Local Instance heap_map_op_comm : Comm (=) (op: heap_map → heap_map → heap_map). 185 | Proof. 186 | intros hm1 hm2. 187 | rewrite /op /heap_map_op_instance. 188 | destruct hm1 as [[[m1|] m1']|], hm2 as [[[m2|] m2']|]; 189 | simpl; auto; 190 | repeat (destruct (decide _); 191 | destruct_and?; subst); 192 | try map_solver. 193 | all: rewrite map_union_comm //. 194 | Qed. 195 | 196 | Local Instance heap_map_op_assoc : Assoc (=) (op: heap_map → heap_map → heap_map). 197 | Proof. 198 | intros hm1 hm2 hm3. 199 | rewrite /op /heap_map_op_instance. 200 | destruct hm1 as [[[m1|] m1']|], hm2 as [[[m2|] m2']|], hm3 as [[[m3|] m3']|]; 201 | simpl; auto; 202 | repeat (destruct (decide _); 203 | destruct_and?; subst); 204 | try map_solver. 205 | Qed. 206 | 207 | Lemma heap_map_valid_op : ∀ (x y: heap_map), ✓ (x ⋅ y) → ✓ x. 208 | Proof. 209 | intros [[mm1 m1']|] [[mm2 m2']|]; 210 | rewrite /valid /heap_map_valid_instance /=; 211 | rewrite /op /heap_map_op_instance /=; 212 | auto. 213 | - destruct mm1 as [m1|], mm2 as [m2|]; try solve [ intuition auto ]. 214 | destruct (decide _) as [Hdisj | ]; [ | done ]. 215 | rewrite map_subseteq_spec => H. 216 | apply map_subseteq_spec. 217 | intros k v Hlookup. 218 | apply H. 219 | rewrite lookup_union_l //. 220 | apply map_disjoint_dom in Hdisj. 221 | apply elem_of_dom_2 in Hlookup. 222 | apply not_elem_of_dom. 223 | set_solver. 224 | - destruct mm1 as [m1|]; intuition eauto. 225 | Qed. 226 | 227 | Definition heap_map_ra_mixin : RAMixin heap_map. 228 | Proof. 229 | (* Several obligations are proven due to the above [Local Instance]s. Most 230 | of the remaining ones are about `pcore`, and are quite simple since it is 231 | total. *) 232 | split; try apply _. 233 | - intros hm1 hm2 chm1; rewrite leibniz_equiv_iff; intros <-. 234 | rewrite /pcore /heap_map_pcore_instance /=. 235 | intros [=]; subst. 236 | eexists; eauto. 237 | - intros hm1 chm1. 238 | rewrite /pcore /heap_map_pcore_instance. 239 | intros [=]; subst. 240 | rewrite heap_map_unit_ok //. 241 | - intros hm1 chm1. 242 | rewrite /pcore /heap_map_pcore_instance. 243 | rewrite leibniz_equiv_iff //. 244 | - intros ??? Hincl. 245 | rewrite /pcore /heap_map_pcore_instance => [= <-]. 246 | eexists; intuition eauto. 247 | exists ε; rewrite heap_map_unit_ok //. 248 | - apply heap_map_valid_op. 249 | Qed. 250 | Canonical Structure heap_mapR := discreteR heap_map heap_map_ra_mixin. 251 | 252 | Global Instance heap_map_cmra_discrete : CmraDiscrete heap_mapR. 253 | Proof. apply discrete_cmra_discrete. Qed. 254 | 255 | Global Instance heap_map_cmra_total : CmraTotal heap_mapR. 256 | Proof. intros hm. auto. Qed. 257 | 258 | (** This is the key lemma for creating this RA in the first place. *) 259 | Lemma heap_map_auth_valid m : ✓ (Auth m). 260 | Proof. 261 | rewrite /valid /heap_map_valid_instance /Auth. 262 | apply map_empty_subseteq. 263 | Qed. 264 | 265 | Lemma heap_map_auth_frag_valid m m' : ✓ (Auth m ⋅ Frag m') ↔ m' ⊆ m. 266 | Proof. 267 | rewrite heap_map_op_auth_frag. 268 | rewrite /valid //=. 269 | Qed. 270 | 271 | Lemma heap_map_auth_other_valid m hm : 272 | ✓ (Auth m ⋅ hm) ↔ ∃ m', hm = Frag m' ∧ m' ⊆ m. 273 | Proof. 274 | split. 275 | - rewrite /valid /heap_map_valid_instance. 276 | rewrite /Auth. 277 | destruct hm as [[mm' m'']|]; simpl; try solve [ intuition idtac ]. 278 | destruct mm' as [m'|]; simpl; try solve [ intuition idtac ]. 279 | rewrite /op /=. 280 | rewrite -> decide_True by apply map_disjoint_empty_l. 281 | rewrite map_empty_union. 282 | intros Hsub. 283 | rewrite /Frag. eexists; split; eauto. 284 | - intros [m' [-> Hsub]]. 285 | rewrite heap_map_auth_frag_valid //. 286 | Qed. 287 | 288 | Lemma heap_map_frag_frag_valid m m' : 289 | ✓ (Frag m ⋅ Frag m') ↔ m ##ₘ m'. 290 | Proof. 291 | split. 292 | - rewrite /Frag /valid /heap_map_valid_instance /=. 293 | rewrite /op /=. 294 | destruct (decide _); intuition eauto. 295 | - intros. 296 | rewrite heap_map_frag_op //. 297 | Qed. 298 | 299 | Lemma heap_map_auth_frag_other_valid m m' hm : 300 | ✓ (Auth m ⋅ Frag m' ⋅ hm) ↔ ∃ m'', hm = Frag m'' ∧ m' ∪ m'' ⊆ m ∧ m' ##ₘ m''. 301 | Proof. 302 | split. 303 | - intros Hval. 304 | assert (✓ (Auth m ⋅ hm)) as H. 305 | { rewrite -assoc_L (comm _ (Frag m')) assoc_L in Hval. 306 | apply cmra_valid_op_l in Hval; auto. } 307 | apply heap_map_auth_other_valid in H as [m'' [-> Hsub]]. 308 | assert (m' ##ₘ m''). 309 | { rewrite -assoc_L in Hval; apply cmra_valid_op_r in Hval. 310 | apply heap_map_frag_frag_valid in Hval; auto. } 311 | rewrite -assoc_L in Hval. 312 | rewrite heap_map_frag_op // in Hval. 313 | apply heap_map_auth_frag_valid in Hval. 314 | eexists; intuition eauto. 315 | - intros [m'' (-> & Hsub & Hdisj)]. 316 | rewrite -assoc_L heap_map_frag_op //. 317 | apply heap_map_auth_frag_valid; auto. 318 | Qed. 319 | 320 | (** This local update permits allocating a points-to fact for a fresh 321 | address. It's the key lemma for reasoning about allocating new references. *) 322 | Lemma heap_map_alloc_update m k v : 323 | m !! k = None → 324 | Auth m ~~> Auth (<[k := v]> m) ⋅ Frag {[k := v]}. 325 | Proof. 326 | intros Hnotin. 327 | assert (m ##ₘ {[k := v]}) as Hm. 328 | { apply map_disjoint_dom. 329 | apply not_elem_of_dom in Hnotin. 330 | set_solver. } 331 | apply cmra_discrete_total_update. 332 | intros hm'; rewrite heap_map_auth_other_valid. 333 | intros [m' [-> Hsub]]. 334 | assert ({[k := v]} ##ₘ m'). 335 | { apply map_disjoint_dom in Hm. apply map_disjoint_dom. 336 | apply not_elem_of_dom in Hnotin. 337 | apply subseteq_dom in Hsub. 338 | set_solver. } 339 | rewrite -assoc_L. 340 | rewrite heap_map_frag_op //. 341 | - apply heap_map_auth_frag_valid. 342 | rewrite insert_union_singleton_l. 343 | apply map_union_mono_l; auto. 344 | Qed. 345 | 346 | (** This lemma relates owning a fragment to the authoritative element. It is 347 | specialized to a singleton fragment since that's how we'll eventually use this 348 | RA in proofs. It's the key lemma to reasoning about loading a value with a 349 | points-to fact. *) 350 | Lemma heap_map_singleton_valid k v m : 351 | ✓ (Auth m ⋅ Frag {[k := v]}) → 352 | m !! k = Some v. 353 | Proof. 354 | rewrite heap_map_auth_frag_valid. 355 | intros Hlookup%map_singleton_subseteq_l; auto. 356 | Qed. 357 | 358 | (** This local update permits updating the value of a key using a fragment. 359 | It's the key lemma for storing using a points-to fact. *) 360 | Lemma heap_map_modify_update m k v v' : 361 | Auth m ⋅ Frag {[k := v]} ~~> Auth (<[k := v']> m) ⋅ Frag {[k := v']}. 362 | Proof. 363 | apply cmra_discrete_total_update. 364 | intros hm. 365 | rewrite ?heap_map_auth_frag_other_valid. 366 | intros [m'' (-> & Hsub & Hdisj)]. 367 | exists m''; intuition eauto. 368 | - rewrite insert_union_singleton_l. 369 | apply map_union_mono_l. 370 | apply map_subseteq_spec => k' v'' Hlookup. 371 | rewrite map_subseteq_spec in Hsub. 372 | apply Hsub. 373 | rewrite lookup_union_r //. 374 | apply lookup_singleton_None. 375 | apply map_disjoint_dom in Hdisj. 376 | apply elem_of_dom_2 in Hlookup. 377 | set_solver. 378 | - apply map_disjoint_dom. 379 | apply map_disjoint_dom in Hdisj. 380 | set_solver. 381 | Qed. 382 | 383 | End heap_map. 384 | 385 | Arguments heap_map K {EqDecision0 Countable0} V. 386 | Arguments heap_mapR K {EqDecision0 Countable0} V. 387 | -------------------------------------------------------------------------------- /src/lang.v: -------------------------------------------------------------------------------- 1 | From stdpp Require Export binders strings. 2 | From stdpp Require Import gmap. 3 | From iris.algebra Require Export ofe. 4 | From iris.program_logic Require Export language ectx_language ectxi_language. 5 | From iris.prelude Require Import options. 6 | Open Scope Z. 7 | 8 | (*| 9 | =================================== 10 | Syntax and semantics of simp_lang. 11 | =================================== 12 | 13 | We define a lambda calculus with heap operations and Fork, much as in heap_lang. 14 | Binders are represented as strings and substitution is *not* capture-avoiding; 15 | instead, we only reason about applications to closed terms, in which case this 16 | substitution function is well-behaved. To make this possible, all values are 17 | identified by a dedicated `Val` constructor in `expr`. One result is that there 18 | is a distinction between Rec (a potentially recursive lambda **`expr`** ) and 19 | RecV (a lambda **`val`** ), and similarly for the Pair expression and PairV 20 | value. The expression immediately reduces to the value. 21 | 22 | Values are extremely simple in this language - they are either integers or 23 | booleans. Locations are represented with integers and booleans as 0 and 1. This 24 | makes simp_lang not great for logical relations proofs or a decent type system 25 | (unlike heap_lang), but that's okay for pedagogical purposes here. 26 | 27 | The language semantics is simplified by instantiating the `ectxi_language` 28 | interface, for languages defined in terms of *evaluation context items*, in line 29 | with how a semantics might be written on a blackboard. We'll get to exactly what 30 | that means later, but basically we will write down "head reduction" rules for 31 | each primitive when its arguments are values and evaluation contexts for the 32 | other reductions and `ectxi_language` will do the rest. You could also directly 33 | instantiate `language` and not use evaluation contexts. 34 | 35 | =========== 36 | Syntax 37 | =========== 38 | 39 | We will be careful to minimize the recursive structure of `expr` to make some 40 | administrative stuff later on simpler. For example we group plus, equals, and 41 | pair into a single `expr` constructor that takes two expressions. 42 | |*) 43 | 44 | Inductive base_lit := 45 | | LitInt (n:Z) 46 | | LitUnit. 47 | 48 | Inductive bin_op := 49 | | PlusOp 50 | | EqOp 51 | | PairOp. 52 | 53 | Inductive un_op := 54 | | FstOp 55 | | SndOp. 56 | 57 | Inductive heap_op := 58 | | AllocOp 59 | | LoadOp 60 | | StoreOp 61 | | FaaOp. 62 | 63 | (*| 64 | Expressions are defined mutually recursively with values. As explained above, an 65 | expression is a value iff it uses the Val constructor, which makes defining 66 | reduction and substitution much simpler, but requires duplicating `Rec` and 67 | `Pair` to `val` as `RecV` and `PairV`. 68 | |*) 69 | 70 | Inductive expr := 71 | (* Values *) 72 | | Val (v : val) 73 | (* Base lambda calculus *) 74 | | Var (x : string) 75 | | Rec (f x : binder) (e : expr) 76 | | App (e1 e2 : expr) 77 | (* Pure operations *) 78 | | BinOp (op : bin_op) (e1 e2 : expr) 79 | | UnOp (op : un_op) (e : expr) 80 | | If (e0 e1 e2 : expr) 81 | (* Concurrency *) 82 | | Fork (e : expr) 83 | (* Heap *) 84 | | HeapOp (op : heap_op) (e1 : expr) (e2 : expr) 85 | with val := 86 | | LitV (l : base_lit) 87 | | RecV (f x : binder) (e : expr) 88 | | PairV (v1 v2 : val) 89 | . 90 | 91 | Bind Scope expr_scope with expr. 92 | Bind Scope val_scope with val. 93 | 94 | (*| 95 | These define part of ectxi_language: we need a type for expressions, a type for 96 | values, and a way to go back and forth (where to_val is partial because only 97 | some expressions are values). Here the use of Val to identify all values makes 98 | this easy and non-recursive. 99 | |*) 100 | 101 | Notation of_val := Val (only parsing). 102 | 103 | Definition to_val (e : expr) : option val := 104 | match e with 105 | | Val v => Some v 106 | | _ => None 107 | end. 108 | 109 | (** Equality and other typeclass stuff *) 110 | Lemma to_of_val v : to_val (of_val v) = Some v. 111 | Proof. by destruct v. Qed. 112 | 113 | Lemma of_to_val e v : to_val e = Some v → of_val v = e. 114 | Proof. destruct e=>//=. by intros [= <-]. Qed. 115 | 116 | Global Instance of_val_inj : Inj (=) (=) of_val. 117 | Proof. intros ??. congruence. Qed. 118 | 119 | (*| 120 | We will now do a bunch of boring work to prove that expressions have decidable 121 | equality and are countable for technical reasons. 122 | |*) 123 | 124 | Global Instance base_lit_eq_dec : EqDecision base_lit. 125 | Proof. solve_decision. Defined. 126 | Global Instance bin_op_eq_dec : EqDecision bin_op. 127 | Proof. solve_decision. Defined. 128 | Global Instance un_op_eq_dec : EqDecision un_op. 129 | Proof. solve_decision. Defined. 130 | Global Instance heap_op_eq_dec : EqDecision heap_op. 131 | Proof. solve_decision. Defined. 132 | Lemma expr_eq_dec (e1 e2: expr) : Decision (e1 = e2) 133 | with val_eq_dec (v1 v2 : val) : Decision (v1 = v2). 134 | Proof. 135 | { refine 136 | (match e1, e2 with 137 | | Val v, Val v' => cast_if (decide (v = v')) 138 | | Var x, Var x' => cast_if (decide (x = x')) 139 | | Rec f x e, Rec f' x' e' => cast_if_and3 (decide (f = f')) (decide (x = x')) (decide (e = e')) 140 | | App e1 e2, App e1' e2' => 141 | cast_if_and (decide (e1 = e1')) (decide (e2 = e2')) 142 | | BinOp op e1 e2, BinOp op' e1' e2' | HeapOp op e1 e2, HeapOp op' e1' e2' => 143 | cast_if_and3 (decide (op = op')) (decide (e1 = e1')) (decide (e2 = e2')) 144 | | If e1 e2 e3, If e1' e2' e3' => 145 | cast_if_and3 (decide (e1 = e1')) (decide (e2 = e2')) (decide (e3 = e3')) 146 | | UnOp op e, UnOp op' e' => 147 | cast_if_and (decide (op = op')) (decide (e = e')) 148 | | Fork e, Fork e' => 149 | cast_if (decide (e = e')) 150 | | _, _ => right _ 151 | end); solve [ abstract intuition congruence ]. } 152 | { refine 153 | (match v1, v2 with 154 | | LitV l, LitV l' => cast_if (decide (l = l')) 155 | | RecV f x e, RecV f' x' e' => cast_if_and3 (decide (f = f')) (decide (x = x')) (decide (e = e')) 156 | | PairV e1 e2, PairV e1' e2' => cast_if_and (decide (e1 = e1')) (decide (e2 = e2')) 157 | | _, _ => right _ 158 | end); try solve [ abstract intuition congruence ]. } 159 | Defined. 160 | Global Instance expr_eq_dec' : EqDecision expr := expr_eq_dec. 161 | Global Instance val_eq_dec' : EqDecision val := val_eq_dec. 162 | 163 | Global Instance base_lit_countable : Countable base_lit. 164 | Proof. 165 | refine (inj_countable' 166 | (λ l, match l with | LitInt n => inl n | LitUnit => inr () end) 167 | (λ v, match v with | inl n => _ | inr _ => _ end) _). 168 | intros []; eauto. 169 | Qed. 170 | 171 | Global Instance bin_op_countable : Countable bin_op. 172 | Proof. 173 | refine (inj_countable' 174 | (λ op, match op with | PlusOp => 0 | EqOp => 1 | PairOp => 2 end) 175 | (λ n, match n with | 0 => _ | 1 => _ | 2 => _ 176 | | _ => ltac:(constructor) end) _). 177 | intros []; eauto. 178 | Qed. 179 | 180 | Global Instance un_op_countable : Countable un_op. 181 | Proof. 182 | refine (inj_countable' 183 | (λ op, match op with | FstOp => 0 | SndOp => 1 end) 184 | (λ n, match n with | 0 => _ | 1 => _ | _ => ltac:(constructor) end) _). 185 | intros []; eauto. 186 | Qed. 187 | 188 | Global Instance heap_op_countable : Countable heap_op. 189 | Proof. 190 | refine (inj_countable' 191 | (λ op, match op with | AllocOp => 0 | LoadOp => 1 | StoreOp => 2 | FaaOp => 3 end) 192 | (λ n, match n with | 0 => _ | 1 => _ | 2 => _ | 3 => _ 193 | | _ => ltac:(constructor) end) _). 194 | intros []; eauto. 195 | Qed. 196 | 197 | Global Instance expr_countable : Countable expr. 198 | Proof. 199 | set (enc := 200 | fix go e := 201 | match e with 202 | | Val v => GenNode 0 [gov v] 203 | | Var x => GenLeaf (inl (inl x)) 204 | | Rec f x e => GenNode 1 [GenLeaf (inl (inr f)); GenLeaf (inl (inr x)); go e] 205 | | App e1 e2 => GenNode 2 [go e1; go e2] 206 | | UnOp op e => GenNode 3 [GenLeaf (inr (inr (inl op))); go e] 207 | | BinOp op e1 e2 => GenNode 4 [GenLeaf (inr (inr (inr (inl op)))); go e1; go e2] 208 | | HeapOp op e1 e2 => GenNode 5 [GenLeaf (inr (inr (inr (inr op)))); go e1; go e2] 209 | | If e0 e1 e2 => GenNode 6 [go e0; go e1; go e2] 210 | | Fork e => GenNode 7 [go e] 211 | end 212 | with gov v := 213 | match v with 214 | | LitV l => GenLeaf (inr (inl l)) 215 | | RecV f x e => 216 | GenNode 0 [GenLeaf (inl (inr f)); GenLeaf (inl (inr x)); go e] 217 | | PairV v1 v2 => GenNode 1 [gov v1; gov v2] 218 | end 219 | for go). 220 | set (dec := 221 | fix go e := 222 | match e with 223 | | GenNode 0 [v] => Val (gov v) 224 | | GenLeaf (inl (inl x)) => Var x 225 | | GenNode 1 [GenLeaf (inl (inr f)); GenLeaf (inl (inr x)); e] => Rec f x (go e) 226 | | GenNode 2 [e1; e2] => App (go e1) (go e2) 227 | | GenNode 3 [GenLeaf (inr (inr (inl op))); e] => UnOp op (go e) 228 | | GenNode 4 [GenLeaf (inr (inr (inr (inl op)))); e1; e2] => BinOp op (go e1) (go e2) 229 | | GenNode 5 [GenLeaf (inr (inr (inr (inr op)))); e1; e2] => HeapOp op (go e1) (go e2) 230 | | GenNode 6 [e0; e1; e2] => If (go e0) (go e1) (go e2) 231 | | GenNode 7 [e] => Fork (go e) 232 | | _ => Val $ LitV LitUnit (* dummy *) 233 | end 234 | with gov v := 235 | match v with 236 | | GenLeaf (inr (inl l)) => LitV l 237 | | GenNode 0 [GenLeaf (inl (inr f)); GenLeaf (inl (inr x)); e] => RecV f x (go e) 238 | | GenNode 1 [v1; v2] => PairV (gov v1) (gov v2) 239 | | _ => LitV LitUnit (* dummy *) 240 | end 241 | for go). 242 | refine (inj_countable' enc dec _). 243 | refine (fix go (e : expr) {struct e} := _ with gov (v : val) {struct v} := _ for go). 244 | - destruct e as [v| | | | | | | |]; simpl; f_equal; 245 | [exact (gov v)|done..]. 246 | - destruct v; by f_equal. 247 | Qed. 248 | Global Instance val_countable : Countable val. 249 | Proof. refine (inj_countable of_val to_val _); auto using to_of_val. Qed. 250 | 251 | (*| 252 | =========== 253 | Semantics 254 | =========== 255 | 256 | Now we can get to the fun stuff - the semantics! 257 | 258 | Actually before we do anything interesting in the semantics we'll define the 259 | contextual reduction rules, which specify what the next redex is. Evaluation 260 | contexts help us define how evaluation should recurse into expressions. 261 | 262 | We use a right-to-left evaluation order only to match heap_lang - see its 263 | documentation for why that's useful (it has to do with giving specifications for 264 | higher-order functions, nothing that'll come up here). 265 | 266 | Evaluation contexts here are commonly seen in a basic presentation on the lambda 267 | calculus, but we give a brief review anyway. You should read AppRCtx e1 the way 268 | we'd write e1 □ on the board - that is, we can reduce the argument to an 269 | application at any time. AppLCtx v2 is like □ v2, which says that we can reduce 270 | an application's function once the right-hand side is a value. These two rules 271 | are what determine a (deterministic) right-to-left evaluation order, and we 272 | consistently follow that convention for binary operations and heap operations as 273 | well. 274 | 275 | Technically these evaluation contexts items are given meaning by `fill_item` 276 | below, which specifies how to combine an evaluation context with an expression 277 | that goes in the "hole". Note that we define only evaluation context **items** 278 | here, so that the holes are not recursive. To make them recursive 279 | `ectxi_language` defines complete evaluation contexts as `list ectx_item`. 280 | |*) 281 | 282 | (** Evaluation contexts *) 283 | Inductive ectx_item := 284 | | AppLCtx (v2 : val) 285 | | AppRCtx (e1 : expr) 286 | | BinOpLCtx (op : bin_op) (v2 : val) 287 | | BinOpRCtx (op : bin_op) (e1 : expr) 288 | | UnOpCtx (op : un_op) 289 | | IfCtx (e1 e2 : expr) 290 | | HeapOpLCtx (op : heap_op) (v2 : val) 291 | | HeapOpRCtx (op : heap_op) (e1 : expr) 292 | . 293 | 294 | Definition fill_item (Ki : ectx_item) (e : expr) : expr := 295 | match Ki with 296 | | AppLCtx v2 => App e (of_val v2) 297 | | AppRCtx e1 => App e1 e 298 | | BinOpLCtx op v2 => BinOp op e (Val v2) 299 | | BinOpRCtx op e1 => BinOp op e1 e 300 | | UnOpCtx op => UnOp op e 301 | | IfCtx e1 e2 => If e e1 e2 302 | | HeapOpLCtx op v2 => HeapOp op e (Val v2) 303 | | HeapOpRCtx op e1 => HeapOp op e1 e 304 | end. 305 | 306 | (*| 307 | Context reductions define all the "boring" reduction rules; now we need to 308 | define how to reduce each constructor for `expr` when applied to values, the 309 | "head reduction" rules. The most important of these is the beta reduction rule 310 | for `App (RecV f x v1) v2`, namely substitution. Substitution is not capture 311 | avoiding but does not recurse into values (identified by the `Val` constructor, 312 | once again), which works properly as long as the expressions e are closed. This 313 | substitution definition only handles the non-recursive part of functions; the 314 | beta rule will substitute the function for itself recursively. 315 | |*) 316 | 317 | (** Substitution *) 318 | Fixpoint subst (x : string) (v : val) (e : expr) : expr := 319 | match e with 320 | | Val _ => e 321 | | Var y => if decide (x = y) then Val v else Var y 322 | | Rec f y e => 323 | Rec f y $ if decide (BNamed x ≠ f ∧ BNamed x ≠ y) then subst x v e else e 324 | | App e1 e2 => App (subst x v e1) (subst x v e2) 325 | | BinOp op e1 e2 => BinOp op (subst x v e1) (subst x v e2) 326 | | UnOp op e => UnOp op (subst x v e) 327 | | If e0 e1 e2 => If (subst x v e0) (subst x v e1) (subst x v e2) 328 | | Fork e => Fork (subst x v e) 329 | | HeapOp op e1 e2 => HeapOp op (subst x v e1) (subst x v e2) 330 | end. 331 | 332 | Definition subst' (mx : binder) (v : val) : expr → expr := 333 | match mx with BNamed x => subst x v | BAnon => id end. 334 | 335 | (*| 336 | Now we'll give the pure semantics of simp_lang. These two Gallina definitions 337 | `bin_op_eval` and `un_op_eval` define the semantics of all the pure operations, 338 | when the types of their arguments make sense. 339 | |*) 340 | 341 | Definition LitBool (b:bool) : base_lit := 342 | if b then LitInt 1 else LitInt 0. 343 | 344 | Definition bin_op_eval (op: bin_op) (v1 v2: val) : option val := 345 | match op with 346 | | PlusOp => match v1, v2 with 347 | | LitV (LitInt n1), LitV (LitInt n2) => 348 | Some (LitV (LitInt (n1 + n2))) 349 | | _, _ => None 350 | end 351 | | EqOp => Some (LitV $ LitBool $ bool_decide (v1 = v2)) 352 | | PairOp => Some (PairV v1 v2) 353 | end. 354 | 355 | Definition un_op_eval (op: un_op) (v: val) : option val := 356 | match op, v with 357 | | FstOp, PairV v1 v2 => Some v1 358 | | SndOp, PairV v1 v2 => Some v2 359 | | _, _ => None 360 | end. 361 | 362 | (*| 363 | To give a semantics for the heap operations we need some state for them to 364 | manipulate. This is a really simple language so its state is just a finite 365 | mapping from locations (integers) to values. We still wrap it in a record 366 | because this is good practice in case you want to extend it later. 367 | |*) 368 | 369 | Definition loc := Z. 370 | 371 | Record state : Type := { 372 | heap: gmap loc val; 373 | }. 374 | 375 | (** the language interface needs these things to be inhabited, I believe *) 376 | Global Instance state_inhabited : Inhabited state := 377 | populate {| heap := inhabitant; |}. 378 | Global Instance val_inhabited : Inhabited val := populate (LitV LitUnit). 379 | Global Instance expr_inhabited : Inhabited expr := populate (Val inhabitant). 380 | 381 | Definition state_upd_heap (f: gmap loc val → gmap loc val) (σ: state) : state := 382 | {| heap := f σ.(heap) |}. 383 | Global Arguments state_upd_heap _ !_ /. 384 | 385 | (*| 386 | The main semantics of simp_lang. `base_step e σ κs e' σ' efs` says that 387 | expression `e` in state `σ` reduces to `e'` and state `σ'` while forking threads 388 | `efs`. It also produces "observations" `κs`, which are related to prophecy 389 | variables and are unused here (in fact `observation` is an empty inductive so 390 | `κs = []`). 391 | 392 | These are the head reduction rules, which apply when we have a primitive at the 393 | head of the expression applied to values. For any other expression the Iris 394 | ectxi_language implementation will take care of using our evaluation contexts to 395 | find the next head step to take. 396 | |*) 397 | 398 | Inductive observation :=. 399 | Lemma observations_empty (κs: list observation) : κs = []. 400 | Proof. by destruct κs as [ | [] ]. Qed. 401 | 402 | Inductive base_step : expr → state → list observation → expr → state → list expr → Prop := 403 | | RecS f x e σ : 404 | base_step (Rec f x e) σ [] (Val $ RecV f x e) σ [] 405 | | BetaS f x e1 v2 e' σ : 406 | (* this is where recursive functions are implemented, by substituting the 407 | entire expression [RecV f x e1] for [f] *) 408 | e' = subst' x v2 (subst' f (RecV f x e1) e1) → 409 | base_step (App (Val $ RecV f x e1) (Val v2)) σ [] e' σ [] 410 | | BinOpS op v1 v2 v' σ : 411 | bin_op_eval op v1 v2 = Some v' → 412 | base_step (BinOp op (Val v1) (Val v2)) σ [] (Val v') σ [] 413 | | UnOpS op v v' σ : 414 | un_op_eval op v = Some v' → 415 | base_step (UnOp op (Val v)) σ [] (Val v') σ [] 416 | | IfFalseS e1 e2 σ : 417 | base_step (If (Val $ LitV $ LitInt 0) e1 e2) σ [] e2 σ [] 418 | | IfTrueS n e1 e2 σ : 419 | 0 ≠ n → 420 | base_step (If (Val $ LitV $ LitInt n) e1 e2) σ [] e1 σ [] 421 | | ForkS e σ: 422 | (* Fork is the only rule that spawns new threads *) 423 | base_step (Fork e) σ [] (Val $ LitV LitUnit) σ [e] 424 | | AllocS v σ l : 425 | σ.(heap) !! l = None → 426 | base_step (HeapOp AllocOp (Val v) (Val $ LitV LitUnit)) σ 427 | [] 428 | (Val $ LitV $ LitInt l) (state_upd_heap <[l := v]> σ) 429 | [] 430 | | LoadS v σ l : 431 | σ.(heap) !! l = Some v → 432 | base_step (HeapOp LoadOp (Val $ LitV $ LitInt l) (Val $ LitV LitUnit)) σ 433 | [] 434 | (Val $ v) σ 435 | [] 436 | | StoreS v w σ l : 437 | σ.(heap) !! l = Some v → 438 | base_step (HeapOp StoreOp (Val $ LitV $ LitInt l) (Val $ w)) σ 439 | [] 440 | (Val $ LitV $ LitUnit) (state_upd_heap <[l := w]> σ) 441 | [] 442 | | FaaS l n1 n2 σ : 443 | σ.(heap) !! l = Some (LitV $ LitInt $ n1) → 444 | base_step (HeapOp FaaOp (Val $ LitV $ LitInt l) (Val $ LitV $ LitInt $ n2)) σ 445 | [] 446 | (Val $ LitV $ LitInt $ n1) (state_upd_heap <[l:=LitV $ LitInt $ n1+n2]> σ) 447 | [] 448 | . 449 | 450 | (*| 451 | We have to prove a few properties that show `fill_item` and `base_step` is 452 | reasonable for `ectxi_language`. 453 | |*) 454 | 455 | Global Instance fill_item_inj Ki : Inj (=) (=) (fill_item Ki). 456 | Proof. destruct Ki; intros ???; simplify_eq/=; auto with f_equal. Qed. 457 | 458 | Lemma fill_item_val Ki e : 459 | is_Some (to_val (fill_item Ki e)) → is_Some (to_val e). 460 | Proof. intros [v ?]. destruct Ki; simplify_option_eq; eauto. Qed. 461 | 462 | Lemma val_base_stuck e1 σ1 κ e2 σ2 efs : base_step e1 σ1 κ e2 σ2 efs → to_val e1 = None. 463 | Proof. destruct 1; naive_solver. Qed. 464 | 465 | Lemma base_ctx_step_val Ki e σ1 κ e2 σ2 efs : 466 | base_step (fill_item Ki e) σ1 κ e2 σ2 efs → is_Some (to_val e). 467 | Proof. destruct Ki; inversion_clear 1; simplify_option_eq; eauto. Qed. 468 | 469 | Lemma fill_item_no_val_inj Ki1 Ki2 e1 e2 : 470 | to_val e1 = None → to_val e2 = None → 471 | fill_item Ki1 e1 = fill_item Ki2 e2 → Ki1 = Ki2. 472 | Proof. destruct Ki1, Ki2; naive_solver eauto with f_equal. Qed. 473 | 474 | Definition fresh_locs (ls : gset loc) : loc := 475 | set_fold (λ k r, (1 + k) `max` r) 1 ls. 476 | 477 | Lemma fresh_locs_fresh ls : 478 | fresh_locs ls ∉ ls. 479 | Proof. 480 | cut (∀ l, l ∈ ls → l < fresh_locs ls). 481 | { intros help Hf%help. lia. } 482 | apply (set_fold_ind_L (λ r ls, ∀ l, l ∈ ls → l < r)); 483 | set_solver by eauto with lia. 484 | Qed. 485 | 486 | (** this theorem will be needed to show that allocation is never stuck when we 487 | prove a WP for it *) 488 | Lemma alloc_fresh v σ : 489 | (* this invocation of [dom] is for backwards compatibility with Iris 3.6.0 *) 490 | let l := fresh_locs (@dom _ (gset _) _ σ.(heap)) in 491 | base_step (HeapOp AllocOp (Val v) (Val $ LitV $ LitUnit)) σ [] 492 | (Val $ LitV $ LitInt l) (state_upd_heap <[l := v]> σ) []. 493 | Proof. 494 | intros. 495 | apply AllocS. 496 | apply (not_elem_of_dom (D := gset loc)). 497 | by apply fresh_locs_fresh. 498 | Qed. 499 | 500 | (*| 501 | This is really where we instantiate the language, by constructing a "mixin" and 502 | then using some canonical structures to build the full record. 503 | 504 | You can see that the mixin uses `fill_item` and `base_step` as the core of the 505 | semantics. It uses `of_val` and `to_val` to define a number of related notions 506 | like reducible and not-stuck and such. 507 | |*) 508 | 509 | Lemma simp_lang_mixin : EctxiLanguageMixin of_val to_val fill_item base_step. 510 | Proof. 511 | split; apply _ || eauto using to_of_val, of_to_val, val_base_stuck, 512 | fill_item_val, fill_item_no_val_inj, base_ctx_step_val. 513 | Qed. 514 | 515 | (** Language *) 516 | Canonical Structure simp_ectxi_lang := EctxiLanguage simp_lang_mixin. 517 | Canonical Structure simp_ectx_lang := EctxLanguageOfEctxi simp_ectxi_lang. 518 | Canonical Structure simp_lang := LanguageOfEctx simp_ectx_lang. 519 | 520 | (* The [LanguageOfEctx] and [EctxLanguageOfEctxi] constructors together build 521 | the entire language semantics and instantiate the general Iris [language] 522 | interface. The semantics at this level is given as a single transition relation 523 | between configurations [cfg] (along with observations which we're ignoring): *) 524 | Check (@step simp_lang). 525 | (* 526 | step 527 | : cfg simp_lang → list (language.observation simp_lang) → cfg simp_lang → Prop 528 | *) 529 | 530 | (* A [cfg] is a [list expr * state]. The list of expressions is a thread pool 531 | accumulating all the spawned threads, where the first expression is the "main" 532 | thread whose return value we care about, and the type of state comes from our 533 | definition above. *) 534 | Eval compute in cfg simp_lang. 535 | (* 536 | = (list expr * state)%type 537 | : Type 538 | *) 539 | --------------------------------------------------------------------------------