From 3188e355b41539b38db8ceea9a0a89c27d8deb88 Mon Sep 17 00:00:00 2001 From: Kaptch Date: Mon, 29 Jan 2024 15:27:56 +0100 Subject: [PATCH 1/7] simplify ctxdep --- theories/gitree/greifiers.v | 491 +++++++++++++++++++++++++++-------- theories/gitree/reductions.v | 76 +++--- theories/gitree/reify.v | 453 +++++++++++++++++++++----------- theories/gitree/weakestpre.v | 477 +++++++++++++++++++++++----------- 4 files changed, 1048 insertions(+), 449 deletions(-) diff --git a/theories/gitree/greifiers.v b/theories/gitree/greifiers.v index 05ded03..f9ff81e 100644 --- a/theories/gitree/greifiers.v +++ b/theories/gitree/greifiers.v @@ -3,11 +3,15 @@ From iris.base_logic.lib Require Export invariants. From gitrees Require Import prelude. From gitrees.gitree Require Import core reify. - -Section greifiers. +Section greifiers_generic. #[local] Open Scope type. + Context (a : is_ctx_dep). + Notation sReifier := (sReifier a). + Notation sReifier_ops := (sReifier_ops a). + Notation sReifier_state := (sReifier_state a). + Notation sReifier_re := (sReifier_re a). - (** Global reifiers: a collection of reifiers *) + (** Global reifiers: a collection of reifiers *) Inductive gReifiers : nat → Type := | gReifiers_nil : gReifiers 0 | gReifiers_cons {n} : sReifier → gReifiers n → gReifiers (S n) @@ -18,10 +22,11 @@ Section greifiers. Proof. revert P Hcons. refine match v with gReifiers_nil => tt - | gReifiers_cons sR v => λ P Hcons, Hcons sR v end. + | gReifiers_cons sR v => λ P Hcons, Hcons sR v end. Defined. - #[global] Instance gReifiers_lookup_total : ∀ m, LookupTotal (fin m) sReifier (gReifiers m) := + #[global] Instance gReifiers_lookup_total + : ∀ m, LookupTotal (fin m) sReifier (gReifiers m) := fix go m i {struct i} := let _ : ∀ m, LookupTotal _ _ _ := @go in match i in fin m return gReifiers m → sReifier with | 0%fin => grs_S_inv (λ _, sReifier) (λ x _, x) @@ -30,7 +35,7 @@ Section greifiers. Program Definition gReifiers_ops {n} (rs : gReifiers n) : opsInterp := {| opid := { i : fin n & opid (sReifier_ops (rs !!! i)) }; - opsInterp_lookup := λ iop, sReifier_ops (rs !!! projT1 iop) (projT2 iop) + opsInterp_lookup := λ iop, sReifier_ops (rs !!! projT1 iop) (projT2 iop) |}. Next Obligation. intros n rs. @@ -65,14 +70,14 @@ Section greifiers. Fixpoint gState_rest {m} (i : fin m) : gReifiers m → oFunctor := match i in fin m return gReifiers m → oFunctor with | 0%fin => grs_S_inv (λ _, oFunctor) - (λ sR rs, gReifiers_state rs) + (λ sR rs, gReifiers_state rs) | FS j => grs_S_inv (λ _, oFunctor) - (λ sR rs, sReifier_state sR * gState_rest j rs)%OF + (λ sR rs, sReifier_state sR * gState_rest j rs)%OF end. Lemma gState_decomp' {m} (i : fin m) (rs : gReifiers m) {X} `{!Cofe X} : gReifiers_state rs ♯ X ≃ - ((sReifier_state (rs !!! i) ♯ X) * (gState_rest i rs ♯ X))%type. + ((sReifier_state (rs !!! i) ♯ X) * (gState_rest i rs ♯ X))%type. Proof. revert i. induction rs as [|n r rs]=>i. { inversion i. } @@ -105,47 +110,116 @@ Section greifiers. Proof. rewrite ofe_iso_12. reflexivity. Qed. Lemma gState_recomp_decomp {m} (i : fin m) {rs : gReifiers m} {X} `{!Cofe X} (σ : sReifier_state (rs !!! i) ♯ X) rest fs : - gState_decomp i fs ≡ (σ, rest) → - gState_recomp rest σ ≡ fs. + gState_decomp i fs ≡ (σ, rest) → + gState_recomp rest σ ≡ fs. Proof. unfold gState_recomp. simpl. intros <-. rewrite ofe_iso_21//. Qed. Opaque gState_recomp gState_decomp. - Program Definition gReifiers_re {n} (rs : gReifiers n) {X} `{!Cofe X} - (op : opid (gReifiers_ops rs)) : - (Ins (gReifiers_ops rs op) ♯ X) * (gReifiers_state rs ♯ X) * ((Outs (gReifiers_ops rs op) ♯ X) -n> laterO X) -n> - optionO (laterO X * (gReifiers_state rs ♯ X)) +End greifiers_generic. + +Section greifiers. + #[local] Open Scope type. + + Program Definition gReifiers_re_ctx_indep {n} (rs : gReifiers NotCtxDep n) {X} `{!Cofe X} + (op : opid (gReifiers_ops NotCtxDep rs)) : + (Ins (gReifiers_ops NotCtxDep rs op) ♯ X) * (gReifiers_state NotCtxDep rs ♯ X) -n> + optionO ((Outs (gReifiers_ops NotCtxDep rs op) ♯ X) * (gReifiers_state NotCtxDep rs ♯ X)) + := λne xst, + let i := projT1 op in + let op' := projT2 op in + let x := xst.1 in + let st := xst.2 in + let fs := gState_decomp NotCtxDep i st in + let σ := fs.1 in + let rest := fs.2 in + let rx := sReifier_re NotCtxDep (rs !!! i) op' (x, σ) in + optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) rx. + Next Obligation. solve_proper_please. Qed. + + Program Definition gReifiers_re_ctx_dep {n} (rs : gReifiers CtxDep n) {X} `{!Cofe X} + (op : opid (gReifiers_ops CtxDep rs)) : + (Ins (gReifiers_ops CtxDep rs op) ♯ X) * (gReifiers_state CtxDep rs ♯ X) + * ((Outs (gReifiers_ops CtxDep rs op) ♯ X) -n> laterO X) -n> + optionO (laterO X * (gReifiers_state CtxDep rs ♯ X)) := λne xst, let i := projT1 op in let op' := projT2 op in let a := xst.1.1 in let b := xst.1.2 in let c := xst.2 in - let fs := gState_decomp i b in + let fs := gState_decomp CtxDep i b in let σ := fs.1 in let rest := fs.2 in - let rx := sReifier_re (rs !!! i) op' (a, σ, c) in - optionO_map (prodO_map idfun (gState_recomp rest)) rx. + let rx := sReifier_re CtxDep (rs !!! i) op' (a, σ, c) in + optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) rx. Next Obligation. solve_proper_please. Qed. + Program Definition gReifiers_re_type {n} + (a : is_ctx_dep) (rs : gReifiers a n) {X} `{!Cofe X} + (op : opid (gReifiers_ops a rs)) : ofe := + match a with + | CtxDep => + prodO (prodO (Ins (gReifiers_ops a rs op) ♯ X) (gReifiers_state a rs ♯ X)) + (Outs (gReifiers_ops a rs op) ♯ X -n> laterO X) -n> + optionO (prodO (laterO X) (gReifiers_state a rs ♯ X)) + | NotCtxDep => + prodO (Ins (gReifiers_ops a rs op) ♯ X) (gReifiers_state a rs ♯ X) -n> + optionO (prodO (Outs (gReifiers_ops a rs op) ♯ X) (gReifiers_state a rs ♯ X)) + end. + + Program Definition gReifiers_re {n} (a : is_ctx_dep) (rs : gReifiers a n) + {X} `{!Cofe X} + (op : opid (gReifiers_ops a rs)) : + @gReifiers_re_type n a rs X _ op. + Proof. + destruct a. + - apply gReifiers_re_ctx_dep. + - apply gReifiers_re_ctx_indep. + Defined. + (** We can turn a collection of reifiers into a single reifier *) - Definition gReifiers_sReifier {n} (rs : gReifiers n) : sReifier := - {| sReifier_ops := gReifiers_ops rs; - sReifier_state := gReifiers_state rs; - sReifier_re := @gReifiers_re n rs; + Program Definition gReifiers_sReifier {n} (a : is_ctx_dep) (rs : gReifiers a n) + : sReifier a := + {| sReifier_ops := gReifiers_ops a rs; + sReifier_state := gReifiers_state a rs; + sReifier_re X _ op := _; |}. + Next Obligation. + intros n [|] rs X ? op. + - apply (@gReifiers_re n CtxDep rs X _ op). + - apply (@gReifiers_re n NotCtxDep rs X _ op). + Defined. + + Lemma gReifiers_re_idx_ctx_dep {n} (i : fin n) (rs : gReifiers CtxDep n) + {X} `{!Cofe X} (op : opid (sReifier_ops CtxDep (rs !!! i))) + (x : Ins (sReifier_ops CtxDep _ op) ♯ X) + (σ : sReifier_state CtxDep (rs !!! i) ♯ X) + (rest : gState_rest CtxDep i rs ♯ X) + (κ : (Outs (sReifier_ops CtxDep (rs !!! i) op) ♯ X -n> laterO X)) : + gReifiers_re CtxDep rs (existT i op) (x, gState_recomp CtxDep rest σ, κ) ≡ + optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) + (sReifier_re CtxDep (rs !!! i) op (x, σ, κ)). + Proof. + unfold gReifiers_re. cbn-[prodO_map optionO_map]. + f_equiv; last repeat f_equiv. + - eapply optionO_map_proper. + intros [x1 x2]; simpl. f_equiv. + f_equiv. f_equiv. + rewrite gState_decomp_recomp//. + - rewrite gState_decomp_recomp//. + Qed. - Lemma gReifiers_re_idx {n} (i : fin n) (rs : gReifiers n) - {X} `{!Cofe X} (op : opid (sReifier_ops (rs !!! i))) - (x : Ins (sReifier_ops _ op) ♯ X) - (σ : sReifier_state (rs !!! i) ♯ X) - (κ : (Outs (sReifier_ops (rs !!! i) op) ♯ X -n> laterO X)) - (rest : gState_rest i rs ♯ X) : - gReifiers_re rs (existT i op) (x, gState_recomp rest σ, κ) ≡ - optionO_map (prodO_map idfun (gState_recomp rest)) - (sReifier_re (rs !!! i) op (x, σ, κ)). + Lemma gReifiers_re_idx_ctx_indep {n} (i : fin n) (rs : gReifiers NotCtxDep n) + {X} `{!Cofe X} (op : opid (sReifier_ops NotCtxDep (rs !!! i))) + (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) + (σ : sReifier_state NotCtxDep (rs !!! i) ♯ X) + (rest : gState_rest NotCtxDep i rs ♯ X) : + gReifiers_re NotCtxDep rs (existT i op) (x, gState_recomp NotCtxDep rest σ) ≡ + optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) + (sReifier_re NotCtxDep (rs !!! i) op (x, σ)). Proof. unfold gReifiers_re. cbn-[prodO_map optionO_map]. f_equiv; last repeat f_equiv. @@ -156,54 +230,117 @@ Section greifiers. - rewrite gState_decomp_recomp//. Qed. - Class subReifier {n} (r : sReifier) (rs : gReifiers n) := + Program Definition gReifiers_re_idx_type {n} a (i : fin n) (rs : gReifiers a n) + {X} `{!Cofe X} (op : opid (sReifier_ops a (rs !!! i))) + (x : Ins (sReifier_ops a _ op) ♯ X) + (σ : sReifier_state a (rs !!! i) ♯ X) + (rest : gState_rest a i rs ♯ X) : + Type. + Proof. + destruct a. + - apply (∀ (κ : (Outs (sReifier_ops CtxDep (rs !!! i) op) ♯ X -n> laterO X)), + gReifiers_re CtxDep rs (existT i op) (x, gState_recomp CtxDep rest σ, κ) ≡ + optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) + (sReifier_re CtxDep (rs !!! i) op (x, σ, κ))). + - apply (gReifiers_re NotCtxDep rs (existT i op) (x, gState_recomp NotCtxDep rest σ) ≡ + optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) + (sReifier_re NotCtxDep (rs !!! i) op (x, σ))). + Defined. + + Lemma gReifiers_re_idx {n} a (i : fin n) (rs : gReifiers a n) + {X} `{!Cofe X} (op : opid (sReifier_ops a (rs !!! i))) + (x : Ins (sReifier_ops a _ op) ♯ X) + (σ : sReifier_state a (rs !!! i) ♯ X) + (rest : gState_rest a i rs ♯ X) : gReifiers_re_idx_type a i rs op x σ rest. + Proof. + destruct a. + - intros κ. apply gReifiers_re_idx_ctx_dep. + - apply gReifiers_re_idx_ctx_indep. + Qed. + + Program Definition sR_re_type {n} + {X} `{!Cofe X} (a : is_ctx_dep) (r : sReifier a) (rs : gReifiers a n) + (sR_idx : fin n) + (sR_ops : subEff (sReifier_ops a r) (sReifier_ops a (rs !!! sR_idx))) + (sR_state : sReifier_state a r ♯ X ≃ sReifier_state a (rs !!! sR_idx) ♯ X) + (m : nat) (op : opid (sReifier_ops a r)) : Type. + Proof. + destruct a. + - apply (∀ (x : Ins (sReifier_ops CtxDep r op) ♯ X) + (y : laterO X) + (s1 s2 : sReifier_state CtxDep r ♯ X) + (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)), + sReifier_re CtxDep r op (x, s1, k) ≡{m}≡ Some (y, s2) → + sReifier_re CtxDep (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡{m}≡ + Some (y, sR_state s2)). + - apply (∀ (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) + (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) + (s1 s2 : sReifier_state NotCtxDep r ♯ X), + sReifier_re NotCtxDep r op (x, s1) ≡{m}≡ Some (y, s2) → + sReifier_re NotCtxDep (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1) ≡{m}≡ + Some (subEff_outs y, sR_state s2)). + Defined. + + Class subReifier {n} {a : is_ctx_dep} (r : sReifier a) (rs : gReifiers a n) := { sR_idx : fin n; - sR_ops :: subEff (sReifier_ops r) (sReifier_ops (rs !!! sR_idx)); + sR_ops :: subEff (sReifier_ops a r) (sReifier_ops a (rs !!! sR_idx)); sR_state {X} `{!Cofe X} : - sReifier_state r ♯ X ≃ sReifier_state (rs !!! sR_idx) ♯ X; - sR_re (m : nat) {X} `{!Cofe X} (op : opid (sReifier_ops r)) - (x : Ins (sReifier_ops r op) ♯ X) - (y : laterO X) - (s1 s2 : sReifier_state r ♯ X) - (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) : - sReifier_re r op (x, s1, k) ≡{m}≡ Some (y, s2) → - sReifier_re (rs !!! sR_idx) (subEff_opid op) - (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡{m}≡ - Some (y, sR_state s2) + sReifier_state a r ♯ X ≃ sReifier_state a (rs !!! sR_idx) ♯ X; + sR_re (m : nat) {X} `{!Cofe X} (op : opid (sReifier_ops a r)) + : sR_re_type a r rs sR_idx sR_ops (@sR_state X _) m op }. - #[global] Instance subReifier_here {n} (r : sReifier) (rs : gReifiers n) : - subReifier r (gReifiers_cons r rs). + #[global] Instance subReifier_here {n} {a : is_ctx_dep} + (r : sReifier a) (rs : gReifiers a n) : + subReifier r (gReifiers_cons a r rs). Proof. simple refine ({| sR_idx := 0%fin |}). - simpl. apply subEff_id. - simpl. intros. apply iso_ofe_refl. - - intros X ? op x y ? s1 s2 k HEQ; simpl. - unfold ofe_iso_1'; simpl. - rewrite ccompose_id_r HEQ. - reflexivity. + - destruct a. + + intros X ? op x y ? s1 s2 k HEQ; simpl. + unfold ofe_iso_1'; simpl. + rewrite ccompose_id_r HEQ. + reflexivity. + + intros X ? op x y s1 s2. + simpl. eauto. Defined. - #[global] Instance subReifier_there {n} (r r' : sReifier) (rs : gReifiers n) : + #[global] Instance subReifier_there {n} {a : is_ctx_dep} + (r r' : sReifier a) (rs : gReifiers a n) : subReifier r rs → - subReifier r (gReifiers_cons r' rs). + subReifier r (gReifiers_cons a r' rs). Proof. intros SR. simple refine ({| sR_idx := FS sR_idx |}). - simpl. intros. apply sR_state. - - intros X ? op x y s1 s2. - simpl. apply sR_re. + - destruct a. + + intros ? X ? op x y s1 s2. + simpl. intros. + pose proof (@sR_re n CtxDep r rs _ m X _ op) as G. + simpl in G. + by apply G. + + intros ? X ? op x y s1 s2. + simpl. intros. + pose proof (@sR_re n NotCtxDep r rs _ m X _ op) as G. + simpl in G. + by apply G. Defined. - #[local] Definition subR_op {n} {r : sReifier} {rs : gReifiers n} `{!subReifier r rs} : - opid (sReifier_ops r) → opid (gReifiers_ops rs). + #[local] Definition subR_op {n} {a : is_ctx_dep} + {r : sReifier a} {rs : gReifiers a n} `{!subReifier r rs} : + opid (sReifier_ops a r) → opid (gReifiers_ops a rs). Proof. intros op. simpl. refine (existT sR_idx (subEff_opid op)). Defined. - #[export] Instance subReifier_subEff {n} {r : sReifier} {rs : gReifiers n} `{!subReifier r rs} : - subEff (sReifier_ops r) (gReifiers_ops rs). + + #[export] Instance subReifier_subEff {n} {a : is_ctx_dep} + {r : sReifier a} {rs : gReifiers a n} `{!subReifier r rs} : + subEff (sReifier_ops a r) (gReifiers_ops a rs). Proof. simple refine {| subEff_opid := subR_op |}. - intros op X ?. simpl. @@ -212,90 +349,216 @@ Section greifiers. apply subEff_outs. Defined. - Lemma subReifier_reify_idx {n} (r : sReifier) (rs : gReifiers n) - `{!subReifier r rs} {X} `{!Cofe X} (op : opid (sReifier_ops r)) - (x : Ins (sReifier_ops _ op) ♯ X) - (* (y : Outs (sReifier_ops _ op) ♯ X) *) - (y : laterO X) - (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) - (s1 s2 : sReifier_state r ♯ X) : - sReifier_re r op (x, s1, k) ≡ Some (y, s2) → - sReifier_re (rs !!! sR_idx) (subEff_opid op) - (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ - Some (y, sR_state s2). + Program Definition subReifier_reify_idx_type {n} + (a : is_ctx_dep) (r : sReifier a) (rs : gReifiers a n) + `{!subReifier r rs} X `{!Cofe X} (op : opid (sReifier_ops a r)) : Type. + Proof. + destruct a. + - apply (∀ (x : Ins (sReifier_ops CtxDep r op) ♯ X) + (y : laterO X) + (s1 s2 : sReifier_state CtxDep r ♯ X) + (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)), + sReifier_re CtxDep r op (x, s1, k) ≡ Some (y, s2) → + sReifier_re CtxDep (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ + Some (y, sR_state s2)). + - apply (∀ (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) + (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) + (s1 s2 : sReifier_state NotCtxDep r ♯ X), + sReifier_re NotCtxDep r op (x, s1) ≡ Some (y, s2) → + sReifier_re NotCtxDep (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1) ≡ + Some (subEff_outs y, sR_state s2)). + Defined. + + Lemma subReifier_reify_idx {n} {a : is_ctx_dep} + (r : sReifier a) (rs : gReifiers a n) + `{!subReifier r rs} {X} `{!Cofe X} (op : opid (sReifier_ops a r)) + : subReifier_reify_idx_type a r rs X op. Proof. - intros Hx. apply equiv_dist=>m. - apply sR_re. by apply equiv_dist. + destruct a. + - intros Hx. intros. apply equiv_dist=>m. + pose proof (@sR_re n CtxDep r rs _ m X _ op Hx y s1 s2 k) as G. + simpl in G. + rewrite G; first done. + by apply equiv_dist. + - intros Hx. intros. apply equiv_dist=>m. + pose proof (@sR_re n NotCtxDep r rs _ m X _ op Hx y s1 s2) as G. + simpl in G. + rewrite G; first done. + by apply equiv_dist. Qed. - Lemma subReifier_reify {n} (r : sReifier) - (rs : gReifiers n) `{!subReifier r rs} {X} `{!Cofe X} - (op : opid (sReifier_ops r)) - (x : Ins (sReifier_ops _ op) ♯ X) (y : laterO X) - (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) - (σ σ' : sReifier_state r ♯ X) (rest : gState_rest sR_idx rs ♯ X) : - sReifier_re r op (x, σ, k) ≡ Some (y, σ') → - gReifiers_re rs (subEff_opid op) - (subEff_ins x, gState_recomp rest (sR_state σ), k ◎ (subEff_outs ^-1)) - ≡ Some (y, gState_recomp rest (sR_state σ')). + Program Definition subReifier_reify_type {n} (a : is_ctx_dep) (r : sReifier a) + (rs : gReifiers a n) `{!subReifier r rs} X `{!Cofe X} + (op : opid (sReifier_ops a r)) : Type. + Proof. + destruct a. + - apply (∀ (x : Ins (sReifier_ops CtxDep _ op) ♯ X) (y : laterO X) + (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)) + (σ σ' : sReifier_state CtxDep r ♯ X) (rest : gState_rest CtxDep sR_idx rs ♯ X), + sReifier_re CtxDep r op (x, σ, k) ≡ Some (y, σ') → + gReifiers_re CtxDep rs (subEff_opid op) + (subEff_ins x, gState_recomp CtxDep rest (sR_state σ), k ◎ (subEff_outs ^-1)) + ≡ Some (y, gState_recomp CtxDep rest (sR_state σ'))). + - apply (∀ (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) + (σ σ' : sReifier_state NotCtxDep r ♯ X) (rest : gState_rest NotCtxDep sR_idx rs ♯ X), + sReifier_re NotCtxDep r op (x,σ) ≡ Some (y, σ') → + gReifiers_re NotCtxDep rs (subEff_opid op) + (subEff_ins x, gState_recomp NotCtxDep rest (sR_state σ)) + ≡ Some (subEff_outs y, gState_recomp NotCtxDep rest (sR_state σ'))). + Defined. + + Lemma subReifier_reify {n} {a : is_ctx_dep} (r : sReifier a) + (rs : gReifiers a n) `{!subReifier r rs} {X} `{!Cofe X} + (op : opid (sReifier_ops a r)) : subReifier_reify_type a r rs X op. Proof. - intros Hre. - eapply subReifier_reify_idx in Hre. - rewrite gReifiers_re_idx//. - rewrite Hre. simpl. - do 3 f_equiv. + destruct a. + - simpl. + intros x y k σ σ' H Hre. + pose proof (@subReifier_reify_idx n CtxDep r rs _ X _ op x y σ σ' k Hre) as J; clear Hre. + simpl in J. + pose proof (@gReifiers_re_idx n CtxDep sR_idx rs X _ (subEff_opid op)) as J'. + simpl in J'. + rewrite J'; clear J'. + transitivity (prod_map (λ x0 : laterO X, x0) + (λ st : sReifier_state CtxDep (rs !!! sR_idx) ♯ X, + (gState_decomp' CtxDep sR_idx rs ^-1) (st, H)) <$> + (Some (y, sR_state σ'))). + + unfold prod_map. + rewrite option_fmap_proper; [reflexivity | intros ??? | apply J]. + do 2 f_equiv; first assumption. + do 2 f_equiv; assumption. + + simpl; reflexivity. + - simpl. + intros x y σ σ' rest Hre. + pose proof (@subReifier_reify_idx n NotCtxDep r rs _ X _ op x y σ σ' Hre) + as J; clear Hre. + simpl in J. + pose proof (@gReifiers_re_idx n NotCtxDep sR_idx rs X _ (subEff_opid op)) + as J'. + simpl in J'. + rewrite J'; clear J'. + transitivity (prod_map (λ x0 : Outs (sReifier_ops NotCtxDep (rs !!! sR_idx) + (subEff_opid op)) ♯ X, x0) + (λ st : sReifier_state NotCtxDep (rs !!! sR_idx) ♯ X, + (gState_decomp' NotCtxDep sR_idx rs ^-1) (st, rest)) <$> + (Some (subEff_outs y, sR_state σ'))). + + unfold prod_map. + rewrite option_fmap_proper; [reflexivity | intros ??? | apply J]. + do 2 f_equiv; first assumption. + do 2 f_equiv; assumption. + + simpl; reflexivity. Qed. (** Lemma for reasoning internally in iProp *) Context `{!invGS_gen hlc Σ}. Notation iProp := (iProp Σ). Context {sz : nat}. - Variable (rs : gReifiers sz). - Notation sr := (gReifiers_sReifier rs). + Notation sr a rs := (gReifiers_sReifier a rs). - Lemma reify_vis_eqI {A} `{!Cofe A} op i k o σ σ' : - (gReifiers_re rs op (i, σ, k) ≡ Some (o,σ') ⊢@{iProp} reify sr (Vis op i k : IT _ A) σ ≡ (σ', Tau $ o))%I. + Lemma reify_vis_eqI_ctx_dep (rs : gReifiers CtxDep sz) + {A} `{!Cofe A} op i k o σ σ' : + (gReifiers_re CtxDep rs op (i, σ, k) ≡ Some (o,σ') + ⊢@{iProp} reify (sr CtxDep rs) (Vis op i k : IT _ A) σ ≡ (σ', Tau $ o))%I. Proof. apply uPred.internal_eq_entails=>m. - intros H. apply reify_vis_dist. exact H. + intros H. apply reify_vis_dist_ctx_dep. exact H. Qed. - Lemma subReifier_reify_idxI (r : sReifier) - `{!subReifier r rs} {X} `{!Cofe X} (op : opid (sReifier_ops r)) - (x : Ins (sReifier_ops _ op) ♯ X) + Lemma reify_vis_eqI_ctx_indep (rs : gReifiers NotCtxDep sz) + {A} `{!Cofe A} op i k o σ σ' : + (gReifiers_re NotCtxDep rs op (i, σ) ≡ Some (o,σ') + ⊢@{iProp} reify (sr NotCtxDep rs) (Vis op i k : IT _ A) σ ≡ (σ', Tau $ k o))%I. + Proof. + apply uPred.internal_eq_entails=>m. + intros H. apply reify_vis_dist_ctx_indep. exact H. + Qed. + + Lemma subReifier_reify_idxI_ctx_dep (r : sReifier CtxDep) + `{!@subReifier sz CtxDep r rs} {X} `{!Cofe X} (op : opid (sReifier_ops CtxDep r)) + (x : Ins (sReifier_ops CtxDep _ op) ♯ X) (y : laterO X) - (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) - (s1 s2 : sReifier_state r ♯ X) : - sReifier_re r op (x, s1, k) ≡ Some (y, s2) ⊢@{iProp} - sReifier_re (rs !!! sR_idx) (subEff_opid op) - (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ - Some (y, sR_state s2). + (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)) + (s1 s2 : sReifier_state CtxDep r ♯ X) : + sReifier_re CtxDep r op (x, s1, k) ≡ Some (y, s2) ⊢@{iProp} + sReifier_re CtxDep (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ + Some (y, sR_state s2). Proof. apply uPred.internal_eq_entails=>m. - intros H. - rewrite sR_re; last first. - - rewrite H. + intros H'. + rewrite (sR_re (a := CtxDep)); last first. + - rewrite H'. reflexivity. - reflexivity. Qed. - Lemma subReifier_reifyI (r : sReifier) - `{!subReifier r rs} {X} `{!Cofe X} - (op : opid (sReifier_ops r)) - (x : Ins (sReifier_ops _ op) ♯ X) (y : laterO X) - (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) - (σ σ' : sReifier_state r ♯ X) (rest : gState_rest sR_idx rs ♯ X) : - sReifier_re r op (x,σ, k) ≡ Some (y, σ') ⊢@{iProp} - gReifiers_re rs (subEff_opid op) - (subEff_ins x, gState_recomp rest (sR_state σ), k ◎ (subEff_outs ^-1)) - ≡ Some (y, gState_recomp rest (sR_state σ')). + Lemma subReifier_reify_idxI_ctx_indep (r : sReifier NotCtxDep) + `{!@subReifier sz NotCtxDep r rs} {X} `{!Cofe X} (op : opid (sReifier_ops NotCtxDep r)) + (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) + (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) + (s1 s2 : sReifier_state NotCtxDep r ♯ X) : + sReifier_re NotCtxDep r op (x, s1) ≡ Some (y, s2) + ⊢@{iProp} + sReifier_re NotCtxDep (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1) ≡ + Some (subEff_outs y, sR_state s2). + Proof. + apply uPred.internal_eq_entails=>m. + apply (sR_re (a := NotCtxDep)). + Qed. + + Lemma subReifier_reifyI_ctx_dep (r : sReifier CtxDep) + `{!@subReifier sz CtxDep r rs} {X} `{!Cofe X} + (op : opid (sReifier_ops CtxDep r)) + (x : Ins (sReifier_ops CtxDep _ op) ♯ X) (y : laterO X) + (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)) + (σ σ' : sReifier_state CtxDep r ♯ X) (rest : gState_rest CtxDep sR_idx rs ♯ X) : + sReifier_re CtxDep r op (x,σ, k) ≡ Some (y, σ') ⊢@{iProp} + gReifiers_re CtxDep rs (subEff_opid op) + (subEff_ins x, gState_recomp CtxDep rest (sR_state σ), k ◎ (subEff_outs ^-1)) + ≡ Some (y, gState_recomp CtxDep rest (sR_state σ')). Proof. apply uPred.internal_eq_entails=>m. intros He. - eapply sR_re in He. - rewrite gReifiers_re_idx//. + eapply (sR_re (a := CtxDep)) in He. + rewrite (gReifiers_re_idx CtxDep)//. rewrite He. simpl. reflexivity. Qed. + Lemma subReifier_reifyI_ctx_indep (r : sReifier NotCtxDep) + `{!@subReifier sz NotCtxDep r rs} {X} `{!Cofe X} + (op : opid (sReifier_ops NotCtxDep r)) + (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) + (σ σ' : sReifier_state NotCtxDep r ♯ X) (rest : gState_rest NotCtxDep sR_idx rs ♯ X) : + sReifier_re NotCtxDep r op (x,σ) ≡ Some (y, σ') + ⊢@{iProp} + gReifiers_re NotCtxDep rs (subEff_opid op) + (subEff_ins x, gState_recomp NotCtxDep rest (sR_state σ)) + ≡ Some (subEff_outs y, gState_recomp NotCtxDep rest (sR_state σ')). + Proof. + apply uPred.internal_eq_entails=>m. + intros He. + eapply (sR_re (a := NotCtxDep)) in He. + pose proof (@gReifiers_re_idx sz NotCtxDep sR_idx rs X _ (subEff_opid op) + (subEff_ins x)) as J. + simpl in J. + simpl. + rewrite J//; clear J. + transitivity (prod_map (λ x0 : Outs (sReifier_ops NotCtxDep (rs !!! sR_idx) + (subEff_opid op)) ♯ X, x0) + (λ st : sReifier_state NotCtxDep (rs !!! sR_idx) ♯ X, + (gState_decomp' NotCtxDep sR_idx rs ^-1) (st, rest)) <$> + (Some + (subEff_outs y, sR_state σ'))). + - unfold prod_map. + rewrite option_fmap_ne; [reflexivity | intros ??? | apply He]. + do 2 f_equiv; first assumption. + do 2 f_equiv; assumption. + - simpl. + reflexivity. + Qed. + End greifiers. diff --git a/theories/gitree/reductions.v b/theories/gitree/reductions.v index 8036625..cfd14bd 100644 --- a/theories/gitree/reductions.v +++ b/theories/gitree/reductions.v @@ -4,16 +4,13 @@ From gitrees Require Import prelude. From gitrees.gitree Require Import core reify. Section sstep. - Context {A} `{!Cofe A}. - Context (r : sReifier). - Notation F := (sReifier_ops r). - Notation stateF := (sReifier_state r). + Context {A} `{!Cofe A} {a : is_ctx_dep}. + Context (r : sReifier a). + Notation F := (sReifier_ops a r). + Notation stateF := (sReifier_state a r). Notation IT := (IT F A). Notation ITV := (ITV F A). Notation stateO := (stateF ♯ IT). - Implicit Type op : opid F. - Implicit Type α β : IT. - Implicit Type σ : stateO. (** ** Reductions at the Prop level *) Inductive sstep : IT → stateO → IT → stateO → Prop := @@ -25,7 +22,8 @@ Section sstep. α ≡ Vis op i k → reify r α σ1 ≡ (σ2, Tick β) → sstep α σ1 β σ2. - #[export] Instance sstep_proper : Proper ((≡) ==> (≡) ==> (≡) ==> (≡) ==> (iff)) sstep. + #[export] Instance sstep_proper + : Proper ((≡) ==> (≡) ==> (≡) ==> (≡) ==> (iff)) sstep. Proof. intros α1 α2 Ha σ1 σ2 Hs β1 β2 Hb σ'1 σ'2 Hs'. simplify_eq/=. split. @@ -47,13 +45,14 @@ Section sstep. rewrite -H0. repeat f_equiv; eauto. Qed. - Inductive ssteps : IT → stateO → IT → stateO → nat → Prop := + Inductive ssteps : IT → stateO → IT → stateO → nat → Prop := | ssteps_zero α β σ σ' : α ≡ β → σ ≡ σ' → ssteps α σ β σ' 0 | ssteps_many α1 σ1 α2 σ2 α3 σ3 n2 : sstep α1 σ1 α2 σ2 → ssteps α2 σ2 α3 σ3 n2 → ssteps α1 σ1 α3 σ3 (1+n2). - #[export] Instance ssteps_proper : Proper ((≡) ==> (≡) ==> (≡) ==> (≡) ==> (=) ==> (iff)) ssteps. + #[export] Instance ssteps_proper + : Proper ((≡) ==> (≡) ==> (≡) ==> (≡) ==> (=) ==> (iff)) ssteps. Proof. intros α α' Ha σ σ' Hs β β' Hb σ2 σ2' Hs2 n1 n2 Hnm. fold_leibniz. simplify_eq/=. @@ -87,10 +86,10 @@ Section sstep. End sstep. Section istep. - Context {A} `{!Cofe A}. - Context (r : sReifier). - Notation F := (sReifier_ops r). - Notation stateF := (sReifier_state r). + Context {A} `{!Cofe A} {a : is_ctx_dep}. + Context (r : sReifier a). + Notation F := (sReifier_ops a r). + Notation stateF := (sReifier_state a r). Notation IT := (IT F A). Notation ITV := (ITV F A). Notation stateO := (stateF ♯ IT). @@ -105,11 +104,11 @@ Section istep. Solve All Obligations with solve_proper. Program Definition isteps_pre - (self : IT -n> stateO -n> IT -n> stateO -n> natO -n> iProp): + (self : IT -n> stateO -n> IT -n> stateO -n> natO -n> iProp): IT -n> stateO -n> IT -n> stateO -n> natO -n> iProp := λne α σ β σ' n, ((n ≡ 0 ∧ α ≡ β ∧ σ ≡ σ') - ∨ (∃ n' α0 σ0, n ≡ (1+n') ∧ istep α σ α0 σ0 ∧ - ▷ self α0 σ0 β σ' n'))%I. + ∨ (∃ n' α0 σ0, n ≡ (1+n') ∧ istep α σ α0 σ0 ∧ + ▷ self α0 σ0 β σ' n'))%I. Solve All Obligations with solve_proper. Global Instance isteps_pre_ne : NonExpansive isteps_pre. @@ -122,8 +121,8 @@ Section istep. Lemma isteps_unfold α β σ σ' n : isteps α σ β σ' n ≡ ((n ≡ 0 ∧ α ≡ β ∧ σ ≡ σ') - ∨ (∃ n' α0 σ0, n ≡ (1+n') ∧ istep α σ α0 σ0 ∧ - ▷ isteps α0 σ0 β σ' n'))%I. + ∨ (∃ n' α0 σ0, n ≡ (1+n') ∧ istep α σ α0 σ0 ∧ + ▷ isteps α0 σ0 β σ' n'))%I. Proof. unfold isteps. apply (fixpoint_unfold isteps_pre _ _ _ _ _) . Qed. @@ -179,7 +178,7 @@ Section istep. Proof. intros Hprf. destruct (IT_dont_confuse α) - as [[e Ha] | [[n Ha] | [ [g Ha] | [[α' Ha]|[op [i [k Ha]]]] ]]]. + as [[e Ha] | [[n Ha] | [ [g Ha] | [[α' Ha]|[op [i [k Ha]]]] ]]]. + exfalso. eapply uPred.pure_soundness. iPoseProof (Hprf) as "H". iDestruct "H" as (β σ') "[Ha Hs]". rewrite Ha. @@ -207,7 +206,7 @@ Section istep. Proof. intros Hprf. destruct (IT_dont_confuse α) - as [[e Ha] | [[n Ha] | [ [g Ha] | [[α' Ha]|[op [i [k Ha]]]] ]]]. + as [[e Ha] | [[n Ha] | [ [g Ha] | [[α' Ha]|[op [i [k Ha]]]] ]]]. + exfalso. eapply uPred.pure_soundness. iPoseProof (Hprf) as "H". iDestruct "H" as (β σ' op i k) "[Ha _]". rewrite Ha. @@ -246,7 +245,7 @@ Section istep. Local Lemma istep_safe_disj α σ : (∃ β σ', istep α σ β σ') - ⊢ (∃ β σ', α ≡ Tick β ∧ σ ≡ σ') + ⊢ (∃ β σ', α ≡ Tick β ∧ σ ≡ σ') ∨ (∃ β σ', (∃ op i k, α ≡ Vis op i k ∧ reify r α σ ≡ (σ', Tick β))). Proof. rewrite -bi.or_exist. @@ -321,7 +320,8 @@ Section istep. ++ iApply (IT_fun_vis_ne with "Ht"). Qed. Lemma isteps_tick α βv σ σ' k : - isteps (Tick α) σ (IT_of_V βv) σ' k ⊢ ∃ k' : nat, ⌜k = (1 + k')%nat⌝ ∧ ▷ isteps α σ (IT_of_V βv) σ' k'. + isteps (Tick α) σ (IT_of_V βv) σ' k + ⊢ ∃ k' : nat, ⌜k = (1 + k')%nat⌝ ∧ ▷ isteps α σ (IT_of_V βv) σ' k'. Proof. rewrite isteps_unfold. iDestruct 1 as "[[Hk [Ht Hs]] | H]". @@ -337,8 +337,22 @@ Section istep. iRewrite -"Ha". iRewrite "Hs". done. Qed. - Lemma istep_hom (f : IT → IT) `{!IT_hom f} α σ β σ' {G : ∀ o, CtxIndep r IT o} : - istep α σ β σ' ⊢ istep (f α) σ (f β) σ' : iProp. +End istep. + +Section istep_ctx_indep. + Context {A} `{!Cofe A}. + Context (r : sReifier NotCtxDep). + Notation F := (sReifier_ops NotCtxDep r). + Notation stateF := (sReifier_state NotCtxDep r). + Notation IT := (IT F A). + Notation ITV := (ITV F A). + Notation stateO := (stateF ♯ IT). + + Context `{!invGS_gen hlc Σ}. + Notation iProp := (iProp Σ). + + Lemma istep_hom (f : IT → IT) `{!IT_hom f} α σ β σ' : + istep r α σ β σ' ⊢ istep r (f α) σ (f β) σ' : iProp. Proof. iDestruct 1 as "[[Ha Hs]|H]". - iRewrite "Ha". iLeft. iSplit; eauto. iPureIntro. apply hom_tick. @@ -352,14 +366,15 @@ Section istep. iApply (reify_vis_cont with "Hr"). Qed. - Lemma istep_hom_inv α σ β σ' `{!IT_hom f} {G : ∀ o, CtxIndep r IT o} : - istep (f α) σ β σ' ⊢@{iProp} ⌜is_Some (IT_to_V α)⌝ - ∨ (IT_to_V α ≡ None ∧ ∃ α', istep α σ α' σ' ∧ ▷ (β ≡ f α')). + Lemma istep_hom_inv α σ β σ' (f : IT → IT) `{!IT_hom f} : + istep r (f α) σ β σ' ⊢@{iProp} ⌜is_Some (IT_to_V α)⌝ + ∨ (IT_to_V α ≡ None ∧ ∃ α', istep r α σ α' σ' + ∧ ▷ (β ≡ f α')). Proof. iIntros "H". destruct (IT_dont_confuse α) as [[e Ha] | [[n Ha] | [ [g Ha] | [[la Ha]|[op [i [k Ha]]]] ]]]. - - iExFalso. iApply (istep_err σ e β σ'). + - iExFalso. iApply (istep_err r σ e β σ'). iAssert (f α ≡ Err e)%I as "Hf". { iPureIntro. by rewrite Ha hom_err. } iRewrite "Hf" in "H". done. @@ -396,5 +411,4 @@ Section istep. iExists op,i,k. iFrame "Hr". iPureIntro. apply Ha. Qed. - -End istep. +End istep_ctx_indep. diff --git a/theories/gitree/reify.v b/theories/gitree/reify.v index 4050501..49e5deb 100644 --- a/theories/gitree/reify.v +++ b/theories/gitree/reify.v @@ -2,8 +2,12 @@ From iris.proofmode Require Import classes tactics. From gitrees Require Import prelude. From gitrees.gitree Require Import core. -Section reifiers. - Context {A} `{!Cofe A}. +Section reifier. + Inductive is_ctx_dep : Type := + | CtxDep + | NotCtxDep. + + Context {A} `{!Cofe A} (icd : is_ctx_dep). #[local] Open Scope type. (** A single reifier *) @@ -11,57 +15,65 @@ Section reifiers. { sReifier_ops : opsInterp; sReifier_state : oFunctor; sReifier_re {X} `{!Cofe X} : forall (op : opid sReifier_ops), - (Ins (sReifier_ops op) ♯ X) * (sReifier_state ♯ X) * ((Outs (sReifier_ops op) ♯ X) -n> laterO X) - -n> optionO (laterO X * (sReifier_state ♯ X)); + match icd with + | CtxDep => + (Ins (sReifier_ops op) ♯ X) + * (sReifier_state ♯ X) + * ((Outs (sReifier_ops op) ♯ X) -n> laterO X) + -n> optionO (laterO X * (sReifier_state ♯ X)) + | NotCtxDep => + (Ins (sReifier_ops op) ♯ X) + * (sReifier_state ♯ X) + -n> optionO ((Outs (sReifier_ops op) ♯ X) * (sReifier_state ♯ X)) + end; sReifier_inhab :: Inhabited (sReifier_state ♯ unitO); sReifier_cofe X (HX : Cofe X) :: Cofe (sReifier_state ♯ X); }. +End reifier. - Context (r : sReifier). - Notation F := (sReifier_ops r). - Notation stateF := (sReifier_state r). - Notation IT := (IT F A). - Notation ITV := (ITV F A). - Implicit Type op : opid F. - Implicit Type α β : IT. - - Class CtxIndep (X : ofe) `{!Cofe X} (op : opid F) := { - cont_irrelev : - (∃ f : (prodO (Ins (sReifier_ops r _) ♯ X) ((sReifier_state r) ♯ X)) -n> - optionO (prodO (Outs (sReifier_ops r _) ♯ X) (sReifier_state r ♯ X)), - ∀ i σ κ, @sReifier_re _ X _ op (i, σ, κ) ≡ fmap (prodO_map κ idfun) (f (i, σ))); - }. +Section reifier_cofe_inst. + Context {A} `{!Cofe A}. + #[local] Open Scope type. + Notation F a r := (sReifier_ops a r). + Notation stateF a r := (sReifier_state a r). + Notation IT a r := (IT (F a r) A). + Notation ITV a r := (ITV (F a r) A). + Notation stateM a r := ((stateF a r) ♯ (IT a r) -n> ((stateF a r) ♯ (IT a r)) * (IT a r)). - Notation stateM := ((stateF ♯ IT -n> (stateF ♯ IT) * IT)). - #[local] Instance stateT_inhab : Inhabited stateM. + #[global] Instance stateT_inhab {a r} : Inhabited (stateM a r). Proof. simple refine (populate (λne s, (s, inhabitant))). { apply _. } solve_proper. Qed. - #[local] Instance stateM_cofe : Cofe stateM. + #[global] Instance stateM_cofe {a r} : Cofe (stateM a r). Proof. unfold stateM. apply _. Qed. - Opaque laterO_map. - - Program Definition reify_fun : laterO (sumO IT stateM -n> prodO IT stateM) -n> stateM := - λne f s, (s, Fun $ laterO_map (λne f, fstO ◎ f ◎ inlO) f). - Solve All Obligations with solve_proper. +End reifier_cofe_inst. - Program Definition reify_tau : laterO (prodO IT stateM) -n> stateM := - λne x s, (s, Tau $ laterO_map fstO x). - Solve All Obligations with solve_proper. +Section reifier_vis. + Context {A} `{!Cofe A}. + #[local] Open Scope type. + Notation F a r := (sReifier_ops a r). + Notation stateF a r := (sReifier_state a r). + Notation IT a r := (IT (F a r) A). + Notation ITV a r := (ITV (F a r) A). + Notation stateM a r := ((stateF a r) ♯ (IT a r) -n> ((stateF a r) ♯ (IT a r)) * (IT a r)). - Program Definition reify_vis ( op : opid F ) : - oFunctor_car (Ins (F op)) (sumO IT stateM) (prodO IT stateM) -n> - (oFunctor_car (Outs (F op)) (prodO IT stateM) (sumO IT stateM) -n> laterO (prodO IT stateM)) -n> stateM. + Program Definition reify_vis_ctx_dep (r : sReifier CtxDep) ( op : opid (F CtxDep r)) : + oFunctor_car (Ins (F CtxDep r op)) + (sumO (IT CtxDep r) (stateM CtxDep r)) (prodO (IT CtxDep r) (stateM CtxDep r)) -n> + (oFunctor_car (Outs (F CtxDep r op)) (prodO (IT CtxDep r) (stateM CtxDep r)) + (sumO (IT CtxDep r) (stateM CtxDep r)) + -n> laterO (prodO (IT CtxDep r) (stateM CtxDep r))) -n> (stateM CtxDep r). Proof. simpl. - simple refine (λne i (k : _ -n> _) (s : stateF ♯ IT), _). + simple refine (λne i (k : _ -n> _) (s : stateF CtxDep r ♯ (IT CtxDep r)), _). - simple refine - (let ns := sReifier_re r op + (let ns := sReifier_re CtxDep r op (oFunctor_map _ (inlO,fstO) i, s, - (λne o, (laterO_map fstO $ k $ oFunctor_map (Outs (F op)) (fstO, inlO) o))) in _). + (λne o, (laterO_map fstO $ k + $ oFunctor_map (Outs (F CtxDep r op)) (fstO, inlO) o))) in _). + intros m s1 s2 Hs. solve_proper. + simple refine (from_option (λ ns, (ns.2, Tau $ ns.1)) @@ -73,6 +85,56 @@ Section reifiers. - intros m i1 i2 Hi k s. simpl. eapply (from_option_ne (dist m)); solve_proper. Defined. + Program Definition reify_vis_ctx_indep (r : sReifier NotCtxDep) ( op : opid (F NotCtxDep r) ) : + oFunctor_car (Ins (F NotCtxDep r op)) (sumO (IT NotCtxDep r) (stateM NotCtxDep r)) + (prodO (IT NotCtxDep r) (stateM NotCtxDep r)) -n> + (oFunctor_car (Outs (F NotCtxDep r op)) (prodO (IT NotCtxDep r) (stateM NotCtxDep r)) + (sumO (IT NotCtxDep r) (stateM NotCtxDep r)) + -n> laterO (prodO (IT NotCtxDep r) (stateM NotCtxDep r))) -n> (stateM NotCtxDep r). + Proof. + simpl. + simple refine (λne i (k : _ -n> _) (s : stateF NotCtxDep r ♯ (IT NotCtxDep r)), _). + - simple refine (let ns := sReifier_re NotCtxDep r op (oFunctor_map _ (inlO,fstO) i, s) in _). + simple refine (from_option (λ ns, + let out2' := k $ oFunctor_map (Outs (F NotCtxDep r op)) (fstO,inlO) ns.1 in + (ns.2, Tau $ laterO_map fstO out2')) + (s, Err RuntimeErr) ns). + - intros m s1 s2 Hs. simpl. eapply (from_option_ne (dist m)); solve_proper. + - intros m k1 k2 Hk s. simpl. eapply (from_option_ne (dist m)); solve_proper. + - intros m i1 i2 Hi k s. simpl. eapply (from_option_ne (dist m)); solve_proper. + Defined. + + Program Definition reify_vis (a : is_ctx_dep) (r : sReifier a) : ∀ ( op : opid (F a r)), + oFunctor_car (Ins (F a r op)) (sumO (IT a r) (stateM a r)) (prodO (IT a r) (stateM a r)) -n> + (oFunctor_car (Outs (F a r op)) (prodO (IT a r) (stateM a r)) + (sumO (IT a r) (stateM a r)) -n> laterO (prodO (IT a r) (stateM a r))) -n> (stateM a r). + Proof. + destruct a. + - apply reify_vis_ctx_dep. + - apply reify_vis_ctx_indep. + Defined. +End reifier_vis. + +Section reify. + Context {A} `{!Cofe A} {a : is_ctx_dep}. + #[local] Open Scope type. + Context {r : sReifier a}. + Notation F := (sReifier_ops a r). + Notation stateF := (sReifier_state a r). + Notation IT := (IT F A). + Notation ITV := (ITV F A). + Notation stateM := (stateF ♯ IT -n> (stateF ♯ IT) * IT). + + Opaque laterO_map. + + Program Definition reify_fun : laterO (sumO IT stateM -n> prodO IT stateM) -n> stateM := + λne f s, (s, Fun $ laterO_map (λne f, fstO ◎ f ◎ inlO) f). + Solve All Obligations with solve_proper. + + Program Definition reify_tau : laterO (prodO IT stateM) -n> stateM := + λne x s, (s, Tau $ laterO_map fstO x). + Solve All Obligations with solve_proper. + Program Definition reify_err : errorO -n> stateM := λne e s, (s, Err e). Solve All Obligations with solve_proper. @@ -80,35 +142,36 @@ Section reifiers. Solve All Obligations with solve_proper. Program Definition unr : stateM -n> - sumO (sumO (sumO (sumO A (laterO (stateM -n> stateM))) errorO) (laterO stateM)) - (sigTO (λ op : opid F, prodO (oFunctor_apply (Ins (F op)) stateM) (oFunctor_apply (Outs (F op)) stateM -n> laterO stateM))). + sumO (sumO (sumO (sumO A (laterO (stateM -n> stateM))) errorO) (laterO stateM)) + (sigTO (λ op : opid F, prodO (oFunctor_apply (Ins (F op)) stateM) + (oFunctor_apply (Outs (F op)) stateM -n> laterO stateM))). Proof. simple refine (λne d, inl (inl (inr (RuntimeErr)))). Qed. Definition reify : IT -n> stateM := IT_rec1 _ - reify_err - reify_ret - reify_fun - reify_tau - reify_vis - unr. + reify_err + reify_ret + reify_fun + reify_tau + (reify_vis a r) + unr. Definition unreify : stateM -n> IT := IT_rec2 _ - reify_err - reify_ret - reify_fun - reify_tau - reify_vis - unr. + reify_err + reify_ret + reify_fun + reify_tau + (reify_vis a r) + unr. Lemma reify_fun_eq f σ : reify (Fun f) σ ≡ (σ, Fun f). Proof. rewrite /reify/=. trans (reify_fun (laterO_map (sandwich (Perr:=reify_err) (Pret:=reify_ret) - (Parr:=reify_fun) (Ptau:=reify_tau) - (Pvis:=reify_vis) (Punfold:=unr) - stateM) f) σ). + (Parr:=reify_fun) (Ptau:=reify_tau) + (Pvis:=reify_vis a r) (Punfold:=unr) + (stateM)) f) σ). { f_equiv. apply IT_rec1_fun. } simpl. repeat f_equiv. rewrite -laterO_map_compose. @@ -117,22 +180,37 @@ Section reifiers. apply laterO_map_id. Qed. - Lemma reify_vis_dist m op i o k σ σ' : - sReifier_re r op (i, σ, k) ≡{m}≡ Some (o, σ') → +End reify. + +Section reify_props. + Context {A} `{!Cofe A}. + #[local] Open Scope type. + Notation F a r := (sReifier_ops a r). + Notation stateF a r := (sReifier_state a r). + Notation IT a r := (IT (F a r) A). + Notation ITV a r := (ITV (F a r) A). + Notation stateM a r := (stateF a r ♯ (IT a r) -n> (stateF a r ♯ (IT a r)) * (IT a r)). + + Opaque laterO_map. + + Lemma reify_vis_dist_ctx_dep (r : sReifier CtxDep) m op i o k σ σ' : + @sReifier_re CtxDep r (IT CtxDep r) _ op (i, σ, k) ≡{m}≡ Some (o, σ') → reify (Vis op i k) σ ≡{m}≡ (σ', Tau o). Proof. intros Hst. - trans (reify_vis op + trans (reify_vis CtxDep r op (oFunctor_map _ (sumO_rec idfun unreify,prod_in idfun reify) i) - (laterO_map (prod_in idfun reify) ◎ k ◎ (oFunctor_map _ (prod_in idfun reify,sumO_rec idfun unreify))) + (laterO_map (prod_in idfun reify) ◎ k + ◎ (oFunctor_map _ (prod_in idfun reify,sumO_rec idfun unreify))) σ). { f_equiv. rewrite IT_rec1_vis//. } Opaque prod_in. simpl. - pose (rs := (sReifier_re r op - (oFunctor_map (Ins (F op)) (inlO, fstO) - (oFunctor_map (Ins (F op)) (sumO_rec idfun unreify, prod_in idfun reify) i), σ, k))). + pose (rs := (sReifier_re _ r op + (oFunctor_map (Ins (F _ _ op)) (inlO, fstO) + (oFunctor_map (Ins (F _ _ op)) + (sumO_rec idfun unreify, prod_in idfun reify) i), σ, k))). fold rs. - assert (rs ≡ sReifier_re r op (i,σ, k)) as Hr'. + assert (rs ≡ sReifier_re _ r op (i,σ, k)) as Hr'. { unfold rs. f_equiv. f_equiv. rewrite -oFunctor_map_compose. f_equiv. @@ -140,7 +218,7 @@ Section reifiers. repeat f_equiv; intro; done. } assert (rs ≡{m}≡ Some (o, σ')) as Hr. { by rewrite Hr' Hst. } - trans (from_option (λ ns : laterO IT * stateF ♯ IT, (ns.2, Tau ns.1)) + trans (from_option (λ ns : laterO (IT _ _) * (stateF _ _) ♯ (IT _ _), (ns.2, Tau ns.1)) (σ, Err RuntimeErr) rs). - subst rs. @@ -160,40 +238,99 @@ Section reifiers. reflexivity. + by rewrite oFunctor_map_id. - subst rs. - trans (from_option (λ ns : laterO IT * stateF ♯ IT, (ns.2, Tau ns.1)) - (σ, Err RuntimeErr) - (Some (o, σ'))). + trans (from_option (λ ns : laterO (IT _ _) * (stateF _ _) ♯ (IT _ _), (ns.2, Tau ns.1)) + (σ, Err RuntimeErr) + (Some (o, σ'))). + eapply (from_option_ne (dist m)); [solve_proper | solve_proper |]. by rewrite Hr. + reflexivity. Qed. - Lemma reify_vis_eq op i o k σ σ' : - sReifier_re r op (i,σ,k) ≡ Some (o,σ') → + Lemma reify_vis_dist_ctx_indep (r : sReifier NotCtxDep) m op i o k σ σ' : + @sReifier_re NotCtxDep r (IT NotCtxDep r) _ op (i,σ) ≡{m}≡ Some (o,σ') → + reify (Vis op i k) σ ≡{m}≡ (σ', Tau $ k o). + Proof. + intros Hst. + trans (reify_vis _ _ op + (oFunctor_map _ (sumO_rec idfun unreify,prod_in idfun reify) i) + (laterO_map (prod_in idfun reify) ◎ k + ◎ (oFunctor_map _ (prod_in idfun reify,sumO_rec idfun unreify))) + σ). + { f_equiv. rewrite IT_rec1_vis//. } + Opaque prod_in. simpl. + pose (rs := (@sReifier_re NotCtxDep r _ _ op + (oFunctor_map (Ins (F NotCtxDep r op)) (inlO, fstO) + (oFunctor_map (Ins (F NotCtxDep r op)) + (sumO_rec idfun unreify, prod_in idfun reify) i), σ))). + fold rs. + assert (rs ≡ sReifier_re NotCtxDep r op (i,σ)) as Hr'. + { unfold rs. f_equiv. f_equiv. + rewrite -oFunctor_map_compose. + etrans; last by apply oFunctor_map_id. + repeat f_equiv; intro; done. } + assert (rs ≡{m}≡ Some (o,σ')) as Hr. + { by rewrite Hr' Hst. } + trans (from_option (λ ns, + (ns.2, + Tau + (laterO_map fstO + (laterO_map (prod_in idfun reify) + (k + (oFunctor_map (Outs (F NotCtxDep r op)) + (prod_in idfun reify, sumO_rec idfun unreify) + (oFunctor_map (Outs (F NotCtxDep r op)) (fstO, inlO) ns.1))))))) + (σ, Err RuntimeErr) (Some (o,σ'))). + { eapply (from_option_ne (dist m)); solve_proper. } + simpl. repeat f_equiv. + rewrite -laterO_map_compose. + rewrite -oFunctor_map_compose. + trans (laterO_map idfun (k o)); last first. + { by rewrite laterO_map_id. } + repeat f_equiv. + { intro; done. } + trans (oFunctor_map _ (idfun, idfun) o); last first. + { by rewrite oFunctor_map_id. } + simpl. + repeat f_equiv; intro; done. + Qed. + + Lemma reify_vis_eq_ctx_dep r op i o k σ σ' : + @sReifier_re CtxDep r (IT CtxDep r) _ op (i,σ,k) ≡ Some (o,σ') → reify (Vis op i k) σ ≡ (σ', Tau $ o). Proof. intros H. apply equiv_dist=>m. - apply reify_vis_dist. + apply reify_vis_dist_ctx_dep. by apply equiv_dist. Qed. - Lemma reify_vis_None op i k σ : - sReifier_re r op (i,σ,k) ≡ None → + Lemma reify_vis_eq_ctx_indep r op i o k σ σ' : + @sReifier_re NotCtxDep r (IT NotCtxDep r) _ op (i,σ) ≡ Some (o,σ') → + reify (Vis op i k) σ ≡ (σ', Tau $ k o). + Proof. + intros H. apply equiv_dist=>m. + apply reify_vis_dist_ctx_indep. + by apply equiv_dist. + Qed. + + Lemma reify_vis_None_ctx_dep r op i k σ : + @sReifier_re CtxDep r (IT CtxDep r) _ op (i,σ,k) ≡ None → reify (Vis op i k) σ ≡ (σ, Err RuntimeErr). Proof. intros Hs. - trans (reify_vis op + trans (reify_vis _ _ op (oFunctor_map _ (sumO_rec idfun unreify,prod_in idfun reify) i) - (laterO_map (prod_in idfun reify) ◎ k ◎ (oFunctor_map _ (prod_in idfun reify,sumO_rec idfun unreify))) + (laterO_map (prod_in idfun reify) ◎ k + ◎ (oFunctor_map _ (prod_in idfun reify,sumO_rec idfun unreify))) σ). { f_equiv. apply IT_rec1_vis. } simpl. - pose (rs := (sReifier_re r op - (oFunctor_map (Ins (F op)) (inlO, fstO) - (oFunctor_map (Ins (F op)) (sumO_rec idfun unreify, prod_in idfun reify) i), σ, k))). + pose (rs := (sReifier_re CtxDep r op + (oFunctor_map (Ins (F CtxDep r op)) (inlO, fstO) + (oFunctor_map (Ins (F CtxDep r op)) + (sumO_rec idfun unreify, prod_in idfun reify) i), σ, k))). fold rs. - assert (rs ≡ sReifier_re r op (i,σ, k)) as Hr'. + assert (rs ≡ sReifier_re CtxDep r op (i,σ, k)) as Hr'. { unfold rs. f_equiv. f_equiv. rewrite -oFunctor_map_compose. f_equiv. @@ -201,7 +338,7 @@ Section reifiers. repeat f_equiv; intro; done. } assert (rs ≡ None) as Hr. { by rewrite Hr' Hs. } - trans (from_option (λ ns : laterO IT * stateF ♯ IT, (ns.2, Tau ns.1)) + trans (from_option (λ ns : laterO (IT _ _) * (stateF _ _) ♯ (IT _ _), (ns.2, Tau ns.1)) (σ, Err RuntimeErr) rs). { @@ -223,56 +360,99 @@ Section reifiers. reflexivity. - by rewrite oFunctor_map_id. } - trans (from_option (λ ns : laterO IT * stateF ♯ IT, (ns.2, Tau ns.1)) (σ, Err RuntimeErr) None). + trans (from_option (λ ns : laterO (IT CtxDep r) * (stateF _ _) ♯ (IT _ _), + (ns.2, Tau ns.1)) (σ, Err RuntimeErr) None). - f_equiv; [| assumption]. intros [? ?] [? ?] [? ?]; simpl in *; f_equiv; [assumption | f_equiv; assumption]. - reflexivity. Qed. - Lemma reify_is_always_a_tick op x k σ β σ' : - reify (Vis op x k) σ ≡ (σ', β) → (∃ β', β ≡ Tick β') ∨ (β ≡ Err RuntimeErr). + Lemma reify_vis_None_ctx_indep r op i k σ : + @sReifier_re NotCtxDep r (IT NotCtxDep r) _ op (i,σ) ≡ None → + reify (Vis op i k) σ ≡ (σ, Err RuntimeErr). + Proof. + intros Hs. + trans (reify_vis _ _ op + (oFunctor_map _ (sumO_rec idfun unreify,prod_in idfun reify) i) + (laterO_map (prod_in idfun reify) ◎ k + ◎ (oFunctor_map _ + (prod_in idfun reify,sumO_rec idfun unreify))) + σ). + { f_equiv. + apply IT_rec1_vis. } + simpl. + pose (rs := (sReifier_re NotCtxDep r op + (oFunctor_map (Ins (F _ _ op)) (inlO, fstO) + (oFunctor_map (Ins (F _ _ op)) + (sumO_rec idfun unreify, prod_in idfun reify) i), σ))). + fold rs. + assert (rs ≡ sReifier_re _ r op (i,σ)) as Hr'. + { unfold rs. f_equiv. f_equiv. + rewrite -oFunctor_map_compose. + etrans; last by apply oFunctor_map_id. + repeat f_equiv; intro; done. } + assert (rs ≡ None) as Hr. + { by rewrite Hr' Hs. } + trans (from_option (λ ns, + (ns.2, + Tau + (laterO_map fstO + (laterO_map (prod_in idfun reify) + (k + (oFunctor_map (Outs (F _ _ op)) + (prod_in idfun reify, sumO_rec idfun unreify) + (oFunctor_map (Outs (F _ _ op)) (fstO, inlO) ns.1))))))) + (σ, Err RuntimeErr) None). + { apply from_option_proper; solve_proper. } + reflexivity. + Qed. + + Lemma reify_is_always_a_tick {a : is_ctx_dep} r op x k σ β σ' : + reify (a := a) (A := A) (r := r) (Vis op x k) σ ≡ (σ', β) + → (∃ β', β ≡ Tick β') ∨ (β ≡ Err RuntimeErr). Proof. - destruct (sReifier_re r op (x, σ, k)) as [[o σ'']|] eqn:Hre; last first. - - rewrite reify_vis_None; last by rewrite Hre//. - intros [_ ?]. by right. - - rewrite reify_vis_eq;last by rewrite Hre. - intros [? Ho]. - left. - simpl in *. - destruct (Next_uninj o) as [t Ht]. - exists (t). - rewrite <-Ho. - rewrite Ht. - reflexivity. + destruct a. + { + destruct (sReifier_re _ r op (x, σ, k)) as [[o σ'']|] eqn:Hre; last first. + - rewrite reify_vis_None_ctx_dep; last by rewrite Hre//. + intros [_ ?]. by right. + - rewrite reify_vis_eq_ctx_dep;last by rewrite Hre. + intros [? Ho]. + left. + simpl in *. + destruct (Next_uninj o) as [t Ht]. + exists (t). + rewrite <-Ho. + rewrite Ht. + reflexivity. + } + { + destruct (sReifier_re _ r op (x, σ)) as [[o σ'']|] eqn:Hre; last first. + - rewrite reify_vis_None_ctx_indep; last by rewrite Hre//. + intros [_ ?]. by right. + - rewrite reify_vis_eq_ctx_indep;last by rewrite Hre. + intros [? Ho]. + destruct (Next_uninj (k o)) as [lβ Hlb]. + left. exists (lβ). + rewrite Tick_eq. + rewrite -Hlb. symmetry. apply Ho. + } Qed. - Lemma reify_vis_cont op i k1 k2 σ1 σ2 β - {PROP : bi} `{!BiInternalEq PROP} `{H : !(@CtxIndep IT _ op)} : - (reify (Vis op i k1) σ1 ≡ (σ2, Tick β) ⊢ + Lemma reify_vis_cont r op i k1 k2 σ1 σ2 β + {PROP : bi} `{!BiInternalEq PROP} : + (reify (a := NotCtxDep) (A := A) (r := r) (Vis op i k1) σ1 ≡ (σ2, Tick β) ⊢ reify (Vis op i (laterO_map k2 ◎ k1)) σ1 ≡ (σ2, Tick (k2 β)) : PROP)%I. Proof. - destruct (sReifier_re r op (i, σ1, k1)) as [[o σ2']|] eqn:Hre; last first. - - rewrite (reify_vis_None _ _ k1); last by rewrite Hre//. + destruct (sReifier_re _ r op (i,σ1)) as [[o σ2']|] eqn:Hre; last first. + - rewrite reify_vis_None_ctx_indep; last by rewrite Hre//. iIntros "Hr". iExFalso. iPoseProof (prod_equivI with "Hr") as "[_ Hk]". simpl. iApply (IT_tick_err_ne). by iApply internal_eq_sym. - - destruct H as [[f H]]. - pose proof (H i σ1 k1) as H1. - pose proof (H i σ1 (laterO_map k2 ◎ k1)) as H2. - assert (∃ o σ', f (i, σ1) = Some (o, σ')) as [o' [σ' H3]]. - { - destruct (f (i, σ1)) as [[? ?] | ?]; first (do 2 eexists; reflexivity). - simpl in H1. rewrite Hre in H1; inversion H1. - } - rewrite H3 in H1. - simpl in H1. - rewrite H3 in H2. - simpl in H2. - clear f H H3 Hre. - rewrite reify_vis_eq; last first. - { by rewrite H1. } - rewrite reify_vis_eq; last first. - { by rewrite H2. } + - rewrite reify_vis_eq_ctx_indep; last first. + { by rewrite Hre. } + rewrite reify_vis_eq_ctx_indep; last first. + { by rewrite Hre. } iIntros "Hr". iPoseProof (prod_equivI with "Hr") as "[Hs Hk]". iApply prod_equivI. simpl. iSplit; eauto. @@ -281,60 +461,39 @@ Section reifiers. rewrite laterO_map_Next. done. Qed. - Lemma reify_input_cont_inv op i (k1 : _ -n> laterO IT) (k2 : IT -n> IT) σ1 σ2 β - {PROP : bi} `{!BiInternalEq PROP} `{H : !(@CtxIndep IT _ op)} : + Lemma reify_input_cont_inv r op i (k1 : _ -n> laterO (IT NotCtxDep r)) + (k2 : IT _ r -n> IT _ r) σ1 σ2 β + {PROP : bi} `{!BiInternalEq PROP} : (reify (Vis op i (laterO_map k2 ◎ k1)) σ1 ≡ (σ2, Tick β) ⊢ ∃ α, reify (Vis op i k1) σ1 ≡ (σ2, Tick α) ∧ ▷ (β ≡ k2 α) : PROP)%I. Proof. - destruct (sReifier_re r op (i, σ1, (laterO_map k2 ◎ k1))) as [[o σ2']|] eqn:Hre; last first. - - rewrite reify_vis_None; last by rewrite Hre//. + destruct (sReifier_re _ r op (i,σ1)) as [[o σ2']|] eqn:Hre; last first. + - rewrite reify_vis_None_ctx_indep; last by rewrite Hre//. iIntros "Hr". iExFalso. iPoseProof (prod_equivI with "Hr") as "[_ Hk]". simpl. iApply (IT_tick_err_ne). by iApply internal_eq_sym. - - rewrite reify_vis_eq; last first. + - rewrite reify_vis_eq_ctx_indep; last first. { by rewrite Hre. } iIntros "Hr". simpl. iPoseProof (prod_equivI with "Hr") as "[#Hs #Hk]". simpl. iPoseProof (Tau_inj' with "Hk") as "Hk'". - destruct H as [[f H]]. - pose proof (H i σ1 k1) as H1. - pose proof (H i σ1 (laterO_map k2 ◎ k1)) as H2. - assert (∃ o, f (i, σ1) ≡ Some (o, σ2')) as [o' H3]. - { - destruct (f (i, σ1)) as [[? ?] | ?]. - - simpl in H2. - rewrite Hre in H2. - inversion H2 as [? ? H2' |]; subst; inversion H2'; simpl in *; subst. - eexists _; do 2 f_equiv; first reflexivity; symmetry; assumption. - - simpl in H2. - rewrite Hre in H2. - inversion H2. - } - rewrite H3 in H1. - simpl in H1. - rewrite H3 in H2. - simpl in H2. - destruct (Next_uninj (k1 o')) as [a Hk1]. + destruct (Next_uninj (k1 o)) as [a Hk1]. iExists (a). - rewrite reify_vis_eq; last first. - { by rewrite H1. } + rewrite reify_vis_eq_ctx_indep; last first. + { by rewrite Hre. } iSplit. + iApply prod_equivI. simpl. iSplit; eauto. iApply Tau_inj'. done. + iAssert (laterO_map k2 (Next a) ≡ Next β)%I as "Ha". - { - iSimpl in "Hk'". iRewrite -"Hk'". - iPureIntro. rewrite -Hk1. - rewrite Hre in H2. - inversion H2 as [? ? H2' |]; subst; inversion H2'; simpl in *; subst. - symmetry; assumption. - } + { iSimpl in "Hk'". iRewrite -"Hk'". + iPureIntro. rewrite -Hk1. done. } iAssert (Next (k2 a) ≡ Next β)%I as "Hb". { iRewrite -"Ha". iPureIntro. rewrite laterO_map_Next. done. } iNext. by iApply internal_eq_sym. Qed. +End reify_props. -End reifiers. +Arguments reify {_ _ _} _. diff --git a/theories/gitree/weakestpre.v b/theories/gitree/weakestpre.v index 3697856..a3cdb39 100644 --- a/theories/gitree/weakestpre.v +++ b/theories/gitree/weakestpre.v @@ -6,25 +6,32 @@ From gitrees.gitree Require Import core reify greifiers reductions. (** * Ghost state from gReifiers *) -Definition gReifiers_ucmra {n} (rs : gReifiers n) (X : ofe) `{!Cofe X} : ucmra := - discrete_funUR (λ (i : fin n), optionUR (exclR (sReifier_state (rs !!! i) ♯ X))). +Definition gReifiers_ucmra {n} (a : is_ctx_dep) (rs : gReifiers a n) + (X : ofe) `{!Cofe X} : ucmra := + discrete_funUR (λ (i : fin n), + optionUR (exclR (sReifier_state a (rs !!! i) ♯ X))). (** The resource corresponding to the whole global state *) -Definition of_state {n} (rs : gReifiers n) (X : ofe) `{!Cofe X} (st : gReifiers_state rs ♯ X) : gReifiers_ucmra rs X := - λ i, Excl' (fstO (gState_decomp i st)). +Definition of_state {n} (a : is_ctx_dep) (rs : gReifiers a n) + (X : ofe) `{!Cofe X} (st : gReifiers_state a rs ♯ X) + : gReifiers_ucmra a rs X := + λ i, Excl' (fstO (gState_decomp a i st)). (** The resource corresponding to a speicific projection out of the global state *) -Definition of_idx {n} (rs : gReifiers n) (X : ofe) `{!Cofe X} (i : fin n) - (st : sReifier_state (rs !!! i) ♯ X) : gReifiers_ucmra rs X. +Definition of_idx {n} (a : is_ctx_dep) (rs : gReifiers a n) + (X : ofe) `{!Cofe X} (i : fin n) + (st : sReifier_state a (rs !!! i) ♯ X) : gReifiers_ucmra a rs X. Proof. simple refine (λ j, if (decide (j = i)) then _ else None). simpl. induction e. exact (Excl' st). Defined. -Lemma of_state_recomp_lookup_ne {n} (rs : gReifiers n) (X : ofe) `{!Cofe X} - i j (σ1 σ2 : sReifier_state (rs !!! i) ♯ X) rest : +Lemma of_state_recomp_lookup_ne {n} (a : is_ctx_dep) (rs : gReifiers a n) + (X : ofe) `{!Cofe X} + i j (σ1 σ2 : sReifier_state a (rs !!! i) ♯ X) rest : i ≠ j → - of_state rs X (gState_recomp rest σ1) j ≡ of_state rs X (gState_recomp rest σ2) j. + of_state a rs X (gState_recomp a rest σ1) j + ≡ of_state a rs X (gState_recomp a rest σ2) j. Proof. intros Hij. revert σ1 σ2 rest. unfold of_state. @@ -42,31 +49,30 @@ Proof. intro. simplify_eq/=. Qed. - Section ucmra. - Context {n : nat} (rs : gReifiers n). + Context {n : nat} (a : is_ctx_dep) (rs : gReifiers a n). Context (X : ofe) `{!Cofe X}. - Notation gReifiers_ucmra := (gReifiers_ucmra rs X). - Notation of_state := (of_state rs X). - Notation of_idx := (of_idx rs X). + Notation gReifiers_ucmra := (gReifiers_ucmra a rs X). + Notation of_state := (of_state a rs X). + Notation of_idx := (of_idx a rs X). #[export] Instance of_state_ne : NonExpansive of_state. Proof. solve_proper. Qed. #[export] Instance of_state_proper : Proper ((≡) ==> (≡)) of_state. Proof. apply ne_proper, _. Qed. - Lemma of_state_valid (σ : gReifiers_state rs ♯ X) : ✓ (of_state σ). + Lemma of_state_valid (σ : gReifiers_state a rs ♯ X) : ✓ (of_state σ). Proof. intro; done. Qed. - Lemma of_state_recomp_lookup i (σ : sReifier_state (rs !!! i) ♯ X) rest : - of_state (gState_recomp rest σ) i ≡ Excl' σ. + Lemma of_state_recomp_lookup i (σ : sReifier_state a (rs !!! i) ♯ X) rest : + of_state (gState_recomp a rest σ) i ≡ Excl' σ. Proof. unfold of_state. rewrite gState_decomp_recomp. done. Qed. - Lemma of_state_decomp_local_update i (σ σ1 σ2 : sReifier_state (rs !!! i) ♯ X) rest : - (of_state (gState_recomp rest σ1), of_idx i σ2) - ~l~> (of_state (gState_recomp rest σ), of_idx i σ). + Lemma of_state_decomp_local_update i (σ σ1 σ2 : sReifier_state a (rs !!! i) ♯ X) rest : + (of_state (gState_recomp a rest σ1), of_idx i σ2) + ~l~> (of_state (gState_recomp a rest σ), of_idx i σ). Proof. apply discrete_fun_local_update. intros j. @@ -82,7 +88,7 @@ Section ucmra. Qed. Lemma of_state_of_idx_agree i σ1 σ2 rest f Σ : - of_state (gState_recomp rest σ1) ≡ of_idx i σ2 ⋅ f ⊢@{iProp Σ} σ1 ≡ σ2. + of_state (gState_recomp a rest σ1) ≡ of_idx i σ2 ⋅ f ⊢@{iProp Σ} σ1 ≡ σ2. Proof. iIntros "Hs". rewrite discrete_fun_equivI. @@ -100,16 +106,16 @@ Section ucmra. End ucmra. Section weakestpre. - Context {n : nat} (rs : gReifiers n) {A} `{!Cofe A}. - Notation rG := (gReifiers_sReifier rs). - Notation F := (sReifier_ops rG). + Context {n : nat} (a : is_ctx_dep) (rs : gReifiers a n) {A} `{!Cofe A}. + Notation rG := (gReifiers_sReifier a rs). + Notation F := (sReifier_ops a rG). Notation IT := (IT F A). Notation ITV := (ITV F A). - Notation stateF := (gReifiers_state rs). + Notation stateF := (gReifiers_state a rs). Notation stateO := (stateF ♯ IT). - Notation stateR := (gReifiers_ucmra rs IT). - Let of_state := (of_state rs IT). - Let of_idx := (of_idx rs IT). + Notation stateR := (gReifiers_ucmra a rs IT). + Let of_state := (of_state a rs IT). + Let of_idx := (of_idx a rs IT). Notation reify := (reify rG). Notation istep := (istep rG). Notation isteps := (isteps rG). @@ -135,15 +141,16 @@ Section weakestpre. Definition has_full_state `{!stateG Σ} (σ : stateO) : iProp Σ := (own stateG_name (◯ (of_state σ)))%I. Definition has_state_idx `{!stateG Σ} - (i : fin n) (σ : sReifier_state (rs !!! i) ♯ IT) : iProp Σ := + (i : fin n) (σ : sReifier_state a (rs !!! i) ♯ IT) : iProp Σ := (own stateG_name (◯ (of_idx i σ)))%I. - Definition has_substate {sR : sReifier} `{!stateG Σ} `{!subReifier sR rs} - (σ : sReifier_state sR ♯ IT) : iProp Σ := + Definition has_substate {sR : sReifier a} `{!stateG Σ} `{!subReifier sR rs} + (σ : sReifier_state a sR ♯ IT) : iProp Σ := (own stateG_name (◯ (of_idx sR_idx (sR_state σ))))%I. #[export] Instance state_interp_ne `{!stateG Σ} : NonExpansive state_interp. Proof. solve_proper. Qed. - #[export] Instance state_interp_proper `{!stateG Σ} : Proper ((≡) ==> (≡)) state_interp. + #[export] Instance state_interp_proper `{!stateG Σ} + : Proper ((≡) ==> (≡)) state_interp. Proof. solve_proper. Qed. Lemma new_state_interp σ `{!invGS_gen hlc Σ, !statePreG Σ} : @@ -156,9 +163,9 @@ Section weakestpre. Qed. Lemma state_interp_has_state_idx_agree (i : fin n) - (σ1 σ2 : sReifier_state (rs !!! i) ♯ IT) - (rest : gState_rest i rs ♯ IT) `{!stateG Σ} : - state_interp (gState_recomp rest σ1) -∗ has_state_idx i σ2 -∗ σ1 ≡ σ2. + (σ1 σ2 : sReifier_state a (rs !!! i) ♯ IT) + (rest : gState_rest a i rs ♯ IT) `{!stateG Σ} : + state_interp (gState_recomp a rest σ1) -∗ has_state_idx i σ2 -∗ σ1 ≡ σ2. Proof. iIntros "H1 H2". iDestruct (own_valid_2 with "H1 H2") as "Hs". @@ -169,15 +176,15 @@ Section weakestpre. Qed. Lemma state_interp_has_state_idx_update (i : fin n) - (σ σ1 σ2 : sReifier_state (rs !!! i) ♯ IT) - (rest : gState_rest i rs ♯ IT) `{!stateG Σ} : - state_interp (gState_recomp rest σ1) -∗ has_state_idx i σ2 ==∗ - state_interp (gState_recomp rest σ) ∗ has_state_idx i σ. + (σ σ1 σ2 : sReifier_state a (rs !!! i) ♯ IT) + (rest : gState_rest a i rs ♯ IT) `{!stateG Σ} : + state_interp (gState_recomp a rest σ1) -∗ has_state_idx i σ2 ==∗ + state_interp (gState_recomp a rest σ) ∗ has_state_idx i σ. Proof. iIntros "H1 H2". iMod (own_update_2 with "H1 H2") as "H". { apply auth_update. - apply (of_state_decomp_local_update _ _ _ σ). } + apply (of_state_decomp_local_update a _ _ _ σ). } iDestruct "H" as "[$ $]". done. Qed. @@ -209,7 +216,7 @@ Section weakestpre. #[local] Instance wp_pre_contractive s Φ : Contractive (wp_pre s Φ). Proof. unfold wp_pre. - intros m s1 s2 Hs E1 a. simpl. + intros m s1 s2 Hs E1 a'. simpl. (* repeat first [f_contractive | f_equiv; solve_proper *) (* | f_equiv ]. *) f_equiv. f_equiv. f_equiv. @@ -245,7 +252,9 @@ Section weakestpre. format "'WP' α @ s {{ Φ } }") : bi_scope. #[export] Instance wp_ne m : - Proper ((dist m) ==> (pointwise_relation _ (iff)) ==> (dist m) ==> (pointwise_relation _ (dist m)) ==> (dist m)) wp. + Proper ((dist m) ==> + (pointwise_relation _ (iff)) ==> + (dist m) ==> (pointwise_relation _ (dist m)) ==> (dist m)) wp. Proof. intros α1 α2 Ha s s' Hs E1 E2 HE Φ1 Φ2 Hp. assert (E1 = E2) as ->. @@ -267,7 +276,8 @@ Section weakestpre. eapply dist_le; [apply Hp|lia]. Qed. #[export] Instance wp_proper : - Proper ((≡) ==> (pointwise_relation _ (iff)) ==> (≡) ==> (pointwise_relation _ (≡)) ==> (≡)) wp. + Proper ((≡) ==> (pointwise_relation _ (iff)) ==> (≡) ==> + (pointwise_relation _ (≡)) ==> (≡)) wp. Proof. intros α1 α2 Ha s s' Hs E1 E2 HE Φ1 Φ2 Hp. apply equiv_dist=>m. @@ -392,13 +402,13 @@ Section weakestpre. Opaque gState_recomp. (* We can generalize this based on the stuckness bit *) - Lemma wp_reify_idx E1 E2 s Φ i (lop : opid (sReifier_ops (rs !!! i))) : + Lemma wp_reify_idx E1 E2 s Φ i (lop : opid (sReifier_ops a (rs !!! i))) : let op : opid F := (existT i lop) in forall (x : Ins (F op) ♯ IT) (k : Outs (F op) ♯ IT -n> laterO IT), (|={E1,E2}=> ∃ σ σ' β, has_state_idx i σ ∗ - ∀ rest, reify (Vis op x k) (gState_recomp rest σ) - ≡ (gState_recomp rest σ', Tick β) ∗ + ∀ rest, reify (Vis op x k) (gState_recomp a rest σ) + ≡ (gState_recomp a rest σ', Tick β) ∗ ▷ (£ 1 -∗ has_state_idx i σ' -∗ |={E2,E1}=> WP β @ s;E1 {{ Φ }})) -∗ WP (Vis op x k) @ s;E1 {{ Φ }}. Proof. @@ -408,8 +418,8 @@ Section weakestpre. iRight. iSplit. { iPureIntro. apply IT_to_V_Vis. } iIntros (fs) "Hgst". - destruct (gState_decomp i fs) as [σ0 rest] eqn:Hdecomp. - assert (fs ≡ gState_recomp rest σ0) as Hfs. + destruct (gState_decomp a i fs) as [σ0 rest] eqn:Hdecomp. + assert (fs ≡ gState_recomp a rest σ0) as Hfs. { symmetry. apply gState_recomp_decomp. by rewrite Hdecomp. } iMod "H" as (σ σ' β) "[Hlst H]". @@ -424,10 +434,10 @@ Section weakestpre. iSplit. { (* it is safe *) iLeft. - iExists β,(gState_recomp rest σ'). iRight. + iExists β,(gState_recomp a rest σ'). iRight. iExists op,x,k; eauto. } iIntros (fs' α0) "Hst Hlc". rewrite istep_vis. - iAssert (gState_recomp rest σ' ≡ fs' ∧ Tick β ≡ Tick α0)%I + iAssert (gState_recomp a rest σ' ≡ fs' ∧ Tick β ≡ Tick α0)%I with "[Hreify Hst]" as "[Hst Hb]". { iRewrite "Hreify" in "Hst". by rewrite prod_equivI. } @@ -440,39 +450,10 @@ Section weakestpre. iRewrite -"Hb". by iFrame. Qed. - Lemma wp_reify_idx' E1 E2 s Φ i (lop : opid (sReifier_ops (rs !!! i))) : - let op : opid F := (existT i lop) in - forall (x : Ins (F op) ♯ IT) - (k : Outs (F op) ♯ IT -n> laterO IT), - (|={E1,E2}=> ∃ σ y σ' β, has_state_idx i σ ∗ - sReifier_re (rs !!! i) lop (x, σ, k) ≡ Some (y, σ') ∗ - y ≡ Next β ∗ - ▷ (£ 1 -∗ has_state_idx i σ' ={E2,E1}=∗ WP β @ s;E1 {{ Φ }})) - -∗ WP (Vis op x k) @ s;E1 {{ Φ }}. - Proof. - intros op x k. - iIntros "H". - iApply wp_reify_idx. - iMod "H" as (σ y σ' β) "[Hlst [Hreify [Hk H]]]". - iModIntro. iExists σ, σ', β. - iFrame "Hlst". - iIntros (rest). iFrame "H". - iAssert (gReifiers_re rs op (x, gState_recomp rest σ, _) ≡ Some (y, gState_recomp rest σ'))%I - with "[Hreify]" as "Hgreify". - { rewrite gReifiers_re_idx. - iAssert (optionO_map (prodO_map idfun (gState_recomp rest)) (sReifier_re (rs !!! i) lop (x, σ, k)) ≡ optionO_map (prodO_map idfun (gState_recomp rest)) (Some (y, σ')))%I with "[Hreify]" as "H". - - iApply (f_equivI with "Hreify"). - - simpl. iExact "H". - } - iPoseProof (reify_vis_eqI _ _ _ k with "Hgreify") as "Hreify". - iRewrite "Hk" in "Hreify". - by rewrite -Tick_eq. - Qed. - - Lemma wp_reify E1 s Φ i (lop : opid (sReifier_ops (rs !!! i))) + Lemma wp_reify E1 s Φ i (lop : opid (sReifier_ops a (rs !!! i))) x k σ σ' β : let op : opid F := (existT i lop) in - (∀ rest, reify (Vis op x k) (gState_recomp rest σ) ≡ (gState_recomp rest σ', Tick β)) → + (∀ rest, reify (Vis op x k) (gState_recomp a rest σ) ≡ (gState_recomp a rest σ', Tick β)) → has_state_idx i σ -∗ ▷ (£ 1 -∗ has_state_idx i σ' -∗ WP β @ s;E1 {{ Φ }}) -∗ WP (Vis op x k) @ s;E1 {{ Φ }}. @@ -488,62 +469,6 @@ Section weakestpre. iModIntro. by iApply ("H" with "Hlc Hs"). Qed. - Lemma wp_subreify' E1 E2 s Φ sR `{!subReifier sR rs} - (op : opid (sReifier_ops sR)) (x : Ins (sReifier_ops sR op) ♯ IT) - (k : Outs (F (subEff_opid op)) ♯ IT -n> laterO IT) : - (|={E1,E2}=> ∃ σ y σ' β, has_substate σ ∗ - sReifier_re sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') ∗ y ≡ Next β ∗ - ▷ (£ 1 -∗ has_substate σ' ={E2,E1}=∗ WP β @ s;E1 {{ Φ }})) - -∗ WP (Vis (subEff_opid op) (subEff_ins x) k) @ s;E1 {{ Φ }}. - Proof. - iIntros "H". - iApply wp_reify_idx'. - iMod "H" as (σ y σ' β) "[Hlst [Hreify [Hk H]]]". - iModIntro. - iExists (sR_state σ), y, (sR_state σ'), β. - simpl. - iFrame "Hlst H". - rewrite subReifier_reify_idxI. - iFrame "Hk". - iRewrite - "Hreify". - iPureIntro. - do 2 f_equiv. - intros ?; simpl. - by rewrite ofe_iso_12. - Qed. - - Lemma wp_subreify E1 s Φ sR `{!subReifier sR rs} - (op : opid (sReifier_ops sR)) - (x : Ins (sReifier_ops sR op) ♯ IT) (y : laterO IT) - (k : Outs (F (subEff_opid op)) ♯ IT -n> laterO IT) - (σ σ' : sReifier_state sR ♯ IT) β : - sReifier_re sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') → - y ≡ Next β → - has_substate σ -∗ - ▷ (£ 1 -∗ has_substate σ' -∗ WP β @ s;E1 {{ Φ }}) - -∗ - WP (Vis (subEff_opid op) (subEff_ins x) k) @ s;E1 {{ Φ }}. - Proof. - intros HSR Hk. - iIntros "Hlst H". - iApply (wp_reify with "Hlst H"). - intros rest. - rewrite Tick_eq. rewrite -Hk. - rewrite reify_vis_eq //. - pose proof (@subReifier_reify n sR rs _ IT _ op x y (k ◎ subEff_outs) σ σ' rest) as H. - simpl in H. - rewrite <-H. - - simpl. - repeat f_equiv. - + intros ???. - solve_proper. - + intros ?; simpl. - rewrite ofe_iso_12. - reflexivity. - - rewrite HSR. - reflexivity. - Qed. - Lemma wp_err E1 e (s : error → Prop) Φ : s e → ⊢ WP (Err e) @ s;E1 {{ Φ }}. @@ -559,6 +484,7 @@ Section weakestpre. iIntros (σ' β) "Hst". iExFalso. iApply istep_err. done. Qed. + Lemma wp_stuckness_mono α E1 (s1 s2 : error → Prop) Φ : (∀ e, s1 e → s2 e) → WP α @ s1;E1 {{ Φ }} ⊢ WP α @ s2;E1 {{ Φ }}. @@ -736,13 +662,250 @@ Section weakestpre. solve_proper. Qed. - Lemma wp_bind (f : IT → IT) `{!IT_hom f} (α : IT) s Φ `{!NonExpansive Φ} E1 {G : ∀ o : opid F, CtxIndep rG IT o} : +End weakestpre. + +Section weakestpre_specific. + Context {n : nat} {A} `{!Cofe A}. + + Notation rG a rs := (gReifiers_sReifier (n := n) a rs). + Notation F a rs := (sReifier_ops a (rG a rs)). + Notation IT a rs := (IT (F a rs) A). + Notation ITV a rs := (ITV (F a rs) A). + Notation stateF a rs := (gReifiers_state a rs). + Notation stateO a rs := (stateF a rs ♯ IT a rs). + Notation stateR a rs := (gReifiers_ucmra a rs (IT a rs)). + Let of_state a rs := (of_state a rs (IT a rs)). + Let of_idx a rs := (of_idx a rs (IT a rs)). + Notation reify a rs := (reify (rG a rs)). + Notation istep a rs := (istep (rG a rs)). + Notation isteps a rs := (isteps (rG a rs)). + Notation sstep a rs := (sstep (rG a rs)). + Notation ssteps a rs := (ssteps (rG a rs)). + Notation wp a rs := (wp a rs). + + Context `{!invGS Σ}. + Notation iProp := (iProp Σ). + Notation coPsetO := (leibnizO coPset). + + Lemma wp_reify_idx_ctx_dep (rs : gReifiers CtxDep n) + `{!@stateG _ CtxDep rs A _ Σ} E1 E2 s Φ i + (lop : opid (sReifier_ops CtxDep (rs !!! i))) : + let op : opid (F CtxDep rs) := (existT i lop) in + forall (x : Ins (F CtxDep rs op) ♯ IT CtxDep rs) + (k : Outs (F CtxDep rs op) ♯ IT CtxDep rs -n> laterO (IT CtxDep rs)), + (|={E1,E2}=> + ∃ σ y σ' β, has_state_idx CtxDep rs i σ + ∗ sReifier_re CtxDep (rs !!! i) lop (x, σ, k) ≡ Some (y, σ') + ∗ y ≡ Next β + ∗ ▷ (£ 1 -∗ has_state_idx CtxDep rs i σ' ={E2,E1}=∗ wp CtxDep rs β s E1 Φ)) + -∗ wp CtxDep rs (Vis op x k) s E1 Φ. + Proof. + intros op x k. + iIntros "H". + iApply wp_reify_idx. + iMod "H" as (σ y σ' β) "[Hlst [Hreify [Hk H]]]". + iModIntro. iExists σ, σ', β. + iFrame "Hlst". + iIntros (rest). + iFrame "H". + iAssert (gReifiers_re CtxDep rs op (x, gState_recomp CtxDep rest σ, _) + ≡ Some (y, gState_recomp CtxDep rest σ'))%I + with "[Hreify]" as "Hgreify". + { rewrite (gReifiers_re_idx CtxDep). + iAssert (optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) + (sReifier_re CtxDep (rs !!! i) lop (x, σ, k)) + ≡ optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) + (Some (y, σ')))%I with "[Hreify]" as "H". + - iApply (f_equivI with "Hreify"). + - simpl. iExact "H". + } + iPoseProof (reify_vis_eqI_ctx_dep _ _ _ k with "Hgreify") as "Hreify". + iRewrite "Hk" in "Hreify". + by rewrite -Tick_eq. + Qed. + + Lemma wp_reify_idx_ctx_indep (rs : gReifiers NotCtxDep n) + `{!@stateG _ NotCtxDep rs A _ Σ} E1 E2 s Φ i + (lop : opid (sReifier_ops NotCtxDep (rs !!! i))) : + let op : opid (F NotCtxDep rs) := (existT i lop) in + forall (x : Ins (F NotCtxDep rs op) ♯ IT NotCtxDep rs) + (k : Outs (F NotCtxDep rs op) ♯ IT NotCtxDep rs -n> laterO (IT NotCtxDep rs)), + (|={E1,E2}=> ∃ σ y σ' β, has_state_idx NotCtxDep rs i σ + ∗ sReifier_re NotCtxDep (rs !!! i) lop (x, σ) ≡ Some (y, σ') + ∗ k y ≡ Next β + ∗ ▷ (£ 1 -∗ has_state_idx NotCtxDep rs i σ' ={E2,E1}=∗ wp NotCtxDep rs β s E1 Φ)) + -∗ wp NotCtxDep rs (Vis op x k) s E1 Φ. + Proof. + intros op x k. + iIntros "H". + iApply wp_reify_idx. + iMod "H" as (σ y σ' β) "[Hlst [Hreify [Hk H]]]". + iModIntro. iExists σ, σ', β. + iFrame "Hlst". + iIntros (rest). + iFrame "H". + iAssert (gReifiers_re NotCtxDep rs op (x, gState_recomp NotCtxDep rest σ) + ≡ Some (y, gState_recomp NotCtxDep rest σ'))%I + with "[Hreify]" as "Hgreify". + { pose proof (@gReifiers_re_idx n NotCtxDep i rs (IT NotCtxDep rs)) as J. + simpl in J. + simpl. + rewrite J; clear J. + iAssert (optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) + (sReifier_re NotCtxDep (rs !!! i) lop (x, σ)) + ≡ optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) + (Some (y, σ')))%I with "[Hreify]" as "H". + - iApply (f_equivI with "Hreify"). + - simpl. iExact "H". + } + iPoseProof (reify_vis_eqI_ctx_indep _ _ _ k with "Hgreify") as "Hreify". + iRewrite "Hk" in "Hreify". + by rewrite -Tick_eq. + Qed. + + Lemma wp_subreify_ctx_dep' (rs : gReifiers CtxDep n) + `{!@stateG _ CtxDep rs A _ Σ} E1 E2 s Φ sR `{!subReifier sR rs} + (op : opid (sReifier_ops CtxDep sR)) (x : Ins (sReifier_ops CtxDep sR op) ♯ (IT CtxDep rs)) + (k : Outs (F CtxDep rs (subEff_opid op)) ♯ IT CtxDep rs -n> laterO (IT CtxDep rs)) : + (|={E1,E2}=> ∃ σ y σ' β, has_substate CtxDep rs σ ∗ + sReifier_re CtxDep sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') + ∗ y ≡ Next β + ∗ ▷ (£ 1 -∗ has_substate CtxDep rs σ' ={E2,E1}=∗ wp CtxDep rs β s E1 Φ)) + -∗ wp CtxDep rs (Vis (subEff_opid op) (subEff_ins x) k) s E1 Φ. + Proof. + iIntros "H". + iApply wp_reify_idx_ctx_dep. + iMod "H" as (σ y σ' β) "[Hlst [Hreify [Hk H]]]". + iModIntro. + iExists (sR_state σ), y, (sR_state σ'), β. + simpl. + iFrame "Hlst H". + rewrite subReifier_reify_idxI_ctx_dep. + iFrame "Hk". + iRewrite - "Hreify". + iPureIntro. + do 2 f_equiv. + intros ?; simpl. + by rewrite ofe_iso_12. + Qed. + + Lemma wp_subreify_ctx_indep' (rs : gReifiers NotCtxDep n) + `{!@stateG _ NotCtxDep rs A _ Σ} E1 E2 s Φ sR `{!subReifier sR rs} + (op : opid (sReifier_ops NotCtxDep sR)) (x : Ins (sReifier_ops NotCtxDep sR op) ♯ (IT NotCtxDep rs)) + (k : Outs (F NotCtxDep rs (subEff_opid op)) ♯ IT NotCtxDep rs -n> laterO (IT NotCtxDep rs)) : + (|={E1,E2}=> ∃ σ y σ' β, has_substate NotCtxDep rs σ ∗ + sReifier_re NotCtxDep sR op (x, σ) ≡ Some (y, σ') + ∗ k (subEff_outs y) ≡ Next β + ∗ ▷ (£ 1 -∗ has_substate NotCtxDep rs σ' ={E2,E1}=∗ wp NotCtxDep rs β s E1 Φ)) + -∗ wp NotCtxDep rs (Vis (subEff_opid op) (subEff_ins x) k) s E1 Φ. + Proof. + iIntros "H". + iApply wp_reify_idx_ctx_indep. + iMod "H" as (σ y σ' β) "[Hlst [Hreify [Hk H]]]". + iModIntro. + iExists (sR_state σ),(subEff_outs y), (sR_state σ'), β. + iFrame "Hlst H Hk". + by iApply subReifier_reify_idxI_ctx_indep. + Qed. + + Lemma wp_subreify_ctx_dep (rs : gReifiers CtxDep n) + `{!@stateG _ CtxDep rs A _ Σ} E1 s Φ sR `{!subReifier sR rs} + (op : opid (sReifier_ops CtxDep sR)) + (x : Ins (sReifier_ops CtxDep sR op) ♯ IT CtxDep rs) (y : laterO (IT CtxDep rs)) + (k : Outs (F CtxDep rs (subEff_opid op)) ♯ IT CtxDep rs -n> laterO (IT CtxDep rs)) + (σ σ' : sReifier_state CtxDep sR ♯ IT CtxDep rs) β : + sReifier_re CtxDep sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') → + y ≡ Next β → + has_substate CtxDep rs σ -∗ + ▷ (£ 1 -∗ has_substate CtxDep rs σ' -∗ wp CtxDep rs β s E1 Φ) + -∗ + wp CtxDep rs (Vis (subEff_opid op) (subEff_ins x) k) s E1 Φ. + Proof. + intros HSR Hk. + iIntros "Hlst H". + iApply (wp_reify with "Hlst H"). + intros rest. + rewrite Tick_eq. rewrite -Hk. + rewrite reify_vis_eq_ctx_dep //. + pose proof (@subReifier_reify n CtxDep sR rs _ + (IT CtxDep rs) _ op x y (k ◎ subEff_outs) σ σ' rest) as H'. + simpl in H'. + rewrite <-H'. + - simpl. + repeat f_equiv. + + intros ???. + solve_proper. + + intros ?; simpl. + rewrite ofe_iso_12. + reflexivity. + - rewrite HSR. + reflexivity. + Qed. + + Lemma wp_subreify_ctx_indep (rs : gReifiers NotCtxDep n) + `{!@stateG _ NotCtxDep rs A _ Σ} E1 s Φ sR `{!subReifier sR rs} + (op : opid (sReifier_ops NotCtxDep sR)) + (x : Ins (sReifier_ops NotCtxDep sR op) ♯ IT NotCtxDep rs) + (y : Outs (sReifier_ops NotCtxDep sR op) ♯ IT NotCtxDep rs) + (k : Outs (F NotCtxDep rs (subEff_opid op)) ♯ IT NotCtxDep rs -n> laterO (IT NotCtxDep rs)) + (σ σ' : sReifier_state NotCtxDep sR ♯ IT NotCtxDep rs) β : + sReifier_re NotCtxDep sR op (x, σ) ≡ Some (y, σ') → + k (subEff_outs y) ≡ Next β → + has_substate NotCtxDep rs σ -∗ + ▷ (£ 1 -∗ has_substate NotCtxDep rs σ' -∗ wp NotCtxDep rs β s E1 Φ) + -∗ + wp NotCtxDep rs (Vis (subEff_opid op) (subEff_ins x) k) s E1 Φ. + Proof. + intros HSR Hk. + iIntros "Hlst H". + iApply (wp_reify with "Hlst H"). + intros rest. + rewrite Tick_eq. rewrite -Hk. + rewrite reify_vis_eq_ctx_indep //. + by apply (subReifier_reify (a := NotCtxDep)). + Qed. + +End weakestpre_specific. + +Section weakestpre_bind. + Context {n : nat} (rs : gReifiers NotCtxDep n) {A} `{!Cofe A}. + Notation rG := (gReifiers_sReifier NotCtxDep rs). + Notation F := (sReifier_ops NotCtxDep rG). + Notation IT := (IT F A). + Notation ITV := (ITV F A). + Notation stateF := (gReifiers_state NotCtxDep rs). + Notation stateO := (stateF ♯ IT). + Notation stateR := (gReifiers_ucmra NotCtxDep rs IT). + Let of_state := (of_state NotCtxDep rs IT). + Let of_idx := (of_idx NotCtxDep rs IT). + Notation reify := (reify rG). + Notation istep := (istep rG). + Notation isteps := (isteps rG). + Notation sstep := (sstep rG). + Notation ssteps := (ssteps rG). + Notation wp := (wp NotCtxDep rs). + + Implicit Type op : opid F. + Implicit Type α β : IT. + + Context `{!invGS Σ} `{!@stateG _ NotCtxDep rs A _ Σ}. + Notation iProp := (iProp Σ). + Notation coPsetO := (leibnizO coPset). + + Notation "'WP' α @ s ; E {{ Φ } }" := (wp α s E Φ) + (at level 20, α, s, Φ at level 200, only parsing) : bi_scope. + + Notation "'WP' α @ s ; E {{ v , Q } }" := (wp α s E (λ v, Q)) + (at level 20, α, s, Q at level 200, + format "'[hv' 'WP' α '/' @ s ; E '/' {{ '[' v , '/' Q ']' } } ']'") : bi_scope. + + Lemma wp_bind (f : IT → IT) `{!IT_hom f} (α : IT) s Φ `{!NonExpansive Φ} E1 : WP α @ s;E1 {{ βv, WP (f (IT_of_V βv)) @ s;E1 {{ βv, Φ βv }} }} ⊢ WP (f α) @ s;E1 {{ Φ }}. Proof. assert (NonExpansive (λ βv0, WP f (IT_of_V βv0) @ s;E1 {{ βv1, Φ βv1 }})%I). { solve_proper. } iIntros "H". iLöb as "IH" forall (α). - rewrite (wp_unfold (f _)). + rewrite (wp_unfold _ _ (f _)). destruct (IT_to_V (f α)) as [βv|] eqn:Hfa. - iLeft. iExists βv. iSplit; first done. assert (is_Some (IT_to_V α)) as [αv Ha]. @@ -792,15 +955,15 @@ Section weakestpre. iModIntro. iRewrite "Hb". by iApply "IH". Qed. -End weakestpre. +End weakestpre_bind. -Arguments wp {_} rs {_ _ _ _ _} α s E Φ. -Arguments has_full_state {n _ _ _ _ _} σ. -Arguments has_state_idx {n _ _ _ _ _} i σ. -Arguments has_substate {n _ _ _ _ _ _ _} σ. -Arguments stateG {n} rs A {_} Σ. -Arguments statePreG {n} rs A {_} Σ. -Arguments stateΣ {n} rs A {_}. +Arguments wp {_ _} rs {_ _ _ _ _} α s E Φ. +Arguments has_full_state {n _ _ _ _ _ _} σ. +Arguments has_state_idx {n _ _ _ _ _ _} i σ. +Arguments has_substate {n _ _ _ _ _ _ _ _} σ. +Arguments stateG {n _} rs A {_} Σ. +Arguments statePreG {n _} rs A {_} Σ. +Arguments stateΣ {n _} rs A {_}. Definition notStuck : stuckness := λ e, False. @@ -827,10 +990,10 @@ Definition notStuck : stuckness := λ e, False. (at level 20, α, Φ at level 200, format "'WP@{' re } α {{ Φ } }") : bi_scope. - Lemma wp_adequacy cr Σ `{!invGpreS Σ} n (rs : gReifiers n) + Lemma wp_adequacy cr Σ `{!invGpreS Σ} n a (rs : gReifiers a n) {A} `{!Cofe A} `{!statePreG rs A Σ} - (α : IT _ A) σ βv σ' s k (ψ : (ITV (gReifiers_ops rs) A) → Prop) : - ssteps (gReifiers_sReifier rs) α σ (IT_of_V βv) σ' k → + (α : IT _ A) σ βv σ' s k (ψ : (ITV (gReifiers_ops a rs) A) → Prop) : + ssteps (gReifiers_sReifier a rs) α σ (IT_of_V βv) σ' k → (∀ `{H1 : !invGS Σ} `{H2: !stateG rs A Σ}, ∃ Φ, NonExpansive Φ ∧ (∀ βv, Φ βv ⊢ ⌜ψ βv⌝) ∧ (£ cr ∗ has_full_state σ ⊢ WP@{rs} α @ s {{ Φ }})%I) → @@ -841,7 +1004,7 @@ Definition notStuck : stuckness := λ e, False. { intros HH. eapply uPred.pure_soundness; eauto. } eapply (step_fupdN_soundness_lc _ 0 (cr + 3*k)). intros Hinv. iIntros "[Hcr Hlc]". - iMod (new_state_interp rs σ) as (sg) "[Hs Hs2]". + iMod (new_state_interp a rs σ) as (sg) "[Hs Hs2]". destruct (Hprf Hinv sg) as (Φ & HΦ & HΦψ & Hprf'). iPoseProof (Hprf' with "[$Hcr $Hs2]") as "Hic". iPoseProof (wp_ssteps with "[$Hs $Hic]") as "Hphi". @@ -853,20 +1016,20 @@ Definition notStuck : stuckness := λ e, False. by iApply fupd_mask_intro_discard. Qed. - Lemma wp_safety cr Σ `{!invGpreS Σ} n (rs : gReifiers n) + Lemma wp_safety cr Σ `{!invGpreS Σ} n a (rs : gReifiers a n) {A} `{!Cofe A} `{!statePreG rs A Σ} s k - (α β : IT (gReifiers_ops rs) A) (σ σ' : gReifiers_state rs ♯ IT (gReifiers_ops rs) A) : + (α β : IT (gReifiers_ops a rs) A) (σ σ' : gReifiers_state a rs ♯ IT (gReifiers_ops a rs) A) : (∀ Σ P Q, @disjunction_property Σ P Q) → - ssteps (gReifiers_sReifier rs) α σ β σ' k → + ssteps (gReifiers_sReifier a rs) α σ β σ' k → IT_to_V β ≡ None → (∀ `{H1 : !invGS_gen HasLc Σ} `{H2: !stateG rs A Σ}, ∃ Φ, NonExpansive Φ ∧ (£ cr ∗ has_full_state σ ⊢ WP@{rs} α @ s {{ Φ }})%I) → - ((∃ β1 σ1, sstep (gReifiers_sReifier rs) β σ' β1 σ1) + ((∃ β1 σ1, sstep (gReifiers_sReifier a rs) β σ' β1 σ1) ∨ (∃ e, β ≡ Err e ∧ s e)). Proof. Opaque istep. intros Hdisj Hstep Hbv Hwp. - cut (⊢@{iProp Σ} (∃ β1 σ1, istep (gReifiers_sReifier rs) β σ' β1 σ1) + cut (⊢@{iProp Σ} (∃ β1 σ1, istep (gReifiers_sReifier a rs) β σ' β1 σ1) ∨ (∃ e, β ≡ Err e ∧ ⌜s e⌝))%I. { intros [Hprf | Hprf]%Hdisj. - left. @@ -900,7 +1063,7 @@ Definition notStuck : stuckness := λ e, False. iApply (IT_vis_err_ne with "Ha"). } eapply (step_fupdN_soundness_lc _ 0 (cr + (3*k+2))). intros Hinv. iIntros "[Hcr Hlc]". - iMod (new_state_interp rs σ) as (sg) "[Hs Hs2]". + iMod (new_state_interp a rs σ) as (sg) "[Hs Hs2]". destruct (Hwp Hinv sg) as (Φ & HΦ & Hprf'). iPoseProof (Hprf' with "[$Hs2 $Hcr]") as "Hic". iPoseProof (wp_ssteps_isafe with "[$Hs $Hic]") as "H". From 733ff1dc6fb0c52122a439e96646f4d6f965a301 Mon Sep 17 00:00:00 2001 From: Kaptch Date: Mon, 29 Jan 2024 21:17:52 +0100 Subject: [PATCH 2/7] simplify ctxdep reifiers --- _CoqProject | 2 +- theories/affine_lang/lang.v | 16 +- theories/affine_lang/logrel1.v | 188 +------ theories/affine_lang/logrel2.v | 301 ++++------ theories/examples/factorial.v | 18 +- theories/examples/iter.v | 9 +- theories/examples/store.v | 78 ++- theories/input_lang/interp.v | 822 +++++++++++++++------------- theories/input_lang/lang.v | 723 +++++++++++++----------- theories/input_lang/logpred.v | 231 ++++---- theories/input_lang/logrel.v | 406 ++++++-------- theories/input_lang_callcc/hom.v | 7 +- theories/input_lang_callcc/interp.v | 84 ++- theories/input_lang_callcc/lang.v | 2 - theories/input_lang_callcc/logrel.v | 27 +- theories/lang_affine.v | 245 +++++++++ theories/lang_generic.v | 239 +++----- theories/lang_generic_sem.v | 104 ---- theories/program_logic.v | 17 +- 19 files changed, 1661 insertions(+), 1858 deletions(-) create mode 100644 theories/lang_affine.v delete mode 100644 theories/lang_generic_sem.v diff --git a/_CoqProject b/_CoqProject index 018b6cf..e8cf1c7 100644 --- a/_CoqProject +++ b/_CoqProject @@ -16,7 +16,7 @@ vendor/Binding/Resolver.v theories/prelude.v theories/lang_generic.v -theories/lang_generic_sem.v +theories/lang_affine.v theories/gitree/core.v theories/gitree/subofe.v diff --git a/theories/affine_lang/lang.v b/theories/affine_lang/lang.v index 3096ec2..9616cb1 100644 --- a/theories/affine_lang/lang.v +++ b/theories/affine_lang/lang.v @@ -7,12 +7,14 @@ Module io_lang. Definition state := input_lang.lang.state. Definition ty := input_lang.lang.ty. Definition expr := input_lang.lang.expr. - Definition tyctx := tyctx ty. - Definition typed {S} := input_lang.lang.typed (S:=S). - Definition interp_closed {sz} (rs : gReifiers sz) `{!subReifier reify_io rs} (e : expr []) {R} `{!Cofe R, !SubOfe natO R} : IT (gReifiers_ops rs) R := - input_lang.interp.interp_expr rs e (). + Definition tyctx {S : Set} := S → ty. + Definition typed {S : Set} := input_lang.lang.typed (S:=S). + Program Definition ı_scope {sz} (rs : gReifiers NotCtxDep sz) `{!subReifier reify_io rs} {R} `{!Cofe R} : @interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. + Definition interp_closed {sz} (rs : gReifiers NotCtxDep sz) `{!subReifier reify_io rs} (e : expr ∅) {R} `{!Cofe R, !SubOfe natO R} : IT (gReifiers_ops NotCtxDep rs) R := + input_lang.interp.interp_expr rs e (ı_scope rs). End io_lang. +From gitrees Require Export lang_affine. Inductive ty := tBool | tInt | tUnit @@ -42,15 +44,15 @@ Inductive expr : scope → Type := | Alloc {S} : expr S → expr S | Replace {S1 S2} : expr S1 → expr S2 → expr (S1++S2) | Dealloc {S} : expr S → expr S -| EEmbed {τ1 τ1' S} : io_lang.expr [] → ty_conv τ1 τ1' → expr S +| EEmbed {τ1 τ1' S} : io_lang.expr Empty_set → ty_conv τ1 τ1' → expr S . Section affine. Context {sz : nat}. - Variable rs : gReifiers sz. + Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_store rs}. Context `{!subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Context {R : ofe}. Context `{!Cofe R, !SubOfe unitO R, !SubOfe natO R, !SubOfe locO R}. Notation IT := (IT F R). diff --git a/theories/affine_lang/logrel1.v b/theories/affine_lang/logrel1.v index 9710d44..b5757b3 100644 --- a/theories/affine_lang/logrel1.v +++ b/theories/affine_lang/logrel1.v @@ -1,6 +1,6 @@ (** Unary (Kripke) logical relation for the affine lang *) From Equations Require Import Equations. -From gitrees Require Export lang_generic gitree program_logic. +From gitrees Require Export lang_affine gitree program_logic. From gitrees.affine_lang Require Import lang. From gitrees.examples Require Import store pairs. Require Import iris.algebra.gmap. @@ -51,10 +51,10 @@ Inductive typed : forall {S}, tyctx S → expr S → ty → Prop := Section logrel. Context {sz : nat}. - Variable rs : gReifiers sz. + Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_store rs}. Context `{!subReifier input_lang.interp.reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Context `{!SubOfe unitO R}. @@ -67,12 +67,8 @@ Section logrel. (* parameters for the kripke logical relation *) Variable s : stuckness. Context `{A:ofe}. - Variable (P : A → iProp). - Context `{!NonExpansive P}. + Variable (P : A -n> iProp). Local Notation expr_pred := (expr_pred s rs P). - Context {HCI : - ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) IT o}. (* interpreting tys *) Program Definition protected (Φ : ITV -n> iProp) : ITV -n> iProp := λne αv, @@ -110,7 +106,7 @@ Section logrel. end. Definition ssubst_valid {S} (Ω : tyctx S) ss := - lang_generic.ssubst_valid rs (λ τ, protected (interp_ty τ)) Ω ss. + lang_affine.ssubst_valid rs (λ τ, protected (interp_ty τ)) Ω ss. Definition valid1 {S} (Ω : tyctx S) (α : interp_scope S -n> IT) (τ : ty) : iProp := ∀ ss, heap_ctx -∗ ssubst_valid Ω ss -∗ expr_pred (α (interp_ssubst ss)) (interp_ty τ). @@ -119,7 +115,7 @@ Section logrel. ⊢ valid1 Ω1 α τ1 -∗ valid1 Ω2 β τ2 -∗ valid1 (tyctx_app Ω1 Ω2) (interp_pair α β ◎ interp_scope_split) (tPair τ1 τ2). - Proof using HCI. + Proof. Opaque pairITV. iIntros "H1 H2". iIntros (αs) "#Hctx Has". @@ -144,7 +140,7 @@ Section logrel. ⊢ valid1 Ω1 α (tPair τ1 τ2) -∗ valid1 (consC τ1 $ consC τ2 Ω2) β τ -∗ valid1 (tyctx_app Ω1 Ω2) (interp_destruct α β ◎ interp_scope_split) τ. - Proof using HCI. + Proof. Opaque pairITV thunked thunkedV projIT1 projIT2. iIntros "H1 H2". iIntros (αs) "#Hctx Has". @@ -211,7 +207,7 @@ Section logrel. Lemma compat_alloc {S} (Ω : tyctx S) α τ: ⊢ valid1 Ω α τ -∗ valid1 Ω (interp_alloc α) (tRef τ). - Proof using HCI. + Proof. iIntros "H". iIntros (αs) "#Hctx Has". iSpecialize ("H" with "Hctx Has"). @@ -229,7 +225,7 @@ Section logrel. ⊢ valid1 Ω1 α (tRef τ) -∗ valid1 Ω2 β τ' -∗ valid1 (tyctx_app Ω1 Ω2) (interp_replace α β ◎ interp_scope_split) (tPair τ (tRef τ')). - Proof using HCI. + Proof. Opaque pairITV. iIntros "H1 H2". iIntros (αs) "#Hctx Has". @@ -270,7 +266,7 @@ Section logrel. Lemma compat_dealloc {S} (Ω : tyctx S) α τ: ⊢ valid1 Ω α (tRef τ) -∗ valid1 Ω (interp_dealloc α) tUnit. - Proof using HCI. + Proof. iIntros "H". iIntros (αs) "#Hctx Has". iSpecialize ("H" with "Hctx Has"). @@ -328,7 +324,7 @@ Section logrel. ⊢ valid1 Ω1 α (tArr τ1 τ2) -∗ valid1 Ω2 β τ1 -∗ valid1 (tyctx_app Ω1 Ω2) (interp_app α β ◎ interp_scope_split) τ2. - Proof using HCI. + Proof. iIntros "H1 H2". iIntros (αs) "#Hctx Has". iEval(cbn-[interp_app]). @@ -352,7 +348,7 @@ Section logrel. Lemma compat_lam {S} (Ω : tyctx S) τ1 τ2 α : ⊢ valid1 (consC τ1 Ω) α τ2 -∗ valid1 Ω (interp_lam α) (tArr τ1 τ2). - Proof using HCI. + Proof. iIntros "H". iIntros (αs) "#Hctx Has". iIntros (x) "Hx". @@ -397,7 +393,7 @@ Section logrel. Lemma fundamental_affine {S} (Ω : tyctx S) (e : expr S) τ : typed Ω e τ → ⊢ valid1 Ω (interp_expr _ e) τ. - Proof using HCI. + Proof. induction 1; simpl. - by iApply compat_var. - by iApply compat_lam. @@ -420,157 +416,25 @@ Arguments interp_tnat {_ _ _ _ _ _}. Arguments interp_tunit {_ _ _ _ _ _}. Arguments interp_ty {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _} τ. -Local Definition rs : gReifiers 2 := gReifiers_cons reify_store (gReifiers_cons input_lang.interp.reify_io gReifiers_nil). - -Local Instance CtxIndepInputLang R `{!Cofe R} (o : opid (sReifier_ops (gReifiers_sReifier rs))) : - CtxIndep (gReifiers_sReifier rs) (IT (gReifiers_ops rs) R) o. -Proof. - destruct o as [x o]. - inv_fin x. - - simpl. intros [[]| [[]| [[] | [| []]]]]. - + constructor. - unshelve eexists (λne '(l,(σ, σ')), x ← σ !! l; - Some (x, (σ, σ'))). - * apply _. - * apply _. - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]; simpl in *. - apply (option_mbind_ne _ (λ n, Some (n, _)) (λ n, Some (n, _))). - -- intros ? ? ?; repeat f_equiv; [done | |]; apply H. - -- rewrite lookup_ne; last apply H. - simpl. - f_equiv. - apply H. - * intros. - simpl. - destruct σ as [? [? ?]]. - simpl. - match goal with - | |- context G [@mbind option option_bind _ _ ?a ?b] => set (x := b) - end. - symmetry. - match goal with - | |- context G [@mbind option option_bind _ _ ?a ?b] => set (y := b) - end. - assert (y = x) as ->. - { reflexivity. } - destruct x as [x |]; reflexivity. - + constructor. - unshelve eexists (λne '((l,n),(s, s'')), let s' := <[l:=n]>s - in Some ((), (s', s''))). - * apply _. - * solve_proper_prepare. - destruct x as [[? ?] [? ?]]; destruct y as [[? ?] [? ?]]; simpl in *. - do 3 f_equiv; last apply H. - rewrite insert_ne; [| apply H | apply H]. - simpl. - f_equiv. - apply H. - * intros. - simpl. - destruct i as [? ?]. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - + constructor. - unshelve eexists (λne '(n,(s, s'')), let l := Loc.fresh (dom s) in - let s' := <[l:=n]>s in - Some (l, (s', s''))). - * apply _. - * apply _. - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]; simpl in *. - do 2 f_equiv. - -- f_equiv. - destruct H as [_ [H _]]; simpl in H. - apply gmap_dom_ne in H. - apply H. - -- f_equiv; last apply H. - rewrite insert_ne; [| apply H | apply H]. - simpl. - f_equiv. - destruct H as [_ [H _]]; simpl in H. - apply gmap_dom_ne in H. - by rewrite H. - * intros. - simpl. - destruct i as [? ?]. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - + constructor. - simpl. - unshelve eexists (λne '(l,(σ, σ')), Some ((), (delete l σ, σ'))). - * apply _. - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]; simpl in *. - do 2 f_equiv. - f_equiv; last apply H. - rewrite delete_ne; last apply H. - simpl. - f_equiv. - apply H. - * intros. - simpl. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - - intros x; inv_fin x. - + simpl. intros [[]| [[]| []]]. - * constructor. - unshelve eexists (λne '(_, (a, (b, c))), SomeO (_, (_, (_, c)))). - -- simpl in *. - apply ((input_lang.lang.update_input b).1). - -- apply a. - -- apply ((input_lang.lang.update_input b).2). - -- solve_proper_prepare. - destruct x as [? [? [? ?]]]; destruct y as [? [? [? ?]]]. - simpl in *. - do 2 f_equiv. - ++ do 2 f_equiv. - apply H. - ++ f_equiv; first apply H. - f_equiv; last apply H. - do 2 f_equiv; apply H. - -- intros. - simpl. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - * constructor. - unshelve eexists (λne '(x, (y, z)), SomeO ((), _)). - -- simpl in *. - apply (y, ((input_lang.lang.update_output x (fstO z)), ())). - -- solve_proper_prepare. - destruct x as [? [? [? ?]]]; destruct y as [? [? [? ?]]]. - simpl in *. - do 2 f_equiv. - apply pair_ne. - ++ apply H. - ++ do 2 f_equiv; apply H. - -- intros. - simpl. - destruct σ as [σ1 [? []]]; simpl in *. - reflexivity. - + intros i; by apply fin_0_inv. -Qed. +Local Definition rs : gReifiers NotCtxDep 2 := + gReifiers_cons NotCtxDep reify_store (gReifiers_cons NotCtxDep input_lang.interp.reify_io (gReifiers_nil NotCtxDep)). Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. Lemma logrel1_adequacy cr Σ R `{!Cofe R, !SubOfe natO R, !SubOfe unitO R, !SubOfe locO R} `{!invGpreS Σ} `{!statePreG rs R Σ} `{!heapPreG rs R Σ} τ - (α : unitO -n> IT (gReifiers_ops rs) R) (β : IT (gReifiers_ops rs) R) st st' k : + (α : unitO -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ}, - (£ cr ⊢ valid1 rs notStuck (λ _:unitO, True)%I empC α τ)%I) → - ssteps (gReifiers_sReifier rs) (α ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + (£ cr ⊢ valid1 rs notStuck (λne _: unitO, True)%I empC α τ)%I) → + ssteps (gReifiers_sReifier NotCtxDep rs) (α ()) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. intros Hlog Hst. destruct (IT_to_V β) as [βv|] eqn:Hb. { right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } left. - cut ((∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ e, β ≡ Err e ∧ notStuck e)). { intros [?|He]; first done. destruct He as [? [? []]]. } @@ -586,9 +450,9 @@ Proof. iMod (new_heapG rs σ) as (H3) "H". iAssert (has_substate σ ∗ has_substate ios)%I with "[Hst]" as "[Hs Hio]". { unfold has_substate, has_full_state. - assert (of_state rs (IT (gReifiers_ops rs) _) st ≡ - of_idx rs (IT (gReifiers_ops rs) _) sR_idx (sR_state σ) - ⋅ of_idx rs (IT (gReifiers_ops rs) _) sR_idx (sR_state ios)) as ->; last first. + assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) st ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ) + ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state ios)) as ->; last first. { rewrite -own_op. done. } unfold sR_idx. simpl. intro j. @@ -615,10 +479,10 @@ Qed. Definition R := sumO locO (sumO unitO natO). -Lemma logrel1_safety e τ (β : IT (gReifiers_ops rs) R) st st' k : +Lemma logrel1_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : typed empC e τ → - ssteps (gReifiers_sReifier rs) (interp_expr rs e ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ()) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. intros Hty Hst. diff --git a/theories/affine_lang/logrel2.v b/theories/affine_lang/logrel2.v index 6c5b827..a45dac8 100644 --- a/theories/affine_lang/logrel2.v +++ b/theories/affine_lang/logrel2.v @@ -1,18 +1,20 @@ From Equations Require Import Equations. From iris.base_logic.lib Require Import na_invariants. -From gitrees Require Export lang_generic gitree program_logic. +From gitrees Require Export lang_affine gitree program_logic. From gitrees.input_lang Require Import lang interp logpred. From gitrees.affine_lang Require Import lang logrel1. From gitrees.examples Require Import store pairs. Require Import iris.algebra.gmap. +Require Import Binding.Lib Binding.Set Binding.Env. + Local Notation tyctx := (tyctx ty). Inductive typed_glued : forall {S}, tyctx S → expr S → ty → Type := (** FFI *) | typed_Glue {S} (Ω : tyctx S) τ' τ e (tconv : ty_conv τ τ') : - io_lang.typed empC e τ' → + io_lang.typed □ e τ' → typed_glued Ω (EEmbed e tconv) τ (** functions *) | typed_VarG {S} (Ω : tyctx S) (τ : ty) (v : var S) : @@ -57,10 +59,10 @@ Inductive typed_glued : forall {S}, tyctx S → expr S → ty → Type := Section glue. Context {sz : nat}. - Variable rs : gReifiers sz. + Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_store rs}. Context `{!subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Context `{!SubOfe unitO R}. @@ -70,30 +72,30 @@ Section glue. Context `{!invGS Σ, !stateG rs R Σ, !heapG rs R Σ, !na_invG Σ}. Notation iProp := (iProp Σ). - Context {HCI : - ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) IT o}. - Definition s : stuckness := λ e, e = OtherError. Variable p : na_inv_pool_name. Definition valid2 {S} (Ω : tyctx S) (α : interp_scope (E:=F) S -n> IT) (τ : ty) : iProp := - valid1 rs s (λ σ, has_substate σ ∗ na_own p ⊤)%I Ω α τ. + valid1 rs s (λne σ, has_substate σ ∗ na_own p ⊤)%I Ω α τ. - Definition io_valid {S} (Γ : io_lang.tyctx S) α (τ' : io_lang.ty) : iProp := - input_lang.logpred.valid1 rs s (λ _ : unitO, na_own p ⊤) Γ α τ'. + Definition io_valid {S : Set} (Γ : S → io_lang.ty) α (τ' : io_lang.ty) : iProp := + input_lang.logpred.valid1 rs s (λne _ : unitO, na_own p ⊤) Γ α τ'. Local Opaque thunked thunkedV Thunk. + + Program Definition ı_scope : @lang_generic.interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. + Lemma compat_glue_to_affine_bool {S} (Ω : tyctx S) α : - io_valid empC α Tnat ⊢ - valid2 Ω (constO (glue2_bool _ (α ()))) tBool. - Proof using HCI. + io_valid □ α Tnat ⊢ + valid2 Ω (constO (glue2_bool _ (α ı_scope))) tBool. + Proof. iIntros "H". iIntros (ss) "#Hctx Has". simpl. iIntros (σ) "[Hs Hp]". - iSpecialize ("H" $! σ emp_ssubst with "Hs []"). + iSpecialize ("H" $! σ ı_scope with "Hs []"). { unfold logpred.ssubst_valid. - iApply ssubst_valid_nil. } + iIntros ([]). + } iSpecialize ("H" $! tt with "Hp"). simp interp_ssubst. simpl. iApply (wp_bind _ (IFSCtx _ _)). @@ -111,48 +113,58 @@ Section glue. iApply wp_val; eauto with iFrame. Qed. Lemma compat_glue_to_affine_nat {S} (Ω : tyctx S) α : - io_valid empC α Tnat ⊢ - valid2 Ω (constO (α ())) tInt. + io_valid □ α Tnat ⊢ + valid2 Ω (constO (α ı_scope)) tInt. Proof. iIntros "H". iIntros (ss) "#Hctx Has". simpl. iIntros (σ) "[Hs Hp]". - iSpecialize ("H" $! σ emp_ssubst with "Hs []"). + iSpecialize ("H" $! σ ı_scope with "Hs []"). { unfold logpred.ssubst_valid. - iApply ssubst_valid_nil. } + iIntros ([]). + } iSpecialize ("H" $! tt with "Hp"). simp interp_ssubst. simpl. iApply (wp_wand with "H"). iIntros (αv). iDestruct 1 as (_) "[Ha Hp]". iDestruct "Ha" as (σ') "[Ha Hs]". iModIntro. eauto with iFrame. - Qed. + + Lemma IT_move_from_affine (α : @interp_scope (@gReifiers_ops NotCtxDep sz rs) R _ [] -n> IT) + : @lang_generic.interp_scope (@gReifiers_ops NotCtxDep sz rs) R _ ∅ -n> IT. + Proof. + unshelve econstructor. + - intros g. apply α. + constructor. + - repeat intro. + f_equiv. + Defined. + Lemma compat_glue_from_affine_bool α : valid2 empC α tBool ⊢ - heap_ctx -∗ io_valid empC α Tnat. + heap_ctx -∗ io_valid □ (IT_move_from_affine α) Tnat. Proof. iIntros "H #Hctx". iIntros (σ ss) "Hs Hss". - iIntros (_) "Hp". + iIntros (?) "Hp". iSpecialize ("H" $! emp_ssubst with "Hctx [] [$Hs $Hp]"). { iApply ssubst_valid_nil. } - dependent elimination ss as [emp_ssubst]. iApply (wp_wand with "H"). iIntros (αv) "Ha". iDestruct "Ha" as (σ') "[Ha [Hs Hp]]". iModIntro. simpl. iFrame. iExists tt,_; iFrame. iDestruct "Ha" as "[Ha|Ha]"; iExists _; eauto. Qed. + Lemma compat_glue_from_affine_nat α : valid2 empC α tInt ⊢ - heap_ctx -∗ io_valid empC α Tnat. + heap_ctx -∗ io_valid □ (IT_move_from_affine α) Tnat. Proof. iIntros "H #Hctx". iIntros (σ ss) "Hs Hss". - iIntros (_) "Hp". + iIntros (?) "Hp". iSpecialize ("H" $! emp_ssubst with "Hctx [] [$Hs $Hp]"). { iApply ssubst_valid_nil. } - dependent elimination ss as [emp_ssubst]. iApply (wp_wand with "H"). iIntros (αv) "Ha". iDestruct "Ha" as (σ') "[Ha [Hs Hp]]". iModIntro. iExists tt. eauto with iFrame. @@ -160,11 +172,11 @@ Section glue. Lemma compat_glue_from_affine_unit α : valid2 empC α tUnit ⊢ - heap_ctx -∗ io_valid empC (constO (glue_from_affine _ ty_conv_unit (α ()))) Tnat. + heap_ctx -∗ io_valid □ (constO (glue_from_affine _ ty_conv_unit (α ()))) Tnat. Proof. iIntros "H #Hctx". iIntros (σ ss) "Hs Hss". - iIntros (_) "Hp". + iIntros (?) "Hp". simpl. iApply wp_val. iModIntro. iExists tt. iFrame. simpl. eauto with iFrame. @@ -174,19 +186,23 @@ Section glue. Lemma compat_glue_from_affine_fun (τ1 τ2 : ty) (τ1' τ2' : io_lang.ty) α (glue_to_affine glue_from_affine : IT -n> IT) : - (∀ α, io_valid empC α τ1' - ⊢ valid2 empC (constO (glue_to_affine (α ()))) τ1) → + (∀ α, io_valid □ α τ1' + ⊢ valid2 empC (constO (glue_to_affine (α ı_scope))) τ1) → (∀ α, valid2 empC (constO α) τ2 - ⊢ heap_ctx -∗ io_valid empC (constO (glue_from_affine α)) τ2') → + ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine α)) τ2') → valid2 empC (constO α) (tArr τ1 τ2) ⊢ heap_ctx -∗ - io_valid empC + io_valid □ (constO (glue_from_affine_fun _ glue_from_affine glue_to_affine α)) (Tarr (Tarr Tnat τ1') τ2'). - Proof using HCI. + Proof. intros G1 G2. - iIntros "H #Hctx". iIntros (σ ss) "Hs _ _ Hp". - simpl. clear ss. + iIntros "H #Hctx". + unfold io_valid. + unfold logpred.valid1. + iIntros (σ ss) "Hs ?". + simpl. + iIntros (?) "Hp". iSpecialize ("H" $! emp_ssubst with "Hctx [] [$Hs $Hp]"). { iApply ssubst_valid_nil. } simpl. iApply wp_let. @@ -209,7 +225,7 @@ Section glue. iModIntro. simpl. iApply wp_val. iModIntro. iExists tt. iFrame. iExists σ'. iFrame. iModIntro. clear σ σ'. iIntros (σ βv) "Hs #Hb". - iIntros (_) "Hp". + iIntros (?) "Hp". iApply wp_lam. iNext. simpl. iApply wp_let. { solve_proper. } @@ -249,7 +265,7 @@ Section glue. { eauto with iFrame. } iSpecialize ("Hb" $! tt with "Hp"). iApply (wp_wand with "Hb"). - iIntros (γv). iDestruct 1 as (_) "[Hg Hp]". + iIntros (γv). iDestruct 1 as (?) "[Hg Hp]". iDestruct "Hg" as (σ') "[Hg Hst]". iModIntro. simpl. iApply wp_let. @@ -275,28 +291,32 @@ Section glue. iSpecialize ("G1" with "[Hg] Hctx"). { iIntros (ss0) "_ _". by iApply expr_pred_ret. } - iSpecialize ("G1" $! _ emp_ssubst with "Hst []"). - { iApply ssubst_valid_nil. } + iSpecialize ("G1" $! _ ı_scope with "Hst []"). + { + iIntros ([]). + } iApply ("G1" $! tt with "Hp"). Qed. Lemma compat_glue_to_affine_fun {S} (Ω : tyctx S) (τ1 τ2 : ty) (τ1' τ2' : io_lang.ty) α (glue_to_affine glue_from_affine : IT -n> IT) : - (∀ α, io_valid empC α τ2' - ⊢ valid2 Ω (constO (glue_to_affine (α ()))) τ2) → + (∀ α, io_valid □ α τ2' + ⊢ valid2 Ω (constO (glue_to_affine (α ı_scope))) τ2) → (∀ α, valid2 empC (constO α) τ1 - ⊢ heap_ctx -∗ io_valid empC (constO (glue_from_affine α)) τ1') → - io_valid empC α (Tarr (Tarr Tnat τ1') τ2') + ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine α)) τ1') → + io_valid □ α (Tarr (Tarr Tnat τ1') τ2') ⊢ valid2 Ω - (constO (glue_to_affine_fun _ glue_from_affine glue_to_affine (α ()))) + (constO (glue_to_affine_fun _ glue_from_affine glue_to_affine (α ı_scope))) (tArr τ1 τ2). - Proof using HCI. + Proof. intros G1 G2. iIntros "H". iIntros (αs) "#Hctx Has". iIntros (σ) "[Hs Hp]". simpl. - iSpecialize ("H" $! _ emp_ssubst with "Hs []"). - { iApply ssubst_valid_nil. } + iSpecialize ("H" $! _ ı_scope with "Hs []"). + { + iIntros ([]). + } iSpecialize ("H" $! tt with "Hp"). simp interp_ssubst. simpl. iApply wp_let. @@ -345,12 +365,14 @@ Section glue. iIntros (σ0) "Hs". simpl. iApply wp_val. eauto with iFrame. } iSpecialize ("G2" with "Hctx"). - iSpecialize ("G2" $! _ emp_ssubst with "Hs []"). - { iApply ssubst_valid_nil. } + iSpecialize ("G2" $! _ ı_scope with "Hs []"). + { + iIntros ([]). + } iSpecialize ("G2" $! tt with "Hp"). iApply (wp_wand with "G2"). iIntros (β'v). - iDestruct 1 as (_) "[Hb Hp]". iModIntro. + iDestruct 1 as (?) "[Hb Hp]". iModIntro. simpl. clear σ. iDestruct "Hb" as (σ) "[#Hb Hs]". (* calling the original function *) iApply wp_let. @@ -369,7 +391,7 @@ Section glue. iPoseProof ("Ha" $! _ (thunkedV (IT_of_V β'v) l') with "Hs [-Has Hp]") as "H1". { iModIntro. iIntros (σ' βn) "Hs Hbm". iDestruct "Hbm" as (m) "Hbm". - iIntros (_) "Hp". + iIntros (?) "Hp". iApply wp_lam. iNext. iSimpl. iApply (wp_bind _ (IFSCtx _ _)). { solve_proper. } @@ -401,7 +423,7 @@ Section glue. iModIntro. iSpecialize ("H1" $! tt with "Hp"). iApply (wp_wand with "H1"). - iIntros (γv). iDestruct 1 as (_) "[H2 Hp]". + iIntros (γv). iDestruct 1 as (?) "[H2 Hp]". iModIntro. simpl. iDestruct "H2" as (σ') "[#H2 Hs]". iPoseProof (G1 (constO (IT_of_V γv))) as "G1". iSpecialize ("G1" with "[H2]"). @@ -415,11 +437,11 @@ Section glue. Lemma glue_to_affine_compatibility {S} (Ω : tyctx S) (τ1 : ty) (τ1' : io_lang.ty) (Hconv : ty_conv τ1 τ1') α : - io_valid empC α τ1' ⊢ valid2 Ω (constO (glue_to_affine _ Hconv (α ()))) τ1 + io_valid □ α τ1' ⊢ valid2 Ω (constO (glue_to_affine _ Hconv (α ı_scope))) τ1 with glue_from_affine_compatibility (τ1 : ty) (τ1' : io_lang.ty) (Hconv : ty_conv τ1 τ1') (α : IT) : - valid2 empC (constO α) τ1 ⊢ heap_ctx -∗ io_valid empC (constO (glue_from_affine _ Hconv α)) τ1'. - Proof using HCI. + valid2 empC (constO α) τ1 ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine _ Hconv α)) τ1'. + Proof. - destruct Hconv. + by iApply compat_glue_to_affine_bool. + by iApply compat_glue_to_affine_nat. @@ -440,7 +462,7 @@ Section glue. Lemma fundamental_affine_glued {S} (Ω : tyctx S) (e : expr S) τ : typed_glued Ω e τ → ⊢ valid2 Ω (interp_expr _ e) τ. - Proof using HCI. + Proof. intros typed. induction typed; simpl. - iApply glue_to_affine_compatibility. by iApply fundamental. @@ -459,161 +481,32 @@ Section glue. End glue. -Local Definition rs : gReifiers 2 := gReifiers_cons reify_store (gReifiers_cons input_lang.interp.reify_io gReifiers_nil). - -Local Instance CtxIndepInputLang R `{!Cofe R} (o : opid (sReifier_ops (gReifiers_sReifier rs))) : - CtxIndep (gReifiers_sReifier rs) - (IT (sReifier_ops (gReifiers_sReifier rs)) R) o. -Proof. - destruct o as [x o]. - inv_fin x. - - simpl. intros [[]| [[]| [[] | [| []]]]]. - + constructor. - unshelve eexists (λne '(l,(σ, σ')), x ← σ !! l; - Some (x, (σ, σ'))). - * apply _. - * apply _. - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]; simpl in *. - apply (option_mbind_ne _ (λ n, Some (n, _)) (λ n, Some (n, _))). - -- intros ? ? ?; repeat f_equiv; [done | |]; apply H. - -- rewrite lookup_ne; last apply H. - simpl. - f_equiv. - apply H. - * intros. - simpl. - destruct σ as [? [? ?]]. - simpl. - match goal with - | |- context G [@mbind option option_bind _ _ ?a ?b] => set (x := b) - end. - symmetry. - match goal with - | |- context G [@mbind option option_bind _ _ ?a ?b] => set (y := b) - end. - assert (y = x) as ->. - { reflexivity. } - destruct x as [x |]; reflexivity. - + constructor. - unshelve eexists (λne '((l,n),(s, s'')), let s' := <[l:=n]>s - in Some ((), (s', s''))). - * apply _. - * solve_proper_prepare. - destruct x as [[? ?] [? ?]]; destruct y as [[? ?] [? ?]]; simpl in *. - do 3 f_equiv; last apply H. - rewrite insert_ne; [| apply H | apply H]. - simpl. - f_equiv. - apply H. - * intros. - simpl. - destruct i as [? ?]. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - + constructor. - unshelve eexists (λne '(n,(s, s'')), let l := Loc.fresh (dom s) in - let s' := <[l:=n]>s in - Some (l, (s', s''))). - * apply _. - * apply _. - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]; simpl in *. - do 2 f_equiv. - -- f_equiv. - destruct H as [_ [H _]]; simpl in H. - apply gmap_dom_ne in H. - apply H. - -- f_equiv; last apply H. - rewrite insert_ne; [| apply H | apply H]. - simpl. - f_equiv. - destruct H as [_ [H _]]; simpl in H. - apply gmap_dom_ne in H. - by rewrite H. - * intros. - simpl. - destruct i as [? ?]. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - + constructor. - simpl. - unshelve eexists (λne '(l,(σ, σ')), Some ((), (delete l σ, σ'))). - * apply _. - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]; simpl in *. - do 2 f_equiv. - f_equiv; last apply H. - rewrite delete_ne; last apply H. - simpl. - f_equiv. - apply H. - * intros. - simpl. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - - intros x; inv_fin x. - + simpl. intros [[]| [[]| []]]. - * constructor. - unshelve eexists (λne '(_, (a, (b, c))), SomeO (_, (_, (_, c)))). - -- simpl in *. - apply ((input_lang.lang.update_input b).1). - -- apply a. - -- apply ((input_lang.lang.update_input b).2). - -- solve_proper_prepare. - destruct x as [? [? [? ?]]]; destruct y as [? [? [? ?]]]. - simpl in *. - do 2 f_equiv. - ++ do 2 f_equiv. - apply H. - ++ f_equiv; first apply H. - f_equiv; last apply H. - do 2 f_equiv; apply H. - -- intros. - simpl. - destruct σ as [? [? ?]]. - simpl. - reflexivity. - * constructor. - unshelve eexists (λne '(x, (y, z)), SomeO ((), _)). - -- simpl in *. - apply (y, ((input_lang.lang.update_output x (fstO z)), ())). - -- solve_proper_prepare. - destruct x as [? [? [? ?]]]; destruct y as [? [? [? ?]]]. - simpl in *. - do 2 f_equiv. - apply pair_ne. - ++ apply H. - ++ do 2 f_equiv; apply H. - -- intros. - simpl. - destruct σ as [σ1 [? []]]; simpl in *. - reflexivity. - + intros i; by apply fin_0_inv. -Qed. +Local Definition rs : gReifiers NotCtxDep 2 + := gReifiers_cons NotCtxDep reify_store + (gReifiers_cons NotCtxDep input_lang.interp.reify_io (gReifiers_nil NotCtxDep)). Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. -Lemma logrel2_adequacy cr R `{!Cofe R, !SubOfe locO R, !SubOfe natO R, !SubOfe unitO R} Σ `{!invGpreS Σ}`{!statePreG rs R Σ} `{!heapPreG rs R Σ} `{!na_invG Σ} - τ (α : unitO -n> IT (gReifiers_ops rs) R) (β : IT (gReifiers_ops rs) R) st st' k : +Require Import gitrees.gitree.greifiers. + +Lemma logrel2_adequacy (cr : nat) R `{!Cofe R, !SubOfe locO R, !SubOfe natO R, !SubOfe unitO R} + Σ `{!invGpreS Σ}`{!statePreG rs R Σ} `{!heapPreG rs R Σ} `{!na_invG Σ} + (τ : ty) (α : unitO -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ} p, (£ cr ⊢ valid2 rs p empC α τ)%I) → - ssteps (gReifiers_sReifier rs) (α ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + ssteps (gReifiers_sReifier NotCtxDep rs) (α ()) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (β ≡ Err OtherError) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. intros Hlog Hst. destruct (IT_to_V β) as [βv|] eqn:Hb. { right. right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } - cut ((∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ e, β ≡ Err e ∧ s e)). { intros [?|He]; first eauto. right. left. destruct He as [? [? ->]]. done. } - eapply (wp_safety (S cr) _ _ rs s); eauto. + eapply (wp_safety (S cr) _ _ NotCtxDep rs s); eauto. { apply Hdisj. } { by rewrite Hb. } intros H1 H2. @@ -625,9 +518,9 @@ Proof. iMod (new_heapG rs σ) as (H3) "H". iAssert (has_substate σ ∗ has_substate ios)%I with "[Hst]" as "[Hs Hio]". { unfold has_substate, has_full_state. - assert (of_state rs (IT (gReifiers_ops rs) _) st ≡ - of_idx rs (IT (gReifiers_ops rs) _) sR_idx (sR_state σ) - ⋅ of_idx rs (IT (gReifiers_ops rs) _) sR_idx (sR_state ios)) as ->; last first. + assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) st ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ) + ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state ios)) as ->; last first. { rewrite -own_op. done. } unfold sR_idx. simpl. intro j. @@ -656,10 +549,10 @@ Qed. Definition R := sumO locO (sumO natO unitO). -Lemma logrel2_safety e τ (β : IT (gReifiers_ops rs) R) st st' k : +Lemma logrel2_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : typed_glued empC e τ → - ssteps (gReifiers_sReifier rs) (interp_expr rs e ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ()) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (β ≡ Err OtherError) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. diff --git a/theories/examples/factorial.v b/theories/examples/factorial.v index 3ee54dd..3d2cc4e 100644 --- a/theories/examples/factorial.v +++ b/theories/examples/factorial.v @@ -4,18 +4,14 @@ From gitrees.input_lang Require Import lang interp. From gitrees.examples Require Import store while. Section fact. - Definition rs : gReifiers 2 := - gReifiers_cons reify_io (gReifiers_cons reify_store gReifiers_nil). - Notation F := (gReifiers_ops rs). + Definition rs : gReifiers NotCtxDep 2 := + gReifiers_cons NotCtxDep reify_io (gReifiers_cons NotCtxDep reify_store (gReifiers_nil NotCtxDep)). + Notation F := (gReifiers_ops NotCtxDep rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R, !SubOfe unitO R}. Notation IT := (IT F R). Notation ITV := (ITV F R). - Context {HCI : ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) R) o}. - Context `{!invGS Σ, !stateG rs R Σ, !heapG rs R Σ}. Notation iProp := (iProp Σ). @@ -40,7 +36,7 @@ Section fact. heap_ctx -∗ pointsto acc (Ret m) -∗ pointsto ℓ (Ret n) -∗ WP@{rs} fact_imp_body acc ℓ {{ _, pointsto acc (Ret (m * fact n)) }}. - Proof using HCI. + Proof. iIntros "#Hctx Hacc Hl". iLöb as "IH" forall (n m). unfold fact_imp_body. @@ -100,7 +96,7 @@ Section fact. Lemma wp_fact_imp (n : nat) : heap_ctx ⊢ WP@{rs} fact_imp ⊙ (Ret n) {{ βv, βv ≡ RetV (fact n) }}. - Proof using HCI. + Proof. iIntros "#Hctx". iApply wp_lam. iNext. simpl. rewrite get_ret_ret. @@ -114,7 +110,7 @@ Section fact. iIntros (ℓ) "Hl". simpl. iApply wp_seq. { solve_proper. } - iApply (wp_wand _ (λ _, pointsto acc (Ret $ fact n)) with "[-]"); last first. + iApply (wp_wand _ _ (λ _, pointsto acc (Ret $ fact n)) with "[-]"); last first. { simpl. iIntros (_) "Hacc". iModIntro. iApply (wp_read with "Hctx Hacc"). iNext. iNext. iIntros "Hacc". @@ -128,7 +124,7 @@ Section fact. Lemma wp_fact_io (n : nat) : heap_ctx ∗ has_substate (State [n] []) ⊢ WP@{rs} get_ret OUTPUT fact_io {{ _, has_substate (State [] [fact n]) }}. - Proof using HCI. + Proof. iIntros "[#Hctx Htape]". unfold fact_io. iApply (wp_bind _ (get_ret _)). diff --git a/theories/examples/iter.v b/theories/examples/iter.v index 18ef4f9..e19dc18 100644 --- a/theories/examples/iter.v +++ b/theories/examples/iter.v @@ -59,23 +59,20 @@ End iter. Section iter_wp. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers NotCtxDep sz). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ}. Notation iProp := (iProp Σ). - Context {HCI : ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) R) o}. Lemma wp_iter f (m : nat) β Ψ `{!AsVal f} `{!NonExpansive Ψ} : ⊢ WP@{rs} β {{ Ψ }} -∗ □ (∀ βv, Ψ βv -∗ WP@{rs} (f ⊙ (IT_of_V βv)) {{ Ψ }}) -∗ WP@{rs} (ITER ⊙ f ⊙ (Ret m) ⊙ β) {{ Ψ }}. - Proof using HCI. + Proof. iIntros "Hb #H". iApply (wp_bind _ (AppRSCtx (ITER ⊙ f ⊙ (Ret m)))). iApply (wp_wand with "Hb"). diff --git a/theories/examples/store.v b/theories/examples/store.v index 425446f..490ad68 100644 --- a/theories/examples/store.v +++ b/theories/examples/store.v @@ -17,47 +17,47 @@ Proof. apply _. Qed. #[local] Instance state_cofe X `{!Cofe X} : Cofe (stateF ♯ X). Proof. apply _. Qed. -Definition state_read X `{!Cofe X} : loc * (stateF ♯ X) * (laterO X -n> laterO X) → option (laterO X * (stateF ♯ X)) - := λ '(l,σ,κ), x ← σ !! l; - Some (κ x, σ). +Definition state_read X `{!Cofe X} : loc * (stateF ♯ X) → option (laterO X * (stateF ♯ X)) + := λ '(l,σ), x ← σ !! l; + Some (x, σ). #[export] Instance state_read_ne X `{!Cofe X} : - NonExpansive (state_read X : prodO (prodO locO (stateF ♯ X)) (laterO X -n> laterO X) → optionO (prodO (laterO X) (stateF ♯ X))). + NonExpansive (state_read X : prodO locO (stateF ♯ X) → optionO (prodO (laterO X) (stateF ♯ X))). Proof. - intros n [[l1 s1] κ1] [[l2 s2] κ2]. simpl. intros [[-> Hs'] Hs]. - apply (option_mbind_ne _ (λ n, Some (κ1 n, s1)) (λ n, Some (κ2 n, s2))); + intros n [l1 s1] [l2 s2]. simpl. intros [-> Hs]. + apply (option_mbind_ne _ (λ n, Some (n, s1)) (λ n, Some (n, s2))); solve_proper. Qed. -Definition state_dealloc X `{!Cofe X} : loc * (stateF ♯ X) * (unitO -n> laterO X) → option (laterO X * (stateF ♯ X)) - := λ '(l,σ,κ), Some (κ (), delete l σ). +Definition state_dealloc X `{!Cofe X} : loc * (stateF ♯ X) → option (unitO * (stateF ♯ X)) + := λ '(l,σ), Some ((), delete l σ). #[export] Instance state_dealloc_ne X `{!Cofe X} : - NonExpansive (state_dealloc X : prodO (prodO locO (stateF ♯ X)) (unitO -n> laterO X) → optionO (prodO (laterO X) (stateF ♯ X))). + NonExpansive (state_dealloc X : prodO locO (stateF ♯ X) → optionO (prodO unitO (stateF ♯ X))). Proof. - intros n [[l1 s1] κ1] [[l2 s2] κ2]. simpl. intros [[-> Hs'] Hs]. + intros n [l1 s1] [l2 s2]. simpl. intros [-> Hs]. solve_proper. Qed. Definition state_write X `{!Cofe X} : - (loc * (laterO X)) * (stateF ♯ X) * (unitO -n> laterO X) → option (laterO X * (stateF ♯ X)) - := λ '((l,n),s,κ), let s' := <[l:=n]>s - in Some (κ (), s'). + (loc * (laterO X)) * (stateF ♯ X) → option (unit * (stateF ♯ X)) + := λ '((l,n),s), let s' := <[l:=n]>s + in Some ((), s'). #[export] Instance state_write_ne X `{!Cofe X} : - NonExpansive (state_write X : prodO (prodO (prodO locO _) (stateF ♯ _)) (unitO -n> laterO X) → optionO (prodO (laterO X) (stateF ♯ X))). + NonExpansive (state_write X : prodO (prodO locO _) (stateF ♯ _) → optionO (prodO unitO (stateF ♯ X))). Proof. - intros n [[[l1 m1] s1] κ1] [[[l2 m2] s2] κ2]. simpl. - intros [[[Hl%leibnizO_leibniz Hm] Hs] Hκ]. simpl in Hl. + intros n [[l1 m1] s1] [[l2 m2] s2]. simpl. + intros [[Hl%leibnizO_leibniz Hm] Hs]. simpl in Hl. rewrite Hl. solve_proper. Qed. -Definition state_alloc X `{!Cofe X} : (laterO X) * (stateF ♯ X) * (loc -n> laterO X) → option ((laterO X) * (stateF ♯ X)) - := λ '(n,s,κ), let l := Loc.fresh (dom s) in +Definition state_alloc X `{!Cofe X} : (laterO X) * (stateF ♯ X) → option (loc * (stateF ♯ X)) + := λ '(n,s), let l := Loc.fresh (dom s) in let s' := <[l:=n]>s in - Some (κ l, s'). + Some (l, s'). #[export] Instance state_alloc_ne X `{!Cofe X} : - NonExpansive (state_alloc X : prodO (prodO _ (stateF ♯ X)) (locO -n> laterO X) → optionO (prodO (laterO X) (stateF ♯ X))). + NonExpansive (state_alloc X : prodO _ (stateF ♯ X) → optionO (prodO locO (stateF ♯ X))). Proof. - intros n [[m1 s1] κ1] [[m2 s2] κ2]. simpl. - intros [[Hm Hs] Hκ]. simpl in *. + intros n [m1 s1] [m2 s2]. simpl. + intros [Hm Hs]. simpl in *. set (l1 := Loc.fresh (dom s1)). set (l2 := Loc.fresh (dom s2)). assert (l1 = l2) as ->. @@ -83,7 +83,7 @@ Program Definition DeallocE : opInterp := {| |}. Definition storeE : opsInterp := @[ReadE;WriteE;AllocE;DeallocE]. -Canonical Structure reify_store : sReifier. +Canonical Structure reify_store : sReifier NotCtxDep. Proof. simple refine {| sReifier_ops := storeE |}. intros X HX op. @@ -129,11 +129,11 @@ End constructors. Section wp. Context {n : nat}. - Variable (rs : gReifiers n). + Variable (rs : gReifiers NotCtxDep n). Context {R} `{!Cofe R}. Context `{!SubOfe unitO R}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Notation stateO := (stateF ♯ IT). @@ -229,7 +229,7 @@ Section wp. match goal with | |- context G [Vis ?a ?b ?c] => assert (c ≡ idfun ◎ (subEff_outs ^-1)) as -> end; first solve_proper. - iApply wp_subreify'. + iApply wp_subreify_ctx_indep'. iInv (nroot.@"storeE") as (σ) "[>Hlc [Hs Hh]]" "Hcl". iApply (fupd_mask_weaken E1). { set_solver. } @@ -261,10 +261,7 @@ Section wp. simpl. iPureIntro. f_equiv; last done. - intros ???. - do 2 f_equiv. - by rewrite ofe_iso_21. - - done. + - iPureIntro. apply ofe_iso_21. - iNext. iIntros "Hlc Hs". iMod ("Hback" with "Hp") as "Hback". iMod "Hwk" . @@ -296,7 +293,7 @@ Section wp. Proof. iIntros (Hee) "#Hcxt H". unfold READ. simpl. - iApply wp_subreify'. + iApply wp_subreify_ctx_indep'. iInv (nroot.@"storeE") as (σ) "[>Hlc [Hs Hh]]" "Hcl". iApply (fupd_mask_weaken E1). { set_solver. } @@ -309,7 +306,7 @@ Section wp. destruct (Next_uninj x) as [α' Ha']. iApply (lc_fupd_elim_later with "Hlc"). iNext. - iExists σ,(Next (Ret ())),(<[l:=Next β]>σ),(Ret ()). + iExists σ,(),(<[l:=Next β]>σ),(Ret ()). iFrame "Hs". iSimpl. repeat iSplit; [ done | done | ]. iNext. iIntros "Hlc". @@ -342,22 +339,17 @@ Section wp. WP@{rs} ALLOC α k @ s {{ Φ }}. Proof. iIntros "Hh H". - iApply wp_subreify'. + iApply wp_subreify_ctx_indep'. iInv (nroot.@"storeE") as (σ) "[>Hlc [Hs Hh]]" "Hcl". iApply (lc_fupd_elim_later with "Hlc"). iModIntro. set (l:=Loc.fresh (dom σ)). - iExists σ,(Next (k l)),_,(k l). + iExists σ,l,_,(k l). iFrame "Hs". simpl. change (Loc.fresh (dom σ)) with l. - iSplit. - { - iPureIntro. - do 2 f_equiv; last reflexivity. - f_equiv. - by rewrite ofe_iso_21. - } iSplit; first done. + iSplit. + { simpl. rewrite ofe_iso_21. done. } iNext. iIntros "Hlc Hs". iMod (istate_alloc α l with "Hh") as "[Hh Hl]". { apply (not_elem_of_dom_1 (M:=gmap loc)). @@ -376,7 +368,7 @@ Section wp. Proof. iIntros (Hee) "#Hcxt H". unfold DEALLOC. simpl. - iApply wp_subreify'. + iApply wp_subreify_ctx_indep'. iInv (nroot.@"storeE") as (σ) "[>Hlc [Hs Hh]]" "Hcl". iApply (fupd_mask_weaken E1). { set_solver. } @@ -388,7 +380,7 @@ Section wp. { iApply (istate_loc_dom with "Hh Hp"). } destruct Hdom as [x Hx]. destruct (Next_uninj x) as [β' Hb']. - iExists σ,(Next (Ret ())),(delete l σ),(Ret ()). + iExists σ,(),(delete l σ),(Ret ()). iFrame "Hs". repeat iSplit; simpl; eauto. iNext. iIntros "Hlc Hs". diff --git a/theories/input_lang/interp.v b/theories/input_lang/interp.v index dd001c1..41bf8af 100644 --- a/theories/input_lang/interp.v +++ b/theories/input_lang/interp.v @@ -1,7 +1,9 @@ -From Equations Require Import Equations. -From gitrees Require Import gitree. +From gitrees Require Import gitree lang_generic. From gitrees.input_lang Require Import lang. +Require Import Binding.Lib. +Require Import Binding.Set. + Notation stateO := (leibnizO state). Program Definition inputE : opInterp := @@ -18,69 +20,29 @@ Program Definition outputE : opInterp := Definition ioE := @[inputE;outputE]. -Definition wrap_reifier X `{Cofe X} (A B : ofe) : - (A * stateO -n> option (B * stateO))%type -> - (A * stateO * (B -n> laterO X) → option (laterO X * stateO))%type := - λ f, - λ x, let '(i, σ, k) := x in - fmap (prodO_map k idfun) (f (i, σ)). -#[export] Instance wrap_reifier_ne X `{Cofe X} (A B : ofe) f : - NonExpansive (wrap_reifier X A B f). -Proof. - intros n [[a1 σ1] k1] [[a2 σ2] k2] [[Ha Hσ] Hk]. simpl. - solve_proper. -Qed. - (* INPUT *) -Definition reify_input' X `{Cofe X} : unitO * stateO → +Definition reify_input X `{Cofe X} : unitO * stateO → option (natO * stateO) := λ '(o, σ), Some (update_input σ : prodO natO stateO). -#[export] Instance reify_input'_ne X `{Cofe X} : - NonExpansive (reify_input' X). +#[export] Instance reify_input_ne X `{Cofe X} : + NonExpansive (reify_input X). Proof. intros ?[[]][[]][_?]. simpl in *. f_equiv. repeat f_equiv. done. Qed. -Definition reify_input X `{Cofe X} : unitO * stateO * (natO -n> laterO X) → - option (laterO X * stateO) := - λ '(o, σ, k), fmap (prodO_map k idfun) (reify_input' X (o, σ)). -#[export] Instance reify_input_ne X `{Cofe X} : - NonExpansive (reify_input X : prodO (prodO unitO stateO) - (natO -n> laterO X) → - optionO (prodO (laterO X) stateO)). -Proof. - intros n [[? σ1] k1] [[? σ2] k2]. simpl. - intros [[_ ->] Hk]. simpl in *. - repeat f_equiv. assumption. -Qed. - (* OUTPUT *) -Definition reify_output' X `{Cofe X} : (natO * stateO) → +Definition reify_output X `{Cofe X} : (natO * stateO) → option (unitO * stateO) := λ '(n, σ), Some((), update_output n σ : stateO). -#[export] Instance reify_output'_ne X `{Cofe X} : - NonExpansive (reify_output' X). +#[export] Instance reify_output_ne X `{Cofe X} : + NonExpansive (reify_output X). Proof. intros ?[][][]. simpl in *. repeat f_equiv; done. Qed. - -Definition reify_output X `{Cofe X} : (natO * stateO * (unitO -n> laterO X)) → - optionO (prodO (laterO X) stateO) := - λ '(n, σ, k), fmap (prodO_map k idfun) - (reify_output' X (n, σ)). -#[export] Instance reify_output_ne X `{Cofe X} : - NonExpansive (reify_output X : prodO (prodO natO stateO) - (unitO -n> laterO X) → - optionO (prodO (laterO X) stateO)). -Proof. - intros ? [[]] [[]] []; simpl in *. - repeat f_equiv; first assumption; apply H0. -Qed. - -Canonical Structure reify_io : sReifier. +Canonical Structure reify_io : sReifier NotCtxDep. Proof. simple refine {| sReifier_ops := ioE; sReifier_state := stateO @@ -98,15 +60,6 @@ Section constructors. Notation IT := (IT E A). Notation ITV := (ITV E A). - Global Instance ioEctx_indep : - ∀ (o : opid ioE), CtxIndep reify_io IT o. - Proof. - intros op. - destruct op as [[] | [ | []]]. - - constructor. by exists (OfeMor (reify_input' IT)). - - constructor. by exists (OfeMor (reify_output' IT)). - Qed. - Program Definition INPUT : (nat -n> IT) -n> IT := λne k, Vis (E:=E) (subEff_opid (inl ())) (subEff_ins (F:=ioE) (op:=(inl ())) ()) (NextO ◎ k ◎ (subEff_outs (F:=ioE) (op:=(inl ())))^-1). @@ -137,9 +90,9 @@ End constructors. Section weakestpre. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers NotCtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Notation IT := (IT F R). @@ -155,7 +108,7 @@ Section weakestpre. Proof. intros Hs. iIntros "Hs Ha". unfold INPUT. simpl. - iApply (wp_subreify with "Hs"). + iApply (wp_subreify_ctx_indep with "Hs"). { simpl. rewrite Hs//=. } { simpl. by rewrite ofe_iso_21. } iModIntro. done. @@ -169,7 +122,7 @@ Section weakestpre. Proof. intros Hs. iIntros "Hs Ha". unfold OUTPUT. simpl. - iApply (wp_subreify rs with "Hs"). + iApply (wp_subreify_ctx_indep rs with "Hs"). { simpl. by rewrite Hs. } { simpl. done. } iModIntro. iIntros "H1 H2". @@ -180,14 +133,20 @@ End weakestpre. Section interp. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers NotCtxDep sz). Context {subR : subReifier reify_io rs}. - Context {R} `{!Cofe R}. + Context {R} `{CR : !Cofe R}. Context `{!SubOfe natO R}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Notation IT := (IT F R). Notation ITV := (ITV F R). + Global Instance denot_cont_ne (κ : IT -n> IT) : + NonExpansive (λ x : IT, Tau (laterO_map κ (Next x))). + Proof. + solve_proper. + Qed. + (** Interpreting individual operators *) Program Definition interp_input {A} : A -n> IT := λne env, INPUT Ret. @@ -205,18 +164,49 @@ Section interp. Typeclasses Opaque interp_natop. Opaque laterO_map. - Program Definition interp_rec_pre {A} (body : prodO IT (prodO IT A) -n> IT) - : laterO (A -n> IT) -n> A -n> IT := - λne self env, Fun $ laterO_map (λne (self : A -n> IT) (a : IT), - body (self env,(a,env))) self. - Solve All Obligations with first [ solve_proper | solve_proper_please ]. + Program Definition interp_rec_pre {S : Set} (body : @interp_scope F R _ (inc (inc S)) -n> IT) + : laterO (@interp_scope F R _ S -n> IT) -n> @interp_scope F R _ S -n> IT := + λne self env, Fun $ laterO_map (λne (self : @interp_scope F R _ S -n> IT) (a : IT), + body (@extend_scope F R _ _ (@extend_scope F R _ _ env (self env)) a)) self. + Next Obligation. + intros. + solve_proper_prepare. + f_equiv; intros [| [| y']]; simpl; solve_proper. + Qed. + Next Obligation. + intros. + solve_proper_prepare. + f_equiv; intros [| [| y']]; simpl; solve_proper. + Qed. + Next Obligation. + intros. + solve_proper_prepare. + do 3 f_equiv; intros ??; simpl; f_equiv; + intros [| [| y']]; simpl; solve_proper. + Qed. + Next Obligation. + intros. + solve_proper_prepare. + by do 2 f_equiv. + Qed. - Definition interp_rec {A} body : A -n> IT := mmuu (interp_rec_pre body). - Program Definition ir_unf {A} (body : prodO IT (prodO IT A) -n> IT) env : IT -n> IT := - λne a, body (interp_rec body env, (a,env)). - Solve All Obligations with first [ solve_proper | solve_proper_please ]. + Program Definition interp_rec {S : Set} + (body : @interp_scope F R _ (inc (inc S)) -n> IT) : + @interp_scope F R _ S -n> IT := + mmuu (interp_rec_pre body). - Lemma interp_rec_unfold {A} (body : prodO IT (prodO IT A) -n> IT) env : + Program Definition ir_unf {S : Set} + (body : @interp_scope F R _ (inc (inc S)) -n> IT) env : IT -n> IT := + λne a, body (@extend_scope F R _ _ + (@extend_scope F R _ _ env (interp_rec body env)) + a). + Next Obligation. + intros. + solve_proper_prepare. + f_equiv. intros [| [| y']]; simpl; solve_proper. + Qed. + + Lemma interp_rec_unfold {S : Set} (body : @interp_scope F R _ (inc (inc S)) -n> IT) env : interp_rec body env ≡ Fun $ Next $ ir_unf body env. Proof. trans (interp_rec_pre body (Next (interp_rec body)) env). @@ -242,412 +232,458 @@ Section interp. Program Definition interp_nat (n : nat) {A} : A -n> IT := λne env, Ret n. - Program Definition interp_cont {A} (K : A -n> (IT -n> IT)) : A -n> IT := λne env, Fun (Next (K env)). - Next Obligation. - solve_proper. - Qed. - - Program Definition interp_applk {A} (q : A -n> IT) (K : A -n> (IT -n> IT)) : A -n> (IT -n> IT) := λne env t, interp_app q (λne env, K env t) env. - Next Obligation. - solve_proper. - Qed. - Next Obligation. - solve_proper. - Qed. - Next Obligation. - solve_proper. - Qed. - - Program Definition interp_apprk {A} (K : A -n> (IT -n> IT)) (q : A -n> IT) : A -n> (IT -n> IT) := λne env t, interp_app (λne env, K env t) q env. - Next Obligation. - solve_proper. - Qed. - Next Obligation. - solve_proper. - Qed. - Next Obligation. - solve_proper. - Qed. + Program Definition interp_applk {A} + (K : A -n> (IT -n> IT)) + (q : A -n> IT) + : A -n> (IT -n> IT) := + λne env t, interp_app (λne env, K env t) q env. + Solve All Obligations with solve_proper. + + Program Definition interp_apprk {A} + (q : A -n> IT) + (K : A -n> (IT -n> IT)) + : A -n> (IT -n> IT) := + λne env t, interp_app q (λne env, K env t) env. + Solve All Obligations with solve_proper. + + Program Definition interp_natoprk {A} (op : nat_op) + (q : A -n> IT) + (K : A -n> (IT -n> IT)) : A -n> (IT -n> IT) := + λne env t, interp_natop op q (λne env, K env t) env. + Solve All Obligations with solve_proper. + + Program Definition interp_natoplk {A} (op : nat_op) + (K : A -n> (IT -n> IT)) + (q : A -n> IT) : A -n> (IT -n> IT) := + λne env t, interp_natop op (λne env, K env t) q env. + Solve All Obligations with solve_proper. + + Program Definition interp_ifk {A} (K : A -n> (IT -n> IT)) (q : A -n> IT) + (p : A -n> IT) : A -n> (IT -n> IT) := + λne env t, interp_if (λne env, K env t) q p env. + Solve All Obligations with solve_proper. + + Program Definition interp_outputk {A} (K : A -n> (IT -n> IT)) : + A -n> (IT -n> IT) := + λne env t, interp_output (λne env, K env t) env. + Solve All Obligations with solve_proper. (** Interpretation for all the syntactic categories: values, expressions, contexts *) Fixpoint interp_val {S} (v : val S) : interp_scope S -n> IT := match v with - | Lit n => interp_nat n + | LitV n => interp_nat n | RecV e => interp_rec (interp_expr e) end with interp_expr {S} (e : expr S) : interp_scope S -n> IT := - match e with - | Val v => interp_val v - | Var x => interp_var x - | Rec e => interp_rec (interp_expr e) - | App e1 e2 => interp_app (interp_expr e1) (interp_expr e2) - | NatOp op e1 e2 => interp_natop op (interp_expr e1) (interp_expr e2) - | If e e1 e2 => interp_if (interp_expr e) (interp_expr e1) (interp_expr e2) - | Input => interp_input - | Output e => interp_output (interp_expr e) - end. - - Program Definition interp_ctx_item {S : scope} (K : ectx_item S) : interp_scope S -n> IT -n> IT := - match K with - | AppLCtx v2 => λne env t, interp_app (constO t) (interp_val v2) env - | AppRCtx e1 => λne env t, interp_app (interp_expr e1) (constO t) env - | NatOpLCtx op v2 => λne env t, interp_natop op (constO t) (interp_val v2) env - | NatOpRCtx op e1 => λne env t, interp_natop op (interp_expr e1) (constO t) env - | IfCtx e1 e2 => λne env t, interp_if (constO t) (interp_expr e1) (interp_expr e2) env - | OutputCtx => λne env t, interp_output (constO t) env - end. + match e with + | Val v => interp_val v + | Var x => interp_var x + | App e1 e2 => interp_app (interp_expr e1) (interp_expr e2) + | NatOp op e1 e2 => interp_natop op (interp_expr e1) (interp_expr e2) + | If e e1 e2 => interp_if (interp_expr e) (interp_expr e1) (interp_expr e2) + | Input => interp_input + | Output e => interp_output (interp_expr e) + end + with interp_ectx {S} (K : ectx S) : interp_scope S -n> (IT -n> IT) := + match K with + | EmptyK => λne env, idfun + | AppRK e1 K => interp_apprk (interp_expr e1) (interp_ectx K) + | AppLK K v2 => interp_applk (interp_ectx K) (interp_val v2) + | NatOpRK op e1 K => interp_natoprk op (interp_expr e1) (interp_ectx K) + | NatOpLK op K v2 => interp_natoplk op (interp_ectx K) (interp_val v2) + | IfK K e1 e2 => interp_ifk (interp_ectx K) (interp_expr e1) (interp_expr e2) + | OutputK K => interp_outputk (interp_ectx K) + end. Solve All Obligations with first [ solve_proper | solve_proper_please ]. - #[global] Instance interp_val_asval {S} (v : val S) D : AsVal (interp_val v D). - Proof. - destruct v; simpl; first apply _. - rewrite interp_rec_unfold. apply _. - Qed. - Program Fixpoint interp_ectx {S} (K : ectx S) : interp_scope S -n> IT -n> IT - := - match K with - | [] => λne env, idfun - | Ki::K => λne env, interp_ectx K env ◎ interp_ctx_item Ki env - end. - Next Obligation. solve_proper. Defined. (* XXX why can't i qed here? *) - - Lemma interp_ctx_item_fill {S} (Ki : ectx_item S) e env : - interp_expr (fill_item Ki e) env ≡ interp_ctx_item Ki env (interp_expr e env). - Proof. destruct Ki; reflexivity. Qed. - - Lemma interp_ectx_fill {S} (K : ectx S) e env : - interp_expr (fill K e) env ≡ interp_ectx K env (interp_expr e env). + Global Instance interp_val_asval {S} {D : interp_scope S} (v : val S) + : AsVal (interp_val v D). Proof. - revert e; induction K as [|Ki K]=>e; first done. - rewrite IHK. simpl. rewrite interp_ctx_item_fill. done. + destruct v; simpl. + - apply _. + - rewrite interp_rec_unfold. apply _. Qed. - (** Applying renamings and subsitutions to the interpretation of scopes *) - Equations interp_rens_scope {S S' : scope} - (E : interp_scope (E:=F) (R:=R) S') (s : rens S S') : interp_scope (E:=F) (R:=R) S := - interp_rens_scope (S:=[]) E s := tt : interp_scope []; - interp_rens_scope (S:=_::_) E s := - (interp_var (hd_ren s) E, interp_rens_scope E (tl_ren s)). + Global Instance ArrEquiv {A B : Set} : Equiv (A [→] B) := + fun f g => ∀ x, f x = g x. - Equations interp_subs_scope {S S' : scope} - (E : interp_scope (E:=F) (R:=R) S') (s : subs S S') : interp_scope (E:=F) (R:=R) S := - interp_subs_scope (S:=[]) E s := tt : interp_scope []; - interp_subs_scope (S:=_::_) E s := - (interp_expr (hd_sub s) E, interp_subs_scope E (tl_sub s)). + Global Instance ArrDist {A B : Set} `{Dist B} : Dist (A [→] B) := + fun n => fun f g => ∀ x, f x ≡{n}≡ g x. - - Global Instance interp_rens_scope_ne S S2 n : - Proper ((dist n) ==> (≡) ==> (dist n)) (@interp_rens_scope S S2). + Global Instance ren_scope_proper {S S'} : + Proper ((≡) ==> (≡) ==> (≡)) (@ren_scope F _ CR S S'). Proof. intros D D' HE s1 s2 Hs. - induction S as [|τ' S]; simp interp_rens_scope; auto. + intros x; simpl. f_equiv. - - unfold hd_ren; rewrite Hs. by f_equiv. - - apply IHS. intros v. unfold tl_ren; by rewrite Hs. + - apply Hs. + - apply HE. + Qed. + + Lemma interp_expr_ren {S S'} env + (δ : S [→] S') (e : expr S) : + interp_expr (fmap δ e) env ≡ interp_expr e (ren_scope δ env) + with interp_val_ren {S S'} env + (δ : S [→] S') (e : val S) : + interp_val (fmap δ e) env ≡ interp_val e (ren_scope δ env) + with interp_ectx_ren {S S'} env + (δ : S [→] S') (e : ectx S) : + interp_ectx (fmap δ e) env ≡ interp_ectx e (ren_scope δ env). + Proof. + - destruct e; simpl. + + by apply interp_val_ren. + + reflexivity. + + repeat f_equiv; by apply interp_expr_ren. + + repeat f_equiv; by apply interp_expr_ren. + + repeat f_equiv; by apply interp_expr_ren. + + repeat f_equiv; by apply interp_expr_ren. + + repeat f_equiv; by apply interp_expr_ren. + - destruct e; simpl. + + reflexivity. + + clear -interp_expr_ren. + apply bi.siProp.internal_eq_soundness. + iLöb as "IH". + rewrite {2}interp_rec_unfold. + rewrite {2}(interp_rec_unfold (interp_expr e)). + do 1 iApply f_equivI. iNext. + iApply internal_eq_pointwise. + rewrite /ir_unf. iIntros (x). simpl. + rewrite interp_expr_ren. + iApply f_equivI. + iApply internal_eq_pointwise. + iIntros (y'). + destruct y' as [| [| y]]; simpl; first done. + * by iRewrite - "IH". + * done. + - destruct e; simpl; intros ?; simpl. + + reflexivity. + + repeat f_equiv; by apply interp_ectx_ren. + + repeat f_equiv; [by apply interp_ectx_ren | by apply interp_expr_ren | by apply interp_expr_ren]. + + repeat f_equiv; [by apply interp_ectx_ren | by apply interp_val_ren]. + + repeat f_equiv; [by apply interp_expr_ren | by apply interp_ectx_ren]. + + repeat f_equiv; [by apply interp_expr_ren | by apply interp_ectx_ren]. + + repeat f_equiv; [by apply interp_ectx_ren | by apply interp_val_ren]. Qed. - Global Instance interp_subs_scope_ne S S2 n : - Proper ((dist n) ==> (≡) ==> (dist n)) (@interp_subs_scope S S2). + + Lemma interp_comp {S} (e : expr S) (env : interp_scope S) (K : ectx S): + interp_expr (fill K e) env ≡ (interp_ectx K) env ((interp_expr e) env). Proof. - intros D D' HE s1 s2 Hs. - induction S as [|τ' S]; simp interp_subs_scope; auto. - f_equiv. - - unfold hd_sub; by rewrite Hs HE. - - apply IHS. intros v. unfold tl_sub; by rewrite Hs. + revert env. + induction K; simpl; intros env; first reflexivity; try (by rewrite IHK). + - repeat f_equiv. + by rewrite IHK. + - repeat f_equiv. + by rewrite IHK. + - repeat f_equiv. + by rewrite IHK. Qed. - Global Instance interp_rens_scope_proper S S2 : - Proper ((≡) ==> (≡) ==> (≡)) (@interp_rens_scope S S2). + + Program Definition sub_scope {S S'} (δ : S [⇒] S') (env : interp_scope S') + : interp_scope S := λne x, interp_expr (δ x) env. + + Global Instance SubEquiv {A B : Set} : Equiv (A [⇒] B) := fun f g => ∀ x, f x = g x. + + Global Instance sub_scope_proper {S S'} : + Proper ((≡) ==> (≡) ==> (≡)) (@sub_scope S S'). Proof. intros D D' HE s1 s2 Hs. - induction S as [|τ' S]; simp interp_rens_scope; auto. + intros x; simpl. f_equiv. - - unfold hd_ren; rewrite Hs. - by rewrite HE. - - apply IHS. intros v. unfold tl_ren; by rewrite Hs. + - f_equiv. + apply HE. + - apply Hs. + Qed. + + Lemma interp_expr_subst {S S'} (env : interp_scope S') + (δ : S [⇒] S') e : + interp_expr (bind δ e) env ≡ interp_expr e (sub_scope δ env) + with interp_val_subst {S S'} (env : interp_scope S') + (δ : S [⇒] S') e : + interp_val (bind δ e) env ≡ interp_val e (sub_scope δ env) + with interp_ectx_subst {S S'} (env : interp_scope S') + (δ : S [⇒] S') e : + interp_ectx (bind δ e) env ≡ interp_ectx e (sub_scope δ env). + Proof. + - destruct e; simpl. + + by apply interp_val_subst. + + term_simpl. + reflexivity. + + repeat f_equiv; by apply interp_expr_subst. + + repeat f_equiv; by apply interp_expr_subst. + + repeat f_equiv; by apply interp_expr_subst. + + f_equiv. + + repeat f_equiv; by apply interp_expr_subst. + - destruct e; simpl. + + reflexivity. + + clear -interp_expr_subst. + apply bi.siProp.internal_eq_soundness. + iLöb as "IH". + rewrite {2}interp_rec_unfold. + rewrite {2}(interp_rec_unfold (interp_expr e)). + do 1 iApply f_equivI. iNext. + iApply internal_eq_pointwise. + rewrite /ir_unf. iIntros (x). simpl. + rewrite interp_expr_subst. + iApply f_equivI. + iApply internal_eq_pointwise. + iIntros (y'). + destruct y' as [| [| y]]; simpl; first done. + * by iRewrite - "IH". + * do 2 rewrite interp_expr_ren. + iApply f_equivI. + iApply internal_eq_pointwise. + iIntros (z). + done. + - destruct e; simpl; intros ?; simpl. + + reflexivity. + + repeat f_equiv; by apply interp_ectx_subst. + + repeat f_equiv; [by apply interp_ectx_subst | by apply interp_expr_subst | by apply interp_expr_subst]. + + repeat f_equiv; [by apply interp_ectx_subst | by apply interp_val_subst]. + + repeat f_equiv; [by apply interp_expr_subst | by apply interp_ectx_subst]. + + repeat f_equiv; [by apply interp_expr_subst | by apply interp_ectx_subst]. + + repeat f_equiv; [by apply interp_ectx_subst | by apply interp_val_subst]. Qed. - Global Instance interp_subs_scope_proper S S2 : - Proper ((≡) ==> (≡) ==> (≡)) (@interp_subs_scope S S2). + + (** ** Interpretation is a homomorphism (for some constructors) *) + + #[global] Instance interp_ectx_hom_emp {S} env : + IT_hom (interp_ectx (EmptyK : ectx S) env). Proof. - intros D D' HE s1 s2 Hs. - induction S as [|τ' S]; simp interp_subs_scope; auto. - f_equiv. - - unfold hd_sub; by rewrite Hs HE. - - apply IHS. intros v. unfold tl_sub; by rewrite Hs. + simple refine (IT_HOM _ _ _ _ _); intros; auto. + simpl. fold (@idfun IT). f_equiv. intro. simpl. + by rewrite laterO_map_id. Qed. - (** ** The substituion lemma, for renamings and substitutions *) - Lemma interp_rens_scope_tl_ren {S S2} x D (r : rens S S2) : - interp_rens_scope ((x, D) : interp_scope (()::S2)) (tl_ren (rens_lift r)) - ≡ interp_rens_scope D r. + #[global] Instance interp_ectx_hom_output {S} (K : ectx S) env : + IT_hom (interp_ectx K env) -> + IT_hom (interp_ectx (OutputK K) env). Proof. - induction S as [|τ' S]; simp interp_rens_scope; eauto. - f_equiv. - { unfold hd_ren, tl_ren. simp rens_lift interp_var. - done. } - { rewrite -IHS. f_equiv. clear. - intros v. dependent elimination v; - unfold hd_ren, tl_ren; simp rens_lift; auto. } + intros. simple refine (IT_HOM _ _ _ _ _); intros; simpl. + - by rewrite !hom_tick. + - rewrite !hom_vis. + f_equiv. intro. simpl. rewrite -laterO_map_compose. + do 2 f_equiv. by intro. + - by rewrite !hom_err. Qed. - Lemma interp_rens_scope_idren {S} (D : interp_scope S) : - interp_rens_scope D (@idren S) ≡ D. + #[global] Instance interp_ectx_hom_if {S} + (K : ectx S) (e1 e2 : expr S) env : + IT_hom (interp_ectx K env) -> + IT_hom (interp_ectx (IfK K e1 e2) env). Proof. - induction S as [|[] S]; simp interp_rens_scope. - { by destruct D. } - destruct D as [x D]. simp interp_var. simpl. - f_equiv. - trans (interp_rens_scope ((x, D) : interp_scope (()::S)) (tl_ren (rens_lift idren))). - { f_equiv. intros v. unfold tl_ren. - reflexivity. } - rewrite interp_rens_scope_tl_ren. - apply IHS. + intros. simple refine (IT_HOM _ _ _ _ _); intros; simpl. + - rewrite -IF_Tick. do 3 f_equiv. apply hom_tick. + - assert ((interp_ectx K env (Vis op i ko)) ≡ + (Vis op i (laterO_map (λne y, interp_ectx K env y) ◎ ko))). + { by rewrite hom_vis. } + trans (IF (Vis op i (laterO_map (λne y : IT, interp_ectx K env y) ◎ ko)) + (interp_expr e1 env) (interp_expr e2 env)). + { do 3 f_equiv. by rewrite hom_vis. } + rewrite IF_Vis. f_equiv. simpl. + intro. simpl. by rewrite -laterO_map_compose. + - trans (IF (Err e) (interp_expr e1 env) (interp_expr e2 env)). + { repeat f_equiv. apply hom_err. } + apply IF_Err. Qed. - Lemma interp_expr_ren {S D : scope} (M : expr S) (r : rens S D) : - ∀ (E : interp_scope D), - interp_expr (ren_expr M r) E ≡ interp_expr M (interp_rens_scope E r) - with interp_val_ren {S D : scope} (v : val S) (r : rens S D) : - ∀ (E : interp_scope D), - interp_val (ren_val v r) E ≡ interp_val v (interp_rens_scope E r). + #[global] Instance interp_ectx_hom_appr {S} (K : ectx S) + (e : expr S) env : + IT_hom (interp_ectx K env) -> + IT_hom (interp_ectx (AppRK e K) env). Proof. - - revert D r. induction M=> D r D2; simpl; simp ren_expr. - all: try by (simpl; repeat intro; simpl; repeat f_equiv; eauto). - + (* variable *) revert r. - induction v=>r. - * simp interp_var interp_rens_scope. done. - * simp interp_var interp_rens_scope. simpl. - apply (IHv (tl_ren r)). - + (* recursive functions *) simp ren_expr. simpl. - apply bi.siProp.internal_eq_soundness. - iLöb as "IH". - rewrite {2}interp_rec_unfold. - rewrite {2}(interp_rec_unfold (interp_expr M)). - iApply f_equivI. iNext. iApply internal_eq_pointwise. - rewrite /ir_unf. iIntros (x). simpl. - rewrite interp_expr_ren. - iApply f_equivI. - simp interp_rens_scope interp_var. simpl. - rewrite !interp_rens_scope_tl_ren. - iRewrite "IH". - done. - - revert D r. induction v=> D r D2; simpl; simp ren_val; eauto. - (* recursive functions *) - simp ren_expr. simpl. - apply bi.siProp.internal_eq_soundness. - iLöb as "IH". - rewrite {2}interp_rec_unfold. - rewrite {2}(interp_rec_unfold (interp_expr e)). - iApply f_equivI. iNext. iApply internal_eq_pointwise. - rewrite /ir_unf. iIntros (x). simpl. - rewrite interp_expr_ren. - iApply f_equivI. - simp interp_rens_scope interp_var. simpl. - rewrite !interp_rens_scope_tl_ren. - iRewrite "IH". - done. + intros. simple refine (IT_HOM _ _ _ _ _); intros; simpl. + - by rewrite !hom_tick. + - rewrite !hom_vis. f_equiv. intro x. simpl. + by rewrite laterO_map_compose. + - by rewrite !hom_err. Qed. - Lemma interp_subs_scope_tl_sub {S S2} x D (s : subs S S2) : - interp_subs_scope ((x, D) : interp_scope (()::S2)) (tl_sub (subs_lift s)) - ≡ interp_subs_scope D s. + #[global] Instance interp_ectx_hom_appl {S} (K : ectx S) + (v : val S) (env : interp_scope S) : + IT_hom (interp_ectx K env) -> + IT_hom (interp_ectx (AppLK K v) env). Proof. - induction S as [|[] S]; simp interp_subs_scope; first done. - f_equiv. - { unfold hd_sub, tl_sub. simp subs_lift interp_var. - unfold expr_lift. rewrite interp_expr_ren. f_equiv. - trans (interp_rens_scope ((x, D) : interp_scope (()::S2)) (tl_ren (rens_lift idren))). - { f_equiv. intros v. unfold tl_ren. - simp rens_lift idren. done. } - rewrite interp_rens_scope_tl_ren. - apply interp_rens_scope_idren. } - { rewrite -IHS. f_equiv. clear. - intros v. dependent elimination v; - unfold hd_sub, tl_sub; simp subs_lift; auto. } + intros H. simple refine (IT_HOM _ _ _ _ _); intros; simpl. + - rewrite -APP'_Tick_l. do 2 f_equiv. apply hom_tick. + - trans (APP' (Vis op i (laterO_map (interp_ectx K env) ◎ ko)) + (interp_val v env)). + + do 2f_equiv. rewrite hom_vis. do 3 f_equiv. by intro. + + rewrite APP'_Vis_l. f_equiv. intro x. simpl. + by rewrite -laterO_map_compose. + - trans (APP' (Err e) (interp_val v env)). + { do 2f_equiv. apply hom_err. } + apply APP'_Err_l, interp_val_asval. Qed. - Lemma interp_subs_scope_idsub {S} (env : interp_scope S) : - interp_subs_scope env idsub ≡ env. + #[global] Instance interp_ectx_hom_natopr {S} (K : ectx S) + (e : expr S) op env : + IT_hom (interp_ectx K env) -> + IT_hom (interp_ectx (NatOpRK op e K) env). Proof. - induction S as [|[] S]; simp interp_subs_scope. - { by destruct env. } - destruct env as [x env]. - unfold hd_sub, idsub. simpl. - simp interp_var. simpl. f_equiv. - etrans; last first. - { apply IHS. } - rewrite -(interp_subs_scope_tl_sub x env idsub). - repeat f_equiv. intro v. unfold tl_sub, idsub; simpl. - simp subs_lift. unfold expr_lift. simp ren_expr. done. + intros H. simple refine (IT_HOM _ _ _ _ _); intros; simpl. + - by rewrite !hom_tick. + - rewrite !hom_vis. f_equiv. intro x. simpl. + by rewrite laterO_map_compose. + - by rewrite !hom_err. Qed. - Lemma interp_expr_subst {S D : scope} (M : expr S) (s : subs S D) : - ∀ (E : interp_scope D), - interp_expr (subst_expr M s) E ≡ interp_expr M (interp_subs_scope E s) - with interp_val_subst {S D : scope} (v : val S) (s : subs S D) : - ∀ (E : interp_scope D), - interp_val (subst_val v s) E ≡ interp_val v (interp_subs_scope E s). + #[global] Instance interp_ectx_hom_natopl {S} (K : ectx S) + (v : val S) op (env : interp_scope S) : + IT_hom (interp_ectx K env) -> + IT_hom (interp_ectx (NatOpLK op K v) env). Proof. - - revert D s. induction M=> D r D2; simpl; simp subst_expr. - all: try by (simpl; repeat intro; simpl; repeat f_equiv; eauto). - + (* variable *) revert r. - induction v=>r. - * simp interp_var interp_rens_scope. done. - * simp interp_var interp_rens_scope. simpl. - apply (IHv (tl_sub r)). - + (* recursive functions *) simpl. - apply bi.siProp.internal_eq_soundness. - iLöb as "IH". - rewrite {2}interp_rec_unfold. - rewrite {2}(interp_rec_unfold (interp_expr M)). - iApply f_equivI. iNext. iApply internal_eq_pointwise. - rewrite /ir_unf. iIntros (x). simpl. - rewrite interp_expr_subst. - iApply f_equivI. - simp interp_subs_scope interp_var. simpl. - rewrite !interp_subs_scope_tl_sub. - iRewrite "IH". - done. - - revert D s. induction v=> D r D2; simpl; simp subst_val; eauto. - (* recursive functions *) - simp subst_expr. simpl. - apply bi.siProp.internal_eq_soundness. - iLöb as "IH". - rewrite {2}interp_rec_unfold. - rewrite {2}(interp_rec_unfold (interp_expr e)). - iApply f_equivI. iNext. iApply internal_eq_pointwise. - rewrite /ir_unf. iIntros (x). simpl. - rewrite interp_expr_subst. - iApply f_equivI. - simp interp_subs_scope interp_var. simpl. - rewrite !interp_subs_scope_tl_sub. - iRewrite "IH". - done. + intros H. simple refine (IT_HOM _ _ _ _ _); intros; simpl. + - rewrite -NATOP_ITV_Tick_l. do 2 f_equiv. apply hom_tick. + - trans (NATOP (do_natop op) + (Vis op0 i (laterO_map (interp_ectx K env) ◎ ko)) + (interp_val v env)). + { do 2 f_equiv. rewrite hom_vis. f_equiv. by intro. } + rewrite NATOP_ITV_Vis_l. f_equiv. intro x. simpl. + by rewrite -laterO_map_compose. + - trans (NATOP (do_natop op) (Err e) (interp_val v env)). + + do 2 f_equiv. apply hom_err. + + by apply NATOP_Err_l, interp_val_asval. + Qed. + + Lemma get_fun_ret' E A `{Cofe A} n : (∀ f, @get_fun E A _ f (core.Ret n) ≡ Err RuntimeErr). + Proof. + intros. + by rewrite IT_rec1_ret. Qed. - (** ** Interpretation is a homomorphism *) - #[global] Instance interp_ectx_item_hom {S} (Ki : ectx_item S) env : - IT_hom (interp_ctx_item Ki env). - Proof. destruct Ki; simpl; apply _. Qed. - #[global] Instance interp_ectx_hom {S} (K : ectx S) env : + #[global] Instance interp_ectx_hom {S} + (K : ectx S) env : IT_hom (interp_ectx K env). - Proof. induction K; simpl; apply _. Qed. + Proof. + induction K; apply _. + Qed. (** ** Finally, preservation of reductions *) - Lemma interp_expr_head_step {S} env (e : expr S) e' σ σ' n : - head_step e σ e' σ' (n,0) → + Lemma interp_expr_head_step {S : Set} (env : interp_scope S) (e : expr S) e' σ σ' n : + head_step e σ e' σ' (n, 0) → interp_expr e env ≡ Tick_n n $ interp_expr e' env. Proof. inversion 1; cbn-[IF APP' INPUT Tick get_ret2]. - - (*fun->val*) - reflexivity. - (* app lemma *) - rewrite APP_APP'_ITV. + subst. + erewrite APP_APP'_ITV; last apply _. trans (APP (Fun (Next (ir_unf (interp_expr e1) env))) (Next $ interp_val v2 env)). { repeat f_equiv. apply interp_rec_unfold. } rewrite APP_Fun. simpl. rewrite Tick_eq. do 2 f_equiv. simplify_eq. - rewrite interp_expr_subst. f_equiv. - simp interp_subs_scope. unfold hd_sub, tl_sub. simp conssub. - simpl. repeat f_equiv. - generalize (Val (RecV e1)). - generalize (Val v2). - clear. - intros e1 e2. - trans (interp_subs_scope env idsub); last first. - { f_equiv. intro v. simp conssub. done. } - symmetry. - apply interp_subs_scope_idsub. + rewrite !interp_expr_subst. + f_equiv. + intros [| [| x]]; simpl; [| reflexivity | reflexivity]. + rewrite interp_val_ren. + f_equiv. + intros ?; simpl; reflexivity. - (* the natop stuff *) simplify_eq. destruct v1,v2; try naive_solver. simpl in *. rewrite NATOP_Ret. destruct op; simplify_eq/=; done. - - by rewrite IF_True. - - rewrite IF_False; eauto. lia. + - rewrite IF_True; last lia. + reflexivity. + - rewrite IF_False; last lia. + reflexivity. Qed. - Lemma interp_expr_fill_no_reify {S} K env (e e' : expr S) σ σ' n : - head_step e σ e' σ' (n,0) → - interp_expr (fill K e) env ≡ Tick_n n $ interp_expr (fill K e') env. + Lemma interp_expr_fill_no_reify {S} K (env : interp_scope S) (e e' : expr S) σ σ' n : + head_step e σ e' σ' (n, 0) → + interp_expr (fill K e) env + ≡ + Tick_n n $ interp_expr (fill K e') env. Proof. intros He. - trans (interp_ectx K env (interp_expr e env)). - { apply interp_ectx_fill. } - trans (interp_ectx K env (Tick_n n (interp_expr e' env))). - { f_equiv. apply (interp_expr_head_step env) in He. apply He. } - trans (Tick_n n $ interp_ectx K env (interp_expr e' env)); last first. - { f_equiv. symmetry. apply interp_ectx_fill. } - apply hom_tick_n. apply _. + rewrite !interp_comp. + erewrite <-hom_tick_n. + - apply (interp_expr_head_step env) in He. + rewrite He. + reflexivity. + - apply _. Qed. Opaque INPUT OUTPUT_. + Opaque extend_scope. Opaque Ret. Lemma interp_expr_fill_yes_reify {S} K env (e e' : expr S) - (σ σ' : stateO) (σr : gState_rest sR_idx rs ♯ IT) n : - head_step e σ e' σ' (n,1) → - reify (gReifiers_sReifier rs) - (interp_expr (fill K e) env) (gState_recomp σr (sR_state σ)) - ≡ (gState_recomp σr (sR_state σ'), Tick_n n $ interp_expr (fill K e') env). + (σ σ' : stateO) (σr : gState_rest NotCtxDep sR_idx rs ♯ IT) n : + head_step e σ e' σ' (n, 1) → + reify (gReifiers_sReifier NotCtxDep rs) + (interp_expr (fill K e) env) (gState_recomp NotCtxDep σr (sR_state σ)) + ≡ (gState_recomp NotCtxDep σr (sR_state σ'), Tick_n n $ interp_expr (fill K e') env). Proof. intros Hst. - trans (reify (gReifiers_sReifier rs) (interp_ectx K env (interp_expr e env)) - (gState_recomp σr (sR_state σ))). - { f_equiv. by rewrite interp_ectx_fill. } + trans (reify (gReifiers_sReifier NotCtxDep rs) (interp_ectx K env (interp_expr e env)) + (gState_recomp NotCtxDep σr (sR_state σ))). + { f_equiv. by rewrite interp_comp. } inversion Hst; simplify_eq; cbn-[gState_recomp]. - - trans (reify (gReifiers_sReifier rs) (INPUT (interp_ectx K env ◎ Ret)) (gState_recomp σr (sR_state σ))). - { repeat f_equiv; eauto. - rewrite hom_INPUT. f_equiv. by intro. } - rewrite reify_vis_eq //; last first. - { rewrite subReifier_reify//= H4//=. } - repeat f_equiv. - rewrite Tick_eq/=. repeat f_equiv. - rewrite interp_ectx_fill. - simpl. - done. - - trans (reify (gReifiers_sReifier rs) (interp_ectx K env (OUTPUT n0)) (gState_recomp σr (sR_state σ))). - { do 3 f_equiv; eauto. - rewrite get_ret_ret//. } - trans (reify (gReifiers_sReifier rs) (OUTPUT_ n0 (interp_ectx K env (Ret 0))) (gState_recomp σr (sR_state σ))). - { do 2 f_equiv; eauto. - rewrite hom_OUTPUT_//. } - rewrite reify_vis_eq //; last first. + - trans (reify (gReifiers_sReifier NotCtxDep rs) (INPUT (interp_ectx K env ◎ Ret)) (gState_recomp NotCtxDep σr (sR_state σ))). { - simpl. - pose proof (@subReifier_reify sz reify_io rs subR IT _ ((inr (inl ()))) n0) as H. + repeat f_equiv; eauto. + rewrite hom_INPUT. + do 2 f_equiv. by intro. + } + rewrite reify_vis_eq_ctx_indep //; first last. + { + epose proof (@subReifier_reify sz NotCtxDep reify_io rs _ IT _ (inl ()) () _ σ σ' σr) as H. simpl in H. - specialize (H (Next (interp_ectx K env (Ret 0))) (λne _, Next (interp_ectx K env (Ret 0))) σ (update_output n0 σ) σr). + simpl. + erewrite <-H; last first. + - rewrite H4. reflexivity. + - f_equiv; + solve_proper. + } + repeat f_equiv. rewrite Tick_eq/=. repeat f_equiv. + rewrite interp_comp. + rewrite ofe_iso_21. + simpl. + reflexivity. + - trans (reify (gReifiers_sReifier NotCtxDep rs) (interp_ectx K env (OUTPUT n0)) (gState_recomp NotCtxDep σr (sR_state σ))). + { + do 3 f_equiv; eauto. + rewrite get_ret_ret//. + } + trans (reify (gReifiers_sReifier NotCtxDep rs) (OUTPUT_ n0 (interp_ectx K env (Ret 0))) (gState_recomp NotCtxDep σr (sR_state σ))). + { + do 2 f_equiv; eauto. + by rewrite hom_OUTPUT_. + } + rewrite reify_vis_eq_ctx_indep //; last first. + { + epose proof (@subReifier_reify sz NotCtxDep reify_io rs _ IT _ (inr (inl ())) n0 _ σ (update_output n0 σ) σr) as H. simpl in H. - rewrite <-H; last done. + simpl. + erewrite <-H; last reflexivity. f_equiv. - - intros [? ?] [? ?] [? ?]; simpl in *. - solve_proper. - - do 2 f_equiv. - intros ?; simpl. - reflexivity. + intros ???. by rewrite /prod_map H0. } repeat f_equiv. rewrite Tick_eq/=. repeat f_equiv. - rewrite interp_ectx_fill. - simpl. done. + rewrite interp_comp. + reflexivity. Qed. - Lemma soundness {S} (e1 e2 : expr S) σ1 σ2 (σr : gState_rest sR_idx rs ♯ IT) n m env : + Lemma soundness {S} (e1 e2 : expr S) σ1 σ2 (σr : gState_rest NotCtxDep sR_idx rs ♯ IT) n m (env : interp_scope S) : prim_step e1 σ1 e2 σ2 (n,m) → - ssteps (gReifiers_sReifier rs) - (interp_expr e1 env) (gState_recomp σr (sR_state σ1)) - (interp_expr e2 env) (gState_recomp σr (sR_state σ2)) n. + ssteps (gReifiers_sReifier NotCtxDep rs) + (interp_expr e1 env) (gState_recomp NotCtxDep σr (sR_state σ1)) + (interp_expr e2 env) (gState_recomp NotCtxDep σr (sR_state σ2)) n. Proof. Opaque gState_decomp gState_recomp. inversion 1; simplify_eq/=. destruct (head_step_io_01 _ _ _ _ _ _ H2); subst. - assert (σ1 = σ2) as ->. { eapply head_step_no_io; eauto. } - eapply (interp_expr_fill_no_reify K) in H2. - rewrite H2. eapply ssteps_tick_n. + unshelve eapply (interp_expr_fill_no_reify K) in H2; first apply env. + rewrite H2. + rewrite interp_comp. + eapply ssteps_tick_n. - inversion H2;subst. + eapply (interp_expr_fill_yes_reify K env _ _ _ _ σr) in H2. - rewrite interp_ectx_fill. + rewrite interp_comp. rewrite hom_INPUT. - change 1 with (1+0). econstructor; last first. + change 1 with (Nat.add 1 0). econstructor; last first. { apply ssteps_zero; reflexivity. } eapply sstep_reify. { Transparent INPUT. unfold INPUT. simpl. @@ -655,13 +691,13 @@ Section interp. simpl in H2. rewrite -H2. repeat f_equiv; eauto. - rewrite interp_ectx_fill hom_INPUT. + rewrite interp_comp hom_INPUT. eauto. + eapply (interp_expr_fill_yes_reify K env _ _ _ _ σr) in H2. - rewrite interp_ectx_fill. simpl. + rewrite interp_comp. simpl. rewrite get_ret_ret. rewrite hom_OUTPUT_. - change 1 with (1+0). econstructor; last first. + change 1 with (Nat.add 1 0). econstructor; last first. { apply ssteps_zero; reflexivity. } eapply sstep_reify. { Transparent OUTPUT_. unfold OUTPUT_. simpl. @@ -670,7 +706,7 @@ Section interp. rewrite -H2. repeat f_equiv; eauto. Opaque OUTPUT_. - rewrite interp_ectx_fill /= get_ret_ret hom_OUTPUT_. + rewrite interp_comp /= get_ret_ret hom_OUTPUT_. eauto. Qed. diff --git a/theories/input_lang/lang.v b/theories/input_lang/lang.v index 82fac6c..bd6939f 100644 --- a/theories/input_lang/lang.v +++ b/theories/input_lang/lang.v @@ -1,281 +1,284 @@ -From stdpp Require Export strings. -From gitrees Require Export prelude lang_generic. -From Equations Require Import Equations. +From gitrees Require Export prelude. Require Import List. Import ListNotations. -Delimit Scope expr_scope with E. +Require Import Binding.Resolver Binding.Lib Binding.Set Binding.Auto Binding.Env. Inductive nat_op := Add | Sub | Mult. -Inductive expr : scope → Type := +Inductive expr {X : Set} : Type := (* Values *) - | Val : forall {S}, val S → expr S - (* Base lambda calculus *) - | Var : forall {S}, var S → expr S - | Rec : forall {S}, expr (()::()::S) → expr S - | App : forall {S}, expr S → expr S → expr S - (* Base types and their operations *) - | NatOp : forall {S}, - nat_op → expr S → expr S → expr S - | If : forall {S}, - expr S → expr S → expr S → expr S - (* The effects *) - | Input : forall {S}, expr S - | Output : forall {S}, expr S → expr S -with val : scope → Type := - | Lit : forall {S}, nat → val S - | RecV : forall {S}, expr (()::()::S) → val S. - -Bind Scope expr_scope with expr. -Notation of_val := Val (only parsing). - -Definition to_val {S} (e : expr S) : option (val S) := +| Val (v : val) : expr +| Var (x : X) : expr +(* Base lambda calculus *) +| App (e₁ : expr) (e₂ : expr) : expr +(* Base types and their operations *) +| NatOp (op : nat_op) (e₁ : expr) (e₂ : expr) : expr +| If (e₁ : expr) (e₂ : expr) (e₃ : expr) : expr +(* The effects *) +| Input : expr +| Output (e : expr) : expr +with val {X : Set} := +| LitV (n : nat) : val +| RecV (e : @expr (inc (inc X))) : val +with ectx {X : Set} := +| EmptyK : ectx +| OutputK (K : ectx) : ectx +| IfK (K : ectx) (e₁ : expr) (e₂ : expr) : ectx +| AppLK (K : ectx) (v : val) : ectx +| AppRK (e : expr) (K : ectx) : ectx +| NatOpRK (op : nat_op) (e : expr) (K : ectx) : ectx +| NatOpLK (op : nat_op) (K : ectx) (v : val) : ectx. + +Arguments val X%bind : clear implicits. +Arguments expr X%bind : clear implicits. +Arguments ectx X%bind : clear implicits. + +Local Open Scope bind_scope. + +Fixpoint emap {A B : Set} (f : A [→] B) (e : expr A) : expr B := match e with - | Val v => Some v - | _ => None - end. - -Definition do_natop (op : nat_op) (x y : nat) : nat := - match op with - | Add => x+y - | Sub => x-y - | Mult => x+y - end. - -Definition nat_op_interp {S} (n : nat_op) (x y : val S) : option (val S) := - match x, y with - | Lit x, Lit y => Some $ Lit $ do_natop n x y - | _,_ => None - end. - -(** substitution stuff *) -Definition rens S S' := var S → var S'. -Definition subs S S' := var S → expr S'. - -Definition idren {S} : rens S S := fun v => v. -Definition idsub {S} : subs S S := Var. - -Equations conssub {S S' τ} (M : expr S') (s : subs S S') : subs (τ::S) S' := - conssub M s Vz := M; - conssub M s (Vs v) := s v. - -Notation "{/ e ; .. ; f /}" := (conssub e .. (conssub f idsub) ..). - -Definition tl_sub {S S' τ} : subs (τ::S) S' → subs S S' := λ s v, s (Vs v). -Definition hd_sub {S S' τ} : subs (τ::S) S' → expr S' := λ s, s Vz. -Definition tl_ren {S S' τ} : rens (τ::S) S' → rens S S' := λ s v, s (Vs v). -Definition hd_ren {S S' τ} : rens (τ::S) S' → var S' := λ s, s Vz. - -(* Lifting a renaming, renaming terms, and lifting substitutions *) -Equations rens_lift {S S'} (s : rens S S') : rens (()::S) (()::S') := - rens_lift s Vz := Vz; - rens_lift s (Vs v) := Vs $ s v. - -Equations ren_expr {S S'} (M : expr S) (r : rens S S') : expr S' := -ren_expr (Val v) r := Val $ ren_val v r; -ren_expr (Var v) r := Var (r v); -ren_expr (Rec M) r := Rec (ren_expr M (rens_lift (rens_lift r))); -ren_expr (App M N) r := App (ren_expr M r) (ren_expr N r); -ren_expr (NatOp op e1 e2) r := NatOp op (ren_expr e1 r) (ren_expr e2 r); -ren_expr (If e0 e1 e2) r := If (ren_expr e0 r) (ren_expr e1 r) (ren_expr e2 r); -ren_expr Input r := Input; -ren_expr (Output e) r := Output (ren_expr e r); -with ren_val {S S'} (M : val S) (r : rens S S') : val S' := -ren_val (Lit n) _ := Lit n; -ren_val (RecV e) r := RecV (ren_expr e (rens_lift (rens_lift r))). - - -Definition expr_lift {S} (M : expr S) : expr (()::S) := ren_expr M Vs. - -Equations subs_lift {S S'} (s : subs S S') : subs (()::S) (()::S') := - subs_lift s Vz := Var Vz; - subs_lift s (Vs v) := expr_lift $ s v. - -(* We can now define the substitution operation *) -Equations subst_expr {S S'} (M : expr S) (s : subs S S') : expr S' := -subst_expr (Val v) r := Val $ subst_val v r; -subst_expr (Var v) r := r v; -subst_expr (Rec M) r := Rec (subst_expr M (subs_lift (subs_lift r))); -subst_expr (App M N) r := App (subst_expr M r) (subst_expr N r); -subst_expr (NatOp op e1 e2) r := NatOp op (subst_expr e1 r) (subst_expr e2 r); -subst_expr (If e0 e1 e2) r := If (subst_expr e0 r) (subst_expr e1 r) (subst_expr e2 r); -subst_expr (Input) r := Input; -subst_expr (Output e) r := Output (subst_expr e r); -with subst_val {S S'} (M : val S) (r : subs S S') : val S' := -subst_val (Lit n) _ := Lit n; -subst_val (RecV e) r := RecV (subst_expr e (subs_lift (subs_lift r))). - -Definition subst1 {S : scope} {τ} (M : expr (τ::S)) (N : expr S) : expr S - := subst_expr M {/ N /}. -Definition subst2 {S : scope} {i j} (M : expr (i::j::S)) (N1 : expr S) (N2 : expr S) : expr S - := subst_expr M {/ N1; N2 /}. - -Definition appsub {S1 S2 S3} (s : subs S1 S2) (s' : subs S2 S3) : subs S1 S3 := - λ v, subst_expr (s v) s'. - -Global Instance rens_equiv S S' : Equiv (rens S S') := λ s1 s2, ∀ v, s1 v = s2 v. -Global Instance subs_equiv S S' : Equiv (subs S S') := λ s1 s2, ∀ v, s1 v = s2 v. - -Global Instance rens_lift_proper S S' : Proper ((≡) ==> (≡)) (@rens_lift S S'). + | Val v => Val (vmap f v) + | Var x => Var (f x) + | App e₁ e₂ => App (emap f e₁) (emap f e₂) + | NatOp o e₁ e₂ => NatOp o (emap f e₁) (emap f e₂) + | If e₁ e₂ e₃ => If (emap f e₁) (emap f e₂) (emap f e₃) + | Input => Input + | Output e => Output (emap f e) + end +with vmap {A B : Set} (f : A [→] B) (v : val A) : val B := + match v with + | LitV n => LitV n + | RecV e => RecV (emap ((f ↑) ↑) e) + end +with kmap {A B : Set} (f : A [→] B) (K : ectx A) : ectx B := + match K with + | EmptyK => EmptyK + | OutputK K => OutputK (kmap f K) + | IfK K e₁ e₂ => IfK (kmap f K) (emap f e₁) (emap f e₂) + | AppLK K v => AppLK (kmap f K) (vmap f v) + | AppRK e K => AppRK (emap f e) (kmap f K) + | NatOpRK op e K => NatOpRK op (emap f e) (kmap f K) + | NatOpLK op K v => NatOpLK op (kmap f K) (vmap f v) + end. +#[export] Instance FMap_expr : FunctorCore expr := @emap. +#[export] Instance FMap_val : FunctorCore val := @vmap. +#[export] Instance FMap_ectx : FunctorCore ectx := @kmap. + +#[export] Instance SPC_expr : SetPureCore expr := @Var. + +Fixpoint ebind {A B : Set} (f : A [⇒] B) (e : expr A) : expr B := + match e with + | Val v => Val (vbind f v) + | Var x => f x + | App e₁ e₂ => App (ebind f e₁) (ebind f e₂) + | NatOp o e₁ e₂ => NatOp o (ebind f e₁) (ebind f e₂) + | If e₁ e₂ e₃ => If (ebind f e₁) (ebind f e₂) (ebind f e₃) + | Input => Input + | Output e => Output (ebind f e) + end +with vbind {A B : Set} (f : A [⇒] B) (v : val A) : val B := + match v with + | LitV n => LitV n + | RecV e => RecV (ebind ((f ↑) ↑) e) + end +with kbind {A B : Set} (f : A [⇒] B) (K : ectx A) : ectx B := + match K with + | EmptyK => EmptyK + | OutputK K => OutputK (kbind f K) + | IfK K e₁ e₂ => IfK (kbind f K) (ebind f e₁) (ebind f e₂) + | AppLK K v => AppLK (kbind f K) (vbind f v) + | AppRK e K => AppRK (ebind f e) (kbind f K) + | NatOpRK op e K => NatOpRK op (ebind f e) (kbind f K) + | NatOpLK op K v => NatOpLK op (kbind f K) (vbind f v) + end. + +#[export] Instance BindCore_expr : BindCore expr := @ebind. +#[export] Instance BindCore_val : BindCore val := @vbind. +#[export] Instance BindCore_ectx : BindCore ectx := @kbind. + +#[export] Instance IP_typ : SetPure expr. Proof. - intros s1 s2 Hs v. dependent elimination v; simp rens_lift; eauto. - f_equiv. apply Hs. + split; intros; reflexivity. Qed. -Lemma ren_expr_proper {S S'} (e : expr S) : Proper ((≡) ==> (=)) (@ren_expr S S' e) - with ren_val_proper {S S'} v : Proper ((≡) ==> (=)) (@ren_val S S' v). +Fixpoint vmap_id X (δ : X [→] X) (v : val X) : δ ≡ ı → fmap δ v = v +with emap_id X (δ : X [→] X) (e : expr X) : δ ≡ ı → fmap δ e = e +with kmap_id X (δ : X [→] X) (e : ectx X) : δ ≡ ı → fmap δ e = e. Proof. - - revert S'. - induction e; intros S' s1 s2 Hs; simp ren_expr; - f_equiv; try solve [eauto | apply ren_expr_proper; eauto ]. - + by apply ren_val_proper. - + apply ren_expr_proper. by repeat f_equiv. - - revert S'. - induction v; intros S' s1 s2 Hs; simp ren_expr; - f_equiv; try solve [eauto | apply ren_expr_proper; eauto ]. - apply ren_expr_proper. by repeat f_equiv. + - auto_map_id. + - auto_map_id. + - auto_map_id. Qed. -#[export] Existing Instance ren_expr_proper. -#[export] Existing Instance ren_val_proper. - -#[export] Instance subs_lift_proper S S' : Proper ((≡) ==> (≡)) (@subs_lift S S'). +Fixpoint vmap_comp (A B C : Set) (f : B [→] C) (g : A [→] B) h (v : val A) : + f ∘ g ≡ h → fmap f (fmap g v) = fmap h v +with emap_comp (A B C : Set) (f : B [→] C) (g : A [→] B) h (e : expr A) : + f ∘ g ≡ h → fmap f (fmap g e) = fmap h e +with kmap_comp (A B C : Set) (f : B [→] C) (g : A [→] B) h (e : ectx A) : + f ∘ g ≡ h → fmap f (fmap g e) = fmap h e. Proof. - intros s1 s2 Hs v. dependent elimination v; simp subs_lift; eauto. - f_equiv. apply Hs. + - auto_map_comp. + - auto_map_comp. + - auto_map_comp. Qed. -Lemma subst_expr_proper {S S'} (e : expr S) : Proper ((≡) ==> (=)) (@subst_expr S S' e) - with subst_val_proper {S S'} v : Proper ((≡) ==> (=)) (@subst_val S S' v). +#[export] Instance Functor_val : Functor val. +Proof. + split; [exact vmap_id | exact vmap_comp]. +Qed. +#[export] Instance Functor_expr : Functor expr. +Proof. + split; [exact emap_id | exact emap_comp]. +Qed. +#[export] Instance Functor_ectx : Functor ectx. Proof. - - revert S'. - induction e; intros S' s1 s2 Hs; simp subst_expr; - f_equiv; try solve [eauto | apply subst_expr_proper; eauto ]. - + by apply subst_val_proper. - + apply subst_expr_proper. by repeat f_equiv. - - revert S'. - induction v; intros S' s1 s2 Hs; simp subst_expr; - f_equiv; try solve [eauto | apply subst_expr_proper; eauto ]. - apply subst_expr_proper. by repeat f_equiv. + split; [exact kmap_id | exact kmap_comp]. Qed. -#[export] Existing Instance subst_expr_proper. -#[export] Existing Instance subst_val_proper. -Lemma subst_ren_expr {S1 S2 S3} e (s : subs S2 S3) (r : rens S1 S2) : - subst_expr (ren_expr e r) s = subst_expr e (compose s r) -with subst_ren_val {S1 S2 S3} v (s : subs S2 S3) (r : rens S1 S2) : - subst_val (ren_val v r) s = subst_val v (compose s r). +Fixpoint vmap_vbind_pure (A B : Set) (f : A [→] B) (g : A [⇒] B) (v : val A) : + f ̂ ≡ g → fmap f v = bind g v +with emap_ebind_pure (A B : Set) (f : A [→] B) (g : A [⇒] B) (e : expr A) : + f ̂ ≡ g → fmap f e = bind g e +with kmap_kbind_pure (A B : Set) (f : A [→] B) (g : A [⇒] B) (e : ectx A) : + f ̂ ≡ g → fmap f e = bind g e. Proof. - - revert S2 S3 r s. - induction e=>S2 S3 r s; simp ren_expr; simp subst_expr; try f_equiv; eauto. - rewrite IHe. apply subst_expr_proper. - intro v. simpl. - dependent elimination v; simp rens_lift; simp subs_lift; eauto. - f_equiv. dependent elimination v; simp rens_lift; simp subs_lift; eauto. - - revert S2 S3 r s. - induction v=>S2 S3 r s; simpl; simp ren_val; simp subst_val; try f_equiv. - rewrite subst_ren_expr. - apply subst_expr_proper. - intro v. simpl. - dependent elimination v; simp rens_lift; simp subs_lift; eauto. - f_equiv. dependent elimination v; simp rens_lift; simp subs_lift; eauto. + - auto_map_bind_pure. + erewrite emap_ebind_pure; [reflexivity |]. + intros [| [| x]]; term_simpl; [reflexivity | reflexivity |]. + rewrite <-(EQ x). + reflexivity. + - auto_map_bind_pure. + - auto_map_bind_pure. Qed. -Lemma ren_ren_expr {S1 S2 S3} e (s : rens S2 S3) (r : rens S1 S2) : - ren_expr (ren_expr e r) s = ren_expr e (compose s r) -with ren_ren_val {S1 S2 S3} v (s : rens S2 S3) (r : rens S1 S2) : - ren_val (ren_val v r) s = ren_val v (compose s r). +#[export] Instance BindMapPure_val : BindMapPure val. +Proof. + split; intros; now apply vmap_vbind_pure. +Qed. +#[export] Instance BindMapPure_expr : BindMapPure expr. +Proof. + split; intros; now apply emap_ebind_pure. +Qed. +#[export] Instance BindMapPure_ectx : BindMapPure ectx. Proof. - - revert S2 S3 r s. - induction e=>S2 S3 r s; simp ren_expr; try f_equiv; eauto. - rewrite IHe. apply ren_expr_proper. - intro v. simpl. - dependent elimination v; simp rens_lift; simp subs_lift; eauto. - f_equiv. dependent elimination v; simp rens_lift; simp subs_lift; eauto. - - revert S2 S3 r s. - induction v=>S2 S3 r s; simpl; simp ren_val; simp subst_val; try f_equiv. - rewrite ren_ren_expr. - apply ren_expr_proper. - intro v. simpl. - dependent elimination v; simp rens_lift; simp subs_lift; eauto. - f_equiv. dependent elimination v; simp rens_lift; simp subs_lift; eauto. + split; intros; now apply kmap_kbind_pure. Qed. -Definition rcompose {S1 S2 S3} (r : rens S2 S3) (s : subs S1 S2) : subs S1 S3 := - λ v, ren_expr (s v) r. +Fixpoint vmap_vbind_comm (A B₁ B₂ C : Set) (f₁ : B₁ [→] C) (f₂ : A [→] B₂) + (g₁ : A [⇒] B₁) (g₂ : B₂ [⇒] C) (v : val A) : + g₂ ∘ f₂ ̂ ≡ f₁ ̂ ∘ g₁ → bind g₂ (fmap f₂ v) = fmap f₁ (bind g₁ v) +with emap_ebind_comm (A B₁ B₂ C : Set) (f₁ : B₁ [→] C) (f₂ : A [→] B₂) + (g₁ : A [⇒] B₁) (g₂ : B₂ [⇒] C) (e : expr A) : + g₂ ∘ f₂ ̂ ≡ f₁ ̂ ∘ g₁ → bind g₂ (fmap f₂ e) = fmap f₁ (bind g₁ e) +with kmap_kbind_comm (A B₁ B₂ C : Set) (f₁ : B₁ [→] C) (f₂ : A [→] B₂) + (g₁ : A [⇒] B₁) (g₂ : B₂ [⇒] C) (e : ectx A) : + g₂ ∘ f₂ ̂ ≡ f₁ ̂ ∘ g₁ → bind g₂ (fmap f₂ e) = fmap f₁ (bind g₁ e). +Proof. + - auto_map_bind_comm. + erewrite emap_ebind_comm; [reflexivity |]. + erewrite lift_comm; [reflexivity |]. + erewrite lift_comm; [reflexivity | assumption]. + - auto_map_bind_comm. + - auto_map_bind_comm. +Qed. -Lemma ren_subst_expr {S1 S2 S3} e (s : subs S1 S2) (r : rens S2 S3) : - ren_expr (subst_expr e s) r = subst_expr e (rcompose r s) -with ren_subst_val {S1 S2 S3} v (s : subs S1 S2) (r : rens S2 S3) : - ren_val (subst_val v s) r = subst_val v (rcompose r s). +#[export] Instance BindMapComm_val : BindMapComm val. +Proof. + split; intros; now apply vmap_vbind_comm. +Qed. +#[export] Instance BindMapComm_expr : BindMapComm expr. +Proof. + split; intros; now apply emap_ebind_comm. +Qed. +#[export] Instance BindMapComm_ectx : BindMapComm ectx. Proof. - - revert S2 S3 r s. - induction e=>S2 S3 r s; simp subst_expr; simp ren_expr; try f_equiv; eauto. - rewrite IHe. apply subst_expr_proper. - intro v. simpl. unfold rcompose. - dependent elimination v; eauto. - dependent elimination v; eauto. - simp subs_lift. unfold expr_lift. - rewrite !ren_ren_expr. apply ren_expr_proper. - intro x. dependent elimination v; eauto. - - revert S2 S3 r s. - induction v=>S2 S3 r s; simp subst_expr; simp ren_expr; try f_equiv; eauto. - rewrite ren_subst_expr. apply subst_expr_proper. - intro v. simpl. unfold rcompose. - dependent elimination v; eauto. - dependent elimination v; eauto. - simp subs_lift. unfold expr_lift. - rewrite !ren_ren_expr. apply ren_expr_proper. - intro x. dependent elimination v; eauto. + split; intros; now apply kmap_kbind_comm. Qed. -Lemma appsub_lift {S1 S2 S3} (s : subs S1 S2) (s' : subs S2 S3) : - subs_lift (appsub s s') ≡ appsub (subs_lift s) (subs_lift s'). +Fixpoint vbind_id (A : Set) (f : A [⇒] A) (v : val A) : + f ≡ ı → bind f v = v +with ebind_id (A : Set) (f : A [⇒] A) (e : expr A) : + f ≡ ı → bind f e = e +with kbind_id (A : Set) (f : A [⇒] A) (e : ectx A) : + f ≡ ı → bind f e = e. Proof. - unfold appsub. - intro v. dependent elimination v; simp subs_lift; eauto. - unfold expr_lift. rewrite subst_ren_expr. - rewrite ren_subst_expr. apply subst_expr_proper. - intro x. unfold rcompose. simpl. simp subs_lift. done. + - auto_bind_id. + rewrite ebind_id; [reflexivity |]. + apply lift_id, lift_id; assumption. + - auto_bind_id. + - auto_bind_id. Qed. -Lemma subst_expr_appsub {S1 S2 S3} (s1 : subs S1 S2) (s2 : subs S2 S3) e : - subst_expr (subst_expr e s1) s2 = subst_expr e (appsub s1 s2) -with subst_val_appsub {S1 S2 S3} (s1 : subs S1 S2) (s2 : subs S2 S3) v : - subst_val (subst_val v s1) s2 = subst_val v (appsub s1 s2). +Fixpoint vbind_comp (A B C : Set) (f : B [⇒] C) (g : A [⇒] B) h (v : val A) : + f ∘ g ≡ h → bind f (bind g v) = bind h v +with ebind_comp (A B C : Set) (f : B [⇒] C) (g : A [⇒] B) h (e : expr A) : + f ∘ g ≡ h → bind f (bind g e) = bind h e +with kbind_comp (A B C : Set) (f : B [⇒] C) (g : A [⇒] B) h (e : ectx A) : + f ∘ g ≡ h → bind f (bind g e) = bind h e. Proof. - - revert S2 S3 s1 s2. - induction e=>S2 S3 s1 s2; simp subst_expr; try f_equiv; eauto. - rewrite !appsub_lift. apply IHe. - - revert S3 s2. - induction v=>S3 s2; simpl; f_equiv; eauto. - rewrite !appsub_lift. apply subst_expr_appsub. + - auto_bind_comp. + erewrite ebind_comp; [reflexivity |]. + erewrite lift_comp; [reflexivity |]. + erewrite lift_comp; [reflexivity | assumption]. + - auto_bind_comp. + - auto_bind_comp. Qed. -Lemma subst_expr_lift {S S'} e e1 (s : subs S S') : - subst_expr (expr_lift e) (conssub e1 s) = subst_expr e s. +#[export] Instance Bind_val : Bind val. Proof. - unfold expr_lift. - rewrite subst_ren_expr. apply subst_expr_proper. - intro v. simpl. simp conssub. done. + split; intros; [now apply vbind_id | now apply vbind_comp]. +Qed. +#[export] Instance Bind_expr : Bind expr. +Proof. + split; intros; [now apply ebind_id | now apply ebind_comp]. +Qed. +#[export] Instance Bind_ectx : Bind ectx. +Proof. + split; intros; [now apply kbind_id | now apply kbind_comp]. Qed. -Lemma subst_expr_idsub {S} (e : expr S) : - subst_expr e idsub = e -with subst_val_idsub {S} (v : val S) : - subst_val v idsub = v. +Definition to_val {S} (e : expr S) : option (val S) := + match e with + | Val v => Some v + | _ => None + end. + +Definition do_natop (op : nat_op) (x y : nat) : nat := + match op with + | Add => plus x y + | Sub => minus x y + | Mult => mult x y + end. + +Definition nat_op_interp {S} (n : nat_op) (x y : val S) : option (val S) := + match x, y with + | LitV x, LitV y => Some $ LitV $ do_natop n x y + | _,_ => None + end. + +Fixpoint fill {X : Set} (K : ectx X) (e : expr X) : expr X := + match K with + | EmptyK => e + | OutputK K => Output (fill K e) + | IfK K e₁ e₂ => If (fill K e) e₁ e₂ + | AppLK K v => App (fill K e) (Val v) + | AppRK e' K => App e' (fill K e) + | NatOpRK op e' K => NatOp op e' (fill K e) + | NatOpLK op K v => NatOp op (fill K e) (Val v) + end. + +Lemma fill_emap {X Y : Set} (f : X [→] Y) (K : ectx X) (e : expr X) + : fmap f (fill K e) = fill (fmap f K) (fmap f e). Proof. - - induction e; simp subst_expr; simpl; try f_equiv; eauto. - assert ((subs_lift (subs_lift idsub)) ≡ idsub) as ->; last auto. - intro v. - dependent elimination v; simp subs_lift; auto. - dependent elimination v; simp subs_lift; auto. - - induction v; simp subst_val; simpl; try f_equiv; eauto. - assert ((subs_lift (subs_lift idsub)) ≡ idsub) as ->; last auto. - intro v. - dependent elimination v; simp subs_lift; auto. - dependent elimination v; simp subs_lift; auto. + revert f. + induction K as [| ?? IH + | ?? IH + | ?? IH + | ??? IH + | ???? IH + | ??? IH]; + intros f; term_simpl; first done; rewrite IH; reflexivity. Qed. (*** Operational semantics *) @@ -286,7 +289,6 @@ Record state := State { }. #[export] Instance state_inhabited : Inhabited state := populate (State [] []). - Definition update_input (s : state) : nat * state := match s.(inputs) with | [] => (0, s) @@ -296,30 +298,26 @@ Definition update_input (s : state) : nat * state := Definition update_output (n:nat) (s : state) : state := {| inputs := s.(inputs); outputs := n::s.(outputs) |}. - Inductive head_step {S} : expr S → state → expr S → state → nat*nat → Prop := -| RecS e σ : - head_step (Rec e) σ (Val $ RecV e) σ (0,0) -| BetaS e1 v2 e' σ : - e' = subst2 e1 (Val $ RecV e1) (Val v2) → - head_step (App (Val $ RecV e1) (Val v2)) σ e' σ (1,0) +| BetaS e1 v2 σ : + head_step (App (Val $ RecV e1) (Val v2)) σ (subst (Inc := inc) ((subst (F := expr) (Inc := inc) e1) (Val (shift (Inc := inc) v2))) (Val (RecV e1))) σ (1,0) | InputS σ n σ' : update_input σ = (n,σ') → - head_step Input σ (Val (Lit n)) σ' (1,1) + head_step Input σ (Val (LitV n)) σ' (1,1) | OutputS σ n σ' : update_output n σ = σ' → - head_step (Output (Val (Lit n))) σ (Val (Lit 0)) σ' (1,1) + head_step (Output (Val (LitV n))) σ (Val (LitV 0)) σ' (1,1) | NatOpS op v1 v2 v3 σ : nat_op_interp op v1 v2 = Some v3 → head_step (NatOp op (Val v1) (Val v2)) σ (Val v3) σ (0,0) | IfTrueS n e1 e2 σ : n > 0 → - head_step (If (Val (Lit n)) e1 e2) σ + head_step (If (Val (LitV n)) e1 e2) σ e1 σ (0,0) | IfFalseS n e1 e2 σ : n = 0 → - head_step (If (Val (Lit n)) e1 e2) σ + head_step (If (Val (LitV n)) e1 e2) σ e2 σ (0,0) . @@ -333,59 +331,52 @@ Lemma head_step_no_io {S} (e1 e2 : expr S) σ1 σ2 n : head_step e1 σ1 e2 σ2 (n,0) → σ1 = σ2. Proof. inversion 1; eauto. Qed. -Inductive ectx_item {S} := - | AppLCtx (v2 : val S) - | AppRCtx (e1 : expr S) - | NatOpLCtx (op : nat_op) (v2 : val S) - | NatOpRCtx (op : nat_op) (e1 : expr S) - | IfCtx (e1 e2 : expr S) - | OutputCtx -. -Arguments ectx_item S : clear implicits. - -Definition fill_item {S} (Ki : ectx_item S) (e : expr S) : expr S := - match Ki with - | AppLCtx v2 => App e (of_val v2) - | AppRCtx e1 => App e1 e - | NatOpLCtx op v2 => NatOp op e (Val v2) - | NatOpRCtx op e1 => NatOp op e1 e - | IfCtx e1 e2 => If e e1 e2 - | OutputCtx => Output e - end. - (** Carbonara from heap lang *) -Global Instance fill_item_inj {S} (Ki : ectx_item S) : Inj (=) (=) (fill_item Ki). +Global Instance fill_item_inj {S} (Ki : ectx S) : Inj (=) (=) (fill Ki). Proof. induction Ki; intros ???; simplify_eq/=; auto with f_equal. Qed. Lemma fill_item_val {S} Ki (e : expr S) : - is_Some (to_val (fill_item Ki e)) → is_Some (to_val e). + is_Some (to_val (fill Ki e)) → is_Some (to_val e). Proof. intros [v ?]. induction Ki; simplify_option_eq; eauto. Qed. Lemma val_head_stuck {S} (e1 : expr S) σ1 e2 σ2 m : head_step e1 σ1 e2 σ2 m → to_val e1 = None. Proof. destruct 1; naive_solver. Qed. -Lemma head_ctx_item_step_val {S} Ki (e : expr S) σ1 e2 σ2 m : - head_step (fill_item Ki e) σ1 e2 σ2 m → is_Some (to_val e). -Proof. revert m e2. induction Ki; simpl; inversion 1; simplify_option_eq; eauto. Qed. - -Lemma fill_item_no_val_inj {S} Ki1 Ki2 (e1 e2 : expr S) : - to_val e1 = None → to_val e2 = None → - fill_item Ki1 e1 = fill_item Ki2 e2 → Ki1 = Ki2. +Fixpoint ectx_compose {S} (K1 K2 : ectx S) : ectx S + := match K1 with + | EmptyK => K2 + | OutputK K => OutputK (ectx_compose K K2) + | IfK K e₁ e₂ => IfK (ectx_compose K K2) e₁ e₂ + | AppLK K v => AppLK (ectx_compose K K2) v + | AppRK e K => AppRK e (ectx_compose K K2) + | NatOpRK op e K => NatOpRK op e (ectx_compose K K2) + | NatOpLK op K v => NatOpLK op (ectx_compose K K2) v + end. + +Lemma fill_app {S} (K1 K2 : ectx S) e : fill (ectx_compose K1 K2) e = fill K1 (fill K2 e). Proof. - revert Ki1. induction Ki2; intros Ki1; induction Ki1; naive_solver eauto with f_equal. + revert K2. + revert e. + induction K1 as [| ?? IH + | ?? IH + | ?? IH + | ??? IH + | ???? IH + | ??? IH]; + simpl; first done; intros e' K2; rewrite IH; reflexivity. Qed. -(** Lifting the head step **) - -Definition ectx S := (list (ectx_item S)). -Definition fill {S} (K : ectx S) (e : expr S) : expr S := foldl (flip fill_item) e K. - -Lemma fill_app {S} (K1 K2 : ectx S) e : fill (K1 ++ K2) e = fill K2 (fill K1 e). -Proof. apply foldl_app. Qed. - - Lemma fill_val : ∀ {S} K (e : expr S), is_Some (to_val (fill K e)) → is_Some (to_val e). -Proof. intros S K. induction K as [|Ki K IH]=> e //=. by intros ?%IH%fill_item_val. Qed. +Proof. + intros S K. + induction K as [| ?? IH + | ?? IH + | ?? IH + | ??? IH + | ???? IH + | ??? IH]=> e' //=; + inversion 1 as [? HH]; inversion HH. +Qed. Lemma fill_not_val : ∀ {S} K (e : expr S), to_val e = None → to_val (fill K e) = None. Proof. @@ -393,13 +384,20 @@ Proof. eauto using fill_val. Qed. -Lemma fill_empty {S} (e : expr S) : fill [] e = e. +Lemma fill_empty {S} (e : expr S) : fill EmptyK e = e. Proof. reflexivity. Qed. -Lemma fill_comp {S} K1 K2 (e : expr S) : fill K1 (fill K2 e) = fill (K2 ++ K1) e. +Lemma fill_comp {S} K1 K2 (e : expr S) : fill K2 (fill K1 e) = fill (ectx_compose K2 K1) e. Proof. by rewrite fill_app. Qed. -Global Instance fill_inj {S} (K:ectx S) : Inj (=) (=) (fill K). -Proof. induction K as [|Ki K IH]; rewrite /Inj; naive_solver. Qed. - +Global Instance fill_inj {S} (K : ectx S) : Inj (=) (=) (fill K). +Proof. + induction K as [| ?? IH + | ?? IH + | ?? IH + | ??? IH + | ???? IH + | ??? IH]; + rewrite /Inj; naive_solver. +Qed. Inductive prim_step {S} (e1 : expr S) (σ1 : state) (e2 : expr S) (σ2 : state) (n : nat*nat) : Prop:= @@ -469,41 +467,140 @@ Qed. Inductive ty := | Tnat : ty | Tarr : ty → ty → ty. -Local Notation tyctx := (tyctx ty). - -Inductive typed : forall {S}, tyctx S → expr S → ty → Prop := -| typed_Val {S} (Γ : tyctx S) (τ : ty) (v : val S) : +Inductive typed {S : Set} (Γ : S -> ty) : expr S → ty → Prop := +| typed_Val (τ : ty) (v : val S) : typed_val Γ v τ → typed Γ (Val v) τ -| typed_Var {S} (Γ : tyctx S) (τ : ty) (v : var S) : - typed_var Γ v τ → +| typed_Var (τ : ty) (v : S) : + Γ v = τ → typed Γ (Var v) τ -| typed_Rec {S} (Γ : tyctx S) (τ1 τ2 : ty) (e : expr (()::()::S) ) : - typed (consC (Tarr τ1 τ2) (consC τ1 Γ)) e τ2 → - typed Γ (Rec e) (Tarr τ1 τ2) -| typed_App {S} (Γ : tyctx S) (τ1 τ2 : ty) e1 e2 : +| typed_App (τ1 τ2 : ty) e1 e2 : typed Γ e1 (Tarr τ1 τ2) → typed Γ e2 τ1 → typed Γ (App e1 e2) τ2 -| typed_NatOp {S} (Γ : tyctx S) e1 e2 op : +| typed_NatOp e1 e2 op : typed Γ e1 Tnat → typed Γ e2 Tnat → typed Γ (NatOp op e1 e2) Tnat -| typed_If {S} (Γ : tyctx S) e0 e1 e2 τ : +| typed_If e0 e1 e2 τ : typed Γ e0 Tnat → typed Γ e1 τ → typed Γ e2 τ → typed Γ (If e0 e1 e2) τ -| typed_Input {S} (Γ : tyctx S) : +| typed_Input : typed Γ Input Tnat -| typed_Output {S} (Γ : tyctx S) e : +| typed_Output e : typed Γ e Tnat → typed Γ (Output e) Tnat -with typed_val : forall {S}, tyctx S → val S → ty → Prop := -| typed_Lit {S} (Γ : tyctx S) n : - typed_val Γ (Lit n) Tnat -| typed_RecV {S} (Γ : tyctx S) (τ1 τ2 : ty) (e : expr (()::()::S) ) : - typed (consC (Tarr τ1 τ2) (consC τ1 Γ)) e τ2 → +with typed_val {S : Set} (Γ : S -> ty) : val S → ty → Prop := +| typed_Lit n : + typed_val Γ (LitV n) Tnat +| typed_RecV (τ1 τ2 : ty) (e : expr (inc (inc S))) : + typed (Γ ▹ (Tarr τ1 τ2) ▹ τ1) e τ2 → typed_val Γ (RecV e) (Tarr τ1 τ2) . +Declare Scope syn_scope. +Delimit Scope syn_scope with syn. + +Coercion Val : val >-> expr. + +Coercion App : expr >-> Funclass. +Coercion AppLK : ectx >-> Funclass. +Coercion AppRK : expr >-> Funclass. + +Class AsSynExpr (F : Set -> Type) := { __asSynExpr : ∀ S, F S -> expr S }. + +Arguments __asSynExpr {_} {_} {_}. + +Global Instance AsSynExprValue : AsSynExpr val := { + __asSynExpr _ v := Val v + }. +Global Instance AsSynExprExpr : AsSynExpr expr := { + __asSynExpr _ e := e + }. + +Class OpNotation (A B C D : Type) := { __op : A -> B -> C -> D }. + +Global Instance OpNotationExpr {S : Set} {F G : Set -> Type} `{AsSynExpr F, AsSynExpr G} : OpNotation (F S) nat_op (G S) (expr S) := { + __op e₁ op e₂ := NatOp op (__asSynExpr e₁) (__asSynExpr e₂) + }. + +Global Instance OpNotationLK {S : Set} : OpNotation (ectx S) (nat_op) (val S) (ectx S) := { + __op K op v := NatOpLK op K v + }. + +Global Instance OpNotationRK {S : Set} {F : Set -> Type} `{AsSynExpr F} : OpNotation (F S) (nat_op) (ectx S) (ectx S) := { + __op e op K := NatOpRK op (__asSynExpr e) K + }. + +Class IfNotation (A B C D : Type) := { __if : A -> B -> C -> D }. + +Global Instance IfNotationExpr {S : Set} {F G H : Set -> Type} `{AsSynExpr F, AsSynExpr G, AsSynExpr H} : IfNotation (F S) (G S) (H S) (expr S) := { + __if e₁ e₂ e₃ := If (__asSynExpr e₁) (__asSynExpr e₂) (__asSynExpr e₃) + }. + +Global Instance IfNotationK {S : Set} {F G : Set -> Type} `{AsSynExpr F, AsSynExpr G} : IfNotation (ectx S) (F S) (G S) (ectx S) := { + __if K e₂ e₃ := IfK K (__asSynExpr e₂) (__asSynExpr e₃) + }. + +Class OutputNotation (A B : Type) := { __output : A -> B }. + +Global Instance OutputNotationExpr {S : Set} {F : Set -> Type} `{AsSynExpr F} : OutputNotation (F S) (expr S) := { + __output e := Output (__asSynExpr e) + }. + +Global Instance OutputNotationK {S : Set} : OutputNotation (ectx S) (ectx S) := { + __output K := OutputK K + }. + +Class AppNotation (A B C : Type) := { __app : A -> B -> C }. + +Global Instance AppNotationExpr {S : Set} {F G : Set -> Type} `{AsSynExpr F, AsSynExpr G} : AppNotation (F S) (G S) (expr S) := { + __app e₁ e₂ := App (__asSynExpr e₁) (__asSynExpr e₂) + }. + +Global Instance AppNotationLK {S : Set} : AppNotation (ectx S) (val S) (ectx S) := { + __app K v := AppLK K v + }. + +Global Instance AppNotationRK {S : Set} {F : Set -> Type} `{AsSynExpr F} : AppNotation (F S) (ectx S) (ectx S) := { + __app e K := AppRK (__asSynExpr e) K + }. + +Notation of_val := Val (only parsing). + +Notation "x '⋆' y" := (__app x%syn y%syn) (at level 40, y at next level, left associativity) : syn_scope. +Notation "x '+' y" := (__op x%syn Add y%syn) : syn_scope. +Notation "x '-' y" := (__op x%syn Sub y%syn) : syn_scope. +Notation "x '*' y" := (__op x%syn Mult y%syn) : syn_scope. +Notation "'if' x 'then' y 'else' z" := (__if x%syn y%syn z%syn) : syn_scope. +Notation "'output' x" := (__output x%syn) (at level 60) : syn_scope. +Notation "'#' n" := (LitV n) (at level 60) : syn_scope. +Notation "'input'" := (Input) : syn_scope. +Notation "'rec' e" := (RecV e%syn) (at level 60) : syn_scope. +Notation "'$' fn" := (set_pure_resolver fn) (at level 60) : syn_scope. +Notation "□" := (EmptyK) : syn_scope. +Notation "K '⟪' e '⟫'" := (fill K%syn e%syn) (at level 60) : syn_scope. + +Definition LamV {S : Set} (e : expr (inc S)) : val S := + RecV (shift e). + +Notation "'λ' . e" := (LamV e%syn) (at level 60) : syn_scope. + +Definition LetE {S : Set} (e : expr S) (e' : expr (inc S)) : expr S := + App (LamV e') (e). + +Notation "'let_' e₁ 'in' e₂" := (LetE e₁%syn e₂%syn) (at level 60, right associativity) : syn_scope. + +Definition SeqE {S : Set} (e e' : expr S) : expr S := + App (LamV (shift e)) e'. + +Notation "e₁ ';;' e₂" := (SeqE e₁%syn e₂%syn) : syn_scope. + +Declare Scope typ_scope. +Delimit Scope typ_scope with typ. + +Notation "'ℕ'" := (Tnat) (at level 1) : typ_scope. +Notation "A →ₜ B" := (Tarr A%typ B%typ) + (right associativity, at level 60) : typ_scope. diff --git a/theories/input_lang/logpred.v b/theories/input_lang/logpred.v index 34888c2..ea1b802 100644 --- a/theories/input_lang/logpred.v +++ b/theories/input_lang/logpred.v @@ -1,38 +1,38 @@ (** Unary (Kripke) logical relation for the IO lang *) -From Equations Require Import Equations. -From gitrees Require Import gitree program_logic. +From gitrees Require Import gitree program_logic lang_generic. From gitrees.input_lang Require Import lang interp. +Require Import Binding.Lib Binding.Set Binding.Env. Section io_lang. Context {sz : nat}. - Variable rs : gReifiers sz. + Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ, !na_invG Σ}. Notation iProp := (iProp Σ). - Context {HCI : ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) R) o}. + + Canonical Structure exprO S := leibnizO (expr S). + Canonical Structure valO S := leibnizO (val S). Variable s : stuckness. - Context {A:ofe}. - Variable (P : A → iProp). - Context `{!NonExpansive P}. + Context {A : ofe}. + Variable (P : A -n> iProp). - Local Notation tyctx := (tyctx ty). Local Notation expr_pred := (expr_pred s rs P). Program Definition interp_tnat : ITV -n> iProp := λne αv, (∃ n : nat, αv ≡ RetV n)%I. Solve All Obligations with solve_proper. + Program Definition interp_tarr (Φ1 Φ2 : ITV -n> iProp) := λne αv, - (□ ∀ σ βv, has_substate σ -∗ - Φ1 βv -∗ - expr_pred (IT_of_V αv ⊙ (IT_of_V βv)) (λne v, ∃ σ', Φ2 v ∗ has_substate σ'))%I. + (□ ∀ σ βv, has_substate σ + -∗ Φ1 βv + -∗ expr_pred (IT_of_V αv ⊙ (IT_of_V βv)) + (λne v, ∃ σ', Φ2 v ∗ has_substate σ'))%I. Solve All Obligations with solve_proper. Fixpoint interp_ty (τ : ty) : ITV -n> iProp := @@ -41,55 +41,58 @@ Section io_lang. | Tarr τ1 τ2 => interp_tarr (interp_ty τ1) (interp_ty τ2) end. - Definition ssubst_valid {S} (Γ : tyctx S) ss := ssubst_valid rs interp_ty Γ ss. + Definition ssubst_valid {S : Set} + (Γ : S -> ty) + (ss : @interp_scope F R _ S) : iProp := + (∀ x, □ expr_pred (ss x) (interp_ty (Γ x)))%I. #[global] Instance io_lang_interp_ty_pers τ βv : Persistent (io_lang.interp_ty τ βv). Proof. induction τ; apply _. Qed. - #[global] Instance ssubst_valid_pers {S} (Γ : tyctx S) ss : Persistent (ssubst_valid Γ ss). + #[global] Instance ssubst_valid_pers {S : Set} (Γ : S → ty) ss : Persistent (ssubst_valid Γ ss). Proof. apply _. Qed. - Program Definition valid1 {S} (Γ : tyctx S) (α : interp_scope S -n> IT) (τ : ty) : iProp := + Program Definition valid1 {S : Set} (Γ : S → ty) (α : interp_scope S -n> IT) (τ : ty) : iProp := (∀ σ ss, has_substate σ -∗ ssubst_valid Γ ss -∗ - expr_pred (α (interp_ssubst ss)) (λne v, ∃ σ', interp_ty τ v ∗ has_substate σ'))%I. + expr_pred (α ss) (λne v, ∃ σ', interp_ty τ v ∗ has_substate σ'))%I. Solve Obligations with solve_proper. - Lemma compat_nat {S} n (Ω : tyctx S) : + Lemma compat_nat {S : Set} n (Ω : S → ty) : ⊢ valid1 Ω (interp_nat rs n) Tnat. Proof. iIntros (σ αs) "Hs Has". simpl. iApply expr_pred_ret. simpl. eauto with iFrame. Qed. - Lemma compat_var {S} Ω τ (v : var S) : - typed_var Ω v τ → - ⊢ valid1 Ω (interp_var v) τ. + + Lemma compat_var {S : Set} (Ω : S → ty) (v : S) : + ⊢ valid1 Ω (interp_var v) (Ω v). Proof. - intros Hv. iIntros (σ ss) "Hs Has". simpl. - unfold ssubst_valid. - iInduction Hv as [|? ? ? Ω v] "IH" forall (ss); simpl. - - dependent elimination ss as [cons_ssubst αv ss]. - rewrite ssubst_valid_cons. - simp interp_var. simpl. - iDestruct "Has" as "[H _]". - iApply expr_pred_ret; simpl; eauto with iFrame. - - dependent elimination ss as [cons_ssubst αv ss]. - rewrite ssubst_valid_cons. - simp interp_var. simpl. - iDestruct "Has" as "[_ H]". - by iApply ("IH" with "Hs H"). + iIntros (x) "HP". + simpl. + iSpecialize ("Has" $! v x with "HP"). + iApply (wp_wand with "Has"). + iIntros (v') "HH". + simpl. + iDestruct "HH" as "(%y & HH & HP')". + iModIntro. + iExists y. + iFrame "HP'". + iExists σ. + iFrame. Qed. - Lemma compat_if {S} (Γ : tyctx S) τ α β1 β2 : + Lemma compat_if {S : Set} (Γ : S → ty) τ α β1 β2 : ⊢ valid1 Γ α Tnat -∗ valid1 Γ β1 τ -∗ valid1 Γ β2 τ -∗ valid1 Γ (interp_if rs α β1 β2) τ. - Proof using HCI. + Proof. iIntros "H0 H1 H2". iIntros (σ ss) "Hs #Has". iSpecialize ("H0" with "Hs Has"). - simpl. iApply (expr_pred_bind (IFSCtx _ _) with "H0"). + simpl. + iApply (expr_pred_bind (IFSCtx _ _) with "H0"). iIntros (αv) "Ha/=". iDestruct "Ha" as (σ') "[Ha Hs]". iDestruct "Ha" as (n) "Hn". @@ -102,7 +105,7 @@ Section io_lang. iApply ("H1" with "Hs Has Hx"). Qed. - Lemma compat_input {S} (Γ : tyctx S) : + Lemma compat_input {S : Set} (Γ : S → ty) : ⊢ valid1 Γ (interp_input rs) Tnat. Proof. iIntros (σ ss) "Hs #Has". @@ -114,9 +117,9 @@ Section io_lang. iApply wp_val. simpl. eauto with iFrame. Qed. - Lemma compat_output {S} (Γ : tyctx S) α : + Lemma compat_output {S : Set} (Γ : S → ty) α : ⊢ valid1 Γ α Tnat → valid1 Γ (interp_output rs α) Tnat. - Proof using HCI. + Proof. iIntros "H". iIntros (σ ss) "Hs #Has". iSpecialize ("H" with "Hs Has"). @@ -134,11 +137,11 @@ Section io_lang. eauto with iFrame. Qed. - Lemma compat_app {S} (Γ : tyctx S) α β τ1 τ2 : + Lemma compat_app {S : Set} (Γ : S → ty) α β τ1 τ2 : ⊢ valid1 Γ α (Tarr τ1 τ2) -∗ valid1 Γ β τ1 -∗ valid1 Γ (interp_app rs α β) τ2. - Proof using HCI. + Proof. iIntros "H1 H2". iIntros (σ ss) "Hs #Has". simpl. iSpecialize ("H2" with "Hs Has"). @@ -154,15 +157,14 @@ Section io_lang. iApply ("Ha" with "Hs Hb"). Qed. - Lemma compat_rec {S} (Γ : tyctx S) τ1 τ2 α : - ⊢ □ valid1 (consC (Tarr τ1 τ2) (consC τ1 Γ)) α τ2 -∗ + Lemma compat_rec {S : Set} (Γ : S → ty) τ1 τ2 α : + ⊢ □ valid1 (Γ ▹ (Tarr τ1 τ2) ▹ τ1) α τ2 -∗ valid1 Γ (interp_rec rs α) (Tarr τ1 τ2). Proof. iIntros "#H". iIntros (σ ss) "Hs #Hss". - pose (env := (interp_ssubst ss)). fold env. - simp subst_expr. - pose (f := (ir_unf rs α env)). - iAssert (interp_rec rs α env ≡ IT_of_V $ FunV (Next f))%I as "Hf". + term_simpl. + pose (f := (ir_unf rs α ss)). + iAssert (interp_rec rs α ss ≡ IT_of_V $ FunV (Next f))%I as "Hf". { iPureIntro. apply interp_rec_unfold. } iRewrite "Hf". iApply expr_pred_ret. simpl. iExists _. iFrame. iModIntro. @@ -172,24 +174,34 @@ Section io_lang. iIntros (x) "Hx". iApply wp_lam. iNext. - pose (ss' := cons_ssubst (FunV (Next f)) (cons_ssubst βv ss)). + pose (ss' := (extend_scope (extend_scope ss (interp_rec rs α ss)) (IT_of_V βv))). iSpecialize ("H" $! _ ss' with "Hs []"). - { unfold ssubst_valid. - unfold ss'. - rewrite !ssubst_valid_cons. - by iFrame "IH Hw Hss". } - unfold f. simpl. - unfold ss'. simp interp_ssubst. - iAssert (IT_of_V (FunV (Next f)) ≡ interp_rec rs α env)%I as "Heq". + { + unfold ssubst_valid. + iIntros ([| [|]]); term_simpl. + - iModIntro; by iApply expr_pred_ret. + - iModIntro. + iRewrite "Hf". + iIntros (x') "Hx". + iApply wp_val. + iModIntro. + iExists x'. + iFrame "Hx". + iModIntro. + iApply "IH". + - iApply "Hss". + } + unfold f. + iAssert (IT_of_V (FunV (Next f)) ≡ interp_rec rs α ss)%I as "Heq". { rewrite interp_rec_unfold. done. } iRewrite -"Heq". by iApply "H". Qed. - Lemma compat_natop {S} (Γ : tyctx S) op α β : + Lemma compat_natop {S : Set} (Γ : S → ty) op α β : ⊢ valid1 Γ α Tnat -∗ valid1 Γ β Tnat -∗ valid1 Γ (interp_natop _ op α β) Tnat. - Proof using HCI. + Proof. iIntros "H1 H2". iIntros (σ ss) "Hs #Has". simpl. iSpecialize ("H2" with "Hs Has"). @@ -210,15 +222,14 @@ Section io_lang. eauto with iFrame. Qed. - Lemma fundamental {S} (Γ : tyctx S) e τ : + Lemma fundamental {S : Set} (Γ : S → ty) e τ : typed Γ e τ → ⊢ valid1 Γ (interp_expr rs e) τ - with fundamental_val {S} (Γ : tyctx S) v τ : + with fundamental_val {S : Set} (Γ : S → ty) v τ : typed_val Γ v τ → ⊢ valid1 Γ (interp_val rs v) τ. - Proof using HCI. + Proof. - destruct 1. + by iApply fundamental_val. - + by iApply compat_var. - + iApply compat_rec; iApply fundamental; eauto. + + subst. by iApply compat_var. + iApply compat_app; iApply fundamental; eauto. + iApply compat_natop; iApply fundamental; eauto. + iApply compat_if; iApply fundamental; eauto. @@ -228,77 +239,40 @@ Section io_lang. + iApply compat_nat. + iApply compat_rec; iApply fundamental; eauto. Qed. - Lemma fundmanetal_closed (e : expr []) (τ : ty) : - typed empC e τ → - ⊢ valid1 empC (interp_expr rs e) τ. - Proof using HCI. apply fundamental. Qed. + + Lemma fundmanetal_closed (e : expr ∅) (τ : ty) : + typed □ e τ → + ⊢ valid1 □ (interp_expr rs e) τ. + Proof. apply fundamental. Qed. End io_lang. Arguments interp_ty {_ _ _ _ _ _ _ _ _ _ _ _} τ. Arguments interp_tarr {_ _ _ _ _ _ _ _ _ _ _} Φ1 Φ2. -Local Definition rs : gReifiers _ := gReifiers_cons reify_io gReifiers_nil. - -Local Instance CtxIndepInputLang R `{!Cofe R} (o : opid (sReifier_ops (gReifiers_sReifier rs))) : - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) R) o. -Proof. - destruct o as [x o]. - inv_fin x. - - simpl. intros [[]| [[]| []]]. - + constructor. - unshelve eexists (λne '(_, (a, b)), SomeO (_, (_, b))). - * simpl in *. - apply ((update_input a).1). - * simpl in *. - apply ((update_input a).2). - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]. - simpl in *. - do 2 f_equiv. - -- do 2 f_equiv. - apply H. - -- f_equiv; last apply H. - do 2 f_equiv. - apply H. - * intros. - simpl. - destruct σ. - simpl. - reflexivity. - + constructor. - unshelve eexists (λne '(x, y), SomeO ((), _)). - * simpl in *. - apply ((update_output x (fstO y)), ()). - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]. - simpl in *. - do 4 f_equiv. - -- apply H. - -- apply H. - * intros. - simpl. - destruct σ as [σ1 []]; simpl in *. - reflexivity. - - intros i; by apply fin_0_inv. -Qed. +Local Definition rs : gReifiers NotCtxDep _ := gReifiers_cons NotCtxDep reify_io (gReifiers_nil NotCtxDep). Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. -Lemma logpred_adequacy cr Σ R `{!Cofe R, SubOfe natO R}`{!invGpreS Σ}`{!statePreG rs R Σ} τ (α : unitO -n> IT (gReifiers_ops rs) R) (β : IT (gReifiers_ops rs) R) st st' k - : +Require Import gitrees.gitree.greifiers. + +Program Definition ı_scope R `{!Cofe R} : @interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. + +Lemma logpred_adequacy cr Σ R `{!Cofe R, SubOfe natO R} + `{!invGpreS Σ} `{!statePreG rs R Σ} τ + (α : interp_scope ∅ -n> IT (gReifiers_ops NotCtxDep rs) R) + (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ}, - (£ cr ⊢ valid1 rs notStuck (λ _:unitO, True)%I empC α τ)%I) → - ssteps (gReifiers_sReifier rs) (α ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + (£ cr ⊢ valid1 rs notStuck (λne _ : unitO, True)%I □ α τ)%I) → + ssteps (gReifiers_sReifier NotCtxDep rs) (α (ı_scope _)) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. intros Hlog Hst. destruct (IT_to_V β) as [βv|] eqn:Hb. { right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } left. - cut ((∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) + cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ e, β ≡ Err e ∧ notStuck e)). { intros [?|He]; first done. destruct He as [? [? []]]. } @@ -306,15 +280,15 @@ Proof. { apply Hdisj. } { by rewrite Hb. } intros H1 H2. - exists (interp_ty (s:=notStuck) (P:=(λ _:unitO, True)) τ)%I. split. + exists (interp_ty (s:=notStuck) (P:=(λne _:unitO, True)) τ)%I. split. { apply _. } iIntros "[Hcr Hst]". iPoseProof (Hlog with "Hcr") as "Hlog". destruct st as [σ []]. iAssert (has_substate σ) with "[Hst]" as "Hs". { unfold has_substate, has_full_state. - assert (of_state rs (IT (gReifiers_ops rs) _) (σ,()) ≡ - of_idx rs (IT (gReifiers_ops rs) _) sR_idx (sR_state σ)) as ->; last done. + assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) (σ,()) ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ)) as ->; last done. intro j. unfold sR_idx. simpl. unfold of_state, of_idx. destruct decide as [Heq|]; last first. @@ -325,7 +299,10 @@ Proof. rewrite (eq_pi _ _ Heq eq_refl)//. } iSpecialize ("Hlog" $! σ with "Hs []"). - { iApply ssubst_valid_nil. } + { + iIntros (x). + destruct x. + } iSpecialize ("Hlog" $! tt with "[//]"). iApply (wp_wand with"Hlog"). iIntros ( βv). simpl. iDestruct 1 as (_) "[H _]". @@ -333,10 +310,10 @@ Proof. done. Qed. -Lemma io_lang_safety e τ σ st' (β : IT (sReifier_ops (gReifiers_sReifier rs)) natO) k : - typed empC e τ → - ssteps (gReifiers_sReifier rs) (interp_expr _ e ()) (σ,()) β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier rs) β st' β1 st1) +Lemma io_lang_safety e τ σ st' (β : IT (sReifier_ops NotCtxDep (gReifiers_sReifier NotCtxDep rs)) natO) k : + typed □ e τ → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e (ı_scope _)) (σ, ()) β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. intros Htyped Hsteps. diff --git a/theories/input_lang/logrel.v b/theories/input_lang/logrel.v index 169a7df..a1f0017 100644 --- a/theories/input_lang/logrel.v +++ b/theories/input_lang/logrel.v @@ -1,25 +1,21 @@ (** Logical relation for adequacy for the IO lang *) -From Equations Require Import Equations. -From gitrees Require Import gitree. +From gitrees Require Import gitree lang_generic. From gitrees.input_lang Require Import lang interp. +Require Import Binding.Lib Binding.Set Binding.Env. Section logrel. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers NotCtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops NotCtxDep rs). Notation IT := (IT F natO). Notation ITV := (ITV F natO). Context `{!invGS Σ, !stateG rs natO Σ}. Notation iProp := (iProp Σ). Notation restO := (gState_rest sR_idx rs ♯ IT). - Variable (HCi : ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) natO) o). Canonical Structure exprO S := leibnizO (expr S). Canonical Structure valO S := leibnizO (val S). - Local Notation tyctx := (tyctx ty). Notation "'WP' α {{ β , Φ } }" := (wp rs α notStuck ⊤ (λ β, Φ)) (at level 20, α, Φ at level 200, @@ -35,7 +31,7 @@ Section logrel. WP α {{ βv, ∃ m v σ', ⌜prim_steps e σ (Val v) σ' m⌝ ∗ V βv v ∗ has_substate σ' }})%I. Definition logrel_nat {S} (βv : ITV) (v : val S) : iProp := - (∃ n, βv ≡ RetV n ∧ ⌜v = Lit n⌝)%I. + (∃ n, βv ≡ RetV n ∧ ⌜v = LitV n⌝)%I. Definition logrel_arr {S} V1 V2 (βv : ITV) (vf : val S) : iProp := (∃ f, IT_of_V βv ≡ Fun f ∧ □ ∀ αv v, V1 αv v -∗ logrel_expr V2 (APP' (Fun f) (IT_of_V αv)) (App (Val vf) (Val v)))%I. @@ -76,7 +72,7 @@ Section logrel. (∀ v βv, logrel_val τ1 βv v -∗ logrel_expr V2 (f (IT_of_V βv)) (fill K (Val v))) -∗ logrel_expr V2 (f α) (fill K e). - Proof using HCi. + Proof. iIntros "H1 H2". iLöb as "IH" forall (α e). iIntros (σ) "Hs". @@ -102,8 +98,8 @@ Section logrel. by econstructor. Qed. - Lemma logrel_step_pure {S} (e' e : expr S) α V : - (∀ σ, prim_step e σ e' σ (0,0)) → + Lemma logrel_step_pure {S} n (e' e : expr S) α V : + (∀ σ, prim_step e σ e' σ (n,0)) → logrel_expr V α e' ⊢ logrel_expr V α e. Proof. intros Hpure. @@ -112,219 +108,167 @@ Section logrel. iSpecialize ("H" with "Hs"). iApply (wp_wand with "H"). iIntros (βv). iDestruct 1 as ([m m'] v σ' Hsteps) "[H2 Hs]". - iExists (m,m'),v,σ'. iFrame "H2 Hs". + iExists ((Nat.add n m),m'),v,σ'. iFrame "H2 Hs". iPureIntro. - eapply (prim_steps_app (0,0) (m,m')); eauto. + eapply (prim_steps_app (n,0) (m,m')); eauto. { eapply prim_step_steps, Hpure. } Qed. - (* a matching list of closing substitutions *) - Inductive subs2 : scope → Type := - | emp_subs2 : subs2 [] - | cons_subs2 {S} : val [] → ITV → subs2 S → subs2 (()::S) - . - - Equations subs_of_subs2 {S} (ss : subs2 S) : subs S [] := - subs_of_subs2 emp_subs2 v => idsub v; - subs_of_subs2 (cons_subs2 t α ss) Vz := Val t; - subs_of_subs2 (cons_subs2 t α ss) (Vs v) := subs_of_subs2 ss v. - - Equations its_of_subs2 {S} (ss : subs2 S) : interp_scope (E:=F) (R:=natO) S := - its_of_subs2 emp_subs2 := (); - its_of_subs2 (cons_subs2 t α ss) := (IT_of_V α, its_of_subs2 ss). - - Equations list_of_subs2 {S} (ss : subs2 S) : list (val []*ITV) := - list_of_subs2 emp_subs2 := []; - list_of_subs2 (cons_subs2 v α ss) := (v,α)::(list_of_subs2 ss). + Definition ssubst2_valid {S : Set} + (Γ : S -> ty) + (ss : @interp_scope F natO _ S) + (γ : S [⇒] Empty_set) : iProp := + (∀ x, □ logrel (Γ x) (ss x) (γ x))%I. - Lemma subs_of_emp_subs2 : subs_of_subs2 emp_subs2 ≡ idsub. - Proof. intros v. dependent elimination v. Qed. + Definition logrel_valid {S : Set} + (Γ : S -> ty) + (e : expr S) + (α : @interp_scope F natO _ S -n> IT) + (τ : ty) : iProp := + (□ ∀ (ss : @interp_scope F natO _ S) + (γ : S [⇒] Empty_set), + ssubst2_valid Γ ss γ → logrel τ (α ss) (bind γ e))%I. - Definition subs2_valid {S} (Γ : tyctx S) (ss : subs2 S) : iProp := - ([∗ list] τx ∈ zip (list_of_tyctx Γ) (list_of_subs2 ss), - logrel_val (τx.1) (τx.2.2) (τx.2.1))%I. - - Definition logrel_valid {S} (Γ : tyctx S) (e : expr S) (α : interp_scope S -n> IT) (τ : ty) : iProp := - (∀ ss, subs2_valid Γ ss → logrel τ - (α (its_of_subs2 ss)) - (subst_expr e (subs_of_subs2 ss)))%I. - - Lemma compat_var {S} (Γ : tyctx S) (x : var S) τ : - typed_var Γ x τ → ⊢ logrel_valid Γ (Var x) (interp_var x) τ. + Lemma compat_var {S : Set} (Γ : S → ty) (x : S): + ⊢ logrel_valid Γ (Var x) (interp_var x) (Γ x). Proof. - intros Hx. iIntros (ss) "Hss". - simp subst_expr. - iInduction Hx as [|Hx] "IH". - - dependent elimination ss. simp subs_of_subs2. - simp interp_var. rewrite /subs2_valid. - simp list_of_tyctx list_of_subs2 its_of_subs2. simpl. - iDestruct "Hss" as "[Hv Hss]". - iApply (logrel_of_val with "Hv"). - - dependent elimination ss. simp subs_of_subs2. - simp interp_var. rewrite /subs2_valid. - simp list_of_tyctx list_of_subs2 its_of_subs2. simpl. - iDestruct "Hss" as "[Hv Hss]". by iApply "IH". + iModIntro. iIntros (ss γ) "Hss". iApply "Hss". Qed. - Lemma compat_if {S} (Γ : tyctx S) (e0 e1 e2 : expr S) α0 α1 α2 τ : + Lemma compat_if {S : Set} (Γ : S → ty) (e0 e1 e2 : expr S) α0 α1 α2 τ : ⊢ logrel_valid Γ e0 α0 Tnat -∗ logrel_valid Γ e1 α1 τ -∗ logrel_valid Γ e2 α2 τ -∗ logrel_valid Γ (If e0 e1 e2) (interp_if rs α0 α1 α2) τ. - Proof using HCi. - iIntros "H0 H1 H2". iIntros (ss) "#Hss". - simpl. simp subst_expr. - pose (s := (subs_of_subs2 ss)). fold s. + Proof. + iIntros "#H0 #H1 #H2". + iModIntro. + iIntros (ss γ) "#Hss". + simpl. iSpecialize ("H0" with "Hss"). - iApply (logrel_bind (IFSCtx (α1 (its_of_subs2 ss)) (α2 (its_of_subs2 ss))) - [IfCtx (subst_expr e1 s) (subst_expr e2 s)] - with "H0"). + term_simpl. + iApply (@logrel_bind Empty_set + (IFSCtx (α1 ss) (α2 ss)) + (IfK EmptyK (bind γ e1) (bind γ e2)) _ (bind γ e0) (α0 ss) Tnat with "H0"). iIntros (v βv). iDestruct 1 as (n) "[Hb ->]". iRewrite "Hb". simpl. unfold IFSCtx. destruct (decide (0 < n)). - - rewrite IF_True//. + - rewrite IF_True; last done. iSpecialize ("H1" with "Hss"). iApply (logrel_step_pure with "H1"). - intros ?. apply (Ectx_step' []). + intros ?. apply (Ectx_step' EmptyK). econstructor; eauto. - rewrite IF_False; last lia. iSpecialize ("H2" with "Hss"). iApply (logrel_step_pure with "H2"). - intros ?. apply (Ectx_step' []). + intros ?. apply (Ectx_step' EmptyK). econstructor; eauto. lia. Qed. - Lemma compat_recV {S} Γ (e : expr (()::()::S)) τ1 τ2 α : - ⊢ □ logrel_valid (consC (Tarr τ1 τ2) (consC τ1 Γ)) e α τ2 -∗ + Lemma compat_recV {S : Set} (Γ : S -> ty) (e : expr (inc (inc S))) τ1 τ2 α : + ⊢ □ logrel_valid ((Γ ▹ (Tarr τ1 τ2) ▹ τ1)) e α τ2 -∗ logrel_valid Γ (Val $ RecV e) (interp_rec rs α) (Tarr τ1 τ2). Proof. - iIntros "#H". iIntros (ss) "#Hss". - pose (s := (subs_of_subs2 ss)). fold s. - pose (env := (its_of_subs2 ss)). fold env. - simp subst_expr. - pose (f := (ir_unf rs α env)). + iIntros "#H !> %env %γ #Henv". + set (f := (ir_unf rs α env)). iAssert (interp_rec rs α env ≡ IT_of_V $ FunV (Next f))%I as "Hf". { iPureIntro. apply interp_rec_unfold. } iRewrite "Hf". - iApply logrel_of_val. iLöb as "IH". iSimpl. - iExists (Next f). iSplit; eauto. + Opaque IT_of_V. + iApply logrel_of_val; term_simpl. + iExists _. iSplit. + { iPureIntro. apply into_val. } iModIntro. - iIntros (βv w) "#Hw". - iAssert ((APP' (Fun $ Next f) (IT_of_V βv)) ≡ (Tick (ir_unf rs α env (IT_of_V βv))))%I - as "Htick". - { iPureIntro. rewrite APP_APP'_ITV. - rewrite APP_Fun. simpl. done. } - iRewrite "Htick". iClear "Htick". + iLöb as "IH". + iIntros (αv v) "#Hw". + rewrite APP_APP'_ITV APP_Fun laterO_map_Next -Tick_eq. + pose (ss' := (extend_scope (extend_scope env (interp_rec rs α env)) (IT_of_V αv))). + set (γ' := ((mk_subst (Val (rec bind ((γ ↑) ↑)%bind e)%syn)) + ∘ ((mk_subst (shift (Val v))) ∘ ((γ ↑) ↑)))%bind). + rewrite /logrel. + iSpecialize ("H" $! ss' γ'). + set (γ1 := ((γ ↑) ↑)%bind). + iApply (logrel_step_pure 1 ((bind γ' e)%syn) with "[]"). + { + intros ?. eapply (Ectx_step' EmptyK). term_simpl. subst γ1 γ'. + rewrite -!bind_bind_comp'. + apply BetaS. + } + rewrite {2}/ss'. rewrite /f. iIntros (σ) "Hs". - iApply wp_tick. iNext. simpl. - pose (ss' := cons_subs2 (RecV (subst_expr e (subs_lift (subs_lift s)))) (FunV (Next (ir_unf rs α env))) (cons_subs2 w βv ss)). - iSpecialize ("H" $! ss' with "[Hss]"). - { rewrite {2}/subs2_valid /ss'. simp list_of_tyctx list_of_subs2. - cbn-[logrel_val]. iFrame "Hss Hw". fold f. iRewrite -"Hf". - by iApply "IH". } - iSpecialize ("H" with "Hs"). - iClear "IH Hss Hw". - unfold ss'. simpl. simp its_of_subs2. fold f env. - iRewrite "Hf". simpl. - iApply (wp_wand with "H"). - iIntros (v). - iDestruct 1 as ([m m'] v0 σ0 Hsteps) "[Hv Hs]". - iExists (1+m,0+m'),v0,σ0. iFrame "Hv Hs". - iPureIntro. econstructor; eauto. - apply (Ectx_step' []). - apply BetaS. - clear. - unfold subst2. - rewrite subst_expr_appsub. - apply subst_expr_proper. - intro v. - dependent elimination v. - { simp subs_of_subs2. unfold appsub. - simp subs_lift. simp subst_expr. - simp conssub. reflexivity. } - dependent elimination v. - { simp subs_of_subs2. unfold appsub. - simp subs_lift. unfold expr_lift. - simp ren_expr. simp subst_expr. - simp conssub. reflexivity. } - { simp subs_of_subs2. unfold appsub. - simp subs_lift. unfold expr_lift. - fold s. remember (s v) as e1. - rewrite ren_ren_expr. - rewrite subst_ren_expr. - trans (subst_expr e1 idsub). - - symmetry. apply subst_expr_idsub. - - apply subst_expr_proper. - intro v'. simpl. simp conssub. - reflexivity. } - Qed. - - Lemma compat_rec {S} Γ (e : expr (()::()::S)) τ1 τ2 α : - ⊢ □ logrel_valid (consC (Tarr τ1 τ2) (consC τ1 Γ)) e α τ2 -∗ - logrel_valid Γ (Rec e) (interp_rec rs α) (Tarr τ1 τ2). - Proof. - iIntros "#H". iIntros (ss) "#Hss". - pose (s := (subs_of_subs2 ss)). fold s. - pose (env := (its_of_subs2 ss)). fold env. - simp subst_expr. - iApply (logrel_step_pure (Val (RecV (subst_expr e (subs_lift (subs_lift s)))))). - { intros ?. eapply (Ectx_step' []). econstructor. } - iPoseProof (compat_recV with "H") as "H2". - iSpecialize ("H2" with "Hss"). - simp subst_expr. iApply "H2". + iApply wp_tick. iNext. + iApply "H"; eauto; iClear "H". + rewrite /ss' /γ'. + iIntros (x'); destruct x' as [| [| x']]; term_simpl; iModIntro. + * by iApply logrel_of_val. + * iRewrite "Hf". + iApply logrel_of_val. + simpl. + iExists (Next (ir_unf rs α env)). + iSplit; first done. + iModIntro. + iApply "IH". + * iApply "Henv". Qed. - Lemma compat_app {S} Γ (e1 e2 : expr S) τ1 τ2 α1 α2 : + Lemma compat_app {S : Set} (Γ : S → ty) (e1 e2 : expr S) τ1 τ2 α1 α2 : ⊢ logrel_valid Γ e1 α1 (Tarr τ1 τ2) -∗ logrel_valid Γ e2 α2 τ1 -∗ logrel_valid Γ (App e1 e2) (interp_app rs α1 α2) τ2. - Proof using HCi. - iIntros "H1 H2". iIntros (ss) "#Hss". + Proof. + iIntros "#H1 #H2". + iIntros (ss). + iModIntro. + iIntros (γ). + iIntros "#Hss". iSpecialize ("H1" with "Hss"). iSpecialize ("H2" with "Hss"). - pose (s := (subs_of_subs2 ss)). fold s. - pose (env := its_of_subs2 ss). fold env. - simp subst_expr. simpl. - iApply (logrel_bind (AppRSCtx (α1 env)) [AppRCtx (subst_expr e1 s)] with "H2"). - iIntros (v2 β2) "H2". iSimpl. - iApply (logrel_bind (AppLSCtx (IT_of_V β2)) [AppLCtx v2] with "H1"). - iIntros (v1 β1) "H1". simpl. - iDestruct "H1" as (f) "[Hα H1]". + unfold interp_app. simpl. - unfold AppLSCtx. iRewrite "Hα". (** XXX why doesn't simpl work here? *) - iApply ("H1" with "H2"). + assert ((bind γ (App e1 e2))%syn = (fill (AppRK (bind γ e1) EmptyK) (bind γ e2))) as ->. + { reflexivity. } + iApply (logrel_bind (AppRSCtx (α1 ss)) (AppRK (bind γ e1) EmptyK) with "H2"). + iIntros (v2 β2) "#H2'". iSimpl. + iApply (logrel_bind (AppLSCtx (IT_of_V β2)) (AppLK EmptyK v2) with "H1"). + iIntros (v1 β1) "#H1'". iSimpl. + iDestruct "H1'" as (f) "[Hα H1']". + simpl. + unfold AppLSCtx. iRewrite "Hα". + iApply ("H1'" with "H2'"). Qed. - Lemma compat_input {S} Γ : + Lemma compat_input {S : Set} (Γ : S → ty) : ⊢ logrel_valid Γ (Input : expr S) (interp_input rs) Tnat. Proof. - iIntros (ss) "Hss". + iModIntro. + iIntros (ss γ) "Hss". iIntros (σ) "Hs". destruct (update_input σ) as [n σ'] eqn:Hinp. iApply (wp_input with "Hs []"); first eauto. iNext. iIntros "Hlc Hs". iApply wp_val. - iExists (1,1),(Lit n),σ'. + iExists (1,1),(LitV n),σ'. iFrame "Hs". iModIntro. iSplit. { iPureIntro. - simp subst_expr. + term_simpl. apply prim_step_steps. - apply (Ectx_step' []). - by constructor. } + apply (Ectx_step' EmptyK). + by constructor. + } iExists n. eauto. Qed. - Lemma compat_output {S} Γ (e: expr S) α : + + Lemma compat_output {S : Set} (Γ : S → ty) (e: expr S) α : ⊢ logrel_valid Γ e α Tnat -∗ logrel_valid Γ (Output e) (interp_output rs α) Tnat. - Proof using HCi. - iIntros "H1". - iIntros (ss) "Hss". + Proof. + iIntros "#H1". + iModIntro. + iIntros (ss γ) "#Hss". iSpecialize ("H1" with "Hss"). - pose (s := (subs_of_subs2 ss)). fold s. - pose (env := its_of_subs2 ss). fold env. - simp subst_expr. simpl. - iApply (logrel_bind (get_ret _) [OutputCtx] with "H1"). + term_simpl. + iApply (logrel_bind (get_ret _) (OutputK EmptyK) with "H1"). iIntros (v βv). iDestruct 1 as (m) "[Hb ->]". iRewrite "Hb". simpl. @@ -332,55 +276,53 @@ Section logrel. rewrite get_ret_ret. iApply (wp_output with "Hs []"); first done. iNext. iIntros "Hlc Hs". - iExists (1,1),(Lit 0),_. + iExists (1,1),(LitV 0),_. iFrame "Hs". iSplit. { iPureIntro. apply prim_step_steps. - apply (Ectx_step' []). + apply (Ectx_step' EmptyK). by constructor. } iExists 0. eauto. Qed. - Lemma compat_natop {S} (Γ : tyctx S) e1 e2 α1 α2 op : + Lemma compat_natop {S : Set} (Γ : S → ty) e1 e2 α1 α2 op : ⊢ logrel_valid Γ e1 α1 Tnat -∗ logrel_valid Γ e2 α2 Tnat -∗ logrel_valid Γ (NatOp op e1 e2) (interp_natop rs op α1 α2) Tnat. - Proof using HCi. - iIntros "H1 H2". iIntros (ss) "#Hss". + Proof. + iIntros "#H1 #H2". iModIntro. iIntros (ss γ) "#Hss". iSpecialize ("H1" with "Hss"). iSpecialize ("H2" with "Hss"). - pose (s := (subs_of_subs2 ss)). fold s. - pose (env := its_of_subs2 ss). fold env. - simp subst_expr. simpl. - iApply (logrel_bind (NatOpRSCtx (do_natop op) (α1 env)) [NatOpRCtx op (subst_expr e1 s)] with "H2"). - iIntros (v2 β2) "H2". iSimpl. - iApply (logrel_bind (NatOpLSCtx (do_natop op) (IT_of_V β2)) [NatOpLCtx op v2] with "H1"). - iIntros (v1 β1) "H1". simpl. - iDestruct "H1" as (n1) "[Hn1 ->]". - iDestruct "H2" as (n2) "[Hn2 ->]". + term_simpl. + iApply (logrel_bind (NatOpRSCtx (do_natop op) (α1 ss)) (NatOpRK op (bind γ e1) EmptyK) with "H2"). + iIntros (v2 β2) "H2'". iSimpl. + iApply (logrel_bind (NatOpLSCtx (do_natop op) (IT_of_V β2)) (NatOpLK op EmptyK v2) with "H1"). + iIntros (v1 β1) "H1'". simpl. + iDestruct "H1'" as (n1) "[Hn1 ->]". + iDestruct "H2'" as (n2) "[Hn2 ->]". unfold NatOpLSCtx. iAssert ((NATOP (do_natop op) (IT_of_V β1) (IT_of_V β2)) ≡ Ret (do_natop op n1 n2))%I with "[Hn1 Hn2]" as "Hr". { iRewrite "Hn1". simpl. iRewrite "Hn2". simpl. iPureIntro. by rewrite NATOP_Ret. } - iApply (logrel_step_pure (Val (Lit (do_natop op n1 n2)))). - { intro. apply (Ectx_step' []). constructor. + iApply (logrel_step_pure _ (Val (LitV (do_natop op n1 n2)))). + { intro. apply (Ectx_step' EmptyK). constructor. destruct op; simpl; eauto. } iRewrite "Hr". iApply (logrel_of_val (RetV $ do_natop op n1 n2)). iExists _. iSplit; eauto. Qed. - Lemma fundamental {S} (Γ : tyctx S) τ e : + Lemma fundamental {S : Set} (Γ : S → ty) τ e : typed Γ e τ → ⊢ logrel_valid Γ e (interp_expr rs e) τ - with fundamental_val {S} (Γ : tyctx S) τ v : + with fundamental_val {S : Set} (Γ : S → ty) τ v : typed_val Γ v τ → ⊢ logrel_valid Γ (Val v) (interp_val rs v) τ. - Proof using HCi. + Proof. - induction 1; simpl. + by apply fundamental_val. - + by apply compat_var. - + iApply compat_rec. iApply IHtyped. + + subst. + by apply compat_var. + iApply compat_app. ++ iApply IHtyped1. ++ iApply IHtyped2. @@ -395,7 +337,7 @@ Section logrel. + iApply compat_output. iApply IHtyped. - induction 1; simpl. - + iIntros (ss) "Hss". simp subst_expr. simpl. + + iModIntro. iIntros (ss γ) "Hss". term_simpl. iApply (logrel_of_val (RetV n)). iExists n. eauto. + iApply compat_recV. by iApply fundamental. Qed. @@ -404,74 +346,32 @@ End logrel. Definition κ {S} {E} : ITV E natO → val S := λ x, match x with - | core.RetV n => Lit n - | _ => Lit 0 + | core.RetV n => LitV n + | _ => LitV 0 end. -Lemma κ_Ret {S} {E} n : κ ((RetV n) : ITV E natO) = (Lit n : val S). +Lemma κ_Ret {S} {E} n : κ ((RetV n) : ITV E natO) = (LitV n : val S). Proof. Transparent RetV. unfold RetV. simpl. done. Opaque RetV. Qed. -Definition rs : gReifiers 1 := gReifiers_cons reify_io gReifiers_nil. +Definition rs : gReifiers NotCtxDep 1 := gReifiers_cons NotCtxDep reify_io (gReifiers_nil NotCtxDep). -Local Instance CtxIndepInputLang R `{!Cofe R} (o : opid (sReifier_ops (gReifiers_sReifier rs))) : - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) R) o. -Proof. - destruct o as [x o]. - inv_fin x. - - simpl. intros [[]| [[]| []]]. - + constructor. - unshelve eexists (λne '(_, (a, b)), SomeO (_, (_, b))). - * simpl in *. - apply ((update_input a).1). - * simpl in *. - apply ((update_input a).2). - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]. - simpl in *. - do 2 f_equiv. - -- do 2 f_equiv. - apply H. - -- f_equiv; last apply H. - do 2 f_equiv. - apply H. - * intros. - simpl. - destruct σ. - simpl. - reflexivity. - + constructor. - unshelve eexists (λne '(x, y), SomeO ((), _)). - * simpl in *. - apply ((update_output x (fstO y)), ()). - * solve_proper_prepare. - destruct x as [? [? ?]]; destruct y as [? [? ?]]. - simpl in *. - do 4 f_equiv. - -- apply H. - -- apply H. - * intros. - simpl. - destruct σ as [σ1 []]; simpl in *. - reflexivity. - - intros i; by apply fin_0_inv. -Qed. +Require Import gitrees.gitree.greifiers. -Lemma logrel_nat_adequacy Σ `{!invGpreS Σ}`{!statePreG rs natO Σ} {S} (α : IT (gReifiers_ops rs) natO) (e : expr S) n σ σ' k : +Lemma logrel_nat_adequacy Σ `{!invGpreS Σ}`{!statePreG rs natO Σ} {S} (α : IT (gReifiers_ops NotCtxDep rs) natO) (e : expr S) n σ σ' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs natO Σ}, (True ⊢ logrel rs Tnat α e)%I) → - ssteps (gReifiers_sReifier rs) α (σ,()) (Ret n) σ' k → ∃ m σ', prim_steps e σ (Val $ Lit n) σ' m. + ssteps (gReifiers_sReifier NotCtxDep rs) α (σ,()) (Ret n) σ' k → ∃ m σ', prim_steps e σ (Val $ LitV n) σ' m. Proof. intros Hlog Hst. - pose (ϕ := λ (βv : ITV (gReifiers_ops rs) natO), + pose (ϕ := λ (βv : ITV (gReifiers_ops NotCtxDep rs) natO), ∃ m σ', prim_steps e σ (Val $ κ βv) σ' m). cut (ϕ (RetV n)). { destruct 1 as ( m' & σ2 & Hm). exists m', σ2. revert Hm. by rewrite κ_Ret. } eapply (wp_adequacy 0); eauto. intros Hinv1 Hst1. - pose (Φ := (λ (βv : ITV (gReifiers_ops rs) natO), ∃ n, logrel_val rs Tnat (Σ:=Σ) (S:=S) βv (Lit n) - ∗ ⌜∃ m σ', prim_steps e σ (Val $ Lit n) σ' m⌝)%I). + pose (Φ := (λ (βv : ITV (gReifiers_ops NotCtxDep rs) natO), ∃ n, logrel_val rs Tnat (Σ:=Σ) (S:=S) βv (LitV n) + ∗ ⌜∃ m σ', prim_steps e σ (Val $ LitV n) σ' m⌝)%I). assert (NonExpansive Φ). { unfold Φ. intros l a1 a2 Ha. repeat f_equiv. done. } @@ -492,8 +392,8 @@ Proof. iPoseProof (Hlog with "[//]") as "Hlog". iAssert (has_substate σ) with "[Hs]" as "Hs". { unfold has_substate, has_full_state. - assert (of_state rs (IT (gReifiers_ops rs) natO) (σ, ()) ≡ - of_idx rs (IT (gReifiers_ops rs) natO) sR_idx (sR_state σ)) as -> ; last done. + assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) natO) (σ, ()) ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) natO) 0 σ) as ->; last done. intro j. unfold sR_idx. simpl. unfold of_state, of_idx. destruct decide as [Heq|]; last first. @@ -512,23 +412,25 @@ Proof. iExists l. iSplit; eauto. Qed. -Theorem adequacy (e : expr []) (k : nat) σ σ' n : - typed empC e Tnat → - ssteps (gReifiers_sReifier rs) (interp_expr rs e ()) (σ,()) (Ret k : IT _ natO) σ' n → - ∃ mm σ', prim_steps e σ (Val $ Lit k) σ' mm. +Program Definition ı_scope : @interp_scope (gReifiers_ops NotCtxDep rs) natO _ Empty_set := λne (x : ∅), match x with end. + +Theorem adequacy (e : expr ∅) (k : nat) σ σ' n : + typed □ e Tnat → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ı_scope) (σ,()) (Ret k : IT _ natO) σ' n → + ∃ mm σ', prim_steps e σ (Val $ LitV k) σ' mm. Proof. intros Hty Hst. pose (Σ:=#[invΣ;stateΣ rs natO]). - eapply (logrel_nat_adequacy Σ (interp_expr rs e ())); last eassumption. + eapply (logrel_nat_adequacy Σ (interp_expr rs e ı_scope)); last eassumption. intros ? ?. iPoseProof (fundamental rs) as "H". { apply Hty. } unfold logrel_valid. iIntros "_". - iSpecialize ("H" $! (emp_subs2 rs)). - simp its_of_subs2. - rewrite subs_of_emp_subs2. - rewrite subst_expr_idsub. + unshelve iSpecialize ("H" $! ı_scope _ with "[]"). + { apply ı%bind. } + { iIntros (x); destruct x. } + rewrite ebind_id; first last. + { intros ?; reflexivity. } iApply "H". - unfold subs2_valid. done. Qed. diff --git a/theories/input_lang_callcc/hom.v b/theories/input_lang_callcc/hom.v index 7250497..16d7ad0 100644 --- a/theories/input_lang_callcc/hom.v +++ b/theories/input_lang_callcc/hom.v @@ -1,16 +1,15 @@ -From Equations Require Import Equations. From gitrees Require Import gitree. From gitrees.input_lang_callcc Require Import lang interp. -Require Import gitrees.lang_generic_sem. +Require Import gitrees.lang_generic. Require Import Binding.Lib Binding.Set Binding.Env. Open Scope stdpp_scope. Section hom. Context {sz : nat}. - Context {rs : gReifiers sz}. + Context {rs : gReifiers CtxDep sz}. Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops CtxDep rs). Notation IT := (IT F natO). Notation ITV := (ITV F natO). diff --git a/theories/input_lang_callcc/interp.v b/theories/input_lang_callcc/interp.v index 73ad941..ae97290 100644 --- a/theories/input_lang_callcc/interp.v +++ b/theories/input_lang_callcc/interp.v @@ -1,7 +1,6 @@ -From Equations Require Import Equations. From gitrees Require Import gitree. From gitrees.input_lang_callcc Require Import lang. -Require Import gitrees.lang_generic_sem. +Require Import gitrees.lang_generic. Require Import Binding.Lib. Require Import Binding.Set. @@ -84,7 +83,7 @@ Proof. repeat f_equiv; apply H0. Qed. -Canonical Structure reify_io : sReifier. +Canonical Structure reify_io : sReifier CtxDep. Proof. simple refine {| sReifier_ops := ioE; sReifier_state := stateO @@ -175,9 +174,9 @@ End constructors. Section weakestpre. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers CtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops CtxDep rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Notation IT := (IT F R). @@ -194,7 +193,7 @@ Section weakestpre. Proof. iIntros (Hσ) "Hs Ha". rewrite hom_INPUT. simpl. - iApply (wp_subreify with "Hs"). + iApply (wp_subreify_ctx_dep with "Hs"). + simpl. by rewrite Hσ. + by rewrite ofe_iso_21. + done. @@ -209,20 +208,6 @@ Section weakestpre. eapply (wp_input' σ σ' n k idfun). Qed. - (* Lemma wp_input (σ σ' : stateO) (n : nat) (k : natO -n> IT) Φ s : *) - (* update_input σ = (n, σ') → *) - (* has_substate σ -∗ *) - (* ▷ (£ 1 -∗ has_substate σ' -∗ WP@{rs} (k n) @ s {{ Φ }}) -∗ *) - (* WP@{rs} (INPUT k) @ s {{ Φ }}. *) - (* Proof. *) - (* intros Hs. iIntros "Hs Ha". *) - (* unfold INPUT. simpl. *) - (* iApply (wp_subreify with "Hs"). *) - (* { simpl. by rewrite Hs. } *) - (* { simpl. by rewrite ofe_iso_21. } *) - (* iModIntro. done. *) - (* Qed. *) - Lemma wp_output' (σ σ' : stateO) (n : nat) (κ : IT -n> IT) `{!IT_hom κ} Φ s : update_output n σ = σ' → @@ -232,13 +217,12 @@ Section weakestpre. Proof. iIntros (Hσ) "Hs Ha". rewrite /OUTPUT hom_OUTPUT_. - iApply (wp_subreify with "Hs"). + iApply (wp_subreify_ctx_dep with "Hs"). + simpl. by rewrite Hσ. + done. + done. Qed. - Lemma wp_output (σ σ' : stateO) (n : nat) Φ s : update_output n σ = σ' → has_substate σ -∗ @@ -259,7 +243,7 @@ Section weakestpre. Proof. iIntros "Hs Ha". rewrite /THROW. simpl. rewrite hom_vis. - iApply (wp_subreify with "Hs"); simpl; done. + iApply (wp_subreify_ctx_dep with "Hs"); simpl; done. Qed. Lemma wp_throw (σ : stateO) (f : laterO (IT -n> IT)) (x : IT) Φ s : @@ -278,7 +262,7 @@ Section weakestpre. iIntros "Hs Ha". unfold CALLCC. simpl. rewrite hom_vis. - iApply (wp_subreify _ _ _ _ _ _ _ ((later_map k ((f (laterO_map k))))) with "Hs"). + iApply (wp_subreify_ctx_dep _ _ _ _ _ _ _ ((later_map k ((f (laterO_map k))))) with "Hs"). { simpl. repeat f_equiv. @@ -301,11 +285,11 @@ End weakestpre. Section interp. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers CtxDep sz). Context {subR : subReifier reify_io rs}. Context {R} `{CR : !Cofe R}. Context `{!SubOfe natO R}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops CtxDep rs). Notation IT := (IT F R). Notation ITV := (ITV F R). @@ -946,26 +930,26 @@ Section interp. Opaque Ret. Lemma interp_expr_fill_yes_reify {S} K env (e e' : expr S) - (σ σ' : stateO) (σr : gState_rest sR_idx rs ♯ IT) n : + (σ σ' : stateO) (σr : gState_rest CtxDep sR_idx rs ♯ IT) n : head_step e σ e' σ' K (n, 1) → - reify (gReifiers_sReifier rs) - (interp_expr (fill K e) env) (gState_recomp σr (sR_state σ)) - ≡ (gState_recomp σr (sR_state σ'), Tick_n n $ interp_expr (fill K e') env). + reify (gReifiers_sReifier CtxDep rs) + (interp_expr (fill K e) env) (gState_recomp CtxDep σr (sR_state σ)) + ≡ (gState_recomp CtxDep σr (sR_state σ'), Tick_n n $ interp_expr (fill K e') env). Proof. intros Hst. - trans (reify (gReifiers_sReifier rs) (interp_ectx K env (interp_expr e env)) - (gState_recomp σr (sR_state σ))). + trans (reify (gReifiers_sReifier CtxDep rs) (interp_ectx K env (interp_expr e env)) + (gState_recomp CtxDep σr (sR_state σ))). { f_equiv. by rewrite interp_comp. } inversion Hst; simplify_eq; cbn-[gState_recomp]. - - trans (reify (gReifiers_sReifier rs) (INPUT (interp_ectx K env ◎ Ret)) (gState_recomp σr (sR_state σ))). + - trans (reify (gReifiers_sReifier CtxDep rs) (INPUT (interp_ectx K env ◎ Ret)) (gState_recomp CtxDep σr (sR_state σ))). { repeat f_equiv; eauto. rewrite hom_INPUT. do 2 f_equiv. by intro. } - rewrite reify_vis_eq //; first last. + rewrite reify_vis_eq_ctx_dep //; first last. { - epose proof (@subReifier_reify sz reify_io rs _ IT _ (inl ()) () (Next (interp_ectx K env (Ret n0))) (NextO ◎ (interp_ectx K env ◎ Ret)) σ σ' σr) as H. + epose proof (@subReifier_reify sz CtxDep reify_io rs _ IT _ (inl ()) () (Next (interp_ectx K env (Ret n0))) (NextO ◎ (interp_ectx K env ◎ Ret)) σ σ' σr) as H. simpl in H. simpl. erewrite <-H; last first. @@ -976,19 +960,19 @@ Section interp. repeat f_equiv. rewrite Tick_eq/=. repeat f_equiv. rewrite interp_comp. reflexivity. - - trans (reify (gReifiers_sReifier rs) (interp_ectx K env (OUTPUT n0)) (gState_recomp σr (sR_state σ))). + - trans (reify (gReifiers_sReifier CtxDep rs) (interp_ectx K env (OUTPUT n0)) (gState_recomp CtxDep σr (sR_state σ))). { do 3 f_equiv; eauto. rewrite get_ret_ret//. } - trans (reify (gReifiers_sReifier rs) (OUTPUT_ n0 (interp_ectx K env (Ret 0))) (gState_recomp σr (sR_state σ))). + trans (reify (gReifiers_sReifier CtxDep rs) (OUTPUT_ n0 (interp_ectx K env (Ret 0))) (gState_recomp CtxDep σr (sR_state σ))). { do 2 f_equiv; eauto. by rewrite hom_OUTPUT_. } - rewrite reify_vis_eq //; last first. + rewrite reify_vis_eq_ctx_dep //; last first. { - epose proof (@subReifier_reify sz reify_io rs _ IT _ (inr (inl ())) n0 (Next (interp_ectx K env ((Ret 0)))) (constO (Next (interp_ectx K env ((Ret 0))))) σ (update_output n0 σ) σr) as H. + epose proof (@subReifier_reify sz CtxDep reify_io rs _ IT _ (inr (inl ())) n0 (Next (interp_ectx K env ((Ret 0)))) (constO (Next (interp_ectx K env ((Ret 0))))) σ (update_output n0 σ) σr) as H. simpl in H. simpl. erewrite <-H; last reflexivity. @@ -1007,17 +991,17 @@ Section interp. Transparent CALLCC. unfold CALLCC. simpl. - set (subEff1 := @subReifier_subEff sz reify_io rs subR). - trans (reify (gReifiers_sReifier rs) (CALLCC_ f (laterO_map (interp_ectx K env))) gσ). + set (subEff1 := @subReifier_subEff sz CtxDep reify_io rs subR). + trans (reify (gReifiers_sReifier CtxDep rs) (CALLCC_ f (laterO_map (interp_ectx K env))) gσ). { do 2 f_equiv. rewrite hom_CALLCC_. f_equiv. by intro. } - rewrite reify_vis_eq//; last first. + rewrite reify_vis_eq_ctx_dep//; last first. { simpl. - epose proof (@subReifier_reify sz reify_io rs subR IT _ + epose proof (@subReifier_reify sz CtxDep reify_io rs subR IT _ (inr (inr (inl ()))) f _ (laterO_map (interp_ectx K env)) σ' σ' σr) as H. simpl in H. @@ -1037,11 +1021,11 @@ Section interp. do 2 f_equiv. by intro. Qed. - Lemma soundness {S} (e1 e2 : expr S) σ1 σ2 (σr : gState_rest sR_idx rs ♯ IT) n m (env : interp_scope S) : + Lemma soundness {S} (e1 e2 : expr S) σ1 σ2 (σr : gState_rest CtxDep sR_idx rs ♯ IT) n m (env : interp_scope S) : prim_step e1 σ1 e2 σ2 (n,m) → - ssteps (gReifiers_sReifier rs) - (interp_expr e1 env) (gState_recomp σr (sR_state σ1)) - (interp_expr e2 env) (gState_recomp σr (sR_state σ2)) n. + ssteps (gReifiers_sReifier CtxDep rs) + (interp_expr e1 env) (gState_recomp CtxDep σr (sR_state σ1)) + (interp_expr e2 env) (gState_recomp CtxDep σr (sR_state σ2)) n. Proof. Opaque gState_decomp gState_recomp. inversion 1; simplify_eq/=. @@ -1113,7 +1097,7 @@ Section interp. match goal with | |- context G [ofe_mor_car _ _ _ (Next ?f)] => set (f' := f) end. - trans (reify (gReifiers_sReifier rs) (THROW (interp_val v env) (Next f')) (gState_recomp σr (sR_state σ2))). + trans (reify (gReifiers_sReifier CtxDep rs) (THROW (interp_val v env) (Next f')) (gState_recomp CtxDep σr (sR_state σ2))). { f_equiv; last done. f_equiv. @@ -1125,12 +1109,12 @@ Section interp. intros x; simpl. destruct ((subEff_outs ^-1) x). } - rewrite reify_vis_eq; first (rewrite Tick_eq; reflexivity). + rewrite reify_vis_eq_ctx_dep; first (rewrite Tick_eq; reflexivity). simpl. match goal with | |- context G [(_, _, ?a)] => set (κ := a) end. - epose proof (@subReifier_reify sz reify_io rs subR IT _ + epose proof (@subReifier_reify sz CtxDep reify_io rs subR IT _ (inr (inr (inr (inl ())))) (Next (interp_val v env), Next f') (Next (Tau (Next ((interp_ectx K' env) (interp_val v env))))) (Empty_setO_rec _) σ2 σ2 σr) as H'. diff --git a/theories/input_lang_callcc/lang.v b/theories/input_lang_callcc/lang.v index d917ecb..e7cb712 100644 --- a/theories/input_lang_callcc/lang.v +++ b/theories/input_lang_callcc/lang.v @@ -1,6 +1,4 @@ -From stdpp Require Export strings. From gitrees Require Export prelude. -From Equations Require Import Equations. Require Import List. Import ListNotations. diff --git a/theories/input_lang_callcc/logrel.v b/theories/input_lang_callcc/logrel.v index 1adfe97..b88fac4 100644 --- a/theories/input_lang_callcc/logrel.v +++ b/theories/input_lang_callcc/logrel.v @@ -1,17 +1,16 @@ (** Logical relation for adequacy for the IO lang *) -From Equations Require Import Equations. From gitrees Require Import gitree. From gitrees.input_lang_callcc Require Import lang interp hom. -Require Import gitrees.lang_generic_sem. +Require Import gitrees.lang_generic. Require Import Binding.Lib Binding.Set Binding.Env. Open Scope stdpp_scope. Section logrel. Context {sz : nat}. - Variable (rs : gReifiers sz). + Variable (rs : gReifiers CtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops CtxDep rs). Notation IT := (IT F natO). Notation ITV := (ITV F natO). Context `{!invGS Σ, !stateG rs natO Σ}. @@ -273,7 +272,7 @@ Section logrel. iApply "HK'". simpl. unfold logrel_arr. - _iExists (Next (ir_unf rs α env)). + iExists (Next (ir_unf rs α env)). iSplit; first done. iModIntro. iApply "IH". @@ -685,19 +684,19 @@ Lemma κ_Ret {S} {E} n : κ ((RetV n) : ITV E natO) = (LitV n : val S). Proof. Transparent RetV. unfold RetV. simpl. done. Opaque RetV. Qed. -Definition rs : gReifiers 1 := gReifiers_cons reify_io gReifiers_nil. +Definition rs : gReifiers CtxDep 1 := gReifiers_cons CtxDep reify_io (gReifiers_nil CtxDep). Require Import gitrees.gitree.greifiers. Lemma logrel_nat_adequacy Σ `{!invGpreS Σ} `{!statePreG rs natO Σ} {S} - (α : IT (gReifiers_ops rs) natO) + (α : IT (gReifiers_ops CtxDep rs) natO) (e : expr S) n σ σ' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs natO Σ}, (⊢ logrel rs Tnat α e)%I) → - ssteps (gReifiers_sReifier rs) α (σ, ()) (Ret n) σ' k → + ssteps (gReifiers_sReifier CtxDep rs) α (σ, ()) (Ret n) σ' k → ∃ m σ', prim_steps e σ (Val $ LitV n) σ' m. Proof. intros Hlog Hst. - pose (ϕ := λ (βv : ITV (gReifiers_ops rs) natO), + pose (ϕ := λ (βv : ITV (gReifiers_ops CtxDep rs) natO), ∃ m σ', prim_steps e σ (Val $ κ βv) σ' m). cut (ϕ (RetV n)). { @@ -706,7 +705,7 @@ Proof. } eapply (wp_adequacy 0); eauto. intros Hinv1 Hst1. - pose (Φ := (λ (βv : ITV (gReifiers_ops rs) natO), + pose (Φ := (λ (βv : ITV (gReifiers_ops CtxDep rs) natO), ∃ n, logrel_val rs Tnat (Σ:=Σ) (S:=S) βv (LitV n) ∗ ⌜∃ m σ', prim_steps e σ (Val $ LitV n) σ' m⌝)%I). assert (NonExpansive Φ). @@ -732,8 +731,8 @@ Proof. iAssert (has_substate σ) with "[Hs]" as "Hs". { unfold has_substate, has_full_state. - assert ((of_state rs (IT (sReifier_ops (gReifiers_sReifier rs)) natO) (σ, ())) ≡ - (of_idx rs (IT (sReifier_ops (gReifiers_sReifier rs)) natO) sR_idx (sR_state σ))) + assert ((of_state CtxDep rs (IT (sReifier_ops CtxDep (gReifiers_sReifier CtxDep rs)) natO) (σ, ())) ≡ + (of_idx CtxDep rs (IT (sReifier_ops CtxDep (gReifiers_sReifier CtxDep rs)) natO) sR_idx (sR_state σ))) as -> ; last done. intros j. unfold sR_idx. simpl. unfold of_state, of_idx. @@ -766,11 +765,11 @@ Proof. iExists l. iSplit; eauto. Qed. -Program Definition ı_scope : @interp_scope (gReifiers_ops rs) natO _ Empty_set := λne (x : ∅), match x with end. +Program Definition ı_scope : @interp_scope (gReifiers_ops CtxDep rs) natO _ Empty_set := λne (x : ∅), match x with end. Theorem adequacy (e : expr ∅) (k : nat) σ σ' n : typed □ e Tnat → - ssteps (gReifiers_sReifier rs) (interp_expr rs e ı_scope) (σ, ()) (Ret k : IT _ natO) σ' n → + ssteps (gReifiers_sReifier CtxDep rs) (interp_expr rs e ı_scope) (σ, ()) (Ret k : IT _ natO) σ' n → ∃ mm σ', prim_steps e σ (Val $ LitV k) σ' mm. Proof. intros Hty Hst. diff --git a/theories/lang_affine.v b/theories/lang_affine.v new file mode 100644 index 0000000..b13c0be --- /dev/null +++ b/theories/lang_affine.v @@ -0,0 +1,245 @@ +From gitrees Require Import prelude. +From gitrees Require Import gitree. +From Equations Require Import Equations. +Require Import List. +Import ListNotations. + +(** XXX: We /NEED/ this line for [Equations Derive] to work, *) +(* this flag is globally unset by std++, but Equations need obligations to be transparent. *) +Set Transparent Obligations. + +Derive NoConfusion NoConfusionHom for list. + +Definition scope := (list unit). + +(** Variables in a context *) +Inductive var : scope → Type := +| Vz : forall {S : scope} {s}, var (s::S) +| Vs : forall {S : scope} {s}, var S -> var (s::S) +. +Derive Signature NoConfusion for var. + +Inductive tyctx (ty : Type) : scope → Type := +| empC : tyctx ty [] +| consC : forall{Γ}, ty → tyctx ty Γ → tyctx ty (()::Γ) +. +Arguments empC {_}. +Arguments consC {_ _} _ _. + +Equations list_of_tyctx {S ty} (Γ : tyctx ty S) : list ty := + list_of_tyctx empC := []; + list_of_tyctx (consC τ Γ') := τ::list_of_tyctx Γ'. + +Equations tyctx_app {S1 S2 ty} (c1 : tyctx ty S1) (c2 : tyctx ty S2) : tyctx ty (S1++S2) := + tyctx_app empC c2 := c2; + tyctx_app (consC τ c1) c2 := consC τ (tyctx_app c1 c2). + +Inductive typed_var {ty : Type}: forall {S}, tyctx ty S → var S → ty → Prop := +| typed_var_Z S (τ : ty) (Γ : tyctx ty S) : + typed_var (consC τ Γ) Vz τ +| typed_var_S S (τ τ' : ty) (Γ : tyctx ty S) v : + typed_var Γ v τ → + typed_var (consC τ' Γ) (Vs v) τ +. + +Section interp. + Local Open Scope type. + Context {E: opsInterp}. + Context {R} `{!Cofe R}. + Notation IT := (IT E R). + Notation ITV := (ITV E R). + + Fixpoint interp_scope (S : scope) : ofe := + match S with + | [] => unitO + | τ::Sc => prodO IT (interp_scope Sc) + end. + + Instance interp_scope_cofe S : Cofe (interp_scope S). + Proof. induction S; simpl; apply _. Qed. + + Instance interp_scope_inhab S : Inhabited (interp_scope S). + Proof. induction S; simpl; apply _. Defined. + + Equations interp_var {S : scope} (v : var S) : interp_scope S -n> IT := + interp_var (S:=(_::_)) Vz := fstO; + interp_var (S:=(_::Sc)) (Vs v) := interp_var v ◎ sndO. + + Instance interp_var_ne S (v : var S) : NonExpansive (@interp_var S v). + Proof. + intros n D1 D2 HD12. induction v; simp interp_var. + - by f_equiv. + - eapply IHv. by f_equiv. + Qed. + + Global Instance interp_var_proper S (v : var S) : Proper ((≡) ==> (≡)) (interp_var v). + Proof. apply ne_proper. apply _. Qed. + + Definition interp_scope_split {S1 S2} : + interp_scope (S1 ++ S2) -n> interp_scope S1 * interp_scope S2. + Proof. + induction S1 as [|? S1]; simpl. + - simple refine (λne x, (tt, x)). + solve_proper. + - simple refine (λne xy, let ss := IHS1 xy.2 in ((xy.1, ss.1), ss.2)). + solve_proper. + Defined. + + (** scope substituions *) + Inductive ssubst : scope → Type := + | emp_ssubst : ssubst [] + | cons_ssubst {S} : ITV → ssubst S → ssubst (tt::S) + . + + Equations interp_ssubst {S} (ss : ssubst S) : interp_scope S := + interp_ssubst emp_ssubst := tt; + interp_ssubst (cons_ssubst αv ss) := (IT_of_V αv, interp_ssubst ss). + + Equations list_of_ssubst {S} (ss : ssubst S) : list ITV := + list_of_ssubst emp_ssubst := []; + list_of_ssubst (cons_ssubst αv ss) := αv::(list_of_ssubst ss). + + Equations ssubst_split {S1 S2} (αs : ssubst (S1++S2)) : ssubst S1 * ssubst S2 := + ssubst_split (S1:=[]) αs := (emp_ssubst,αs); + ssubst_split (S1:=u::_) (cons_ssubst αv αs) := + (cons_ssubst αv (ssubst_split αs).1, (ssubst_split αs).2). + Lemma interp_scope_ssubst_split {S1 S2} (αs : ssubst (S1++S2)) : + interp_scope_split (interp_ssubst αs) ≡ + (interp_ssubst (ssubst_split αs).1, interp_ssubst (ssubst_split αs).2). + Proof. + induction S1 as [|u S1]; simpl. + - simp ssubst_split. simpl. + simp interp_ssubst. done. + - dependent elimination αs as [cons_ssubst αv αs]. + simp ssubst_split. simpl. + simp interp_ssubst. repeat f_equiv; eauto; simpl. + + rewrite IHS1//. + + rewrite IHS1//. + Qed. + +End interp. + +(* Common definitions and lemmas for Kripke logical relations *) +Section kripke_logrel. + Variable s : stuckness. + + Context {sz : nat} {a : is_ctx_dep}. + Variable rs : gReifiers a sz. + Context {R} `{!Cofe R}. + + Notation F := (gReifiers_ops a rs). + Notation IT := (IT F R). + Notation ITV := (ITV F R). + Context `{!invGS Σ, !stateG rs R Σ}. + Notation iProp := (iProp Σ). + + Context {A:ofe}. (* The type & predicate for the explicit Kripke worlds *) + Variable (P : A -n> iProp). + + Implicit Types α β : IT. + Implicit Types αv βv : ITV. + Implicit Types Φ Ψ : ITV -n> iProp. + + Program Definition expr_pred (α : IT) (Φ : ITV -n> iProp) : iProp := + (∀ x : A, P x -∗ WP@{rs} α @ s {{ v, ∃ y : A, Φ v ∗ P y }}). + #[export] Instance expr_pred_ne : NonExpansive2 expr_pred. + Proof. solve_proper. Qed. + #[export] Instance expr_pred_proper : Proper ((≡) ==> (≡) ==> (≡)) expr_pred . + Proof. solve_proper. Qed. + + Definition ssubst_valid {ty} (interp_ty : ty → ITV -n> iProp) {S} (Γ : tyctx ty S) (ss : ssubst S) : iProp := + ([∗ list] τx ∈ zip (list_of_tyctx Γ) (list_of_ssubst (E:=F) ss), + interp_ty (τx.1) (τx.2))%I. + + Lemma ssubst_valid_nil {ty} (interp_ty : ty → ITV -n> iProp) : + ⊢ ssubst_valid interp_ty empC emp_ssubst. + Proof. + unfold ssubst_valid. + by simp list_of_tyctx list_of_ssubst. + Qed. + + Lemma ssubst_valid_cons {ty} (interp_ty : ty → ITV -n> iProp) {S} + (Γ : tyctx ty S) (ss : ssubst S) τ αv : + ssubst_valid interp_ty (consC τ Γ) (cons_ssubst αv ss) + ⊣⊢ interp_ty τ αv ∗ ssubst_valid interp_ty Γ ss. + Proof. + unfold ssubst_valid. + by simp list_of_tyctx list_of_ssubst. + Qed. + + Lemma ssubst_valid_app {ty} (interp_ty : ty → ITV -n> iProp) + {S1 S2} (Ω1 : tyctx ty S1) (Ω2 : tyctx ty S2) αs : + ssubst_valid interp_ty (tyctx_app Ω1 Ω2) αs ⊢ + ssubst_valid interp_ty Ω1 (ssubst_split αs).1 + ∗ ssubst_valid interp_ty Ω2 (ssubst_split αs).2. + Proof. + iInduction Ω1 as [|τ Ω1] "IH" forall (Ω2); simp tyctx_app ssubst_split. + - simpl. iIntros "$". iApply ssubst_valid_nil. + - iIntros "H". + rewrite {4 5}/ssubst_valid. + simpl in αs. + dependent elimination αs as [cons_ssubst αv αs]. + simp ssubst_split. simpl. + simp list_of_ssubst list_of_tyctx. + simpl. iDestruct "H" as "[$ H]". + by iApply "IH". + Qed. + + Lemma expr_pred_ret α αv Φ `{!IntoVal α αv} : + Φ αv ⊢ expr_pred α Φ. + Proof. + iIntros "H". + iIntros (x) "Hx". iApply wp_val. + eauto with iFrame. + Qed. + + Lemma expr_pred_frame α Φ : + WP@{rs} α @ s {{ Φ }} ⊢ expr_pred α Φ. + Proof. + iIntros "H". + iIntros (x) "Hx". + iApply (wp_wand with "H"). + eauto with iFrame. + Qed. +End kripke_logrel. + +Section kripke_logrel_ctx_indep. + Variable s : stuckness. + + Context {sz : nat}. + Variable rs : gReifiers NotCtxDep sz. + Context {R} `{!Cofe R}. + + Notation F := (gReifiers_ops NotCtxDep rs). + Notation IT := (IT F R). + Notation ITV := (ITV F R). + Context `{!invGS Σ, !stateG rs R Σ}. + Notation iProp := (iProp Σ). + + Context {A : ofe}. + Variable (P : A -n> iProp). + + Implicit Types α β : IT. + Implicit Types αv βv : ITV. + Implicit Types Φ Ψ : ITV -n> iProp. + + Local Notation expr_pred := (expr_pred s rs P). + + Lemma expr_pred_bind f `{!IT_hom f} α Φ Ψ `{!NonExpansive Φ} + : expr_pred α Ψ ⊢ + (∀ αv, Ψ αv -∗ expr_pred (f (IT_of_V αv)) Φ) + -∗ expr_pred (f α) Φ. + Proof. + iIntros "H1 H2". + iIntros (x) "Hx". + iApply wp_bind. + { solve_proper. } + iSpecialize ("H1" with "Hx"). + iApply (wp_wand with "H1"). + iIntros (βv). iDestruct 1 as (y) "[Hb Hy]". + iModIntro. + iApply ("H2" with "Hb Hy"). + Qed. +End kripke_logrel_ctx_indep. + +Arguments expr_pred_bind {_ _ _ _ _ _ _ _ _ _} f {_ _}. diff --git a/theories/lang_generic.v b/theories/lang_generic.v index 7c27639..144c061 100644 --- a/theories/lang_generic.v +++ b/theories/lang_generic.v @@ -1,46 +1,9 @@ From gitrees Require Import prelude. From gitrees Require Import gitree. -From Equations Require Import Equations. Require Import List. Import ListNotations. -(** XXX: We /NEED/ this line for [Equations Derive] to work, - this flag is globally unset by std++, but Equations need obligations to be transparent. *) -Set Transparent Obligations. - -Derive NoConfusion NoConfusionHom for list. - -Definition scope := (list unit). - -(** Variables in a context *) -Inductive var : scope → Type := -| Vz : forall {S : scope} {s}, var (s::S) -| Vs : forall {S : scope} {s}, var S -> var (s::S) -. -Derive Signature NoConfusion for var. - -Inductive tyctx (ty : Type) : scope → Type := -| empC : tyctx ty [] -| consC : forall{Γ}, ty → tyctx ty Γ → tyctx ty (()::Γ) -. -Arguments empC {_}. -Arguments consC {_ _} _ _. - -Equations list_of_tyctx {S ty} (Γ : tyctx ty S) : list ty := - list_of_tyctx empC := []; - list_of_tyctx (consC τ Γ') := τ::list_of_tyctx Γ'. - -Equations tyctx_app {S1 S2 ty} (c1 : tyctx ty S1) (c2 : tyctx ty S2) : tyctx ty (S1++S2) := - tyctx_app empC c2 := c2; - tyctx_app (consC τ c1) c2 := consC τ (tyctx_app c1 c2). - -Inductive typed_var {ty : Type}: forall {S}, tyctx ty S → var S → ty → Prop := -| typed_var_Z S (τ : ty) (Γ : tyctx ty S) : - typed_var (consC τ Γ) Vz τ -| typed_var_S S (τ τ' : ty) (Γ : tyctx ty S) v : - typed_var Γ v τ → - typed_var (consC τ' Γ) (Vs v) τ -. +Require Import Binding.Lib Binding.Set. Section interp. Local Open Scope type. @@ -49,141 +12,81 @@ Section interp. Notation IT := (IT E R). Notation ITV := (ITV E R). - Fixpoint interp_scope (S : scope) : ofe := - match S with - | [] => unitO - | τ::Sc => prodO IT (interp_scope Sc) - end. + Definition interp_scope (S : Set) : ofe := (leibnizO S) -n> IT. - Instance interp_scope_cofe S : Cofe (interp_scope S). - Proof. induction S; simpl; apply _. Qed. + Global Instance interp_scope_cofe S : Cofe (interp_scope S). + Proof. apply _. Qed. - Instance interp_scope_inhab S : Inhabited (interp_scope S). - Proof. induction S; simpl; apply _. Defined. + Global Instance interp_scope_inhab S : Inhabited (interp_scope S). + Proof. apply _. Defined. - Equations interp_var {S : scope} (v : var S) : interp_scope S -n> IT := - interp_var (S:=(_::_)) Vz := fstO; - interp_var (S:=(_::Sc)) (Vs v) := interp_var v ◎ sndO. - - Instance interp_var_ne S (v : var S) : NonExpansive (@interp_var S v). - Proof. - intros n D1 D2 HD12. induction v; simp interp_var. - - by f_equiv. - - eapply IHv. by f_equiv. + Program Definition interp_var {S : Set} (v : S) : interp_scope S -n> IT := + λne (f : interp_scope S), f v. + Next Obligation. + solve_proper. Qed. - Global Instance interp_var_proper S (v : var S) : Proper ((≡) ==> (≡)) (interp_var v). + Global Instance interp_var_proper {S : Set} (v : S) : Proper ((≡) ==> (≡)) (interp_var v). Proof. apply ne_proper. apply _. Qed. - Definition interp_scope_split {S1 S2} : - interp_scope (S1 ++ S2) -n> interp_scope S1 * interp_scope S2. - Proof. - induction S1 as [|? S1]; simpl. - - simple refine (λne x, (tt, x)). - solve_proper. - - simple refine (λne xy, let ss := IHS1 xy.2 in ((xy.1, ss.1), ss.2)). - solve_proper. - Defined. - - (** scope substituions *) - Inductive ssubst : scope → Type := - | emp_ssubst : ssubst [] - | cons_ssubst {S} : ITV → ssubst S → ssubst (tt::S) - . - - Equations interp_ssubst {S} (ss : ssubst S) : interp_scope S := - interp_ssubst emp_ssubst := tt; - interp_ssubst (cons_ssubst αv ss) := (IT_of_V αv, interp_ssubst ss). - - Equations list_of_ssubst {S} (ss : ssubst S) : list ITV := - list_of_ssubst emp_ssubst := []; - list_of_ssubst (cons_ssubst αv ss) := αv::(list_of_ssubst ss). - - Equations ssubst_split {S1 S2} (αs : ssubst (S1++S2)) : ssubst S1 * ssubst S2 := - ssubst_split (S1:=[]) αs := (emp_ssubst,αs); - ssubst_split (S1:=u::_) (cons_ssubst αv αs) := - (cons_ssubst αv (ssubst_split αs).1, (ssubst_split αs).2). - Lemma interp_scope_ssubst_split {S1 S2} (αs : ssubst (S1++S2)) : - interp_scope_split (interp_ssubst αs) ≡ - (interp_ssubst (ssubst_split αs).1, interp_ssubst (ssubst_split αs).2). - Proof. - induction S1 as [|u S1]; simpl. - - simp ssubst_split. simpl. - simp interp_ssubst. done. - - dependent elimination αs as [cons_ssubst αv αs]. - simp ssubst_split. simpl. - simp interp_ssubst. repeat f_equiv; eauto; simpl. - + rewrite IHS1//. - + rewrite IHS1//. + Program Definition extend_scope {S : Set} : interp_scope S -n> IT -n> interp_scope (inc S) + := λne γ μ x, let x' : inc S := x in + match x' with + | VZ => μ + | VS x'' => γ x'' + end. + Next Obligation. + intros ???? [| x] [| y]; term_simpl; [solve_proper | inversion 1 | inversion 1 | inversion 1; by subst]. + Qed. + Next Obligation. + intros ??????. + intros [| a]; term_simpl; solve_proper. + Qed. + Next Obligation. + intros ??????. + intros [| a]; term_simpl; solve_proper. Qed. + Program Definition ren_scope {S S'} (δ : S [→] S') (env : interp_scope S') + : interp_scope S := λne x, env (δ x). + End interp. (* Common definitions and lemmas for Kripke logical relations *) Section kripke_logrel. Variable s : stuckness. - Context {sz : nat}. - Variable rs : gReifiers sz. + Context {sz : nat} {a : is_ctx_dep}. + Variable rs : gReifiers a sz. Context {R} `{!Cofe R}. - Notation F := (gReifiers_ops rs). + Notation F := (gReifiers_ops a rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ}. Notation iProp := (iProp Σ). Context {A:ofe}. (* The type & predicate for the explicit Kripke worlds *) - Variable (P : A → iProp). - Context `{!NonExpansive P}. + Variable (P : A -n> iProp). Implicit Types α β : IT. Implicit Types αv βv : ITV. Implicit Types Φ Ψ : ITV -n> iProp. Program Definition expr_pred (α : IT) (Φ : ITV -n> iProp) : iProp := - (∀ x : A, P x -∗ WP@{rs} α @ s {{ v, ∃ y : A, Φ v ∗ P y }}). - #[export] Instance expr_pred_ne : NonExpansive2 expr_pred. - Proof. solve_proper. Qed. - #[export] Instance expr_pred_proper : Proper ((≡) ==> (≡) ==> (≡)) expr_pred . - Proof. solve_proper. Qed. - - Definition ssubst_valid {ty} (interp_ty : ty → ITV -n> iProp) {S} (Γ : tyctx ty S) (ss : ssubst S) : iProp := - ([∗ list] τx ∈ zip (list_of_tyctx Γ) (list_of_ssubst (E:=F) ss), - interp_ty (τx.1) (τx.2))%I. - - Lemma ssubst_valid_nil {ty} (interp_ty : ty → ITV -n> iProp) : - ⊢ ssubst_valid interp_ty empC emp_ssubst. - Proof. - unfold ssubst_valid. - by simp list_of_tyctx list_of_ssubst. + (∀ x : A, P x -∗ wp rs α s ⊤ (λne v, ∃ y : A, Φ v ∗ P y)). + Next Obligation. + solve_proper. Qed. - Lemma ssubst_valid_cons {ty} (interp_ty : ty → ITV -n> iProp) {S} - (Γ : tyctx ty S) (ss : ssubst S) τ αv : - ssubst_valid interp_ty (consC τ Γ) (cons_ssubst αv ss) - ⊣⊢ interp_ty τ αv ∗ ssubst_valid interp_ty Γ ss. + Global Instance expr_pred_ne {n} : Proper (dist n ==> dist n ==> dist n) expr_pred. Proof. - unfold ssubst_valid. - by simp list_of_tyctx list_of_ssubst. + solve_proper. Qed. - Lemma ssubst_valid_app {ty} (interp_ty : ty → ITV -n> iProp) - {S1 S2} (Ω1 : tyctx ty S1) (Ω2 : tyctx ty S2) αs : - ssubst_valid interp_ty (tyctx_app Ω1 Ω2) αs ⊢ - ssubst_valid interp_ty Ω1 (ssubst_split αs).1 - ∗ ssubst_valid interp_ty Ω2 (ssubst_split αs).2. + Global Instance expr_pred_proper : Proper (equiv ==> equiv ==> equiv) expr_pred. Proof. - iInduction Ω1 as [|τ Ω1] "IH" forall (Ω2); simp tyctx_app ssubst_split. - - simpl. iIntros "$". iApply ssubst_valid_nil. - - iIntros "H". - rewrite {4 5}/ssubst_valid. - simpl in αs. - dependent elimination αs as [cons_ssubst αv αs]. - simp ssubst_split. simpl. - simp list_of_ssubst list_of_tyctx. - simpl. iDestruct "H" as "[$ H]". - by iApply "IH". + solve_proper. Qed. Lemma expr_pred_ret α αv Φ `{!IntoVal α αv} : @@ -191,35 +94,61 @@ Section kripke_logrel. Proof. iIntros "H". iIntros (x) "Hx". iApply wp_val. - eauto with iFrame. + simpl. + iExists x. + by iFrame. Qed. + Lemma expr_pred_frame α Φ : + WP@{rs} α @ s {{ Φ }} ⊢ expr_pred α Φ. + Proof. + iIntros "H". + iIntros (x) "Hx". + iApply (wp_wand with "H"). + simpl. + iIntros (v) "Hv". + iExists x. + by iFrame. + Qed. + +End kripke_logrel. + +Section kripke_logrel_ctx_indep. + Variable s : stuckness. + + Context {sz : nat}. + Variable rs : gReifiers NotCtxDep sz. + Context {R} `{!Cofe R}. + + Notation F := (gReifiers_ops NotCtxDep rs). + Notation IT := (IT F R). + Notation ITV := (ITV F R). + Context `{!invGS Σ, !stateG rs R Σ}. + Notation iProp := (iProp Σ). + + Context {A : ofe}. + Variable (P : A -n> iProp). + + Implicit Types α β : IT. + Implicit Types αv βv : ITV. + Implicit Types Φ Ψ : ITV -n> iProp. + + Local Notation expr_pred := (expr_pred s rs P). + Lemma expr_pred_bind f `{!IT_hom f} α Φ Ψ `{!NonExpansive Φ} - {G : ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) IT o} : - expr_pred α Ψ ⊢ - (∀ αv, Ψ αv -∗ expr_pred (f (IT_of_V αv)) Φ) -∗ - expr_pred (f α) Φ. + : expr_pred α Ψ ⊢ + (∀ αv, Ψ αv -∗ expr_pred (f (IT_of_V αv)) Φ) + -∗ expr_pred (f α) Φ. Proof. iIntros "H1 H2". iIntros (x) "Hx". iApply wp_bind. - { solve_proper. } iSpecialize ("H1" with "Hx"). iApply (wp_wand with "H1"). iIntros (βv). iDestruct 1 as (y) "[Hb Hy]". iModIntro. iApply ("H2" with "Hb Hy"). Qed. - - Lemma expr_pred_frame α Φ : - WP@{rs} α @ s {{ Φ }} ⊢ expr_pred α Φ. - Proof. - iIntros "H". - iIntros (x) "Hx". - iApply (wp_wand with "H"). - eauto with iFrame. - Qed. -End kripke_logrel. +End kripke_logrel_ctx_indep. Arguments expr_pred_bind {_ _ _ _ _ _ _ _ _ _} f {_ _}. diff --git a/theories/lang_generic_sem.v b/theories/lang_generic_sem.v deleted file mode 100644 index 21816f1..0000000 --- a/theories/lang_generic_sem.v +++ /dev/null @@ -1,104 +0,0 @@ -From gitrees Require Import prelude. -From gitrees Require Import gitree. -Require Import List. -Import ListNotations. - -Require Import Binding.Lib Binding.Set. - -Section interp. - Local Open Scope type. - Context {E: opsInterp}. - Context {R} `{!Cofe R}. - Notation IT := (IT E R). - Notation ITV := (ITV E R). - - Definition interp_scope (S : Set) : ofe := (leibnizO S) -n> IT. - - Global Instance interp_scope_cofe S : Cofe (interp_scope S). - Proof. apply _. Qed. - - Global Instance interp_scope_inhab S : Inhabited (interp_scope S). - Proof. apply _. Defined. - - Program Definition interp_var {S : Set} (v : S) : interp_scope S -n> IT := - λne (f : interp_scope S), f v. - Next Obligation. - solve_proper. - Qed. - - Global Instance interp_var_proper {S : Set} (v : S) : Proper ((≡) ==> (≡)) (interp_var v). - Proof. apply ne_proper. apply _. Qed. - - Program Definition extend_scope {S : Set} : interp_scope S -n> IT -n> interp_scope (inc S) - := λne γ μ x, let x' : inc S := x in - match x' with - | VZ => μ - | VS x'' => γ x'' - end. - Next Obligation. - intros ???? [| x] [| y]; term_simpl; [solve_proper | inversion 1 | inversion 1 | inversion 1; by subst]. - Qed. - Next Obligation. - intros ??????. - intros [| a]; term_simpl; solve_proper. - Qed. - Next Obligation. - intros ??????. - intros [| a]; term_simpl; solve_proper. - Qed. - - Program Definition ren_scope {S S'} (δ : S [→] S') (env : interp_scope S') - : interp_scope S := λne x, env (δ x). - -End interp. - -(* Common definitions and lemmas for Kripke logical relations *) -Section kripke_logrel. - Variable s : stuckness. - - Context {sz : nat}. - Variable rs : gReifiers sz. - Context {R} `{!Cofe R}. - - Notation F := (gReifiers_ops rs). - Notation IT := (IT F R). - Notation ITV := (ITV F R). - Context `{!invGS Σ, !stateG rs R Σ}. - Notation iProp := (iProp Σ). - - Context {A:ofe}. (* The type & predicate for the explicit Kripke worlds *) - Variable (P : A -n> iProp). - - Implicit Types α β : IT. - Implicit Types αv βv : ITV. - Implicit Types Φ Ψ : ITV -n> iProp. - - Program Definition expr_pred (α : IT) (Φ : ITV -n> iProp) : iProp := - (∀ x : A, P x -∗ wp rs α s ⊤ (λne v, ∃ y : A, Φ v ∗ P y)). - Next Obligation. - solve_proper. - Qed. - - Lemma expr_pred_ret α αv Φ `{!IntoVal α αv} : - Φ αv ⊢ expr_pred α Φ. - Proof. - iIntros "H". - iIntros (x) "Hx". iApply wp_val. - simpl. - iExists x. - by iFrame. - Qed. - - Lemma expr_pred_frame α Φ : - WP@{rs} α @ s {{ Φ }} ⊢ expr_pred α Φ. - Proof. - iIntros "H". - iIntros (x) "Hx". - iApply (wp_wand with "H"). - simpl. - iIntros (v) "Hv". - iExists x. - by iFrame. - Qed. - -End kripke_logrel. diff --git a/theories/program_logic.v b/theories/program_logic.v index eb86191..d4d514f 100644 --- a/theories/program_logic.v +++ b/theories/program_logic.v @@ -2,9 +2,9 @@ From gitrees Require Import gitree. Section program_logic. - Context {sz : nat}. - Variable rs : gReifiers sz. - Notation F := (gReifiers_ops rs). + Context {sz : nat} {a : is_ctx_dep}. + Variable rs : gReifiers a sz. + Notation F := (gReifiers_ops a rs). Context {R} `{!Cofe R}. Notation IT := (IT F R). Notation ITV := (ITV F R). @@ -26,21 +26,18 @@ End program_logic. Section program_logic_ctx_indep. Context {sz : nat}. - Variable rs : gReifiers sz. - Notation F := (gReifiers_ops rs). + Variable rs : gReifiers NotCtxDep sz. + Notation F := (gReifiers_ops NotCtxDep rs). Context {R} `{!Cofe R}. Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ}. Notation iProp := (iProp Σ). - Context {HCI : ∀ o : opid (sReifier_ops (gReifiers_sReifier rs)), - CtxIndep (gReifiers_sReifier rs) - (ITF_solution.IT (sReifier_ops (gReifiers_sReifier rs)) R) o}. Lemma wp_seq α β s Φ `{!NonExpansive Φ} : WP@{rs} α @ s {{ _, WP@{rs} β @ s {{ Φ }} }} ⊢ WP@{rs} SEQ α β @ s {{ Φ }}. - Proof using HCI. + Proof. iIntros "H". iApply (wp_bind _ (SEQCtx β)). iApply (wp_wand with "H"). @@ -50,7 +47,7 @@ Section program_logic_ctx_indep. Lemma wp_let α (f : IT -n> IT) s Φ `{!NonExpansive Φ} : WP@{rs} α @ s {{ αv, WP@{rs} f (IT_of_V αv) @ s {{ Φ }} }} ⊢ WP@{rs} (LET α f) @ s {{ Φ }}. - Proof using HCI. + Proof. iIntros "H". iApply (wp_bind _ (LETCTX f)). iApply (wp_wand with "H"). From d5cbc7756e021530cc54b6cc7afe3f163cf9ee31 Mon Sep 17 00:00:00 2001 From: Kaptch Date: Tue, 30 Jan 2024 23:11:17 +0100 Subject: [PATCH 3/7] coercion --- theories/gitree/reify.v | 20 ++++++++++++++++++++ theories/lang_generic.v | 13 +++++++++++++ 2 files changed, 33 insertions(+) diff --git a/theories/gitree/reify.v b/theories/gitree/reify.v index 49e5deb..4ba8243 100644 --- a/theories/gitree/reify.v +++ b/theories/gitree/reify.v @@ -31,6 +31,26 @@ Section reifier. }. End reifier. +Section reifier_coercion. + Context {A} `{!Cofe A}. + #[local] Open Scope type. + Program Definition sReifier_NotCtxDep_CtxDep (r : sReifier NotCtxDep) + : sReifier CtxDep := + {| + sReifier_ops := sReifier_ops _ r; + sReifier_state := sReifier_state _ r; + sReifier_re x xc op := + (λne y, (optionO_map (prodO_map y.2 idfun) + (sReifier_re _ r op (y.1.1, y.1.2)))); + sReifier_inhab := sReifier_inhab _ r; + sReifier_cofe := sReifier_cofe _ r; + |}. + Next Obligation. + intros. + repeat intro; repeat f_equiv; assumption. + Qed. +End reifier_coercion. + Section reifier_cofe_inst. Context {A} `{!Cofe A}. #[local] Open Scope type. diff --git a/theories/lang_generic.v b/theories/lang_generic.v index 144c061..a4a3663 100644 --- a/theories/lang_generic.v +++ b/theories/lang_generic.v @@ -26,6 +26,19 @@ Section interp. solve_proper. Qed. + Definition interp_scope_split {S1 S2 : Set} : + interp_scope (sum S1 S2) -n> interp_scope S1 * interp_scope S2. + Proof. + simple refine (λne (f : interp_scope (sum S1 S2)), _). + - split. + + simple refine (λne x, _). + apply (f (inl x)). + + simple refine (λne x, _). + apply (f (inr x)). + - repeat intro; simpl. + repeat f_equiv; intro; simpl; f_equiv; assumption. + Defined. + Global Instance interp_var_proper {S : Set} (v : S) : Proper ((≡) ==> (≡)) (interp_var v). Proof. apply ne_proper. apply _. Qed. From 3c34ea1e8ebcbc195243dc297f3d46e4eefb224b Mon Sep 17 00:00:00 2001 From: Kaptch Date: Thu, 1 Feb 2024 00:28:03 +0100 Subject: [PATCH 4/7] refactoring (affine) --- _CoqProject | 40 +- theories/affine_lang/logrel1.v | 493 ---------- theories/{examples => effects}/store.v | 1 - theories/{ => examples}/affine_lang/lang.v | 113 ++- theories/examples/affine_lang/logrel1.v | 931 ++++++++++++++++++ theories/{ => examples}/affine_lang/logrel2.v | 182 ++-- theories/{ => examples}/input_lang/interp.v | 2 +- theories/{ => examples}/input_lang/lang.v | 0 theories/{ => examples}/input_lang/logpred.v | 2 +- theories/{ => examples}/input_lang/logrel.v | 2 +- .../{ => examples}/input_lang_callcc/hom.v | 2 +- .../{ => examples}/input_lang_callcc/interp.v | 2 +- .../{ => examples}/input_lang_callcc/lang.v | 0 .../{ => examples}/input_lang_callcc/logrel.v | 2 +- theories/lang_affine.v | 245 ----- theories/{examples => lib}/factorial.v | 6 +- theories/{examples => lib}/iter.v | 0 theories/{examples => lib}/pairs.v | 0 theories/{examples => lib}/while.v | 0 19 files changed, 1128 insertions(+), 895 deletions(-) delete mode 100644 theories/affine_lang/logrel1.v rename theories/{examples => effects}/store.v (99%) rename theories/{ => examples}/affine_lang/lang.v (69%) create mode 100644 theories/examples/affine_lang/logrel1.v rename theories/{ => examples}/affine_lang/logrel2.v (77%) rename theories/{ => examples}/input_lang/interp.v (99%) rename theories/{ => examples}/input_lang/lang.v (100%) rename theories/{ => examples}/input_lang/logpred.v (99%) rename theories/{ => examples}/input_lang/logrel.v (99%) rename theories/{ => examples}/input_lang_callcc/hom.v (98%) rename theories/{ => examples}/input_lang_callcc/interp.v (99%) rename theories/{ => examples}/input_lang_callcc/lang.v (100%) rename theories/{ => examples}/input_lang_callcc/logrel.v (99%) delete mode 100644 theories/lang_affine.v rename theories/{examples => lib}/factorial.v (97%) rename theories/{examples => lib}/iter.v (100%) rename theories/{examples => lib}/pairs.v (100%) rename theories/{examples => lib}/while.v (100%) diff --git a/_CoqProject b/_CoqProject index e8cf1c7..4ef1b32 100644 --- a/_CoqProject +++ b/_CoqProject @@ -16,7 +16,6 @@ vendor/Binding/Resolver.v theories/prelude.v theories/lang_generic.v -theories/lang_affine.v theories/gitree/core.v theories/gitree/subofe.v @@ -29,22 +28,23 @@ theories/gitree.v theories/program_logic.v -theories/input_lang_callcc/lang.v -theories/input_lang_callcc/interp.v -theories/input_lang_callcc/hom.v -theories/input_lang_callcc/logrel.v - -theories/input_lang/lang.v -theories/input_lang/interp.v -theories/input_lang/logpred.v -theories/input_lang/logrel.v - -theories/affine_lang/lang.v -theories/affine_lang/logrel1.v -theories/affine_lang/logrel2.v - -theories/examples/store.v -theories/examples/pairs.v -theories/examples/while.v -theories/examples/factorial.v -theories/examples/iter.v +theories/examples/input_lang_callcc/lang.v +theories/examples/input_lang_callcc/interp.v +theories/examples/input_lang_callcc/hom.v +theories/examples/input_lang_callcc/logrel.v + +theories/examples/input_lang/lang.v +theories/examples/input_lang/interp.v +theories/examples/input_lang/logpred.v +theories/examples/input_lang/logrel.v + +theories/examples/affine_lang/lang.v +theories/examples/affine_lang/logrel1.v +theories/examples/affine_lang/logrel2.v + +theories/effects/store.v + +theories/lib/pairs.v +theories/lib/while.v +theories/lib/factorial.v +theories/lib/iter.v diff --git a/theories/affine_lang/logrel1.v b/theories/affine_lang/logrel1.v deleted file mode 100644 index b5757b3..0000000 --- a/theories/affine_lang/logrel1.v +++ /dev/null @@ -1,493 +0,0 @@ -(** Unary (Kripke) logical relation for the affine lang *) -From Equations Require Import Equations. -From gitrees Require Export lang_affine gitree program_logic. -From gitrees.affine_lang Require Import lang. -From gitrees.examples Require Import store pairs. -Require Import iris.algebra.gmap. - -Local Notation tyctx := (tyctx ty). - -Inductive typed : forall {S}, tyctx S → expr S → ty → Prop := -(** functions *) -| typed_Var {S} (Ω : tyctx S) (τ : ty) (v : var S) : - typed_var Ω v τ → - typed Ω (Var v) τ -| typed_Lam {S} (Ω : tyctx S) (τ1 τ2 : ty) (e : expr (()::S) ) : - typed (consC τ1 Ω) e τ2 → - typed Ω (Lam e) (tArr τ1 τ2) -| typed_App {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 : ty) e1 e2 : - typed Ω1 e1 (tArr τ1 τ2) → - typed Ω2 e2 τ1 → - typed (tyctx_app Ω1 Ω2) (App e1 e2) τ2 -(** pairs *) -| typed_Pair {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 : ty) e1 e2 : - typed Ω1 e1 τ1 → - typed Ω2 e2 τ2 → - typed (tyctx_app Ω1 Ω2) (EPair e1 e2) (tPair τ1 τ2) -| typed_Destruct {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 τ : ty) - (e1 : expr S1) (e2 : expr (()::()::S2)) : - typed Ω1 e1 (tPair τ1 τ2) → - typed (consC τ1 (consC τ2 Ω2)) e2 τ → - typed (tyctx_app Ω1 Ω2) (EDestruct e1 e2) τ -(** references *) -| typed_Alloc {S} (Ω : tyctx S) τ e : - typed Ω e τ → - typed Ω (Alloc e) (tRef τ) -| typed_Replace {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 : ty) e1 e2 : - typed Ω1 e1 (tRef τ1) → - typed Ω2 e2 τ2 → - typed (tyctx_app Ω1 Ω2) (Replace e1 e2) (tPair τ1 (tRef τ2)) -| typed_Dealloc {S} (Ω : tyctx S) e τ : - typed Ω e (tRef τ) → - typed Ω (Dealloc e) tUnit -(** literals *) -| typed_Nat {S} (Ω : tyctx S) n : - typed Ω (LitNat n) tInt -| typed_Bool {S} (Ω : tyctx S) b : - typed Ω (LitBool b) tBool -| typed_Unit {S} (Ω : tyctx S) : - typed Ω LitUnit tUnit -. - -Section logrel. - Context {sz : nat}. - Variable rs : gReifiers NotCtxDep sz. - Context `{!subReifier reify_store rs}. - Context `{!subReifier input_lang.interp.reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). - Context {R} `{!Cofe R}. - Context `{!SubOfe natO R}. - Context `{!SubOfe unitO R}. - Context `{!SubOfe locO R}. - Notation IT := (IT F R). - Notation ITV := (ITV F R). - Context `{!invGS Σ, !stateG rs R Σ, !heapG rs R Σ}. - Notation iProp := (iProp Σ). - - (* parameters for the kripke logical relation *) - Variable s : stuckness. - Context `{A:ofe}. - Variable (P : A -n> iProp). - Local Notation expr_pred := (expr_pred s rs P). - - (* interpreting tys *) - Program Definition protected (Φ : ITV -n> iProp) : ITV -n> iProp := λne αv, - (WP@{rs} Force (IT_of_V αv) @ s {{ Φ }})%I. - Solve All Obligations with solve_proper_please. - Program Definition interp_tbool : ITV -n> iProp := λne αv, - (αv ≡ RetV 0 ∨ αv ≡ RetV 1)%I. - Solve All Obligations with solve_proper_please. - Program Definition interp_tnat : ITV -n> iProp := λne αv, - (∃ n : nat, αv ≡ RetV n)%I. - Solve All Obligations with solve_proper_please. - Program Definition interp_tunit : ITV -n> iProp := λne αv, - (αv ≡ RetV ())%I. - Solve All Obligations with solve_proper_please. - Program Definition interp_tpair (Φ1 Φ2 : ITV -n> iProp) : ITV -n> iProp := λne αv, - (∃ β1v β2v, IT_of_V αv ≡ pairITV (IT_of_V β1v) (IT_of_V β2v) ∗ - Φ1 β1v ∗ Φ2 β2v)%I. - Solve All Obligations with solve_proper_please. - Program Definition interp_tarr (Φ1 Φ2 : ITV -n> iProp) : ITV -n> iProp := λne αv, - (∀ βv, Φ1 βv -∗ expr_pred ((IT_of_V αv) ⊙ (Thunk (IT_of_V βv))) Φ2)%I. - Solve All Obligations with solve_proper_please. - - Program Definition interp_ref (Φ : ITV -n> iProp) : ITV -n> iProp := λne αv, - (∃ (l : loc) βv, αv ≡ RetV l ∗ pointsto l (IT_of_V βv) ∗ Φ βv)%I. - Solve All Obligations with solve_proper_please. - - Fixpoint interp_ty (τ : ty) : ITV -n> iProp := - match τ with - | tBool => interp_tbool - | tUnit => interp_tunit - | tInt => interp_tnat - | tPair τ1 τ2 => interp_tpair (interp_ty τ1) (interp_ty τ2) - | tArr τ1 τ2 => interp_tarr (interp_ty τ1) (interp_ty τ2) - | tRef τ => interp_ref (interp_ty τ) - end. - - Definition ssubst_valid {S} (Ω : tyctx S) ss := - lang_affine.ssubst_valid rs (λ τ, protected (interp_ty τ)) Ω ss. - - Definition valid1 {S} (Ω : tyctx S) (α : interp_scope S -n> IT) (τ : ty) : iProp := - ∀ ss, heap_ctx -∗ ssubst_valid Ω ss -∗ expr_pred (α (interp_ssubst ss)) (interp_ty τ). - - Lemma compat_pair {S1 S2} (Ω1: tyctx S1) (Ω2:tyctx S2) α β τ1 τ2 : - ⊢ valid1 Ω1 α τ1 -∗ - valid1 Ω2 β τ2 -∗ - valid1 (tyctx_app Ω1 Ω2) (interp_pair α β ◎ interp_scope_split) (tPair τ1 τ2). - Proof. - Opaque pairITV. - iIntros "H1 H2". - iIntros (αs) "#Hctx Has". - cbn-[interp_pair]. - unfold ssubst_valid. - rewrite ssubst_valid_app. - rewrite interp_scope_ssubst_split. - iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. - iSpecialize ("H1" with "Hctx Ha1"). - iSpecialize ("H2" with "Hctx Ha2"). - iApply (expr_pred_bind with "H2"). - iIntros (βv) "Hb". simpl. - rewrite -> get_val_ITV. simpl. - iApply (expr_pred_bind with "H1"). - iIntros (αv) "Ha". simpl. - rewrite -> get_val_ITV. simpl. - iApply expr_pred_ret. - iExists _,_. iFrame. done. - Qed. - - Lemma compat_destruct {S1 S2} (Ω1: tyctx S1) (Ω2:tyctx S2) α β τ1 τ2 τ : - ⊢ valid1 Ω1 α (tPair τ1 τ2) -∗ - valid1 (consC τ1 $ consC τ2 Ω2) β τ -∗ - valid1 (tyctx_app Ω1 Ω2) (interp_destruct α β ◎ interp_scope_split) τ. - Proof. - Opaque pairITV thunked thunkedV projIT1 projIT2. - iIntros "H1 H2". - iIntros (αs) "#Hctx Has". - cbn-[interp_destruct]. - unfold ssubst_valid. - rewrite ssubst_valid_app. - rewrite interp_scope_ssubst_split. - iDestruct "Has" as "[Ha1 Ha2]". - iSpecialize ("H1" with "Hctx Ha1"). - simpl. iApply (expr_pred_bind (LETCTX _) with "H1"). - iIntros (αv) "Ha". unfold LETCTX. simpl. - rewrite LET_Val/=. - iDestruct "Ha" as (β1 β2) "[#Ha [Hb1 Hb2]]". - iIntros (x) "Hx". - iApply wp_let. - { solve_proper. } - iApply (wp_Thunk with "Hctx"). - { solve_proper_please. } - iNext. iIntros (l1) "Hl1". simpl. - iApply wp_let. - { solve_proper. } - iApply (wp_Thunk with "Hctx"). - { solve_proper_please. } - iNext. iIntros (l2) "Hl2". simpl. - iSpecialize ("H2" $! (cons_ssubst (thunkedV (projIT1 (IT_of_V αv)) l1) - $ cons_ssubst (thunkedV (projIT2 (IT_of_V αv)) l2) (ssubst_split αs).2) - with "Hctx [-Hx] Hx"). - { unfold ssubst_valid. rewrite !ssubst_valid_cons. - iFrame. Transparent thunkedV thunked. - iSplitL "Hb1 Hl1". - - simpl. iApply wp_lam. simpl. iNext. - iApply (wp_bind _ (IFSCtx _ _)). - iApply (wp_read with "Hctx Hl1"). - iNext. iNext. iIntros "Hl1". - iApply wp_val. iModIntro. unfold IFSCtx. - rewrite IF_False; last lia. - iApply wp_seq. - { solve_proper. } - iApply (wp_write with "Hctx Hl1"). - iNext. iNext. iIntros "Hl1". - iRewrite "Ha". - rewrite projIT1_pairV. simpl. - repeat iApply wp_tick. - repeat iNext. iApply wp_val. done. - - simpl. iApply wp_lam. simpl. iNext. - iApply (wp_bind _ (IFSCtx _ _)). - iApply (wp_read with "Hctx Hl2"). - iNext. iNext. iIntros "Hl2". - iApply wp_val. iModIntro. unfold IFSCtx. - rewrite IF_False; last lia. - iApply wp_seq. - { solve_proper. } - iApply (wp_write with "Hctx Hl2"). - iNext. iNext. iIntros "Hl2". - iRewrite "Ha". - rewrite projIT2_pairV. simpl. - repeat iApply wp_tick. - repeat iNext. iApply wp_val. done. - } - simp interp_ssubst. - iApply "H2". - Qed. - - Lemma compat_alloc {S} (Ω : tyctx S) α τ: - ⊢ valid1 Ω α τ -∗ - valid1 Ω (interp_alloc α) (tRef τ). - Proof. - iIntros "H". - iIntros (αs) "#Hctx Has". - iSpecialize ("H" with "Hctx Has"). - simpl. iApply (expr_pred_bind (LETCTX _) with "H"). - iIntros (αv) "Hav". unfold LETCTX. simpl. - rewrite LET_Val/=. - iApply expr_pred_frame. - iApply (wp_alloc with "Hctx"). - iNext. iNext. iIntros (l) "Hl". - iApply wp_val. iModIntro. simpl. - eauto with iFrame. - Qed. - - Lemma compat_replace {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) α β τ τ' : - ⊢ valid1 Ω1 α (tRef τ) -∗ - valid1 Ω2 β τ' -∗ - valid1 (tyctx_app Ω1 Ω2) (interp_replace α β ◎ interp_scope_split) (tPair τ (tRef τ')). - Proof. - Opaque pairITV. - iIntros "H1 H2". - iIntros (αs) "#Hctx Has". - cbn-[interp_replace]. - unfold ssubst_valid. - rewrite ssubst_valid_app. - rewrite interp_scope_ssubst_split. - iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. - iSpecialize ("H1" with "Hctx Ha1"). - iSpecialize ("H2" with "Hctx Ha2"). - iApply (expr_pred_bind (LETCTX _) with "H2"). - iIntros (βv) "Hb". unfold LETCTX. simpl. - rewrite LET_Val/=. - iApply (expr_pred_bind with "H1"). - iIntros (αv) "Ha". simpl. - iDestruct "Ha" as (l γ) "[Ha [Hl Hg]]". - iApply expr_pred_frame. - iRewrite "Ha". simpl. - rewrite IT_of_V_Ret. - rewrite -> get_ret_ret; simpl. - iApply wp_let. - { solve_proper. } - iApply (wp_read with "Hctx Hl"). - iNext. iNext. iIntros "Hl". - iApply wp_val. iModIntro. - simpl. iApply wp_seq. - { solve_proper. } - iApply (wp_write with "Hctx Hl"). - iNext. iNext. iIntros "Hl". - rewrite get_val_ITV. simpl. - rewrite get_val_ITV. simpl. - iApply wp_val. iModIntro. - iExists γ,(RetV l). - iSplit; first done. - iFrame. eauto with iFrame. - Qed. - - Lemma compat_dealloc {S} (Ω : tyctx S) α τ: - ⊢ valid1 Ω α (tRef τ) -∗ - valid1 Ω (interp_dealloc α) tUnit. - Proof. - iIntros "H". - iIntros (αs) "#Hctx Has". - iSpecialize ("H" with "Hctx Has"). - iApply (expr_pred_bind with "H"). - iIntros (αv) "Ha /=". - iDestruct "Ha" as (l βv) "[Ha [Hl Hb]]". - iRewrite "Ha". iApply expr_pred_frame. simpl. - rewrite IT_of_V_Ret. rewrite -> get_ret_ret. simpl. - iApply (wp_dealloc with "Hctx Hl"). - iNext. iNext. eauto with iFrame. - Qed. - - Lemma compat_bool {S} b (Ω : tyctx S) : - ⊢ valid1 Ω (interp_litbool b) tBool. - Proof. - iIntros (αs) "#Hctx Has". - iApply expr_pred_ret. - destruct b; simpl; eauto. - Qed. - Lemma compat_nat {S} n (Ω : tyctx S) : - ⊢ valid1 Ω (interp_litnat n) tInt. - Proof. - iIntros (αs) "#Hctx Has". - iApply expr_pred_ret. eauto with iFrame. - Qed. - Lemma compat_unit {S} (Ω : tyctx S) : - ⊢ valid1 Ω interp_litunit tUnit. - Proof. - iIntros (αs) "#Hctx Has". - iApply expr_pred_ret. eauto with iFrame. - Qed. - Lemma compat_var {S} Ω τ (v : var S) : - typed_var Ω v τ → - ⊢ valid1 Ω (Force ◎ interp_var v) τ. - Proof. - iIntros (Hv ss) "#Hctx Has". - iApply expr_pred_frame. - unfold ssubst_valid. - iInduction Hv as [|? ? ? Ω v] "IH" forall (ss); simpl. - - dependent elimination ss as [cons_ssubst αv ss]. - rewrite ssubst_valid_cons. - simp interp_var. simpl. - iDestruct "Has" as "[H _]". - simp interp_ssubst. simpl. done. - - dependent elimination ss as [cons_ssubst αv ss]. - rewrite ssubst_valid_cons. - simp interp_var. simpl. - iDestruct "Has" as "[_ H]". - simp interp_ssubst. simpl. - by iApply ("IH" with "H"). - Qed. - - Lemma compat_app {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) - α β τ1 τ2 : - ⊢ valid1 Ω1 α (tArr τ1 τ2) -∗ - valid1 Ω2 β τ1 -∗ - valid1 (tyctx_app Ω1 Ω2) (interp_app α β ◎ interp_scope_split) τ2. - Proof. - iIntros "H1 H2". - iIntros (αs) "#Hctx Has". - iEval(cbn-[interp_app]). - unfold ssubst_valid. - rewrite ssubst_valid_app. - rewrite interp_scope_ssubst_split. - iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. - iSpecialize ("H1" with "Hctx Ha1"). - iSpecialize ("H2" with "Hctx Ha2"). - Local Opaque Thunk. - iSimpl. - iApply (expr_pred_bind (LETCTX _) with "H2"). - iIntros (βv) "H2". unfold LETCTX. iSimpl. - rewrite LET_Val/=. - iApply (expr_pred_bind (LETCTX _) with "H1"). - iIntros (αv) "H1". unfold LETCTX. iSimpl. - rewrite LET_Val/=. - by iApply "H1". - Qed. - - Lemma compat_lam {S} (Ω : tyctx S) τ1 τ2 α : - ⊢ valid1 (consC τ1 Ω) α τ2 -∗ - valid1 Ω (interp_lam α) (tArr τ1 τ2). - Proof. - iIntros "H". - iIntros (αs) "#Hctx Has". - iIntros (x) "Hx". - iApply wp_val. - iModIntro. simpl. - iExists _; iFrame. - iIntros (βv) "Hb". clear x. - iIntros (x) "Hx". - iApply (wp_bind _ (AppRSCtx _)). - { solve_proper. } - Local Transparent Thunk. - Local Opaque thunked thunkedV. - iSimpl. iApply (wp_alloc with "Hctx"). - { solve_proper. } - iNext. iNext. iIntros (l) "Hl". - iApply wp_val. iModIntro. - unfold AppRSCtx. - iApply wp_lam. iNext. - iEval(cbn-[thunked]). - iSpecialize ("H" $! (cons_ssubst (thunkedV (IT_of_V βv) l) αs) - with "Hctx [-Hx] Hx"). - { unfold ssubst_valid. - rewrite ssubst_valid_cons. iFrame. - Local Transparent thunked thunkedV. - iApply wp_lam. iNext. simpl. - iApply (wp_bind _ (IFSCtx _ _)). - iApply (wp_read with "Hctx Hl"). - iNext. iNext. iIntros "Hl". - iApply wp_val. iModIntro. - unfold IFSCtx. simpl. - rewrite IF_False; last lia. - iApply wp_seq. - { solve_proper. } - iApply (wp_write with "Hctx Hl"). - iNext. iNext. iIntros "Hl". - iApply wp_val. iModIntro. - iApply "Hb". } - simp interp_ssubst. - iApply "H". - Qed. - - Lemma fundamental_affine {S} (Ω : tyctx S) (e : expr S) τ : - typed Ω e τ → - ⊢ valid1 Ω (interp_expr _ e) τ. - Proof. - induction 1; simpl. - - by iApply compat_var. - - by iApply compat_lam. - - by iApply compat_app. - - by iApply compat_pair. - - by iApply compat_destruct. - - by iApply compat_alloc. - - by iApply compat_replace. - - by iApply compat_dealloc. - - by iApply compat_nat. - - by iApply compat_bool. - - by iApply compat_unit. - Qed. - -End logrel. - -Arguments interp_tarr {_ _ _ _ _ _ _ _ _ _ _ _ _} Φ1 Φ2. -Arguments interp_tbool {_ _ _ _ _ _}. -Arguments interp_tnat {_ _ _ _ _ _}. -Arguments interp_tunit {_ _ _ _ _ _}. -Arguments interp_ty {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _} τ. - -Local Definition rs : gReifiers NotCtxDep 2 := - gReifiers_cons NotCtxDep reify_store (gReifiers_cons NotCtxDep input_lang.interp.reify_io (gReifiers_nil NotCtxDep)). - -Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. - -Lemma logrel1_adequacy cr Σ R `{!Cofe R, !SubOfe natO R, !SubOfe unitO R, !SubOfe locO R} `{!invGpreS Σ} - `{!statePreG rs R Σ} `{!heapPreG rs R Σ} τ - (α : unitO -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : - (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ}, - (£ cr ⊢ valid1 rs notStuck (λne _: unitO, True)%I empC α τ)%I) → - ssteps (gReifiers_sReifier NotCtxDep rs) (α ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ βv, IT_of_V βv ≡ β). -Proof. - intros Hlog Hst. - destruct (IT_to_V β) as [βv|] eqn:Hb. - { right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } - left. - cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ e, β ≡ Err e ∧ notStuck e)). - { intros [?|He]; first done. - destruct He as [? [? []]]. } - eapply (wp_safety (S cr)); eauto. - { apply Hdisj. } - { by rewrite Hb. } - intros H1 H2. - exists (λ _, True)%I. split. - { apply _. } - iIntros "[[Hone Hcr] Hst]". - pose (σ := st.1). - pose (ios := st.2.1). - iMod (new_heapG rs σ) as (H3) "H". - iAssert (has_substate σ ∗ has_substate ios)%I with "[Hst]" as "[Hs Hio]". - { unfold has_substate, has_full_state. - assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) st ≡ - of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ) - ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state ios)) as ->; last first. - { rewrite -own_op. done. } - unfold sR_idx. simpl. - intro j. - rewrite discrete_fun_lookup_op. - inv_fin j. - { unfold of_state, of_idx. simpl. - erewrite (eq_pi _ _ _ (@eq_refl _ 0%fin)). done. } - intros j. inv_fin j. - { unfold of_state, of_idx. simpl. - erewrite (eq_pi _ _ _ (@eq_refl _ 1%fin)). done. } - intros i. inversion i. } - iApply fupd_wp. - iMod (inv_alloc (nroot.@"storeE") _ (∃ σ, £ 1 ∗ has_substate σ ∗ own (heapG_name rs) (●V σ))%I with "[-Hcr]") as "#Hinv". - { iNext. iExists _. iFrame. } - simpl. - iPoseProof (@Hlog _ _ _ with "Hcr") as "Hlog". - iSpecialize ("Hlog" $! emp_ssubst with "Hinv []"). - { iApply ssubst_valid_nil. } - iSpecialize ("Hlog" $! tt with "[//]"). - iModIntro. - iApply (wp_wand with "Hlog"). - eauto with iFrame. -Qed. - -Definition R := sumO locO (sumO unitO natO). - -Lemma logrel1_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : - typed empC e τ → - ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ()) st β st' k → - (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ βv, IT_of_V βv ≡ β). -Proof. - intros Hty Hst. - pose (Σ:=#[invΣ;stateΣ rs R;heapΣ rs R]). - eapply (logrel1_adequacy 0 Σ); eauto; try apply _. - iIntros (? ? ?) "_". - by iApply fundamental_affine. -Qed. diff --git a/theories/examples/store.v b/theories/effects/store.v similarity index 99% rename from theories/examples/store.v rename to theories/effects/store.v index 490ad68..6b72ad6 100644 --- a/theories/examples/store.v +++ b/theories/effects/store.v @@ -1,4 +1,3 @@ -From Equations Require Import Equations. From iris.algebra Require Import gmap excl auth gmap_view. From iris.proofmode Require Import classes tactics. From iris.base_logic Require Import algebra. diff --git a/theories/affine_lang/lang.v b/theories/examples/affine_lang/lang.v similarity index 69% rename from theories/affine_lang/lang.v rename to theories/examples/affine_lang/lang.v index 9616cb1..3e5505f 100644 --- a/theories/affine_lang/lang.v +++ b/theories/examples/affine_lang/lang.v @@ -1,6 +1,7 @@ From gitrees Require Export lang_generic gitree program_logic. -From gitrees.input_lang Require Import lang interp. -From gitrees.examples Require Import store pairs. +From gitrees.examples.input_lang Require Import lang interp. +From gitrees.effects Require Import store. +From gitrees.lib Require Import pairs. (* for namespace sake *) Module io_lang. @@ -14,15 +15,13 @@ Module io_lang. input_lang.interp.interp_expr rs e (ı_scope rs). End io_lang. -From gitrees Require Export lang_affine. +Require Import Binding.Resolver Binding.Lib Binding.Set Binding.Auto Binding.Env. Inductive ty := tBool | tInt | tUnit | tArr (τ1 τ2 : ty) | tPair (τ1 τ2 : ty) | tRef (τ : ty). -Local Notation tyctx := (tyctx ty). - Inductive ty_conv : ty → io_lang.ty → Type := | ty_conv_bool : ty_conv tBool Tnat | ty_conv_int : ty_conv tInt Tnat @@ -32,19 +31,19 @@ Inductive ty_conv : ty → io_lang.ty → Type := ty_conv (tArr τ1 τ2) (Tarr (Tarr Tnat t1) t2) . -Inductive expr : scope → Type := -| LitBool (b : bool) {S} : expr S -| LitNat (n : nat) {S} : expr S -| LitUnit {S} : expr S -| Lam {S} : expr (tt::S) → expr S -| Var {S} : var S → expr S -| App {S1 S2} : expr S1 → expr S2 → expr (S1++S2) -| EPair {S1 S2} : expr S1 → expr S2 → expr (S1++S2) -| EDestruct {S1 S2} : expr S1 → expr (()::()::S2) → expr (S1++S2) -| Alloc {S} : expr S → expr S -| Replace {S1 S2} : expr S1 → expr S2 → expr (S1++S2) -| Dealloc {S} : expr S → expr S -| EEmbed {τ1 τ1' S} : io_lang.expr Empty_set → ty_conv τ1 τ1' → expr S +Inductive expr : ∀ (S : Set), Type := +| LitBool {S : Set} (b : bool) : expr S +| LitNat {S : Set} (n : nat) : expr S +| LitUnit {S : Set} : expr S +| Lam {S : Set} : expr (inc S) → expr S +| Var {S : Set} : S → expr S +| App {S1 S2 : Set} : expr S1 → expr S2 → expr (sum S1 S2) +| EPair {S1 S2 : Set} : expr S1 → expr S2 → expr (sum S1 S2) +| EDestruct {S1 S2 : Set} : expr S1 → expr (inc (inc S2)) → expr (sum S1 S2) +| Alloc {S : Set} : expr S → expr S +| Replace {S1 S2 : Set} : expr S1 → expr S2 → expr (sum S1 S2) +| Dealloc {S : Set} : expr S → expr S +| EEmbed {S : Set} {τ1 τ1'} : io_lang.expr Empty_set → ty_conv τ1 τ1' → expr S . Section affine. @@ -64,9 +63,11 @@ Section affine. λit _, IF (READ ℓ) (Err OtherError) (SEQ (WRITE ℓ (Ret 1)) e). Solve All Obligations with first [solve_proper|solve_proper_please]. + Program Definition thunkedV : IT -n> locO -n> ITV := λne e ℓ, FunV $ Next (λne _, IF (READ ℓ) (Err OtherError) (SEQ (WRITE ℓ (Ret 1)) e)). Solve All Obligations with first [solve_proper|solve_proper_please]. + #[export] Instance thunked_into_val e l : IntoVal (thunked e l) (thunkedV e l). Proof. unfold IntoVal. simpl. f_equiv. f_equiv. intro. done. @@ -75,45 +76,52 @@ Section affine. Program Definition Thunk : IT -n> IT := λne e, ALLOC (Ret 0) (thunked e). Solve All Obligations with first [solve_proper|solve_proper_please]. + Program Definition Force : IT -n> IT := λne e, e ⊙ (Ret 0). Local Open Scope type. Definition interp_litbool {A} (b : bool) : A -n> IT := λne _, - Ret (if b then 1 else 0). + Ret (if b then 1 else 0). + Definition interp_litnat {A} (n : nat) : A -n> IT := λne _, - Ret n. + Ret n. + Definition interp_litunit {A} : A -n> IT := λne _, Ret tt. + Program Definition interp_pair {A1 A2} (t1 : A1 -n> IT) (t2 : A2 -n> IT) - : A1*A2 -n> IT := λne env, + : A1 * A2 -n> IT := λne env, pairIT (t1 env.1) (t2 env.2). (* we don't need to evaluate the pair here, i.e. lazy pairs? *) Next Obligation. solve_proper_please. Qed. - Program Definition interp_lam {A : ofe} (b : (IT * A) -n> IT) : A -n> IT := λne env, - λit x, b (x,env). - Solve All Obligations with solve_proper_please. + + Program Definition interp_lam {S : Set} (b : @interp_scope F R _ (inc S) -n> IT) : @interp_scope F R _ S -n> IT := λne env, (λit x, (b (@extend_scope F R _ _ env x))). + Next Obligation. + intros; repeat intro; repeat f_equiv; assumption. + Qed. + Next Obligation. + intros; repeat intro; repeat f_equiv; intro; simpl; + f_equiv; intro z; simpl. + destruct z; done. + Qed. + Program Definition interp_app {A1 A2 : ofe} (t1 : A1 -n> IT) (t2 : A2 -n> IT) : A1*A2 -n> IT := λne env, LET (t2 env.2) $ λne x, LET (t1 env.1) $ λne f, APP' f (Thunk x). Solve All Obligations with solve_proper_please. - Program Definition interp_destruct {A1 A2 : ofe} - (ps : A1 -n> IT) (t : IT*(IT*A2) -n> IT) - : A1*A2 -n> IT := λne env, - LET (ps env.1) $ λne ps, - LET (Thunk (projIT1 ps)) $ λne x, - LET (Thunk (projIT2 ps)) $ λne y, - t (x, (y, env.2)). - Solve All Obligations with solve_proper_please. + Program Definition interp_alloc {A} (α : A -n> IT) : A -n> IT := λne env, LET (α env) $ λne α, ALLOC α Ret. Solve All Obligations with solve_proper_please. + Program Definition interp_replace {A1 A2} (α : A1 -n> IT) (β : A2 -n> IT) : A1*A2 -n> IT := λne env, LET (β env.2) $ λne β, flip get_ret (α env.1) $ λne (l : loc), LET (READ l) $ λne γ, SEQ (WRITE l β) (pairIT γ (Ret l)). Solve All Obligations with solve_proper_please. + Program Definition interp_dealloc {A} (α : A -n> IT) : A -n> IT := λne env, get_ret DEALLOC (α env). Solve All Obligations with solve_proper_please. @@ -154,6 +162,35 @@ Section affine. | ty_conv_fun conv1 conv2 => glue_from_affine_fun (glue_from_affine conv2) (glue_to_affine conv1) end. + Program Definition interp_destruct {S1 S2 : Set} + (ps : @interp_scope F R _ S1 -n> IT) + (t : (@interp_scope F R _ (inc (inc S2)) -n> IT)) + : (@interp_scope F R _ S1 * @interp_scope F R _ S2) -n> IT + := λne env, + LET (ps env.1) $ λne z, + LET (Thunk (projIT1 z)) $ λne x, + LET (Thunk (projIT2 z)) $ λne y, + (t (@extend_scope F R _ _ (@extend_scope F R _ _ env.2 y) x)). + Next Obligation. + intros; repeat intro; repeat f_equiv; assumption. + Qed. + Next Obligation. + intros; repeat intro; repeat f_equiv; intro; simpl; f_equiv; intro A; simpl. + destruct A as [| [|]]; [assumption | reflexivity | reflexivity]. + Qed. + Next Obligation. + intros; repeat intro; repeat f_equiv; first assumption. + intro; simpl; f_equiv; intro; simpl. + repeat f_equiv; intro; simpl; repeat f_equiv; intro; simpl. + repeat f_equiv; assumption. + Qed. + Next Obligation. + intros; repeat intro; repeat f_equiv; first assumption. + intro; simpl; f_equiv; intro; simpl. + repeat f_equiv; intro; simpl; repeat f_equiv; intro A; simpl. + destruct A as [| [|]]; [reflexivity | reflexivity |]. + repeat f_equiv; assumption. + Qed. Fixpoint interp_expr {S} (e : expr S) : interp_scope S -n> IT := match e with @@ -162,12 +199,16 @@ Section affine. | LitUnit => interp_litunit | Var v => Force ◎ interp_var v | Lam e => interp_lam (interp_expr e) - | App e1 e2 => interp_app (interp_expr e1) (interp_expr e2) ◎ interp_scope_split - | EPair e1 e2 => interp_pair (interp_expr e1) (interp_expr e2) ◎ interp_scope_split - | EDestruct e1 e2 => interp_destruct (interp_expr e1) (interp_expr e2) ◎ interp_scope_split + | App e1 e2 => + interp_app (interp_expr e1) (interp_expr e2) ◎ interp_scope_split + | EPair e1 e2 => + interp_pair (interp_expr e1) (interp_expr e2) ◎ interp_scope_split + | EDestruct e1 e2 => + interp_destruct (interp_expr e1) (interp_expr e2) ◎ interp_scope_split | Alloc e => interp_alloc (interp_expr e) | Dealloc e => interp_dealloc (interp_expr e) - | Replace e1 e2 => interp_replace (interp_expr e1) (interp_expr e2) ◎ interp_scope_split + | Replace e1 e2 => + interp_replace (interp_expr e1) (interp_expr e2) ◎ interp_scope_split | EEmbed e tconv => constO $ glue_to_affine tconv (io_lang.interp_closed _ e) end. diff --git a/theories/examples/affine_lang/logrel1.v b/theories/examples/affine_lang/logrel1.v new file mode 100644 index 0000000..50b2fcd --- /dev/null +++ b/theories/examples/affine_lang/logrel1.v @@ -0,0 +1,931 @@ +(** Unary (Kripke) logical relation for the affine lang *) +From gitrees Require Export gitree program_logic. +From gitrees.examples.affine_lang Require Import lang. +From gitrees.effects Require Import store. +From gitrees.lib Require Import pairs. +Require Import iris.algebra.gmap. +Require Import stdpp.finite. + +Require Import Binding.Resolver Binding.Lib Binding.Set Binding.Auto Binding.Env. + +Lemma fin_to_set_sum {S1 S2 : Set} `{EqDecision S1} `{EqDecision S2} `{EqDecision (S1 + S2)} + `{Finite S1} `{Finite S2} `{Finite (S1 + S2)} + `{Countable S1} `{Countable S2} `{Countable (S1 + S2)} + : let A1 : gset S1 := (fin_to_set S1) in + let A2 : gset (S1 + S2) := set_map (inl : S1 → S1 + S2) A1 in + let B1 : gset S2 := (fin_to_set S2) in + let B2 : gset (S1 + S2) := set_map (inr : S2 → S1 + S2) B1 in + let C : gset (S1 + S2) := fin_to_set (S1 + S2) in + C = A2 ∪ B2. +Proof. + apply set_eq. + intros [x|x]; simpl; split; intros _. + - apply elem_of_union; left. + apply elem_of_map_2. + apply elem_of_fin_to_set. + - apply elem_of_fin_to_set. + - apply elem_of_union; right. + apply elem_of_map_2. + apply elem_of_fin_to_set. + - apply elem_of_fin_to_set. +Qed. + +Lemma fin_to_set_empty : + let A : gset ∅ := fin_to_set ∅ in + let B : gset ∅ := empty in + A = B. +Proof. + apply set_eq; intros []. +Qed. + +Section InstSum. + Global Instance EqDecisionLeft {S1 S2 : Set} {H : EqDecision (S1 + S2)} : EqDecision S1. + Proof. + intros x y. + destruct (decide (inl x = inl y)) as [G | G]; + [left; by inversion G | right; intros C; by subst]. + Qed. + + Global Instance EqDecisionRight {S1 S2 : Set} {H : EqDecision (S1 + S2)} : EqDecision S2. + Proof. + intros x y. + destruct (decide (inr x = inr y)) as [G | G]; + [left; by inversion G | right; intros C; by subst]. + Qed. + + Global Instance FiniteLeft {S1 S2 : Set} `{EqDecision S1} + `{EqDecision (S1 + S2)} `{Finite (S1 + S2)} + : Finite S1. + Proof. + unshelve econstructor. + - apply (foldr (λ x acc, match x with + | inl x => x :: acc + | inr _ => acc + end) [] (enum (S1 + S2))). + - set (l := enum (S1 + S2)). + assert (NoDup l) as K; first apply NoDup_enum. + clearbody l. + induction l as [| a l IH]; first constructor. + destruct a as [a | a]; simpl. + + constructor. + * intros C. + assert (inl a ∈ l) as C'. + { + clear -C. + induction l as [| b l IH]; first inversion C. + destruct b as [b | b]; simpl. + - rewrite foldr_cons in C. + rewrite elem_of_cons in C. + destruct C as [-> | C]. + + apply elem_of_cons. + by left. + + right. + apply IH. + apply C. + - apply elem_of_cons. + right. + rewrite foldr_cons in C. + apply IH. + apply C. + } + by inversion K. + * apply IH. + by inversion K. + + apply IH. + by inversion K. + - intros x. + set (l := enum (S1 + S2)). + assert (inl x ∈ l) as K; first apply elem_of_enum. + clearbody l. + induction l as [| a l IH]; first inversion K. + destruct a as [a | a]; simpl. + + rewrite elem_of_cons in K. + destruct K as [K | K]. + * inversion K; subst. + apply elem_of_cons; by left. + * apply elem_of_cons; right; by apply IH. + + rewrite elem_of_cons in K. + destruct K as [K | K]; first inversion K. + by apply IH. + Qed. + + Global Instance FiniteRight {S1 S2 : Set} `{EqDecision S2} + `{EqDecision (S1 + S2)} `{H : Finite (S1 + S2)} + : Finite S2. + Proof. + unshelve econstructor. + - apply (foldr (λ x acc, match x with + | inl _ => acc + | inr x => x :: acc + end) [] (enum (S1 + S2))). + - set (l := enum (S1 + S2)). + assert (NoDup l) as K; first apply NoDup_enum. + clearbody l. + induction l as [| a l IH]; first constructor. + destruct a as [a | a]; simpl. + + apply IH. + by inversion K. + + constructor. + * intros C. + assert (inr a ∈ l) as C'. + { + clear -C. + induction l as [| b l IH]; first inversion C. + destruct b as [b | b]; simpl. + - apply elem_of_cons. + right. + rewrite foldr_cons in C. + apply IH. + apply C. + - rewrite foldr_cons in C. + rewrite elem_of_cons in C. + destruct C as [-> | C]. + + apply elem_of_cons. + by left. + + right. + apply IH. + apply C. + } + by inversion K. + * apply IH. + by inversion K. + - intros x. + set (l := enum (S1 + S2)). + assert (inr x ∈ l) as K; first apply elem_of_enum. + clearbody l. + induction l as [| a l IH]; first inversion K. + destruct a as [a | a]; simpl. + + rewrite elem_of_cons in K. + destruct K as [K | K]; first inversion K. + by apply IH. + + rewrite elem_of_cons in K. + destruct K as [K | K]. + * inversion K; subst. + apply elem_of_cons; by left. + * apply elem_of_cons; right; by apply IH. + Qed. + +End InstSum. + +Section InstInc. + Context (S : Set). + + Global Instance EqDecisionIncN {HS : EqDecision S} (n : nat) : EqDecision (Init.Nat.iter n inc S). + Proof using S. + induction n; simpl. + - apply _. + - intros [|x] [|y]. + + by left. + + by right. + + by right. + + destruct (decide (x = y)) as [-> |]; first by left. + right; by inversion 1. + Qed. + + Global Instance EqDecisionInc {HS : EqDecision S} : EqDecision (inc S). + Proof using S. + assert (inc S = Init.Nat.iter 1 inc S) as ->; first done. + by apply EqDecisionIncN. + Qed. + + Global Instance FiniteIncN {HS : EqDecision S} (HF : Finite S) (n : nat) {HS' : EqDecision (Init.Nat.iter n inc S)} : Finite (Init.Nat.iter n inc S). + Proof using S. + induction n. + - apply (@surjective_finite S HS HF _ _ id). + apply _. + - simpl. + unshelve eapply (@surjective_finite (option (Init.Nat.iter n inc S))); simpl in *. + + intros [x |]. + * apply (VS x). + * apply VZ. + + apply _. + + intros [| x]; simpl. + * exists None; reflexivity. + * exists (Some x); reflexivity. + Qed. + + Global Instance FiniteInc {HS : EqDecision S} (HF : Finite S) (HE : EqDecision (inc S)) : Finite (inc S). + Proof using S. + assert (J : @Finite (Init.Nat.iter 1 inc S) HE). + { apply FiniteIncN, HF. } + simpl in J. + apply J. + Qed. + +End InstInc. + +Definition sum_map' {A B C : Set} (f : A → C) (g : B → C) : sum A B → C := + λ x, match x with | inl x' => f x' | inr x' => g x' end. + +Inductive typed : forall {S : Set}, (S → ty) → expr S → ty → Prop := +(** functions *) +| typed_Var {S : Set} (Ω : S → ty) (v : S) : + typed Ω (Var v) (Ω v) +| typed_Lam {S : Set} (Ω : S → ty) (τ1 τ2 : ty) (e : expr (inc S) ) : + typed (Ω ▹ τ1) e τ2 → + typed Ω (Lam e) (tArr τ1 τ2) +| typed_App {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 : ty) e1 e2 : + typed Ω1 e1 (tArr τ1 τ2) → + typed Ω2 e2 τ1 → + typed (sum_map' Ω1 Ω2) (App e1 e2) τ2 +(** pairs *) +| typed_Pair {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 : ty) e1 e2 : + typed Ω1 e1 τ1 → + typed Ω2 e2 τ2 → + typed (sum_map' Ω1 Ω2) (EPair e1 e2) (tPair τ1 τ2) +| typed_Destruct {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 τ : ty) + (e1 : expr S1) (e2 : expr (inc (inc S2))) : + typed Ω1 e1 (tPair τ1 τ2) → + typed ((Ω2 ▹ τ2) ▹ τ1) e2 τ → + typed (sum_map' Ω1 Ω2) (EDestruct e1 e2) τ +(** references *) +| typed_Alloc {S : Set} (Ω : S → ty) τ e : + typed Ω e τ → + typed Ω (Alloc e) (tRef τ) +| typed_Replace {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 : ty) e1 e2 : + typed Ω1 e1 (tRef τ1) → + typed Ω2 e2 τ2 → + typed (sum_map' Ω1 Ω2) (Replace e1 e2) (tPair τ1 (tRef τ2)) +| typed_Dealloc {S : Set} (Ω : S → ty) e τ : + typed Ω e (tRef τ) → + typed Ω (Dealloc e) tUnit +(** literals *) +| typed_Nat {S : Set} (Ω : S → ty) n : + typed Ω (LitNat n) tInt +| typed_Bool {S : Set} (Ω : S → ty) b : + typed Ω (LitBool b) tBool +| typed_Unit {S : Set} (Ω : S → ty) : + typed Ω LitUnit tUnit +. + +Section logrel. + Context {sz : nat}. + Variable rs : gReifiers NotCtxDep sz. + Context `{!subReifier reify_store rs}. + Context `{!subReifier input_lang.interp.reify_io rs}. + Notation F := (gReifiers_ops NotCtxDep rs). + Context {R} `{!Cofe R}. + Context `{!SubOfe natO R}. + Context `{!SubOfe unitO R}. + Context `{!SubOfe locO R}. + Notation IT := (IT F R). + Notation ITV := (ITV F R). + Context `{!invGS Σ, !stateG rs R Σ, !heapG rs R Σ}. + Notation iProp := (iProp Σ). + + (* parameters for the kripke logical relation *) + Variable s : stuckness. + Context `{A:ofe}. + Variable (P : A -n> iProp). + Local Notation expr_pred := (expr_pred s rs P). + + (* interpreting tys *) + Program Definition protected (Φ : ITV -n> iProp) : ITV -n> iProp := λne αv, + (WP@{rs} Force (IT_of_V αv) @ s {{ Φ }})%I. + Solve All Obligations with solve_proper_please. + Program Definition interp_tbool : ITV -n> iProp := λne αv, + (αv ≡ RetV 0 ∨ αv ≡ RetV 1)%I. + Solve All Obligations with solve_proper_please. + Program Definition interp_tnat : ITV -n> iProp := λne αv, + (∃ n : nat, αv ≡ RetV n)%I. + Solve All Obligations with solve_proper_please. + Program Definition interp_tunit : ITV -n> iProp := λne αv, + (αv ≡ RetV ())%I. + Solve All Obligations with solve_proper_please. + Program Definition interp_tpair (Φ1 Φ2 : ITV -n> iProp) : ITV -n> iProp := λne αv, + (∃ β1v β2v, IT_of_V αv ≡ pairITV (IT_of_V β1v) (IT_of_V β2v) ∗ + Φ1 β1v ∗ Φ2 β2v)%I. + Solve All Obligations with solve_proper_please. + Program Definition interp_tarr (Φ1 Φ2 : ITV -n> iProp) : ITV -n> iProp := λne αv, + (∀ βv, Φ1 βv -∗ expr_pred ((IT_of_V αv) ⊙ (Thunk (IT_of_V βv))) Φ2)%I. + Solve All Obligations with solve_proper_please. + + Program Definition interp_ref (Φ : ITV -n> iProp) : ITV -n> iProp := λne αv, + (∃ (l : loc), αv ≡ RetV l ∗ ∃ βv, pointsto l (IT_of_V βv) ∗ Φ βv)%I. + Solve All Obligations with solve_proper_please. + + Fixpoint interp_ty (τ : ty) : ITV -n> iProp := + match τ with + | tBool => interp_tbool + | tUnit => interp_tunit + | tInt => interp_tnat + | tPair τ1 τ2 => interp_tpair (interp_ty τ1) (interp_ty τ2) + | tArr τ1 τ2 => interp_tarr (interp_ty τ1) (interp_ty τ2) + | tRef τ => interp_ref (interp_ty τ) + end. + + Program Definition ssubst_valid {S : Set} `{!EqDecision S} `{!Finite S} + (Ω : S → ty) (ss : interp_scope S) : iProp + := ([∗ set] x ∈ (fin_to_set S), + (expr_pred (ss x) (protected (interp_ty (Ω x))))%I). + + Definition valid1 {S : Set} `{!EqDecision S} `{!Finite S} (Ω : S → ty) + (α : interp_scope S -n> IT) (τ : ty) : iProp := + ∀ ss, heap_ctx + -∗ (ssubst_valid Ω ss) + -∗ expr_pred (α ss) (interp_ty τ). + + Lemma ssubst_valid_empty (αs : interp_scope ∅) : + ⊢ ssubst_valid □ αs. + Proof. + iStartProof. + unfold ssubst_valid. + match goal with + | |- context G [big_opS ?a ?b ?c] => assert (c = empty) as -> + end. + { apply set_eq; intros []. } + by iApply big_sepS_empty. + Qed. + + Lemma ssubst_valid_app + {S1 S2 : Set} `{!EqDecision S1} `{!Finite S1} + `{!EqDecision S2} `{!Finite S2} + `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} + (Ω1 : S1 → ty) (Ω2 : S2 → ty) + (αs : interp_scope (sum S1 S2)) : + (ssubst_valid (sum_map' Ω1 Ω2) αs) ⊢ + (ssubst_valid Ω1 (interp_scope_split αs).1) + ∗ (ssubst_valid Ω2 (interp_scope_split αs).2). + Proof. + iIntros "H". + rewrite /ssubst_valid fin_to_set_sum big_sepS_union; first last. + { + apply elem_of_disjoint. + intros [x | x]. + - rewrite !elem_of_list_to_set. + intros _ H2. + apply elem_of_list_fmap_2 in H2. + destruct H2 as [y [H2 H2']]; inversion H2. + - rewrite !elem_of_list_to_set. + intros H1 _. + apply elem_of_list_fmap_2 in H1. + destruct H1 as [y [H1 H1']]; inversion H1. + } + iDestruct "H" as "(H1 & H2)". + iSplitL "H1". + - rewrite big_opS_list_to_set; first last. + { + apply NoDup_fmap. + - intros ??; by inversion 1. + - apply NoDup_elements. + } + rewrite big_sepL_fmap /=. + rewrite big_sepS_elements. + iFrame "H1". + - rewrite big_opS_list_to_set; first last. + { + apply NoDup_fmap. + - intros ??; by inversion 1. + - apply NoDup_elements. + } + rewrite big_sepL_fmap /=. + rewrite big_sepS_elements. + iFrame "H2". + Qed. + + Lemma ssubst_valid_cons {S : Set} `{!EqDecision S} `{!Finite S} + (Ω : S → ty) (αs : interp_scope S) τ t : + ssubst_valid Ω αs ∗ expr_pred t (protected (interp_ty τ)) ⊢ ssubst_valid (Ω ▹ τ) (extend_scope αs t). + Proof. + iIntros "(H & G)". + rewrite /ssubst_valid. + pose (Y := let A := {[VZ]} : @gset (leibnizO (inc S)) _ finite_countable in + let B := fin_to_set (leibnizO S) : gset (leibnizO S) in + let C := set_map (VS : S → inc S) B + : gset (inc S) in A ∪ C). + assert (fin_to_set (inc S) = Y) as ->. + { + subst Y; simpl. + apply set_eq. + intros [| x]. + - split. + + intros _; apply elem_of_union; left. + by apply elem_of_singleton. + + intros _; apply elem_of_fin_to_set. + - split. + + intros _; apply elem_of_union; right. + apply elem_of_map_2, elem_of_fin_to_set. + + intros H. + apply elem_of_fin_to_set. + } + subst Y; simpl. + rewrite big_sepS_union; first last. + { + apply elem_of_disjoint. + intros [| x]. + - rewrite !elem_of_list_to_set. + intros _ H2. + apply elem_of_list_fmap_2 in H2. + destruct H2 as [y [H2 H2']]; inversion H2. + - rewrite !elem_of_list_to_set. + intros H1 _. + apply elem_of_singleton_1 in H1. + inversion H1. + } + iSplitL "G". + - rewrite big_opS_singleton. + iFrame "G". + - erewrite big_opS_set_map. + + iFrame "H". + + intros ?? H; by inversion H. + Qed. + + Lemma ssubst_valid_lookup {S : Set} `{!EqDecision S} `{!Finite S} + (Ω : S → ty) (αs : interp_scope S) x : + ssubst_valid Ω αs ⊢ expr_pred (αs x) (protected (interp_ty (Ω x))). + Proof. + iIntros "H". + iDestruct (big_sepS_elem_of_acc _ _ x with "H") as "($ & _)"; + first apply elem_of_fin_to_set. + Qed. + + Lemma compat_pair {S1 S2 : Set} + `{!EqDecision S1} `{!Finite S1} + `{!EqDecision S2} `{!Finite S2} + `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} + (Ω1 : S1 → ty) (Ω2 : S2 → ty) α β τ1 τ2 : + ⊢ valid1 Ω1 α τ1 -∗ + valid1 Ω2 β τ2 -∗ + valid1 (sum_map' Ω1 Ω2) (interp_pair α β ◎ interp_scope_split) (tPair τ1 τ2). + Proof. + Opaque pairITV. + iIntros "H1 H2". + iIntros (αs) "#Hctx Has". + cbn-[interp_pair]. + rewrite ssubst_valid_app. + iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. + iSpecialize ("H1" with "Hctx Ha1"). + iSpecialize ("H2" with "Hctx Ha2"). + iApply (expr_pred_bind with "H2"). + iIntros (βv) "Hb". simpl. + rewrite -> get_val_ITV. simpl. + iApply (expr_pred_bind with "H1"). + iIntros (αv) "Ha". simpl. + rewrite -> get_val_ITV. simpl. + iApply expr_pred_ret. + simpl. + iExists _,_. + by iFrame. + Qed. + + Lemma compat_destruct {S1 S2 : Set} + `{!EqDecision S1} `{!Finite S1} + `{!EqDecision S2} `{!Finite S2} + `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} + (Ω1 : S1 → ty) (Ω2 : S2 → ty) + α β τ1 τ2 τ : + ⊢ valid1 Ω1 α (tPair τ1 τ2) + -∗ valid1 (Ω2 ▹ τ2 ▹ τ1) β τ + -∗ valid1 (sum_map' Ω1 Ω2) (interp_destruct α β ◎ interp_scope_split) τ. + Proof. + Opaque pairITV thunked thunkedV projIT1 projIT2. + iIntros "H1 H2". + iIntros (αs) "#Hctx Has". + cbn-[interp_destruct]. + rewrite ssubst_valid_app. + iDestruct "Has" as "[Ha1 Ha2]". + iSpecialize ("H1" with "Hctx Ha1"). + iApply (expr_pred_bind (LETCTX _) with "H1"). + iIntros (αv) "Ha". unfold LETCTX. simpl. + rewrite LET_Val/=. + iDestruct "Ha" as (β1 β2) "[#Ha [Hb1 Hb2]]". + iIntros (x) "Hx". + iApply wp_let. + iApply (wp_Thunk with "Hctx"). + { + repeat intro; simpl. + repeat f_equiv. + intro; simpl. + f_equiv. + intro B; simpl. + destruct B as [| [|]]; [by f_equiv | reflexivity | reflexivity]. + } + iNext. iIntros (l1) "Hl1". simpl. + iApply wp_let. + { solve_proper. } + iApply (wp_Thunk with "Hctx"). + { + repeat intro; simpl. + repeat f_equiv. + intro B; simpl. + destruct B as [| [|]]; [reflexivity | by f_equiv | reflexivity]. + } + iNext. iIntros (l2) "Hl2". simpl. + pose (ss' := (extend_scope + (extend_scope + (interp_scope_split αs).2 + (IT_of_V (thunkedV (projIT2 (IT_of_V αv)) l2))) + (IT_of_V (thunkedV (projIT1 (IT_of_V αv)) l1)))). + iSpecialize ("H2" $! ss' + with "Hctx [-Hx] Hx"). + { + iApply ssubst_valid_cons. + iSplitR "Hl1 Hb1". + - iApply ssubst_valid_cons. + iSplitL "Ha2"; first done. + Transparent thunkedV thunked. + simpl. + iIntros (z) "Hz". + simpl. + iApply wp_val. + iModIntro. + iExists z; iFrame. + iApply wp_lam. + iNext. + simpl. + iApply (wp_bind _ (IFSCtx _ _)). + iApply (wp_read with "Hctx Hl2"). + iNext. iNext. iIntros "Hl2". + iApply wp_val. iModIntro. + unfold IFSCtx. simpl. + rewrite IF_False; last lia. + iApply wp_seq. + { solve_proper. } + iApply (wp_write with "Hctx Hl2"). + iNext. iNext. iIntros "Hl2". + iRewrite "Ha". + simpl. + rewrite projIT2_pairV. + do 3 (iApply wp_tick; iNext). + iApply wp_val. iModIntro. + iApply "Hb2". + - Transparent thunkedV thunked. + simpl. + iIntros (z) "Hz". + simpl. + iApply wp_val. + iModIntro. + iExists z; iFrame. + iApply wp_lam. + iNext. + simpl. + iApply (wp_bind _ (IFSCtx _ _)). + iApply (wp_read with "Hctx Hl1"). + iNext. iNext. iIntros "Hl1". + iApply wp_val. iModIntro. + unfold IFSCtx. simpl. + rewrite IF_False; last lia. + iApply wp_seq. + { solve_proper. } + iApply (wp_write with "Hctx Hl1"). + iNext. iNext. iIntros "Hl1". + iRewrite "Ha". + simpl. + rewrite projIT1_pairV. + do 3 (iApply wp_tick; iNext). + iApply wp_val. iModIntro. + iApply "Hb1". + } + iApply "H2". + Qed. + + Lemma compat_alloc {S : Set} + `{!EqDecision S} `{!Finite S} + (Ω : S → ty) α τ: + ⊢ valid1 Ω α τ -∗ + valid1 Ω (interp_alloc α) (tRef τ). + Proof. + iIntros "H". + iIntros (αs) "#Hctx Has". + iSpecialize ("H" with "Hctx Has"). + simpl. iApply (expr_pred_bind (LETCTX _) with "H"). + iIntros (αv) "Hav". unfold LETCTX. simpl. + rewrite LET_Val/=. + iApply expr_pred_frame. + iApply (wp_alloc with "Hctx"). + iNext. iNext. iIntros (l) "Hl". + iApply wp_val. iModIntro. simpl. + iExists l. + iSplit; first done. + iExists αv. + iFrame "Hl". + iFrame. + Qed. + + Lemma compat_replace {S1 S2 : Set} + `{!EqDecision S1} `{!Finite S1} + `{!EqDecision S2} `{!Finite S2} + `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} + (Ω1 : S1 → ty) (Ω2 : S2 → ty) α β τ τ' : + ⊢ valid1 Ω1 α (tRef τ) -∗ + valid1 Ω2 β τ' -∗ + valid1 (sum_map' Ω1 Ω2) (interp_replace α β ◎ interp_scope_split) (tPair τ (tRef τ')). + Proof. + Opaque pairITV. + iIntros "H1 H2". + iIntros (αs) "#Hctx Has". + cbn-[interp_replace]. + rewrite ssubst_valid_app. + iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. + iSpecialize ("H1" with "Hctx Ha1"). + iSpecialize ("H2" with "Hctx Ha2"). + iApply (expr_pred_bind (LETCTX _) with "H2"). + iIntros (βv) "Hb". unfold LETCTX. simpl. + rewrite LET_Val/=. + iApply (expr_pred_bind with "H1"). + iIntros (αv) "Ha". simpl. + iDestruct "Ha" as (l) "[Ha Ha']". + iDestruct "Ha'" as (γ) "[Hl Hg]". + iApply expr_pred_frame. + iRewrite "Ha". simpl. + rewrite IT_of_V_Ret. + rewrite -> get_ret_ret; simpl. + iApply wp_let. + { solve_proper. } + iApply (wp_read with "Hctx Hl"). + iNext. iNext. iIntros "Hl". + iApply wp_val. iModIntro. + simpl. iApply wp_seq. + { solve_proper. } + iApply (wp_write with "Hctx Hl"). + iNext. iNext. iIntros "Hl". + rewrite get_val_ITV. simpl. + rewrite get_val_ITV. simpl. + iApply wp_val. iModIntro. + iExists γ, (RetV l). + iSplit; first done. + iFrame. eauto with iFrame. + Qed. + + Lemma compat_dealloc {S : Set} + `{!EqDecision S} `{!Finite S} + (Ω : S → ty) α τ: + ⊢ valid1 Ω α (tRef τ) -∗ + valid1 Ω (interp_dealloc α) tUnit. + Proof. + iIntros "H". + iIntros (αs) "#Hctx Has". + iSpecialize ("H" with "Hctx Has"). + iApply (expr_pred_bind with "H"). + iIntros (αv) "Ha /=". + iDestruct "Ha" as (l) "[Ha Ha']". + iDestruct "Ha'" as (βv) "[Hl Hb]". + iRewrite "Ha". iApply expr_pred_frame. simpl. + rewrite IT_of_V_Ret. rewrite -> get_ret_ret. simpl. + iApply (wp_dealloc with "Hctx Hl"). + iNext. iNext. eauto with iFrame. + Qed. + + Lemma compat_bool {S : Set} + `{!EqDecision S} `{!Finite S} + b (Ω : S → ty) : + ⊢ valid1 Ω (interp_litbool b) tBool. + Proof. + iIntros (αs) "#Hctx Has". + iApply expr_pred_ret. + destruct b; simpl; eauto. + Qed. + + Lemma compat_nat {S : Set} + `{!EqDecision S} `{!Finite S} + n (Ω : S → ty) : + ⊢ valid1 Ω (interp_litnat n) tInt. + Proof. + iIntros (αs) "#Hctx Has". + iApply expr_pred_ret. eauto with iFrame. + Qed. + + Lemma compat_unit {S : Set} + `{!EqDecision S} `{!Finite S} + (Ω : S → ty) : + ⊢ valid1 Ω interp_litunit tUnit. + Proof. + iIntros (αs) "#Hctx Has". + iApply expr_pred_ret. eauto with iFrame. + Qed. + + Lemma compat_var {S : Set} + `{!EqDecision S} `{!Finite S} + Ω (v : S) : + ⊢ valid1 Ω (Force ◎ interp_var v) (Ω v). + Proof. + iIntros (ss) "#Hctx Has". + iIntros (x) "Hx". + unfold Force. + simpl. + iApply (wp_bind rs (AppRSCtx (ss v))). + { solve_proper. } + iApply wp_val. + iModIntro. + unfold AppRSCtx. + iApply (wp_bind rs (AppLSCtx (IT_of_V (RetV 0)))). + { solve_proper. } + unfold AppLSCtx. + simpl. + unfold ssubst_valid. + iDestruct (ssubst_valid_lookup _ _ v with "Has Hx") as "Has". + iApply (wp_wand with "Has"). + iIntros (w) "(%y & Hw1 & Hw2)"; simpl. + iModIntro. + rewrite IT_of_V_Ret. + iApply (wp_wand with "Hw1 [Hw2]"). + iIntros (z) "Hz". + iModIntro. + iExists y. + iFrame. + Qed. + + Lemma compat_app {S1 S2 : Set} + `{!EqDecision S1} `{!Finite S1} + `{!EqDecision S2} `{!Finite S2} + `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} + (Ω1 : S1 → ty) (Ω2 : S2 → ty) + α β τ1 τ2 : + ⊢ valid1 Ω1 α (tArr τ1 τ2) -∗ + valid1 Ω2 β τ1 -∗ + valid1 (sum_map' Ω1 Ω2) (interp_app α β ◎ interp_scope_split) τ2. + Proof. + iIntros "H1 H2". + iIntros (αs) "#Hctx Has". + iEval(cbn-[interp_app]). + rewrite ssubst_valid_app. + iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. + iSpecialize ("H1" with "Hctx Ha1"). + iSpecialize ("H2" with "Hctx Ha2"). + Local Opaque Thunk. + iSimpl. + iApply (expr_pred_bind (LETCTX _) with "H2"). + iIntros (βv) "H2". unfold LETCTX. iSimpl. + rewrite LET_Val/=. + iApply (expr_pred_bind (LETCTX _) with "H1"). + iIntros (αv) "H1". unfold LETCTX. iSimpl. + rewrite LET_Val/=. + by iApply "H1". + Qed. + + Lemma compat_lam {S : Set} + `{!EqDecision S} `{!Finite S} + (Ω : S → ty) τ1 τ2 α : + ⊢ valid1 (Ω ▹ τ1) α τ2 -∗ + valid1 Ω (interp_lam α) (tArr τ1 τ2). + Proof. + iIntros "H". + iIntros (αs) "#Hctx Has". + iIntros (x) "Hx". + iApply wp_val. + iModIntro. simpl. + iExists _; iFrame. + iIntros (βv) "Hb". clear x. + iIntros (x) "Hx". + iApply (wp_bind _ (AppRSCtx _)). + Local Transparent Thunk. + Local Opaque thunked thunkedV. + iSimpl. iApply (wp_alloc with "Hctx"). + { solve_proper. } + iNext. iNext. iIntros (l) "Hl". + iApply wp_val. iModIntro. + unfold AppRSCtx. + iApply wp_lam. iNext. + iEval(cbn-[thunked]). + pose (ss' := extend_scope αs (IT_of_V (thunkedV (IT_of_V βv) l))). + iSpecialize ("H" $! ss' + with "Hctx [-Hx] Hx"). + { + iApply ssubst_valid_cons. + iFrame "Has". + Local Transparent thunked thunkedV. + simpl. + iIntros (x') "Hx". + iApply wp_val. + iModIntro. + iExists x'. + iFrame "Hx". + iApply wp_lam. + iNext. + iApply (wp_bind _ (IFSCtx _ _)). + iApply (wp_read with "Hctx Hl"). + iNext. iNext. iIntros "Hl". + iApply wp_val. iModIntro. + unfold IFSCtx. simpl. + rewrite IF_False; last lia. + iApply wp_seq. + { solve_proper. } + iApply (wp_write with "Hctx Hl"). + iNext. iNext. iIntros "Hl". + iApply wp_val. iModIntro. + iApply "Hb". + } + iApply "H". + Qed. + + Lemma fundamental_affine (S : Set) + {HE : EqDecision S} {HF : Finite S} + (Ω : S → ty) + (e : expr S) τ : + typed Ω e τ → + ⊢ valid1 Ω (interp_expr _ e) τ. + Proof. + intros H. + induction H. + - by iApply compat_var. + - iApply compat_lam; + iApply IHtyped. + - iApply (@compat_app S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply (@compat_pair S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply (@compat_destruct S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply compat_alloc; + iApply IHtyped. + - iApply (@compat_replace S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply compat_dealloc; + iApply IHtyped. + - by iApply compat_nat. + - by iApply compat_bool. + - by iApply compat_unit. + Qed. + +End logrel. + +Arguments interp_tarr {_ _ _ _ _ _ _ _ _ _ _ _ _} Φ1 Φ2. +Arguments interp_tbool {_ _ _ _ _ _}. +Arguments interp_tnat {_ _ _ _ _ _}. +Arguments interp_tunit {_ _ _ _ _ _}. +Arguments interp_ty {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _} τ. + +Local Definition rs : gReifiers NotCtxDep 2 := + gReifiers_cons NotCtxDep reify_store (gReifiers_cons NotCtxDep input_lang.interp.reify_io (gReifiers_nil NotCtxDep)). + +Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. + +Require Import gitrees.gitree.greifiers. + +Program Definition ı_scope R `{!Cofe R} : @interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. + +Lemma logrel1_adequacy cr Σ R `{!Cofe R, !SubOfe natO R, !SubOfe unitO R, !SubOfe locO R} `{!invGpreS Σ} + `{!statePreG rs R Σ} `{!heapPreG rs R Σ} τ + (α : interp_scope ∅ -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : + (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ}, + (£ cr ⊢ valid1 rs notStuck (λne _: unitO, True)%I □ α τ)%I) → + ssteps (gReifiers_sReifier NotCtxDep rs) (α (ı_scope _)) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) + ∨ (∃ βv, IT_of_V βv ≡ β). +Proof. + intros Hlog Hst. + destruct (IT_to_V β) as [βv|] eqn:Hb. + { right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } + left. + cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) + ∨ (∃ e, β ≡ Err e ∧ notStuck e)). + { intros [?|He]; first done. + destruct He as [? [? []]]. } + eapply (wp_safety (S cr)); eauto. + { apply Hdisj. } + { by rewrite Hb. } + intros H1 H2. + exists (λ _, True)%I. split. + { apply _. } + iIntros "[[Hone Hcr] Hst]". + pose (σ := st.1). + pose (ios := st.2.1). + iMod (new_heapG rs σ) as (H3) "H". + iAssert (has_substate σ ∗ has_substate ios)%I with "[Hst]" as "[Hs Hio]". + { unfold has_substate, has_full_state. + assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) st ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ) + ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state ios)) as ->; last first. + { rewrite -own_op. done. } + unfold sR_idx. simpl. + intro j. + rewrite discrete_fun_lookup_op. + inv_fin j. + { unfold of_state, of_idx. simpl. + erewrite (eq_pi _ _ _ (@eq_refl _ 0%fin)). done. } + intros j. inv_fin j. + { unfold of_state, of_idx. simpl. + erewrite (eq_pi _ _ _ (@eq_refl _ 1%fin)). done. } + intros i. inversion i. } + iApply fupd_wp. + iMod (inv_alloc (nroot.@"storeE") _ (∃ σ, £ 1 ∗ has_substate σ ∗ own (heapG_name rs) (●V σ))%I with "[-Hcr]") as "#Hinv". + { iNext. iExists _. iFrame. } + simpl. + iPoseProof (@Hlog _ _ _ with "Hcr") as "Hlog". + iSpecialize ("Hlog" $! (ı_scope _) with "Hinv []"). + { iApply ssubst_valid_empty. } + iSpecialize ("Hlog" $! tt with "[//]"). + iModIntro. + iApply (wp_wand with "Hlog"). + eauto with iFrame. +Qed. + +Definition R := sumO locO (sumO unitO natO). + +Lemma logrel1_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : + typed □ e τ → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e (ı_scope _)) st β st' k → + (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) + ∨ (∃ βv, IT_of_V βv ≡ β). +Proof. + intros Hty Hst. + pose (Σ:=#[invΣ;stateΣ rs R;heapΣ rs R]). + eapply (logrel1_adequacy 0 Σ); eauto; try apply _. + iIntros (? ? ?) "_". + iApply (fundamental_affine rs notStuck (λne _ : unitO, True)%I). + apply Hty. +Qed. diff --git a/theories/affine_lang/logrel2.v b/theories/examples/affine_lang/logrel2.v similarity index 77% rename from theories/affine_lang/logrel2.v rename to theories/examples/affine_lang/logrel2.v index a45dac8..0c1cba4 100644 --- a/theories/affine_lang/logrel2.v +++ b/theories/examples/affine_lang/logrel2.v @@ -1,59 +1,57 @@ -From Equations Require Import Equations. From iris.base_logic.lib Require Import na_invariants. -From gitrees Require Export lang_affine gitree program_logic. -From gitrees.input_lang Require Import lang interp logpred. -From gitrees.affine_lang Require Import lang logrel1. -From gitrees.examples Require Import store pairs. +From gitrees Require Export gitree program_logic. +From gitrees.examples.input_lang Require Import lang interp logpred. +From gitrees.examples.affine_lang Require Import lang logrel1. +From gitrees.effects Require Import store. +From gitrees.lib Require Import pairs. Require Import iris.algebra.gmap. +Require Import stdpp.finite. Require Import Binding.Lib Binding.Set Binding.Env. -Local Notation tyctx := (tyctx ty). - -Inductive typed_glued : forall {S}, tyctx S → expr S → ty → Type := +Inductive typed_glued : forall {S : Set}, (S → ty) → expr S → ty → Type := (** FFI *) -| typed_Glue {S} (Ω : tyctx S) τ' τ e +| typed_Glue {S : Set} (Ω : S → ty) τ' τ e (tconv : ty_conv τ τ') : io_lang.typed □ e τ' → typed_glued Ω (EEmbed e tconv) τ (** functions *) -| typed_VarG {S} (Ω : tyctx S) (τ : ty) (v : var S) : - typed_var Ω v τ → - typed_glued Ω (Var v) τ -| typed_LamG {S} (Ω : tyctx S) (τ1 τ2 : ty) (e : expr (()::S) ) : - typed_glued (consC τ1 Ω) e τ2 → +| typed_VarG {S : Set} (Ω : S → ty) (τ : ty) (v : S) : + typed_glued Ω (Var v) (Ω v) +| typed_LamG {S : Set} (Ω : S → ty) (τ1 τ2 : ty) (e : expr (inc S) ) : + typed_glued (Ω ▹ τ1) e τ2 → typed_glued Ω (Lam e) (tArr τ1 τ2) -| typed_AppG {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 : ty) e1 e2 : +| typed_AppG {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 : ty) e1 e2 : typed_glued Ω1 e1 (tArr τ1 τ2) → typed_glued Ω2 e2 τ1 → - typed_glued (tyctx_app Ω1 Ω2) (App e1 e2) τ2 + typed_glued (sum_map' Ω1 Ω2) (App e1 e2) τ2 (** pairs *) -| typed_PairG {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 : ty) e1 e2 : +| typed_PairG {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 : ty) e1 e2 : typed_glued Ω1 e1 τ1 → typed_glued Ω2 e2 τ2 → - typed_glued (tyctx_app Ω1 Ω2) (EPair e1 e2) (tPair τ1 τ2) -| typed_DestructG {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 τ : ty) - (e1 : expr S1) (e2 : expr (()::()::S2)) : + typed_glued (sum_map' Ω1 Ω2) (EPair e1 e2) (tPair τ1 τ2) +| typed_DestructG {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 τ : ty) + (e1 : expr S1) (e2 : expr (inc (inc S2))) : typed_glued Ω1 e1 (tPair τ1 τ2) → - typed_glued (consC τ1 (consC τ2 Ω2)) e2 τ → - typed_glued (tyctx_app Ω1 Ω2) (EDestruct e1 e2) τ + typed_glued ((Ω2 ▹ τ2) ▹ τ1) e2 τ → + typed_glued (sum_map' Ω1 Ω2) (EDestruct e1 e2) τ (** references *) -| typed_AllocG {S} (Ω : tyctx S) τ e : +| typed_AllocG {S : Set} (Ω : S → ty) τ e : typed_glued Ω e τ → typed_glued Ω (Alloc e) (tRef τ) -| typed_ReplaceG {S1 S2} (Ω1 : tyctx S1) (Ω2 : tyctx S2) (τ1 τ2 : ty) e1 e2 : +| typed_ReplaceG {S1 S2 : Set} (Ω1 : S1 → ty) (Ω2 : S2 → ty) (τ1 τ2 : ty) e1 e2 : typed_glued Ω1 e1 (tRef τ1) → typed_glued Ω2 e2 τ2 → - typed_glued (tyctx_app Ω1 Ω2) (Replace e1 e2) (tPair τ1 (tRef τ2)) -| typed_DeallocG {S} (Ω : tyctx S) e τ : + typed_glued (sum_map' Ω1 Ω2) (Replace e1 e2) (tPair τ1 (tRef τ2)) +| typed_DeallocG {S : Set} (Ω : S → ty) e τ : typed_glued Ω e (tRef τ) → typed_glued Ω (Dealloc e) tUnit (** literals *) -| typed_NatG {S} (Ω : tyctx S) n : +| typed_NatG {S : Set} (Ω : S → ty) n : typed_glued Ω (LitNat n) tInt -| typed_BoolG {S} (Ω : tyctx S) b : +| typed_BoolG {S : Set} (Ω : S → ty) b : typed_glued Ω (LitBool b) tBool -| typed_UnitG {S} (Ω : tyctx S) : +| typed_UnitG {S : Set} (Ω : S → ty) : typed_glued Ω LitUnit tUnit . @@ -75,17 +73,19 @@ Section glue. Definition s : stuckness := λ e, e = OtherError. Variable p : na_inv_pool_name. - Definition valid2 {S} (Ω : tyctx S) (α : interp_scope (E:=F) S -n> IT) (τ : ty) : iProp := + Definition valid2 {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) (α : interp_scope (E:=F) S -n> IT) + (τ : ty) : iProp := valid1 rs s (λne σ, has_substate σ ∗ na_own p ⊤)%I Ω α τ. - Definition io_valid {S : Set} (Γ : S → io_lang.ty) α (τ' : io_lang.ty) : iProp := + Definition io_valid {S : Set} (Γ : S → io_lang.ty) α (τ' : io_lang.ty) + : iProp := input_lang.logpred.valid1 rs s (λne _ : unitO, na_own p ⊤) Γ α τ'. Local Opaque thunked thunkedV Thunk. Program Definition ı_scope : @lang_generic.interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. - Lemma compat_glue_to_affine_bool {S} (Ω : tyctx S) α : + Lemma compat_glue_to_affine_bool {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) α : io_valid □ α Tnat ⊢ valid2 Ω (constO (glue2_bool _ (α ı_scope))) tBool. Proof. @@ -97,7 +97,7 @@ Section glue. iIntros ([]). } iSpecialize ("H" $! tt with "Hp"). - simp interp_ssubst. simpl. + simpl. iApply (wp_bind _ (IFSCtx _ _)). { solve_proper. } iApply (wp_wand with "H"). @@ -112,7 +112,8 @@ Section glue. * rewrite IF_True ; last lia. iApply wp_val; eauto with iFrame. Qed. - Lemma compat_glue_to_affine_nat {S} (Ω : tyctx S) α : + + Lemma compat_glue_to_affine_nat {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) α : io_valid □ α Tnat ⊢ valid2 Ω (constO (α ı_scope)) tInt. Proof. @@ -124,32 +125,23 @@ Section glue. iIntros ([]). } iSpecialize ("H" $! tt with "Hp"). - simp interp_ssubst. simpl. + simpl. iApply (wp_wand with "H"). iIntros (αv). iDestruct 1 as (_) "[Ha Hp]". iDestruct "Ha" as (σ') "[Ha Hs]". iModIntro. eauto with iFrame. Qed. - Lemma IT_move_from_affine (α : @interp_scope (@gReifiers_ops NotCtxDep sz rs) R _ [] -n> IT) - : @lang_generic.interp_scope (@gReifiers_ops NotCtxDep sz rs) R _ ∅ -n> IT. - Proof. - unshelve econstructor. - - intros g. apply α. - constructor. - - repeat intro. - f_equiv. - Defined. - Lemma compat_glue_from_affine_bool α : - valid2 empC α tBool ⊢ - heap_ctx -∗ io_valid □ (IT_move_from_affine α) Tnat. + valid2 □ α tBool ⊢ + heap_ctx -∗ io_valid □ α Tnat. Proof. iIntros "H #Hctx". iIntros (σ ss) "Hs Hss". iIntros (?) "Hp". - iSpecialize ("H" $! emp_ssubst with "Hctx [] [$Hs $Hp]"). - { iApply ssubst_valid_nil. } + iSpecialize ("H" $! ss with "Hctx [] [$Hs $Hp]"). + { iApply ssubst_valid_empty. } + simpl. iApply (wp_wand with "H"). iIntros (αv) "Ha". iDestruct "Ha" as (σ') "[Ha [Hs Hp]]". iModIntro. simpl. iFrame. iExists tt,_; iFrame. @@ -157,22 +149,22 @@ Section glue. Qed. Lemma compat_glue_from_affine_nat α : - valid2 empC α tInt ⊢ - heap_ctx -∗ io_valid □ (IT_move_from_affine α) Tnat. + valid2 □ α tInt ⊢ + heap_ctx -∗ io_valid □ α Tnat. Proof. iIntros "H #Hctx". iIntros (σ ss) "Hs Hss". iIntros (?) "Hp". - iSpecialize ("H" $! emp_ssubst with "Hctx [] [$Hs $Hp]"). - { iApply ssubst_valid_nil. } + iSpecialize ("H" $! ss with "Hctx [] [$Hs $Hp]"). + { iApply ssubst_valid_empty. } iApply (wp_wand with "H"). iIntros (αv) "Ha". iDestruct "Ha" as (σ') "[Ha [Hs Hp]]". iModIntro. iExists tt. eauto with iFrame. Qed. Lemma compat_glue_from_affine_unit α : - valid2 empC α tUnit ⊢ - heap_ctx -∗ io_valid □ (constO (glue_from_affine _ ty_conv_unit (α ()))) Tnat. + valid2 □ α tUnit ⊢ + heap_ctx -∗ io_valid □ (constO (glue_from_affine _ ty_conv_unit (α ı_scope))) Tnat. Proof. iIntros "H #Hctx". iIntros (σ ss) "Hs Hss". @@ -187,10 +179,10 @@ Section glue. Lemma compat_glue_from_affine_fun (τ1 τ2 : ty) (τ1' τ2' : io_lang.ty) α (glue_to_affine glue_from_affine : IT -n> IT) : (∀ α, io_valid □ α τ1' - ⊢ valid2 empC (constO (glue_to_affine (α ı_scope))) τ1) → - (∀ α, valid2 empC (constO α) τ2 + ⊢ valid2 □ (constO (glue_to_affine (α ı_scope))) τ1) → + (∀ α, valid2 □ (constO α) τ2 ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine α)) τ2') → - valid2 empC (constO α) (tArr τ1 τ2) + valid2 □ (constO α) (tArr τ1 τ2) ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine_fun _ glue_from_affine glue_to_affine α)) @@ -203,8 +195,8 @@ Section glue. iIntros (σ ss) "Hs ?". simpl. iIntros (?) "Hp". - iSpecialize ("H" $! emp_ssubst with "Hctx [] [$Hs $Hp]"). - { iApply ssubst_valid_nil. } + iSpecialize ("H" $! ss with "Hctx [] [$Hs $Hp]"). + { iApply ssubst_valid_empty. } simpl. iApply wp_let. { solve_proper. } iApply (wp_wand with "H"). @@ -274,8 +266,8 @@ Section glue. iSpecialize ("G1" with "[Hg]"). { iIntros (σ0 ss0) "Hs Has". simpl. iApply expr_pred_ret. simpl. eauto with iFrame. } - iSpecialize ("G1" $! emp_ssubst with "Hctx [] [$Hst $Hp]"). - { iApply ssubst_valid_nil. } + iSpecialize ("G1" $! ss with "Hctx [] [$Hst $Hp]"). + { iApply ssubst_valid_empty. } iApply (wp_wand with "G1"). clear βv σ'. iIntros (βv). iDestruct 1 as (σ') "[Hb [Hst Hp]]". @@ -298,11 +290,11 @@ Section glue. iApply ("G1" $! tt with "Hp"). Qed. - Lemma compat_glue_to_affine_fun {S} (Ω : tyctx S) (τ1 τ2 : ty) + Lemma compat_glue_to_affine_fun {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) (τ1 τ2 : ty) (τ1' τ2' : io_lang.ty) α (glue_to_affine glue_from_affine : IT -n> IT) : (∀ α, io_valid □ α τ2' ⊢ valid2 Ω (constO (glue_to_affine (α ı_scope))) τ2) → - (∀ α, valid2 empC (constO α) τ1 + (∀ α, valid2 □ (constO α) τ1 ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine α)) τ1') → io_valid □ α (Tarr (Tarr Tnat τ1') τ2') ⊢ valid2 Ω @@ -314,11 +306,9 @@ Section glue. iIntros (αs) "#Hctx Has". iIntros (σ) "[Hs Hp]". simpl. iSpecialize ("H" $! _ ı_scope with "Hs []"). - { - iIntros ([]). - } + { iIntros ([]). } iSpecialize ("H" $! tt with "Hp"). - simp interp_ssubst. simpl. + simpl. iApply wp_let. { solve_proper. } iApply (wp_wand with "H"). @@ -332,7 +322,6 @@ Section glue. iSimpl. clear σ σ'. iIntros (σ) "[Hs Hp]". iApply (wp_bind _ (AppRSCtx _)). - { solve_proper. } iApply (wp_Thunk with "Hctx"). { solve_proper. } iNext. iIntros (l) "Hl". @@ -435,20 +424,20 @@ Section glue. simpl. done. Qed. - Lemma glue_to_affine_compatibility {S} (Ω : tyctx S) (τ1 : ty) (τ1' : io_lang.ty) + Lemma glue_to_affine_compatibility {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) (τ1 : ty) (τ1' : io_lang.ty) (Hconv : ty_conv τ1 τ1') α : io_valid □ α τ1' ⊢ valid2 Ω (constO (glue_to_affine _ Hconv (α ı_scope))) τ1 with glue_from_affine_compatibility (τ1 : ty) (τ1' : io_lang.ty) (Hconv : ty_conv τ1 τ1') (α : IT) : - valid2 empC (constO α) τ1 ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine _ Hconv α)) τ1'. + valid2 □ (constO α) τ1 ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine _ Hconv α)) τ1'. Proof. - destruct Hconv. - + by iApply compat_glue_to_affine_bool. - + by iApply compat_glue_to_affine_nat. + + by iApply (@compat_glue_to_affine_bool). + + by iApply (@compat_glue_to_affine_nat). + iIntros "_". simpl. iApply compat_unit. - + simpl. iApply compat_glue_to_affine_fun. - * apply glue_to_affine_compatibility. + + simpl. iApply (@compat_glue_to_affine_fun). + * by apply glue_to_affine_compatibility. * apply glue_from_affine_compatibility. - destruct Hconv. + iApply compat_glue_from_affine_bool. @@ -459,7 +448,7 @@ Section glue. * apply glue_from_affine_compatibility. Qed. - Lemma fundamental_affine_glued {S} (Ω : tyctx S) (e : expr S) τ : + Lemma fundamental_affine_glued {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) (e : expr S) τ : typed_glued Ω e τ → ⊢ valid2 Ω (interp_expr _ e) τ. Proof. @@ -467,13 +456,24 @@ Section glue. - iApply glue_to_affine_compatibility. by iApply fundamental. - by iApply compat_var. - - by iApply compat_lam. - - by iApply compat_app. - - by iApply compat_pair. - - by iApply compat_destruct. - - by iApply compat_alloc. - - by iApply compat_replace. - - by iApply compat_dealloc. + - iApply compat_lam. + iApply IHtyped. + - iApply (@compat_app _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply (@compat_pair _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply (@compat_destruct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply compat_alloc. + iApply IHtyped. + - iApply (@compat_replace _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + + iApply IHtyped1. + + iApply IHtyped2. + - iApply compat_dealloc. + iApply IHtyped. - by iApply compat_nat. - by iApply compat_bool. - by iApply compat_unit. @@ -491,10 +491,10 @@ Require Import gitrees.gitree.greifiers. Lemma logrel2_adequacy (cr : nat) R `{!Cofe R, !SubOfe locO R, !SubOfe natO R, !SubOfe unitO R} Σ `{!invGpreS Σ}`{!statePreG rs R Σ} `{!heapPreG rs R Σ} `{!na_invG Σ} - (τ : ty) (α : unitO -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : + (τ : ty) (α : interp_scope Empty_set -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ} p, - (£ cr ⊢ valid2 rs p empC α τ)%I) → - ssteps (gReifiers_sReifier NotCtxDep rs) (α ()) st β st' k → + (£ cr ⊢ valid2 rs p □ α τ)%I) → + ssteps (gReifiers_sReifier NotCtxDep rs) (α (ı_scope rs)) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (β ≡ Err OtherError) ∨ (∃ βv, IT_of_V βv ≡ β). @@ -538,11 +538,11 @@ Proof. simpl. iMod na_alloc as (p) "Hp". iPoseProof (@Hlog _ _ _ p with "Hcr") as "Hlog". - iSpecialize ("Hlog" $! emp_ssubst with "Hinv []"). - { iApply ssubst_valid_nil. } + iSpecialize ("Hlog" $! (ı_scope _) with "Hinv []"). + { iApply ssubst_valid_empty. } unfold expr_pred. simpl. iSpecialize ("Hlog" $! ios with "[$Hio $Hp]"). - iModIntro. simp interp_ssubst. + iModIntro. simpl. iApply (wp_wand with "Hlog"). eauto with iFrame. Qed. @@ -550,8 +550,8 @@ Qed. Definition R := sumO locO (sumO natO unitO). Lemma logrel2_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : - typed_glued empC e τ → - ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ()) st β st' k → + typed_glued □ e τ → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e (ı_scope _)) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (β ≡ Err OtherError) ∨ (∃ βv, IT_of_V βv ≡ β). diff --git a/theories/input_lang/interp.v b/theories/examples/input_lang/interp.v similarity index 99% rename from theories/input_lang/interp.v rename to theories/examples/input_lang/interp.v index 41bf8af..c0ae876 100644 --- a/theories/input_lang/interp.v +++ b/theories/examples/input_lang/interp.v @@ -1,5 +1,5 @@ From gitrees Require Import gitree lang_generic. -From gitrees.input_lang Require Import lang. +From gitrees.examples.input_lang Require Import lang. Require Import Binding.Lib. Require Import Binding.Set. diff --git a/theories/input_lang/lang.v b/theories/examples/input_lang/lang.v similarity index 100% rename from theories/input_lang/lang.v rename to theories/examples/input_lang/lang.v diff --git a/theories/input_lang/logpred.v b/theories/examples/input_lang/logpred.v similarity index 99% rename from theories/input_lang/logpred.v rename to theories/examples/input_lang/logpred.v index ea1b802..b548bc9 100644 --- a/theories/input_lang/logpred.v +++ b/theories/examples/input_lang/logpred.v @@ -1,6 +1,6 @@ (** Unary (Kripke) logical relation for the IO lang *) From gitrees Require Import gitree program_logic lang_generic. -From gitrees.input_lang Require Import lang interp. +From gitrees.examples.input_lang Require Import lang interp. Require Import Binding.Lib Binding.Set Binding.Env. Section io_lang. diff --git a/theories/input_lang/logrel.v b/theories/examples/input_lang/logrel.v similarity index 99% rename from theories/input_lang/logrel.v rename to theories/examples/input_lang/logrel.v index a1f0017..53f9756 100644 --- a/theories/input_lang/logrel.v +++ b/theories/examples/input_lang/logrel.v @@ -1,6 +1,6 @@ (** Logical relation for adequacy for the IO lang *) From gitrees Require Import gitree lang_generic. -From gitrees.input_lang Require Import lang interp. +From gitrees.examples.input_lang Require Import lang interp. Require Import Binding.Lib Binding.Set Binding.Env. Section logrel. diff --git a/theories/input_lang_callcc/hom.v b/theories/examples/input_lang_callcc/hom.v similarity index 98% rename from theories/input_lang_callcc/hom.v rename to theories/examples/input_lang_callcc/hom.v index 16d7ad0..3e4f03e 100644 --- a/theories/input_lang_callcc/hom.v +++ b/theories/examples/input_lang_callcc/hom.v @@ -1,5 +1,5 @@ From gitrees Require Import gitree. -From gitrees.input_lang_callcc Require Import lang interp. +From gitrees.examples.input_lang_callcc Require Import lang interp. Require Import gitrees.lang_generic. Require Import Binding.Lib Binding.Set Binding.Env. diff --git a/theories/input_lang_callcc/interp.v b/theories/examples/input_lang_callcc/interp.v similarity index 99% rename from theories/input_lang_callcc/interp.v rename to theories/examples/input_lang_callcc/interp.v index ae97290..8e41b93 100644 --- a/theories/input_lang_callcc/interp.v +++ b/theories/examples/input_lang_callcc/interp.v @@ -1,5 +1,5 @@ From gitrees Require Import gitree. -From gitrees.input_lang_callcc Require Import lang. +From gitrees.examples.input_lang_callcc Require Import lang. Require Import gitrees.lang_generic. Require Import Binding.Lib. diff --git a/theories/input_lang_callcc/lang.v b/theories/examples/input_lang_callcc/lang.v similarity index 100% rename from theories/input_lang_callcc/lang.v rename to theories/examples/input_lang_callcc/lang.v diff --git a/theories/input_lang_callcc/logrel.v b/theories/examples/input_lang_callcc/logrel.v similarity index 99% rename from theories/input_lang_callcc/logrel.v rename to theories/examples/input_lang_callcc/logrel.v index b88fac4..89d9c44 100644 --- a/theories/input_lang_callcc/logrel.v +++ b/theories/examples/input_lang_callcc/logrel.v @@ -1,6 +1,6 @@ (** Logical relation for adequacy for the IO lang *) From gitrees Require Import gitree. -From gitrees.input_lang_callcc Require Import lang interp hom. +From gitrees.examples.input_lang_callcc Require Import lang interp hom. Require Import gitrees.lang_generic. Require Import Binding.Lib Binding.Set Binding.Env. diff --git a/theories/lang_affine.v b/theories/lang_affine.v deleted file mode 100644 index b13c0be..0000000 --- a/theories/lang_affine.v +++ /dev/null @@ -1,245 +0,0 @@ -From gitrees Require Import prelude. -From gitrees Require Import gitree. -From Equations Require Import Equations. -Require Import List. -Import ListNotations. - -(** XXX: We /NEED/ this line for [Equations Derive] to work, *) -(* this flag is globally unset by std++, but Equations need obligations to be transparent. *) -Set Transparent Obligations. - -Derive NoConfusion NoConfusionHom for list. - -Definition scope := (list unit). - -(** Variables in a context *) -Inductive var : scope → Type := -| Vz : forall {S : scope} {s}, var (s::S) -| Vs : forall {S : scope} {s}, var S -> var (s::S) -. -Derive Signature NoConfusion for var. - -Inductive tyctx (ty : Type) : scope → Type := -| empC : tyctx ty [] -| consC : forall{Γ}, ty → tyctx ty Γ → tyctx ty (()::Γ) -. -Arguments empC {_}. -Arguments consC {_ _} _ _. - -Equations list_of_tyctx {S ty} (Γ : tyctx ty S) : list ty := - list_of_tyctx empC := []; - list_of_tyctx (consC τ Γ') := τ::list_of_tyctx Γ'. - -Equations tyctx_app {S1 S2 ty} (c1 : tyctx ty S1) (c2 : tyctx ty S2) : tyctx ty (S1++S2) := - tyctx_app empC c2 := c2; - tyctx_app (consC τ c1) c2 := consC τ (tyctx_app c1 c2). - -Inductive typed_var {ty : Type}: forall {S}, tyctx ty S → var S → ty → Prop := -| typed_var_Z S (τ : ty) (Γ : tyctx ty S) : - typed_var (consC τ Γ) Vz τ -| typed_var_S S (τ τ' : ty) (Γ : tyctx ty S) v : - typed_var Γ v τ → - typed_var (consC τ' Γ) (Vs v) τ -. - -Section interp. - Local Open Scope type. - Context {E: opsInterp}. - Context {R} `{!Cofe R}. - Notation IT := (IT E R). - Notation ITV := (ITV E R). - - Fixpoint interp_scope (S : scope) : ofe := - match S with - | [] => unitO - | τ::Sc => prodO IT (interp_scope Sc) - end. - - Instance interp_scope_cofe S : Cofe (interp_scope S). - Proof. induction S; simpl; apply _. Qed. - - Instance interp_scope_inhab S : Inhabited (interp_scope S). - Proof. induction S; simpl; apply _. Defined. - - Equations interp_var {S : scope} (v : var S) : interp_scope S -n> IT := - interp_var (S:=(_::_)) Vz := fstO; - interp_var (S:=(_::Sc)) (Vs v) := interp_var v ◎ sndO. - - Instance interp_var_ne S (v : var S) : NonExpansive (@interp_var S v). - Proof. - intros n D1 D2 HD12. induction v; simp interp_var. - - by f_equiv. - - eapply IHv. by f_equiv. - Qed. - - Global Instance interp_var_proper S (v : var S) : Proper ((≡) ==> (≡)) (interp_var v). - Proof. apply ne_proper. apply _. Qed. - - Definition interp_scope_split {S1 S2} : - interp_scope (S1 ++ S2) -n> interp_scope S1 * interp_scope S2. - Proof. - induction S1 as [|? S1]; simpl. - - simple refine (λne x, (tt, x)). - solve_proper. - - simple refine (λne xy, let ss := IHS1 xy.2 in ((xy.1, ss.1), ss.2)). - solve_proper. - Defined. - - (** scope substituions *) - Inductive ssubst : scope → Type := - | emp_ssubst : ssubst [] - | cons_ssubst {S} : ITV → ssubst S → ssubst (tt::S) - . - - Equations interp_ssubst {S} (ss : ssubst S) : interp_scope S := - interp_ssubst emp_ssubst := tt; - interp_ssubst (cons_ssubst αv ss) := (IT_of_V αv, interp_ssubst ss). - - Equations list_of_ssubst {S} (ss : ssubst S) : list ITV := - list_of_ssubst emp_ssubst := []; - list_of_ssubst (cons_ssubst αv ss) := αv::(list_of_ssubst ss). - - Equations ssubst_split {S1 S2} (αs : ssubst (S1++S2)) : ssubst S1 * ssubst S2 := - ssubst_split (S1:=[]) αs := (emp_ssubst,αs); - ssubst_split (S1:=u::_) (cons_ssubst αv αs) := - (cons_ssubst αv (ssubst_split αs).1, (ssubst_split αs).2). - Lemma interp_scope_ssubst_split {S1 S2} (αs : ssubst (S1++S2)) : - interp_scope_split (interp_ssubst αs) ≡ - (interp_ssubst (ssubst_split αs).1, interp_ssubst (ssubst_split αs).2). - Proof. - induction S1 as [|u S1]; simpl. - - simp ssubst_split. simpl. - simp interp_ssubst. done. - - dependent elimination αs as [cons_ssubst αv αs]. - simp ssubst_split. simpl. - simp interp_ssubst. repeat f_equiv; eauto; simpl. - + rewrite IHS1//. - + rewrite IHS1//. - Qed. - -End interp. - -(* Common definitions and lemmas for Kripke logical relations *) -Section kripke_logrel. - Variable s : stuckness. - - Context {sz : nat} {a : is_ctx_dep}. - Variable rs : gReifiers a sz. - Context {R} `{!Cofe R}. - - Notation F := (gReifiers_ops a rs). - Notation IT := (IT F R). - Notation ITV := (ITV F R). - Context `{!invGS Σ, !stateG rs R Σ}. - Notation iProp := (iProp Σ). - - Context {A:ofe}. (* The type & predicate for the explicit Kripke worlds *) - Variable (P : A -n> iProp). - - Implicit Types α β : IT. - Implicit Types αv βv : ITV. - Implicit Types Φ Ψ : ITV -n> iProp. - - Program Definition expr_pred (α : IT) (Φ : ITV -n> iProp) : iProp := - (∀ x : A, P x -∗ WP@{rs} α @ s {{ v, ∃ y : A, Φ v ∗ P y }}). - #[export] Instance expr_pred_ne : NonExpansive2 expr_pred. - Proof. solve_proper. Qed. - #[export] Instance expr_pred_proper : Proper ((≡) ==> (≡) ==> (≡)) expr_pred . - Proof. solve_proper. Qed. - - Definition ssubst_valid {ty} (interp_ty : ty → ITV -n> iProp) {S} (Γ : tyctx ty S) (ss : ssubst S) : iProp := - ([∗ list] τx ∈ zip (list_of_tyctx Γ) (list_of_ssubst (E:=F) ss), - interp_ty (τx.1) (τx.2))%I. - - Lemma ssubst_valid_nil {ty} (interp_ty : ty → ITV -n> iProp) : - ⊢ ssubst_valid interp_ty empC emp_ssubst. - Proof. - unfold ssubst_valid. - by simp list_of_tyctx list_of_ssubst. - Qed. - - Lemma ssubst_valid_cons {ty} (interp_ty : ty → ITV -n> iProp) {S} - (Γ : tyctx ty S) (ss : ssubst S) τ αv : - ssubst_valid interp_ty (consC τ Γ) (cons_ssubst αv ss) - ⊣⊢ interp_ty τ αv ∗ ssubst_valid interp_ty Γ ss. - Proof. - unfold ssubst_valid. - by simp list_of_tyctx list_of_ssubst. - Qed. - - Lemma ssubst_valid_app {ty} (interp_ty : ty → ITV -n> iProp) - {S1 S2} (Ω1 : tyctx ty S1) (Ω2 : tyctx ty S2) αs : - ssubst_valid interp_ty (tyctx_app Ω1 Ω2) αs ⊢ - ssubst_valid interp_ty Ω1 (ssubst_split αs).1 - ∗ ssubst_valid interp_ty Ω2 (ssubst_split αs).2. - Proof. - iInduction Ω1 as [|τ Ω1] "IH" forall (Ω2); simp tyctx_app ssubst_split. - - simpl. iIntros "$". iApply ssubst_valid_nil. - - iIntros "H". - rewrite {4 5}/ssubst_valid. - simpl in αs. - dependent elimination αs as [cons_ssubst αv αs]. - simp ssubst_split. simpl. - simp list_of_ssubst list_of_tyctx. - simpl. iDestruct "H" as "[$ H]". - by iApply "IH". - Qed. - - Lemma expr_pred_ret α αv Φ `{!IntoVal α αv} : - Φ αv ⊢ expr_pred α Φ. - Proof. - iIntros "H". - iIntros (x) "Hx". iApply wp_val. - eauto with iFrame. - Qed. - - Lemma expr_pred_frame α Φ : - WP@{rs} α @ s {{ Φ }} ⊢ expr_pred α Φ. - Proof. - iIntros "H". - iIntros (x) "Hx". - iApply (wp_wand with "H"). - eauto with iFrame. - Qed. -End kripke_logrel. - -Section kripke_logrel_ctx_indep. - Variable s : stuckness. - - Context {sz : nat}. - Variable rs : gReifiers NotCtxDep sz. - Context {R} `{!Cofe R}. - - Notation F := (gReifiers_ops NotCtxDep rs). - Notation IT := (IT F R). - Notation ITV := (ITV F R). - Context `{!invGS Σ, !stateG rs R Σ}. - Notation iProp := (iProp Σ). - - Context {A : ofe}. - Variable (P : A -n> iProp). - - Implicit Types α β : IT. - Implicit Types αv βv : ITV. - Implicit Types Φ Ψ : ITV -n> iProp. - - Local Notation expr_pred := (expr_pred s rs P). - - Lemma expr_pred_bind f `{!IT_hom f} α Φ Ψ `{!NonExpansive Φ} - : expr_pred α Ψ ⊢ - (∀ αv, Ψ αv -∗ expr_pred (f (IT_of_V αv)) Φ) - -∗ expr_pred (f α) Φ. - Proof. - iIntros "H1 H2". - iIntros (x) "Hx". - iApply wp_bind. - { solve_proper. } - iSpecialize ("H1" with "Hx"). - iApply (wp_wand with "H1"). - iIntros (βv). iDestruct 1 as (y) "[Hb Hy]". - iModIntro. - iApply ("H2" with "Hb Hy"). - Qed. -End kripke_logrel_ctx_indep. - -Arguments expr_pred_bind {_ _ _ _ _ _ _ _ _ _} f {_ _}. diff --git a/theories/examples/factorial.v b/theories/lib/factorial.v similarity index 97% rename from theories/examples/factorial.v rename to theories/lib/factorial.v index 3d2cc4e..259395e 100644 --- a/theories/examples/factorial.v +++ b/theories/lib/factorial.v @@ -1,7 +1,7 @@ -From Equations Require Import Equations. From gitrees Require Import gitree program_logic. -From gitrees.input_lang Require Import lang interp. -From gitrees.examples Require Import store while. +From gitrees.examples.input_lang Require Import lang interp. +From gitrees.effects Require Import store. +From gitrees.lib Require Import while. Section fact. Definition rs : gReifiers NotCtxDep 2 := diff --git a/theories/examples/iter.v b/theories/lib/iter.v similarity index 100% rename from theories/examples/iter.v rename to theories/lib/iter.v diff --git a/theories/examples/pairs.v b/theories/lib/pairs.v similarity index 100% rename from theories/examples/pairs.v rename to theories/lib/pairs.v diff --git a/theories/examples/while.v b/theories/lib/while.v similarity index 100% rename from theories/examples/while.v rename to theories/lib/while.v From ddddc3f0878bc70942f1fdd1a65f3a08acaa41cb Mon Sep 17 00:00:00 2001 From: Kaptch Date: Thu, 1 Feb 2024 12:56:12 +0100 Subject: [PATCH 5/7] minor proof simplification (fundamental for affine) --- theories/examples/affine_lang/logrel1.v | 28 ++++++++---------------- theories/examples/affine_lang/logrel2.v | 29 +++++++++---------------- 2 files changed, 19 insertions(+), 38 deletions(-) diff --git a/theories/examples/affine_lang/logrel1.v b/theories/examples/affine_lang/logrel1.v index 50b2fcd..edecf4b 100644 --- a/theories/examples/affine_lang/logrel1.v +++ b/theories/examples/affine_lang/logrel1.v @@ -816,26 +816,16 @@ Section logrel. ⊢ valid1 Ω (interp_expr _ e) τ. Proof. intros H. - induction H. + iStartProof. + iInduction H as [| | | | | | | | | |] "IH". - by iApply compat_var. - - iApply compat_lam; - iApply IHtyped. - - iApply (@compat_app S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply (@compat_pair S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply (@compat_destruct S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply compat_alloc; - iApply IHtyped. - - iApply (@compat_replace S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply compat_dealloc; - iApply IHtyped. + - by iApply compat_lam. + - by iApply (@compat_app S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (@compat_pair S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (@compat_destruct S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply compat_alloc. + - by iApply (@compat_replace S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply compat_dealloc. - by iApply compat_nat. - by iApply compat_bool. - by iApply compat_unit. diff --git a/theories/examples/affine_lang/logrel2.v b/theories/examples/affine_lang/logrel2.v index 0c1cba4..e855def 100644 --- a/theories/examples/affine_lang/logrel2.v +++ b/theories/examples/affine_lang/logrel2.v @@ -452,28 +452,19 @@ Section glue. typed_glued Ω e τ → ⊢ valid2 Ω (interp_expr _ e) τ. Proof. - intros typed. induction typed; simpl. + intros typed. + iStartProof. + iInduction typed as [| | | | | | | | | | |] "IH". - iApply glue_to_affine_compatibility. by iApply fundamental. - by iApply compat_var. - - iApply compat_lam. - iApply IHtyped. - - iApply (@compat_app _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply (@compat_pair _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply (@compat_destruct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply compat_alloc. - iApply IHtyped. - - iApply (@compat_replace _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - + iApply IHtyped1. - + iApply IHtyped2. - - iApply compat_dealloc. - iApply IHtyped. + - by iApply compat_lam. + - by iApply (@compat_app _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (@compat_pair _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (@compat_destruct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply compat_alloc. + - by iApply (@compat_replace _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply compat_dealloc. - by iApply compat_nat. - by iApply compat_bool. - by iApply compat_unit. From 12c70c5c632c04696d2efdc0689b54788eed1edd Mon Sep 17 00:00:00 2001 From: Kaptch Date: Thu, 1 Feb 2024 22:00:15 +0100 Subject: [PATCH 6/7] cleanup --- _CoqProject | 2 + theories/effects/store.v | 3 +- theories/examples/affine_lang/lang.v | 7 +- theories/examples/affine_lang/logrel1.v | 416 ++----------------- theories/examples/affine_lang/logrel2.v | 75 ++-- theories/examples/input_lang/interp.v | 27 +- theories/examples/input_lang/logpred.v | 30 +- theories/examples/input_lang/logrel.v | 19 +- theories/examples/input_lang_callcc/hom.v | 2 +- theories/examples/input_lang_callcc/interp.v | 4 +- theories/examples/input_lang_callcc/logrel.v | 26 +- theories/gitree/greifiers.v | 202 ++++----- theories/gitree/reductions.v | 12 +- theories/gitree/reify.v | 3 + theories/gitree/weakestpre.v | 272 ++++++------ theories/lang_generic.v | 168 +++++++- theories/lib/factorial.v | 7 +- theories/lib/iter.v | 2 +- theories/prelude.v | 3 + theories/program_logic.v | 4 +- theories/utils/finite_sets.v | 221 ++++++++++ 21 files changed, 777 insertions(+), 728 deletions(-) create mode 100644 theories/utils/finite_sets.v diff --git a/_CoqProject b/_CoqProject index 4ef1b32..6ff13f8 100644 --- a/_CoqProject +++ b/_CoqProject @@ -48,3 +48,5 @@ theories/lib/pairs.v theories/lib/while.v theories/lib/factorial.v theories/lib/iter.v + +theories/utils/finite_sets.v \ No newline at end of file diff --git a/theories/effects/store.v b/theories/effects/store.v index 6b72ad6..7027cba 100644 --- a/theories/effects/store.v +++ b/theories/effects/store.v @@ -125,14 +125,13 @@ Section constructors. (λne _, Next (Ret ())). End constructors. - Section wp. Context {n : nat}. Variable (rs : gReifiers NotCtxDep n). Context {R} `{!Cofe R}. Context `{!SubOfe unitO R}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Notation stateO := (stateF ♯ IT). diff --git a/theories/examples/affine_lang/lang.v b/theories/examples/affine_lang/lang.v index 3e5505f..628d2e9 100644 --- a/theories/examples/affine_lang/lang.v +++ b/theories/examples/affine_lang/lang.v @@ -10,9 +10,8 @@ Module io_lang. Definition expr := input_lang.lang.expr. Definition tyctx {S : Set} := S → ty. Definition typed {S : Set} := input_lang.lang.typed (S:=S). - Program Definition ı_scope {sz} (rs : gReifiers NotCtxDep sz) `{!subReifier reify_io rs} {R} `{!Cofe R} : @interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. - Definition interp_closed {sz} (rs : gReifiers NotCtxDep sz) `{!subReifier reify_io rs} (e : expr ∅) {R} `{!Cofe R, !SubOfe natO R} : IT (gReifiers_ops NotCtxDep rs) R := - input_lang.interp.interp_expr rs e (ı_scope rs). + Definition interp_closed {sz} (rs : gReifiers NotCtxDep sz) `{!subReifier reify_io rs} (e : expr ∅) {R} `{!Cofe R, !SubOfe natO R} : IT (gReifiers_ops rs) R := + input_lang.interp.interp_expr rs e ı_scope. End io_lang. Require Import Binding.Resolver Binding.Lib Binding.Set Binding.Auto Binding.Env. @@ -51,7 +50,7 @@ Section affine. Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_store rs}. Context `{!subReifier reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Context {R : ofe}. Context `{!Cofe R, !SubOfe unitO R, !SubOfe natO R, !SubOfe locO R}. Notation IT := (IT F R). diff --git a/theories/examples/affine_lang/logrel1.v b/theories/examples/affine_lang/logrel1.v index edecf4b..d26e888 100644 --- a/theories/examples/affine_lang/logrel1.v +++ b/theories/examples/affine_lang/logrel1.v @@ -1,221 +1,10 @@ (** Unary (Kripke) logical relation for the affine lang *) -From gitrees Require Export gitree program_logic. +Require Import iris.algebra.gmap. +From gitrees Require Export gitree program_logic greifiers. From gitrees.examples.affine_lang Require Import lang. From gitrees.effects Require Import store. From gitrees.lib Require Import pairs. -Require Import iris.algebra.gmap. -Require Import stdpp.finite. - -Require Import Binding.Resolver Binding.Lib Binding.Set Binding.Auto Binding.Env. - -Lemma fin_to_set_sum {S1 S2 : Set} `{EqDecision S1} `{EqDecision S2} `{EqDecision (S1 + S2)} - `{Finite S1} `{Finite S2} `{Finite (S1 + S2)} - `{Countable S1} `{Countable S2} `{Countable (S1 + S2)} - : let A1 : gset S1 := (fin_to_set S1) in - let A2 : gset (S1 + S2) := set_map (inl : S1 → S1 + S2) A1 in - let B1 : gset S2 := (fin_to_set S2) in - let B2 : gset (S1 + S2) := set_map (inr : S2 → S1 + S2) B1 in - let C : gset (S1 + S2) := fin_to_set (S1 + S2) in - C = A2 ∪ B2. -Proof. - apply set_eq. - intros [x|x]; simpl; split; intros _. - - apply elem_of_union; left. - apply elem_of_map_2. - apply elem_of_fin_to_set. - - apply elem_of_fin_to_set. - - apply elem_of_union; right. - apply elem_of_map_2. - apply elem_of_fin_to_set. - - apply elem_of_fin_to_set. -Qed. - -Lemma fin_to_set_empty : - let A : gset ∅ := fin_to_set ∅ in - let B : gset ∅ := empty in - A = B. -Proof. - apply set_eq; intros []. -Qed. - -Section InstSum. - Global Instance EqDecisionLeft {S1 S2 : Set} {H : EqDecision (S1 + S2)} : EqDecision S1. - Proof. - intros x y. - destruct (decide (inl x = inl y)) as [G | G]; - [left; by inversion G | right; intros C; by subst]. - Qed. - - Global Instance EqDecisionRight {S1 S2 : Set} {H : EqDecision (S1 + S2)} : EqDecision S2. - Proof. - intros x y. - destruct (decide (inr x = inr y)) as [G | G]; - [left; by inversion G | right; intros C; by subst]. - Qed. - - Global Instance FiniteLeft {S1 S2 : Set} `{EqDecision S1} - `{EqDecision (S1 + S2)} `{Finite (S1 + S2)} - : Finite S1. - Proof. - unshelve econstructor. - - apply (foldr (λ x acc, match x with - | inl x => x :: acc - | inr _ => acc - end) [] (enum (S1 + S2))). - - set (l := enum (S1 + S2)). - assert (NoDup l) as K; first apply NoDup_enum. - clearbody l. - induction l as [| a l IH]; first constructor. - destruct a as [a | a]; simpl. - + constructor. - * intros C. - assert (inl a ∈ l) as C'. - { - clear -C. - induction l as [| b l IH]; first inversion C. - destruct b as [b | b]; simpl. - - rewrite foldr_cons in C. - rewrite elem_of_cons in C. - destruct C as [-> | C]. - + apply elem_of_cons. - by left. - + right. - apply IH. - apply C. - - apply elem_of_cons. - right. - rewrite foldr_cons in C. - apply IH. - apply C. - } - by inversion K. - * apply IH. - by inversion K. - + apply IH. - by inversion K. - - intros x. - set (l := enum (S1 + S2)). - assert (inl x ∈ l) as K; first apply elem_of_enum. - clearbody l. - induction l as [| a l IH]; first inversion K. - destruct a as [a | a]; simpl. - + rewrite elem_of_cons in K. - destruct K as [K | K]. - * inversion K; subst. - apply elem_of_cons; by left. - * apply elem_of_cons; right; by apply IH. - + rewrite elem_of_cons in K. - destruct K as [K | K]; first inversion K. - by apply IH. - Qed. - - Global Instance FiniteRight {S1 S2 : Set} `{EqDecision S2} - `{EqDecision (S1 + S2)} `{H : Finite (S1 + S2)} - : Finite S2. - Proof. - unshelve econstructor. - - apply (foldr (λ x acc, match x with - | inl _ => acc - | inr x => x :: acc - end) [] (enum (S1 + S2))). - - set (l := enum (S1 + S2)). - assert (NoDup l) as K; first apply NoDup_enum. - clearbody l. - induction l as [| a l IH]; first constructor. - destruct a as [a | a]; simpl. - + apply IH. - by inversion K. - + constructor. - * intros C. - assert (inr a ∈ l) as C'. - { - clear -C. - induction l as [| b l IH]; first inversion C. - destruct b as [b | b]; simpl. - - apply elem_of_cons. - right. - rewrite foldr_cons in C. - apply IH. - apply C. - - rewrite foldr_cons in C. - rewrite elem_of_cons in C. - destruct C as [-> | C]. - + apply elem_of_cons. - by left. - + right. - apply IH. - apply C. - } - by inversion K. - * apply IH. - by inversion K. - - intros x. - set (l := enum (S1 + S2)). - assert (inr x ∈ l) as K; first apply elem_of_enum. - clearbody l. - induction l as [| a l IH]; first inversion K. - destruct a as [a | a]; simpl. - + rewrite elem_of_cons in K. - destruct K as [K | K]; first inversion K. - by apply IH. - + rewrite elem_of_cons in K. - destruct K as [K | K]. - * inversion K; subst. - apply elem_of_cons; by left. - * apply elem_of_cons; right; by apply IH. - Qed. - -End InstSum. - -Section InstInc. - Context (S : Set). - - Global Instance EqDecisionIncN {HS : EqDecision S} (n : nat) : EqDecision (Init.Nat.iter n inc S). - Proof using S. - induction n; simpl. - - apply _. - - intros [|x] [|y]. - + by left. - + by right. - + by right. - + destruct (decide (x = y)) as [-> |]; first by left. - right; by inversion 1. - Qed. - - Global Instance EqDecisionInc {HS : EqDecision S} : EqDecision (inc S). - Proof using S. - assert (inc S = Init.Nat.iter 1 inc S) as ->; first done. - by apply EqDecisionIncN. - Qed. - - Global Instance FiniteIncN {HS : EqDecision S} (HF : Finite S) (n : nat) {HS' : EqDecision (Init.Nat.iter n inc S)} : Finite (Init.Nat.iter n inc S). - Proof using S. - induction n. - - apply (@surjective_finite S HS HF _ _ id). - apply _. - - simpl. - unshelve eapply (@surjective_finite (option (Init.Nat.iter n inc S))); simpl in *. - + intros [x |]. - * apply (VS x). - * apply VZ. - + apply _. - + intros [| x]; simpl. - * exists None; reflexivity. - * exists (Some x); reflexivity. - Qed. - - Global Instance FiniteInc {HS : EqDecision S} (HF : Finite S) (HE : EqDecision (inc S)) : Finite (inc S). - Proof using S. - assert (J : @Finite (Init.Nat.iter 1 inc S) HE). - { apply FiniteIncN, HF. } - simpl in J. - apply J. - Qed. - -End InstInc. - -Definition sum_map' {A B C : Set} (f : A → C) (g : B → C) : sum A B → C := - λ x, match x with | inl x' => f x' | inr x' => g x' end. +From gitrees.utils Require Import finite_sets. Inductive typed : forall {S : Set}, (S → ty) → expr S → ty → Prop := (** functions *) @@ -263,7 +52,7 @@ Section logrel. Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_store rs}. Context `{!subReifier input_lang.interp.reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Context `{!SubOfe unitO R}. @@ -314,10 +103,7 @@ Section logrel. | tRef τ => interp_ref (interp_ty τ) end. - Program Definition ssubst_valid {S : Set} `{!EqDecision S} `{!Finite S} - (Ω : S → ty) (ss : interp_scope S) : iProp - := ([∗ set] x ∈ (fin_to_set S), - (expr_pred (ss x) (protected (interp_ty (Ω x))))%I). + Notation ssubst_valid := (ssubst_valid_fin1 rs ty (λ x, protected (interp_ty x)) expr_pred). Definition valid1 {S : Set} `{!EqDecision S} `{!Finite S} (Ω : S → ty) (α : interp_scope S -n> IT) (τ : ty) : iProp := @@ -325,123 +111,9 @@ Section logrel. -∗ (ssubst_valid Ω ss) -∗ expr_pred (α ss) (interp_ty τ). - Lemma ssubst_valid_empty (αs : interp_scope ∅) : - ⊢ ssubst_valid □ αs. - Proof. - iStartProof. - unfold ssubst_valid. - match goal with - | |- context G [big_opS ?a ?b ?c] => assert (c = empty) as -> - end. - { apply set_eq; intros []. } - by iApply big_sepS_empty. - Qed. - - Lemma ssubst_valid_app - {S1 S2 : Set} `{!EqDecision S1} `{!Finite S1} - `{!EqDecision S2} `{!Finite S2} - `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} - (Ω1 : S1 → ty) (Ω2 : S2 → ty) - (αs : interp_scope (sum S1 S2)) : - (ssubst_valid (sum_map' Ω1 Ω2) αs) ⊢ - (ssubst_valid Ω1 (interp_scope_split αs).1) - ∗ (ssubst_valid Ω2 (interp_scope_split αs).2). - Proof. - iIntros "H". - rewrite /ssubst_valid fin_to_set_sum big_sepS_union; first last. - { - apply elem_of_disjoint. - intros [x | x]. - - rewrite !elem_of_list_to_set. - intros _ H2. - apply elem_of_list_fmap_2 in H2. - destruct H2 as [y [H2 H2']]; inversion H2. - - rewrite !elem_of_list_to_set. - intros H1 _. - apply elem_of_list_fmap_2 in H1. - destruct H1 as [y [H1 H1']]; inversion H1. - } - iDestruct "H" as "(H1 & H2)". - iSplitL "H1". - - rewrite big_opS_list_to_set; first last. - { - apply NoDup_fmap. - - intros ??; by inversion 1. - - apply NoDup_elements. - } - rewrite big_sepL_fmap /=. - rewrite big_sepS_elements. - iFrame "H1". - - rewrite big_opS_list_to_set; first last. - { - apply NoDup_fmap. - - intros ??; by inversion 1. - - apply NoDup_elements. - } - rewrite big_sepL_fmap /=. - rewrite big_sepS_elements. - iFrame "H2". - Qed. - - Lemma ssubst_valid_cons {S : Set} `{!EqDecision S} `{!Finite S} - (Ω : S → ty) (αs : interp_scope S) τ t : - ssubst_valid Ω αs ∗ expr_pred t (protected (interp_ty τ)) ⊢ ssubst_valid (Ω ▹ τ) (extend_scope αs t). - Proof. - iIntros "(H & G)". - rewrite /ssubst_valid. - pose (Y := let A := {[VZ]} : @gset (leibnizO (inc S)) _ finite_countable in - let B := fin_to_set (leibnizO S) : gset (leibnizO S) in - let C := set_map (VS : S → inc S) B - : gset (inc S) in A ∪ C). - assert (fin_to_set (inc S) = Y) as ->. - { - subst Y; simpl. - apply set_eq. - intros [| x]. - - split. - + intros _; apply elem_of_union; left. - by apply elem_of_singleton. - + intros _; apply elem_of_fin_to_set. - - split. - + intros _; apply elem_of_union; right. - apply elem_of_map_2, elem_of_fin_to_set. - + intros H. - apply elem_of_fin_to_set. - } - subst Y; simpl. - rewrite big_sepS_union; first last. - { - apply elem_of_disjoint. - intros [| x]. - - rewrite !elem_of_list_to_set. - intros _ H2. - apply elem_of_list_fmap_2 in H2. - destruct H2 as [y [H2 H2']]; inversion H2. - - rewrite !elem_of_list_to_set. - intros H1 _. - apply elem_of_singleton_1 in H1. - inversion H1. - } - iSplitL "G". - - rewrite big_opS_singleton. - iFrame "G". - - erewrite big_opS_set_map. - + iFrame "H". - + intros ?? H; by inversion H. - Qed. - - Lemma ssubst_valid_lookup {S : Set} `{!EqDecision S} `{!Finite S} - (Ω : S → ty) (αs : interp_scope S) x : - ssubst_valid Ω αs ⊢ expr_pred (αs x) (protected (interp_ty (Ω x))). - Proof. - iIntros "H". - iDestruct (big_sepS_elem_of_acc _ _ x with "H") as "($ & _)"; - first apply elem_of_fin_to_set. - Qed. - Lemma compat_pair {S1 S2 : Set} - `{!EqDecision S1} `{!Finite S1} - `{!EqDecision S2} `{!Finite S2} + `(!EqDecision S1) `(!Finite S1) + `(!EqDecision S2) `(!Finite S2) `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} (Ω1 : S1 → ty) (Ω2 : S2 → ty) α β τ1 τ2 : ⊢ valid1 Ω1 α τ1 -∗ @@ -452,7 +124,7 @@ Section logrel. iIntros "H1 H2". iIntros (αs) "#Hctx Has". cbn-[interp_pair]. - rewrite ssubst_valid_app. + rewrite ssubst_valid_fin_app1. iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. iSpecialize ("H1" with "Hctx Ha1"). iSpecialize ("H2" with "Hctx Ha2"). @@ -469,8 +141,8 @@ Section logrel. Qed. Lemma compat_destruct {S1 S2 : Set} - `{!EqDecision S1} `{!Finite S1} - `{!EqDecision S2} `{!Finite S2} + `(!EqDecision S1) `(!Finite S1) + `(!EqDecision S2) `(!Finite S2) `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} (Ω1 : S1 → ty) (Ω2 : S2 → ty) α β τ1 τ2 τ : @@ -482,7 +154,7 @@ Section logrel. iIntros "H1 H2". iIntros (αs) "#Hctx Has". cbn-[interp_destruct]. - rewrite ssubst_valid_app. + rewrite ssubst_valid_fin_app1. iDestruct "Has" as "[Ha1 Ha2]". iSpecialize ("H1" with "Hctx Ha1"). iApply (expr_pred_bind (LETCTX _) with "H1"). @@ -519,9 +191,9 @@ Section logrel. iSpecialize ("H2" $! ss' with "Hctx [-Hx] Hx"). { - iApply ssubst_valid_cons. + iApply ssubst_valid_fin_cons1. iSplitR "Hl1 Hb1". - - iApply ssubst_valid_cons. + - iApply ssubst_valid_fin_cons1. iSplitL "Ha2"; first done. Transparent thunkedV thunked. simpl. @@ -603,8 +275,8 @@ Section logrel. Qed. Lemma compat_replace {S1 S2 : Set} - `{!EqDecision S1} `{!Finite S1} - `{!EqDecision S2} `{!Finite S2} + `(!EqDecision S1) `(!Finite S1) + `(!EqDecision S2) `(!Finite S2) `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} (Ω1 : S1 → ty) (Ω2 : S2 → ty) α β τ τ' : ⊢ valid1 Ω1 α (tRef τ) -∗ @@ -615,7 +287,7 @@ Section logrel. iIntros "H1 H2". iIntros (αs) "#Hctx Has". cbn-[interp_replace]. - rewrite ssubst_valid_app. + rewrite ssubst_valid_fin_app1. iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. iSpecialize ("H1" with "Hctx Ha1"). iSpecialize ("H2" with "Hctx Ha2"). @@ -712,8 +384,7 @@ Section logrel. { solve_proper. } unfold AppLSCtx. simpl. - unfold ssubst_valid. - iDestruct (ssubst_valid_lookup _ _ v with "Has Hx") as "Has". + iDestruct (ssubst_valid_fin_lookup1 _ _ _ _ _ _ v with "Has Hx") as "Has". iApply (wp_wand with "Has"). iIntros (w) "(%y & Hw1 & Hw2)"; simpl. iModIntro. @@ -726,8 +397,8 @@ Section logrel. Qed. Lemma compat_app {S1 S2 : Set} - `{!EqDecision S1} `{!Finite S1} - `{!EqDecision S2} `{!Finite S2} + `(!EqDecision S1) `(!Finite S1) + `(!EqDecision S2) `(!Finite S2) `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} (Ω1 : S1 → ty) (Ω2 : S2 → ty) α β τ1 τ2 : @@ -738,7 +409,7 @@ Section logrel. iIntros "H1 H2". iIntros (αs) "#Hctx Has". iEval(cbn-[interp_app]). - rewrite ssubst_valid_app. + rewrite ssubst_valid_fin_app1. iDestruct "Has" as "[Ha1 Ha2]". cbn-[interp_app]. iSpecialize ("H1" with "Hctx Ha1"). iSpecialize ("H2" with "Hctx Ha2"). @@ -781,7 +452,7 @@ Section logrel. iSpecialize ("H" $! ss' with "Hctx [-Hx] Hx"). { - iApply ssubst_valid_cons. + iApply ssubst_valid_fin_cons1. iFrame "Has". Local Transparent thunked thunkedV. simpl. @@ -809,7 +480,7 @@ Section logrel. Qed. Lemma fundamental_affine (S : Set) - {HE : EqDecision S} {HF : Finite S} + (HE : EqDecision S) (HF : Finite S) (Ω : S → ty) (e : expr S) τ : typed Ω e τ → @@ -820,11 +491,11 @@ Section logrel. iInduction H as [| | | | | | | | | |] "IH". - by iApply compat_var. - by iApply compat_lam. - - by iApply (@compat_app S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - - by iApply (@compat_pair S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - - by iApply (@compat_destruct S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_app EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_pair EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_destruct EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - by iApply compat_alloc. - - by iApply (@compat_replace S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_replace EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - by iApply compat_dealloc. - by iApply compat_nat. - by iApply compat_bool. @@ -839,30 +510,31 @@ Arguments interp_tnat {_ _ _ _ _ _}. Arguments interp_tunit {_ _ _ _ _ _}. Arguments interp_ty {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _} τ. +Arguments compat_app {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _}. +Arguments compat_pair {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _}. +Arguments compat_destruct {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _}. +Arguments compat_replace {_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _}. + Local Definition rs : gReifiers NotCtxDep 2 := - gReifiers_cons NotCtxDep reify_store (gReifiers_cons NotCtxDep input_lang.interp.reify_io (gReifiers_nil NotCtxDep)). + gReifiers_cons reify_store (gReifiers_cons input_lang.interp.reify_io gReifiers_nil). Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. -Require Import gitrees.gitree.greifiers. - -Program Definition ı_scope R `{!Cofe R} : @interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. - Lemma logrel1_adequacy cr Σ R `{!Cofe R, !SubOfe natO R, !SubOfe unitO R, !SubOfe locO R} `{!invGpreS Σ} `{!statePreG rs R Σ} `{!heapPreG rs R Σ} τ - (α : interp_scope ∅ -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : + (α : interp_scope ∅ -n> IT (gReifiers_ops rs) R) (β : IT (gReifiers_ops rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ}, (£ cr ⊢ valid1 rs notStuck (λne _: unitO, True)%I □ α τ)%I) → - ssteps (gReifiers_sReifier NotCtxDep rs) (α (ı_scope _)) st β st' k → + ssteps (gReifiers_sReifier NotCtxDep rs) (α ı_scope) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ βv, IT_of_V βv ≡ β). + ∨ (∃ βv, (IT_of_V βv ≡ β)%stdpp). Proof. intros Hlog Hst. destruct (IT_to_V β) as [βv|] eqn:Hb. { right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } left. cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ e, β ≡ Err e ∧ notStuck e)). + ∨ (∃ e, (β ≡ Err e)%stdpp ∧ notStuck e)). { intros [?|He]; first done. destruct He as [? [? []]]. } eapply (wp_safety (S cr)); eauto. @@ -877,9 +549,9 @@ Proof. iMod (new_heapG rs σ) as (H3) "H". iAssert (has_substate σ ∗ has_substate ios)%I with "[Hst]" as "[Hs Hio]". { unfold has_substate, has_full_state. - assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) st ≡ - of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ) - ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state ios)) as ->; last first. + assert (of_state NotCtxDep rs (IT (gReifiers_ops rs) _) st ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops rs) _) sR_idx (sR_state σ) + ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops rs) _) sR_idx (sR_state ios))%stdpp as ->; last first. { rewrite -own_op. done. } unfold sR_idx. simpl. intro j. @@ -896,8 +568,8 @@ Proof. { iNext. iExists _. iFrame. } simpl. iPoseProof (@Hlog _ _ _ with "Hcr") as "Hlog". - iSpecialize ("Hlog" $! (ı_scope _) with "Hinv []"). - { iApply ssubst_valid_empty. } + iSpecialize ("Hlog" $! ı_scope with "Hinv []"). + { iApply ssubst_valid_fin_empty1. } iSpecialize ("Hlog" $! tt with "[//]"). iModIntro. iApply (wp_wand with "Hlog"). @@ -906,11 +578,11 @@ Qed. Definition R := sumO locO (sumO unitO natO). -Lemma logrel1_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : +Lemma logrel1_safety e τ (β : IT (gReifiers_ops rs) R) st st' k : typed □ e τ → - ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e (ı_scope _)) st β st' k → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ı_scope) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ βv, IT_of_V βv ≡ β). + ∨ (∃ βv, (IT_of_V βv ≡ β)%stdpp). Proof. intros Hty Hst. pose (Σ:=#[invΣ;stateΣ rs R;heapΣ rs R]). diff --git a/theories/examples/affine_lang/logrel2.v b/theories/examples/affine_lang/logrel2.v index e855def..3db5de0 100644 --- a/theories/examples/affine_lang/logrel2.v +++ b/theories/examples/affine_lang/logrel2.v @@ -1,11 +1,12 @@ +From stdpp Require Import finite. From iris.base_logic.lib Require Import na_invariants. -From gitrees Require Export gitree program_logic. +From iris.algebra Require Import gmap. +From gitrees Require Export gitree program_logic greifiers. From gitrees.examples.input_lang Require Import lang interp logpred. From gitrees.examples.affine_lang Require Import lang logrel1. From gitrees.effects Require Import store. From gitrees.lib Require Import pairs. -Require Import iris.algebra.gmap. -Require Import stdpp.finite. +From gitrees.utils Require Import finite_sets. Require Import Binding.Lib Binding.Set Binding.Env. @@ -60,7 +61,7 @@ Section glue. Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_store rs}. Context `{!subReifier reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Context `{!SubOfe unitO R}. @@ -83,8 +84,6 @@ Section glue. Local Opaque thunked thunkedV Thunk. - Program Definition ı_scope : @lang_generic.interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. - Lemma compat_glue_to_affine_bool {S : Set} `{HE : EqDecision S} `{!Finite S} (Ω : S → ty) α : io_valid □ α Tnat ⊢ valid2 Ω (constO (glue2_bool _ (α ı_scope))) tBool. @@ -93,9 +92,7 @@ Section glue. iIntros (ss) "#Hctx Has". simpl. iIntros (σ) "[Hs Hp]". iSpecialize ("H" $! σ ı_scope with "Hs []"). - { unfold logpred.ssubst_valid. - iIntros ([]). - } + { iIntros ([]). } iSpecialize ("H" $! tt with "Hp"). simpl. iApply (wp_bind _ (IFSCtx _ _)). @@ -121,9 +118,7 @@ Section glue. iIntros (ss) "#Hctx Has". simpl. iIntros (σ) "[Hs Hp]". iSpecialize ("H" $! σ ı_scope with "Hs []"). - { unfold logpred.ssubst_valid. - iIntros ([]). - } + { iIntros ([]). } iSpecialize ("H" $! tt with "Hp"). simpl. iApply (wp_wand with "H"). @@ -140,7 +135,7 @@ Section glue. iIntros (σ ss) "Hs Hss". iIntros (?) "Hp". iSpecialize ("H" $! ss with "Hctx [] [$Hs $Hp]"). - { iApply ssubst_valid_empty. } + { iApply ssubst_valid_fin_empty1. } simpl. iApply (wp_wand with "H"). iIntros (αv) "Ha". iDestruct "Ha" as (σ') "[Ha [Hs Hp]]". @@ -156,7 +151,7 @@ Section glue. iIntros (σ ss) "Hs Hss". iIntros (?) "Hp". iSpecialize ("H" $! ss with "Hctx [] [$Hs $Hp]"). - { iApply ssubst_valid_empty. } + { iApply ssubst_valid_fin_empty1. } iApply (wp_wand with "H"). iIntros (αv) "Ha". iDestruct "Ha" as (σ') "[Ha [Hs Hp]]". iModIntro. iExists tt. eauto with iFrame. @@ -196,7 +191,7 @@ Section glue. simpl. iIntros (?) "Hp". iSpecialize ("H" $! ss with "Hctx [] [$Hs $Hp]"). - { iApply ssubst_valid_empty. } + { iApply ssubst_valid_fin_empty1. } simpl. iApply wp_let. { solve_proper. } iApply (wp_wand with "H"). @@ -267,7 +262,7 @@ Section glue. { iIntros (σ0 ss0) "Hs Has". simpl. iApply expr_pred_ret. simpl. eauto with iFrame. } iSpecialize ("G1" $! ss with "Hctx [] [$Hst $Hp]"). - { iApply ssubst_valid_empty. } + { iApply ssubst_valid_fin_empty1. } iApply (wp_wand with "G1"). clear βv σ'. iIntros (βv). iDestruct 1 as (σ') "[Hb [Hst Hp]]". @@ -432,11 +427,11 @@ Section glue. valid2 □ (constO α) τ1 ⊢ heap_ctx -∗ io_valid □ (constO (glue_from_affine _ Hconv α)) τ1'. Proof. - destruct Hconv. - + by iApply (@compat_glue_to_affine_bool). - + by iApply (@compat_glue_to_affine_nat). + + by iApply compat_glue_to_affine_bool. + + by iApply compat_glue_to_affine_nat. + iIntros "_". simpl. iApply compat_unit. - + simpl. iApply (@compat_glue_to_affine_fun). + + simpl. iApply compat_glue_to_affine_fun. * by apply glue_to_affine_compatibility. * apply glue_from_affine_compatibility. - destruct Hconv. @@ -459,11 +454,11 @@ Section glue. by iApply fundamental. - by iApply compat_var. - by iApply compat_lam. - - by iApply (@compat_app _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - - by iApply (@compat_pair _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - - by iApply (@compat_destruct _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_app EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_pair EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_destruct EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - by iApply compat_alloc. - - by iApply (@compat_replace _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ S1 S2 EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). + - by iApply (compat_replace EqDecisionLeft FiniteLeft EqDecisionRight FiniteRight). - by iApply compat_dealloc. - by iApply compat_nat. - by iApply compat_bool. @@ -473,28 +468,26 @@ Section glue. End glue. Local Definition rs : gReifiers NotCtxDep 2 - := gReifiers_cons NotCtxDep reify_store - (gReifiers_cons NotCtxDep input_lang.interp.reify_io (gReifiers_nil NotCtxDep)). + := gReifiers_cons reify_store + (gReifiers_cons input_lang.interp.reify_io gReifiers_nil). Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. -Require Import gitrees.gitree.greifiers. - Lemma logrel2_adequacy (cr : nat) R `{!Cofe R, !SubOfe locO R, !SubOfe natO R, !SubOfe unitO R} Σ `{!invGpreS Σ}`{!statePreG rs R Σ} `{!heapPreG rs R Σ} `{!na_invG Σ} - (τ : ty) (α : interp_scope Empty_set -n> IT (gReifiers_ops NotCtxDep rs) R) (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : + (τ : ty) (α : interp_scope Empty_set -n> IT (gReifiers_ops rs) R) (β : IT (gReifiers_ops rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ} `{H3: !heapG rs R Σ} p, (£ cr ⊢ valid2 rs p □ α τ)%I) → - ssteps (gReifiers_sReifier NotCtxDep rs) (α (ı_scope rs)) st β st' k → + ssteps (gReifiers_sReifier NotCtxDep rs) (α ı_scope) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (β ≡ Err OtherError) - ∨ (∃ βv, IT_of_V βv ≡ β). + ∨ (β ≡ Err OtherError)%stdpp + ∨ (∃ βv, (IT_of_V βv ≡ β)%stdpp). Proof. intros Hlog Hst. destruct (IT_to_V β) as [βv|] eqn:Hb. { right. right. exists βv. apply IT_of_to_V'. rewrite Hb; eauto. } cut ((∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (∃ e, β ≡ Err e ∧ s e)). + ∨ (∃ e, (β ≡ Err e)%stdpp ∧ s e)). { intros [?|He]; first eauto. right. left. destruct He as [? [? ->]]. done. } eapply (wp_safety (S cr) _ _ NotCtxDep rs s); eauto. @@ -509,9 +502,9 @@ Proof. iMod (new_heapG rs σ) as (H3) "H". iAssert (has_substate σ ∗ has_substate ios)%I with "[Hst]" as "[Hs Hio]". { unfold has_substate, has_full_state. - assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) st ≡ - of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ) - ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state ios)) as ->; last first. + assert (of_state NotCtxDep rs (IT (gReifiers_ops rs) _) st ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops rs) _) sR_idx (sR_state σ) + ⋅ of_idx NotCtxDep rs (IT (gReifiers_ops rs) _) sR_idx (sR_state ios))%stdpp as ->; last first. { rewrite -own_op. done. } unfold sR_idx. simpl. intro j. @@ -529,8 +522,8 @@ Proof. simpl. iMod na_alloc as (p) "Hp". iPoseProof (@Hlog _ _ _ p with "Hcr") as "Hlog". - iSpecialize ("Hlog" $! (ı_scope _) with "Hinv []"). - { iApply ssubst_valid_empty. } + iSpecialize ("Hlog" $! ı_scope with "Hinv []"). + { iApply ssubst_valid_fin_empty1. } unfold expr_pred. simpl. iSpecialize ("Hlog" $! ios with "[$Hio $Hp]"). iModIntro. simpl. @@ -540,12 +533,12 @@ Qed. Definition R := sumO locO (sumO natO unitO). -Lemma logrel2_safety e τ (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : +Lemma logrel2_safety e τ (β : IT (gReifiers_ops rs) R) st st' k : typed_glued □ e τ → - ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e (ı_scope _)) st β st' k → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ı_scope) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) - ∨ (β ≡ Err OtherError) - ∨ (∃ βv, IT_of_V βv ≡ β). + ∨ (β ≡ Err OtherError)%stdpp + ∨ (∃ βv, (IT_of_V βv ≡ β)%stdpp). Proof. intros Hty Hst. pose (Σ:=#[invΣ;stateΣ rs R;heapΣ rs R;na_invΣ]). diff --git a/theories/examples/input_lang/interp.v b/theories/examples/input_lang/interp.v index c0ae876..0f8c135 100644 --- a/theories/examples/input_lang/interp.v +++ b/theories/examples/input_lang/interp.v @@ -1,8 +1,7 @@ From gitrees Require Import gitree lang_generic. From gitrees.examples.input_lang Require Import lang. -Require Import Binding.Lib. -Require Import Binding.Set. +Require Import Binding.Lib Binding.Set. Notation stateO := (leibnizO state). @@ -92,7 +91,7 @@ Section weakestpre. Context {sz : nat}. Variable (rs : gReifiers NotCtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Notation IT := (IT F R). @@ -137,7 +136,7 @@ Section interp. Context {subR : subReifier reify_io rs}. Context {R} `{CR : !Cofe R}. Context `{!SubOfe natO R}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F R). Notation ITV := (ITV F R). @@ -164,10 +163,10 @@ Section interp. Typeclasses Opaque interp_natop. Opaque laterO_map. - Program Definition interp_rec_pre {S : Set} (body : @interp_scope F R _ (inc (inc S)) -n> IT) - : laterO (@interp_scope F R _ S -n> IT) -n> @interp_scope F R _ S -n> IT := - λne self env, Fun $ laterO_map (λne (self : @interp_scope F R _ S -n> IT) (a : IT), - body (@extend_scope F R _ _ (@extend_scope F R _ _ env (self env)) a)) self. + Program Definition interp_rec_pre {S : Set} (body : interp_scope (inc (inc S)) -n> IT) + : laterO (interp_scope S -n> IT) -n> interp_scope S -n> IT := + λne self env, Fun $ laterO_map (λne (self : interp_scope S -n> IT) (a : IT), + body (extend_scope (extend_scope env (self env)) a)) self. Next Obligation. intros. solve_proper_prepare. @@ -191,14 +190,14 @@ Section interp. Qed. Program Definition interp_rec {S : Set} - (body : @interp_scope F R _ (inc (inc S)) -n> IT) : - @interp_scope F R _ S -n> IT := + (body : interp_scope (inc (inc S)) -n> IT) : + interp_scope S -n> IT := mmuu (interp_rec_pre body). Program Definition ir_unf {S : Set} - (body : @interp_scope F R _ (inc (inc S)) -n> IT) env : IT -n> IT := - λne a, body (@extend_scope F R _ _ - (@extend_scope F R _ _ env (interp_rec body env)) + (body : interp_scope (inc (inc S)) -n> IT) env : IT -n> IT := + λne a, body (extend_scope + (extend_scope env (interp_rec body env)) a). Next Obligation. intros. @@ -206,7 +205,7 @@ Section interp. f_equiv. intros [| [| y']]; simpl; solve_proper. Qed. - Lemma interp_rec_unfold {S : Set} (body : @interp_scope F R _ (inc (inc S)) -n> IT) env : + Lemma interp_rec_unfold {S : Set} (body : interp_scope (inc (inc S)) -n> IT) env : interp_rec body env ≡ Fun $ Next $ ir_unf body env. Proof. trans (interp_rec_pre body (Next (interp_rec body)) env). diff --git a/theories/examples/input_lang/logpred.v b/theories/examples/input_lang/logpred.v index b548bc9..303d9d4 100644 --- a/theories/examples/input_lang/logpred.v +++ b/theories/examples/input_lang/logpred.v @@ -2,12 +2,13 @@ From gitrees Require Import gitree program_logic lang_generic. From gitrees.examples.input_lang Require Import lang interp. Require Import Binding.Lib Binding.Set Binding.Env. +Require Import gitrees.gitree.greifiers. Section io_lang. Context {sz : nat}. Variable rs : gReifiers NotCtxDep sz. Context `{!subReifier reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Notation IT := (IT F R). @@ -41,15 +42,10 @@ Section io_lang. | Tarr τ1 τ2 => interp_tarr (interp_ty τ1) (interp_ty τ2) end. - Definition ssubst_valid {S : Set} - (Γ : S -> ty) - (ss : @interp_scope F R _ S) : iProp := - (∀ x, □ expr_pred (ss x) (interp_ty (Γ x)))%I. + Notation ssubst_valid := (ssubst_valid1 rs ty interp_ty expr_pred). #[global] Instance io_lang_interp_ty_pers τ βv : Persistent (io_lang.interp_ty τ βv). Proof. induction τ; apply _. Qed. - #[global] Instance ssubst_valid_pers {S : Set} (Γ : S → ty) ss : Persistent (ssubst_valid Γ ss). - Proof. apply _. Qed. Program Definition valid1 {S : Set} (Γ : S → ty) (α : interp_scope S -n> IT) (τ : ty) : iProp := (∀ σ ss, has_substate σ -∗ ssubst_valid Γ ss -∗ @@ -250,21 +246,17 @@ End io_lang. Arguments interp_ty {_ _ _ _ _ _ _ _ _ _ _ _} τ. Arguments interp_tarr {_ _ _ _ _ _ _ _ _ _ _} Φ1 Φ2. -Local Definition rs : gReifiers NotCtxDep _ := gReifiers_cons NotCtxDep reify_io (gReifiers_nil NotCtxDep). +Local Definition rs : gReifiers NotCtxDep _ := gReifiers_cons reify_io gReifiers_nil. Variable Hdisj : ∀ (Σ : gFunctors) (P Q : iProp Σ), disjunction_property P Q. -Require Import gitrees.gitree.greifiers. - -Program Definition ı_scope R `{!Cofe R} : @interp_scope (gReifiers_ops NotCtxDep rs) R _ Empty_set := λne (x : ∅), match x with end. - Lemma logpred_adequacy cr Σ R `{!Cofe R, SubOfe natO R} `{!invGpreS Σ} `{!statePreG rs R Σ} τ - (α : interp_scope ∅ -n> IT (gReifiers_ops NotCtxDep rs) R) - (β : IT (gReifiers_ops NotCtxDep rs) R) st st' k : + (α : interp_scope ∅ -n> IT (gReifiers_ops rs) R) + (β : IT (gReifiers_ops rs) R) st st' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs R Σ}, (£ cr ⊢ valid1 rs notStuck (λne _ : unitO, True)%I □ α τ)%I) → - ssteps (gReifiers_sReifier NotCtxDep rs) (α (ı_scope _)) st β st' k → + ssteps (gReifiers_sReifier NotCtxDep rs) (α ı_scope) st β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. @@ -287,8 +279,8 @@ Proof. destruct st as [σ []]. iAssert (has_substate σ) with "[Hst]" as "Hs". { unfold has_substate, has_full_state. - assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) (σ,()) ≡ - of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) _) sR_idx (sR_state σ)) as ->; last done. + assert (of_state NotCtxDep rs (IT (gReifiers_ops rs) _) (σ,()) ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops rs) _) sR_idx (sR_state σ)) as ->; last done. intro j. unfold sR_idx. simpl. unfold of_state, of_idx. destruct decide as [Heq|]; last first. @@ -310,9 +302,9 @@ Proof. done. Qed. -Lemma io_lang_safety e τ σ st' (β : IT (sReifier_ops NotCtxDep (gReifiers_sReifier NotCtxDep rs)) natO) k : +Lemma io_lang_safety e τ σ st' (β : IT (sReifier_ops (gReifiers_sReifier NotCtxDep rs)) natO) k : typed □ e τ → - ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e (ı_scope _)) (σ, ()) β st' k → + ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ı_scope) (σ, ()) β st' k → (∃ β1 st1, sstep (gReifiers_sReifier NotCtxDep rs) β st' β1 st1) ∨ (∃ βv, IT_of_V βv ≡ β). Proof. diff --git a/theories/examples/input_lang/logrel.v b/theories/examples/input_lang/logrel.v index 53f9756..c058fd2 100644 --- a/theories/examples/input_lang/logrel.v +++ b/theories/examples/input_lang/logrel.v @@ -1,13 +1,14 @@ (** Logical relation for adequacy for the IO lang *) From gitrees Require Import gitree lang_generic. From gitrees.examples.input_lang Require Import lang interp. +Require Import gitrees.gitree.greifiers. Require Import Binding.Lib Binding.Set Binding.Env. Section logrel. Context {sz : nat}. Variable (rs : gReifiers NotCtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F natO). Notation ITV := (ITV F natO). Context `{!invGS Σ, !stateG rs natO Σ}. @@ -353,24 +354,22 @@ Lemma κ_Ret {S} {E} n : κ ((RetV n) : ITV E natO) = (LitV n : val S). Proof. Transparent RetV. unfold RetV. simpl. done. Opaque RetV. Qed. -Definition rs : gReifiers NotCtxDep 1 := gReifiers_cons NotCtxDep reify_io (gReifiers_nil NotCtxDep). - -Require Import gitrees.gitree.greifiers. +Definition rs : gReifiers NotCtxDep 1 := gReifiers_cons reify_io gReifiers_nil. -Lemma logrel_nat_adequacy Σ `{!invGpreS Σ}`{!statePreG rs natO Σ} {S} (α : IT (gReifiers_ops NotCtxDep rs) natO) (e : expr S) n σ σ' k : +Lemma logrel_nat_adequacy Σ `{!invGpreS Σ}`{!statePreG rs natO Σ} {S} (α : IT (gReifiers_ops rs) natO) (e : expr S) n σ σ' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs natO Σ}, (True ⊢ logrel rs Tnat α e)%I) → ssteps (gReifiers_sReifier NotCtxDep rs) α (σ,()) (Ret n) σ' k → ∃ m σ', prim_steps e σ (Val $ LitV n) σ' m. Proof. intros Hlog Hst. - pose (ϕ := λ (βv : ITV (gReifiers_ops NotCtxDep rs) natO), + pose (ϕ := λ (βv : ITV (gReifiers_ops rs) natO), ∃ m σ', prim_steps e σ (Val $ κ βv) σ' m). cut (ϕ (RetV n)). { destruct 1 as ( m' & σ2 & Hm). exists m', σ2. revert Hm. by rewrite κ_Ret. } eapply (wp_adequacy 0); eauto. intros Hinv1 Hst1. - pose (Φ := (λ (βv : ITV (gReifiers_ops NotCtxDep rs) natO), ∃ n, logrel_val rs Tnat (Σ:=Σ) (S:=S) βv (LitV n) + pose (Φ := (λ (βv : ITV (gReifiers_ops rs) natO), ∃ n, logrel_val rs Tnat (Σ:=Σ) (S:=S) βv (LitV n) ∗ ⌜∃ m σ', prim_steps e σ (Val $ LitV n) σ' m⌝)%I). assert (NonExpansive Φ). { unfold Φ. @@ -392,8 +391,8 @@ Proof. iPoseProof (Hlog with "[//]") as "Hlog". iAssert (has_substate σ) with "[Hs]" as "Hs". { unfold has_substate, has_full_state. - assert (of_state NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) natO) (σ, ()) ≡ - of_idx NotCtxDep rs (IT (gReifiers_ops NotCtxDep rs) natO) 0 σ) as ->; last done. + assert (of_state NotCtxDep rs (IT (gReifiers_ops rs) natO) (σ, ()) ≡ + of_idx NotCtxDep rs (IT (gReifiers_ops rs) natO) 0 σ)%stdpp as ->; last done. intro j. unfold sR_idx. simpl. unfold of_state, of_idx. destruct decide as [Heq|]; last first. @@ -412,8 +411,6 @@ Proof. iExists l. iSplit; eauto. Qed. -Program Definition ı_scope : @interp_scope (gReifiers_ops NotCtxDep rs) natO _ Empty_set := λne (x : ∅), match x with end. - Theorem adequacy (e : expr ∅) (k : nat) σ σ' n : typed □ e Tnat → ssteps (gReifiers_sReifier NotCtxDep rs) (interp_expr rs e ı_scope) (σ,()) (Ret k : IT _ natO) σ' n → diff --git a/theories/examples/input_lang_callcc/hom.v b/theories/examples/input_lang_callcc/hom.v index 3e4f03e..661374b 100644 --- a/theories/examples/input_lang_callcc/hom.v +++ b/theories/examples/input_lang_callcc/hom.v @@ -9,7 +9,7 @@ Section hom. Context {sz : nat}. Context {rs : gReifiers CtxDep sz}. Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops CtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F natO). Notation ITV := (ITV F natO). diff --git a/theories/examples/input_lang_callcc/interp.v b/theories/examples/input_lang_callcc/interp.v index 8e41b93..5daeec3 100644 --- a/theories/examples/input_lang_callcc/interp.v +++ b/theories/examples/input_lang_callcc/interp.v @@ -176,7 +176,7 @@ Section weakestpre. Context {sz : nat}. Variable (rs : gReifiers CtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops CtxDep rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. Notation IT := (IT F R). @@ -289,7 +289,7 @@ Section interp. Context {subR : subReifier reify_io rs}. Context {R} `{CR : !Cofe R}. Context `{!SubOfe natO R}. - Notation F := (gReifiers_ops CtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F R). Notation ITV := (ITV F R). diff --git a/theories/examples/input_lang_callcc/logrel.v b/theories/examples/input_lang_callcc/logrel.v index 89d9c44..cbfd44f 100644 --- a/theories/examples/input_lang_callcc/logrel.v +++ b/theories/examples/input_lang_callcc/logrel.v @@ -2,6 +2,7 @@ From gitrees Require Import gitree. From gitrees.examples.input_lang_callcc Require Import lang interp hom. Require Import gitrees.lang_generic. +Require Import gitrees.gitree.greifiers. Require Import Binding.Lib Binding.Set Binding.Env. Open Scope stdpp_scope. @@ -10,7 +11,7 @@ Section logrel. Context {sz : nat}. Variable (rs : gReifiers CtxDep sz). Context {subR : subReifier reify_io rs}. - Notation F := (gReifiers_ops CtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F natO). Notation ITV := (ITV F natO). Context `{!invGS Σ, !stateG rs natO Σ}. @@ -646,10 +647,11 @@ Section logrel. with fundamental_val {S : Set} (Γ : S -> ty) τ v : typed_val Γ v τ → ⊢ logrel_valid Γ (Val v) (interp_val rs v) τ. Proof. - - induction 1; simpl. - + by apply fundamental_val. + - intros H. + induction H. + + by iApply fundamental_val. + rewrite -H. - by apply compat_var. + by iApply compat_var. + iApply compat_app. ++ iApply IHtyped1. ++ iApply IHtyped2. @@ -684,19 +686,17 @@ Lemma κ_Ret {S} {E} n : κ ((RetV n) : ITV E natO) = (LitV n : val S). Proof. Transparent RetV. unfold RetV. simpl. done. Opaque RetV. Qed. -Definition rs : gReifiers CtxDep 1 := gReifiers_cons CtxDep reify_io (gReifiers_nil CtxDep). - -Require Import gitrees.gitree.greifiers. +Definition rs : gReifiers CtxDep 1 := gReifiers_cons reify_io gReifiers_nil. Lemma logrel_nat_adequacy Σ `{!invGpreS Σ} `{!statePreG rs natO Σ} {S} - (α : IT (gReifiers_ops CtxDep rs) natO) + (α : IT (gReifiers_ops rs) natO) (e : expr S) n σ σ' k : (∀ `{H1 : !invGS Σ} `{H2: !stateG rs natO Σ}, (⊢ logrel rs Tnat α e)%I) → ssteps (gReifiers_sReifier CtxDep rs) α (σ, ()) (Ret n) σ' k → ∃ m σ', prim_steps e σ (Val $ LitV n) σ' m. Proof. intros Hlog Hst. - pose (ϕ := λ (βv : ITV (gReifiers_ops CtxDep rs) natO), + pose (ϕ := λ (βv : ITV (gReifiers_ops rs) natO), ∃ m σ', prim_steps e σ (Val $ κ βv) σ' m). cut (ϕ (RetV n)). { @@ -705,7 +705,7 @@ Proof. } eapply (wp_adequacy 0); eauto. intros Hinv1 Hst1. - pose (Φ := (λ (βv : ITV (gReifiers_ops CtxDep rs) natO), + pose (Φ := (λ (βv : ITV (gReifiers_ops rs) natO), ∃ n, logrel_val rs Tnat (Σ:=Σ) (S:=S) βv (LitV n) ∗ ⌜∃ m σ', prim_steps e σ (Val $ LitV n) σ' m⌝)%I). assert (NonExpansive Φ). @@ -731,8 +731,8 @@ Proof. iAssert (has_substate σ) with "[Hs]" as "Hs". { unfold has_substate, has_full_state. - assert ((of_state CtxDep rs (IT (sReifier_ops CtxDep (gReifiers_sReifier CtxDep rs)) natO) (σ, ())) ≡ - (of_idx CtxDep rs (IT (sReifier_ops CtxDep (gReifiers_sReifier CtxDep rs)) natO) sR_idx (sR_state σ))) + assert ((of_state CtxDep rs (IT (sReifier_ops (gReifiers_sReifier CtxDep rs)) natO) (σ, ())) ≡ + (of_idx CtxDep rs (IT (sReifier_ops (gReifiers_sReifier CtxDep rs)) natO) sR_idx (sR_state σ))) as -> ; last done. intros j. unfold sR_idx. simpl. unfold of_state, of_idx. @@ -765,8 +765,6 @@ Proof. iExists l. iSplit; eauto. Qed. -Program Definition ı_scope : @interp_scope (gReifiers_ops CtxDep rs) natO _ Empty_set := λne (x : ∅), match x with end. - Theorem adequacy (e : expr ∅) (k : nat) σ σ' n : typed □ e Tnat → ssteps (gReifiers_sReifier CtxDep rs) (interp_expr rs e ı_scope) (σ, ()) (Ret k : IT _ natO) σ' n → diff --git a/theories/gitree/greifiers.v b/theories/gitree/greifiers.v index f9ff81e..1e549a4 100644 --- a/theories/gitree/greifiers.v +++ b/theories/gitree/greifiers.v @@ -7,9 +7,6 @@ Section greifiers_generic. #[local] Open Scope type. Context (a : is_ctx_dep). Notation sReifier := (sReifier a). - Notation sReifier_ops := (sReifier_ops a). - Notation sReifier_state := (sReifier_state a). - Notation sReifier_re := (sReifier_re a). (** Global reifiers: a collection of reifiers *) Inductive gReifiers : nat → Type := @@ -135,7 +132,7 @@ Section greifiers. let fs := gState_decomp NotCtxDep i st in let σ := fs.1 in let rest := fs.2 in - let rx := sReifier_re NotCtxDep (rs !!! i) op' (x, σ) in + let rx := sReifier_re (rs !!! i) op' (x, σ) in optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) rx. Next Obligation. solve_proper_please. Qed. @@ -153,7 +150,7 @@ Section greifiers. let fs := gState_decomp CtxDep i b in let σ := fs.1 in let rest := fs.2 in - let rx := sReifier_re CtxDep (rs !!! i) op' (a, σ, c) in + let rx := sReifier_re (rs !!! i) op' (a, σ, c) in optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) rx. Next Obligation. solve_proper_please. Qed. @@ -194,14 +191,14 @@ Section greifiers. Defined. Lemma gReifiers_re_idx_ctx_dep {n} (i : fin n) (rs : gReifiers CtxDep n) - {X} `{!Cofe X} (op : opid (sReifier_ops CtxDep (rs !!! i))) - (x : Ins (sReifier_ops CtxDep _ op) ♯ X) - (σ : sReifier_state CtxDep (rs !!! i) ♯ X) + {X} `{!Cofe X} (op : opid (sReifier_ops (rs !!! i))) + (x : Ins (sReifier_ops _ op) ♯ X) + (σ : sReifier_state (rs !!! i) ♯ X) (rest : gState_rest CtxDep i rs ♯ X) - (κ : (Outs (sReifier_ops CtxDep (rs !!! i) op) ♯ X -n> laterO X)) : + (κ : (Outs (sReifier_ops (rs !!! i) op) ♯ X -n> laterO X)) : gReifiers_re CtxDep rs (existT i op) (x, gState_recomp CtxDep rest σ, κ) ≡ optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) - (sReifier_re CtxDep (rs !!! i) op (x, σ, κ)). + (sReifier_re (rs !!! i) op (x, σ, κ)). Proof. unfold gReifiers_re. cbn-[prodO_map optionO_map]. f_equiv; last repeat f_equiv. @@ -213,13 +210,13 @@ Section greifiers. Qed. Lemma gReifiers_re_idx_ctx_indep {n} (i : fin n) (rs : gReifiers NotCtxDep n) - {X} `{!Cofe X} (op : opid (sReifier_ops NotCtxDep (rs !!! i))) - (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) - (σ : sReifier_state NotCtxDep (rs !!! i) ♯ X) + {X} `{!Cofe X} (op : opid (sReifier_ops (rs !!! i))) + (x : Ins (sReifier_ops _ op) ♯ X) + (σ : sReifier_state (rs !!! i) ♯ X) (rest : gState_rest NotCtxDep i rs ♯ X) : gReifiers_re NotCtxDep rs (existT i op) (x, gState_recomp NotCtxDep rest σ) ≡ optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) - (sReifier_re NotCtxDep (rs !!! i) op (x, σ)). + (sReifier_re (rs !!! i) op (x, σ)). Proof. unfold gReifiers_re. cbn-[prodO_map optionO_map]. f_equiv; last repeat f_equiv. @@ -231,26 +228,26 @@ Section greifiers. Qed. Program Definition gReifiers_re_idx_type {n} a (i : fin n) (rs : gReifiers a n) - {X} `{!Cofe X} (op : opid (sReifier_ops a (rs !!! i))) - (x : Ins (sReifier_ops a _ op) ♯ X) - (σ : sReifier_state a (rs !!! i) ♯ X) + {X} `{!Cofe X} (op : opid (sReifier_ops (rs !!! i))) + (x : Ins (sReifier_ops _ op) ♯ X) + (σ : sReifier_state (rs !!! i) ♯ X) (rest : gState_rest a i rs ♯ X) : Type. Proof. destruct a. - - apply (∀ (κ : (Outs (sReifier_ops CtxDep (rs !!! i) op) ♯ X -n> laterO X)), + - apply (∀ (κ : (Outs (sReifier_ops (rs !!! i) op) ♯ X -n> laterO X)), gReifiers_re CtxDep rs (existT i op) (x, gState_recomp CtxDep rest σ, κ) ≡ optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) - (sReifier_re CtxDep (rs !!! i) op (x, σ, κ))). + (sReifier_re (rs !!! i) op (x, σ, κ))). - apply (gReifiers_re NotCtxDep rs (existT i op) (x, gState_recomp NotCtxDep rest σ) ≡ optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) - (sReifier_re NotCtxDep (rs !!! i) op (x, σ))). + (sReifier_re (rs !!! i) op (x, σ))). Defined. Lemma gReifiers_re_idx {n} a (i : fin n) (rs : gReifiers a n) - {X} `{!Cofe X} (op : opid (sReifier_ops a (rs !!! i))) - (x : Ins (sReifier_ops a _ op) ♯ X) - (σ : sReifier_state a (rs !!! i) ♯ X) + {X} `{!Cofe X} (op : opid (sReifier_ops (rs !!! i))) + (x : Ins (sReifier_ops _ op) ♯ X) + (σ : sReifier_state (rs !!! i) ♯ X) (rest : gState_rest a i rs ♯ X) : gReifiers_re_idx_type a i rs op x σ rest. Proof. destruct a. @@ -261,34 +258,34 @@ Section greifiers. Program Definition sR_re_type {n} {X} `{!Cofe X} (a : is_ctx_dep) (r : sReifier a) (rs : gReifiers a n) (sR_idx : fin n) - (sR_ops : subEff (sReifier_ops a r) (sReifier_ops a (rs !!! sR_idx))) - (sR_state : sReifier_state a r ♯ X ≃ sReifier_state a (rs !!! sR_idx) ♯ X) - (m : nat) (op : opid (sReifier_ops a r)) : Type. + (sR_ops : subEff (sReifier_ops r) (sReifier_ops (rs !!! sR_idx))) + (sR_state : sReifier_state r ♯ X ≃ sReifier_state (rs !!! sR_idx) ♯ X) + (m : nat) (op : opid (sReifier_ops r)) : Type. Proof. destruct a. - - apply (∀ (x : Ins (sReifier_ops CtxDep r op) ♯ X) + - apply (∀ (x : Ins (sReifier_ops r op) ♯ X) (y : laterO X) - (s1 s2 : sReifier_state CtxDep r ♯ X) - (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)), - sReifier_re CtxDep r op (x, s1, k) ≡{m}≡ Some (y, s2) → - sReifier_re CtxDep (rs !!! sR_idx) (subEff_opid op) + (s1 s2 : sReifier_state r ♯ X) + (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)), + sReifier_re r op (x, s1, k) ≡{m}≡ Some (y, s2) → + @sReifier_re CtxDep (rs !!! sR_idx) _ _ (subEff_opid op) (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡{m}≡ Some (y, sR_state s2)). - - apply (∀ (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) - (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) - (s1 s2 : sReifier_state NotCtxDep r ♯ X), - sReifier_re NotCtxDep r op (x, s1) ≡{m}≡ Some (y, s2) → - sReifier_re NotCtxDep (rs !!! sR_idx) (subEff_opid op) + - apply (∀ (x : Ins (sReifier_ops _ op) ♯ X) + (y : Outs (sReifier_ops _ op) ♯ X) + (s1 s2 : sReifier_state r ♯ X), + sReifier_re r op (x, s1) ≡{m}≡ Some (y, s2) → + @sReifier_re NotCtxDep (rs !!! sR_idx) _ _ (subEff_opid op) (subEff_ins x, sR_state s1) ≡{m}≡ Some (subEff_outs y, sR_state s2)). Defined. Class subReifier {n} {a : is_ctx_dep} (r : sReifier a) (rs : gReifiers a n) := { sR_idx : fin n; - sR_ops :: subEff (sReifier_ops a r) (sReifier_ops a (rs !!! sR_idx)); + sR_ops :: subEff (sReifier_ops r) (sReifier_ops (rs !!! sR_idx)); sR_state {X} `{!Cofe X} : - sReifier_state a r ♯ X ≃ sReifier_state a (rs !!! sR_idx) ♯ X; - sR_re (m : nat) {X} `{!Cofe X} (op : opid (sReifier_ops a r)) + sReifier_state r ♯ X ≃ sReifier_state (rs !!! sR_idx) ♯ X; + sR_re (m : nat) {X} `{!Cofe X} (op : opid (sReifier_ops r)) : sR_re_type a r rs sR_idx sR_ops (@sR_state X _) m op }. @@ -331,7 +328,7 @@ Section greifiers. #[local] Definition subR_op {n} {a : is_ctx_dep} {r : sReifier a} {rs : gReifiers a n} `{!subReifier r rs} : - opid (sReifier_ops a r) → opid (gReifiers_ops a rs). + opid (sReifier_ops r) → opid (gReifiers_ops a rs). Proof. intros op. simpl. @@ -340,7 +337,7 @@ Section greifiers. #[export] Instance subReifier_subEff {n} {a : is_ctx_dep} {r : sReifier a} {rs : gReifiers a n} `{!subReifier r rs} : - subEff (sReifier_ops a r) (gReifiers_ops a rs). + subEff (sReifier_ops r) (gReifiers_ops a rs). Proof. simple refine {| subEff_opid := subR_op |}. - intros op X ?. simpl. @@ -351,29 +348,29 @@ Section greifiers. Program Definition subReifier_reify_idx_type {n} (a : is_ctx_dep) (r : sReifier a) (rs : gReifiers a n) - `{!subReifier r rs} X `{!Cofe X} (op : opid (sReifier_ops a r)) : Type. + `{!subReifier r rs} X `{!Cofe X} (op : opid (sReifier_ops r)) : Type. Proof. destruct a. - - apply (∀ (x : Ins (sReifier_ops CtxDep r op) ♯ X) + - apply (∀ (x : Ins (sReifier_ops r op) ♯ X) (y : laterO X) - (s1 s2 : sReifier_state CtxDep r ♯ X) - (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)), - sReifier_re CtxDep r op (x, s1, k) ≡ Some (y, s2) → - sReifier_re CtxDep (rs !!! sR_idx) (subEff_opid op) + (s1 s2 : sReifier_state r ♯ X) + (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)), + sReifier_re r op (x, s1, k) ≡ Some (y, s2) → + @sReifier_re CtxDep (rs !!! sR_idx) _ _ (subEff_opid op) (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ Some (y, sR_state s2)). - - apply (∀ (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) - (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) - (s1 s2 : sReifier_state NotCtxDep r ♯ X), - sReifier_re NotCtxDep r op (x, s1) ≡ Some (y, s2) → - sReifier_re NotCtxDep (rs !!! sR_idx) (subEff_opid op) + - apply (∀ (x : Ins (sReifier_ops _ op) ♯ X) + (y : Outs (sReifier_ops _ op) ♯ X) + (s1 s2 : sReifier_state r ♯ X), + sReifier_re r op (x, s1) ≡ Some (y, s2) → + @sReifier_re NotCtxDep (rs !!! sR_idx) _ _ (subEff_opid op) (subEff_ins x, sR_state s1) ≡ Some (subEff_outs y, sR_state s2)). Defined. Lemma subReifier_reify_idx {n} {a : is_ctx_dep} (r : sReifier a) (rs : gReifiers a n) - `{!subReifier r rs} {X} `{!Cofe X} (op : opid (sReifier_ops a r)) + `{!subReifier r rs} {X} `{!Cofe X} (op : opid (sReifier_ops r)) : subReifier_reify_idx_type a r rs X op. Proof. destruct a. @@ -391,19 +388,19 @@ Section greifiers. Program Definition subReifier_reify_type {n} (a : is_ctx_dep) (r : sReifier a) (rs : gReifiers a n) `{!subReifier r rs} X `{!Cofe X} - (op : opid (sReifier_ops a r)) : Type. + (op : opid (sReifier_ops r)) : Type. Proof. destruct a. - - apply (∀ (x : Ins (sReifier_ops CtxDep _ op) ♯ X) (y : laterO X) - (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)) - (σ σ' : sReifier_state CtxDep r ♯ X) (rest : gState_rest CtxDep sR_idx rs ♯ X), - sReifier_re CtxDep r op (x, σ, k) ≡ Some (y, σ') → + - apply (∀ (x : Ins (sReifier_ops _ op) ♯ X) (y : laterO X) + (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) + (σ σ' : sReifier_state r ♯ X) (rest : gState_rest CtxDep sR_idx rs ♯ X), + sReifier_re r op (x, σ, k) ≡ Some (y, σ') → gReifiers_re CtxDep rs (subEff_opid op) (subEff_ins x, gState_recomp CtxDep rest (sR_state σ), k ◎ (subEff_outs ^-1)) ≡ Some (y, gState_recomp CtxDep rest (sR_state σ'))). - - apply (∀ (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) - (σ σ' : sReifier_state NotCtxDep r ♯ X) (rest : gState_rest NotCtxDep sR_idx rs ♯ X), - sReifier_re NotCtxDep r op (x,σ) ≡ Some (y, σ') → + - apply (∀ (x : Ins (sReifier_ops _ op) ♯ X) (y : Outs (sReifier_ops _ op) ♯ X) + (σ σ' : sReifier_state r ♯ X) (rest : gState_rest NotCtxDep sR_idx rs ♯ X), + sReifier_re r op (x,σ) ≡ Some (y, σ') → gReifiers_re NotCtxDep rs (subEff_opid op) (subEff_ins x, gState_recomp NotCtxDep rest (sR_state σ)) ≡ Some (subEff_outs y, gState_recomp NotCtxDep rest (sR_state σ'))). @@ -411,7 +408,7 @@ Section greifiers. Lemma subReifier_reify {n} {a : is_ctx_dep} (r : sReifier a) (rs : gReifiers a n) `{!subReifier r rs} {X} `{!Cofe X} - (op : opid (sReifier_ops a r)) : subReifier_reify_type a r rs X op. + (op : opid (sReifier_ops r)) : subReifier_reify_type a r rs X op. Proof. destruct a. - simpl. @@ -422,7 +419,7 @@ Section greifiers. simpl in J'. rewrite J'; clear J'. transitivity (prod_map (λ x0 : laterO X, x0) - (λ st : sReifier_state CtxDep (rs !!! sR_idx) ♯ X, + (λ st : sReifier_state (rs !!! sR_idx) ♯ X, (gState_decomp' CtxDep sR_idx rs ^-1) (st, H)) <$> (Some (y, sR_state σ'))). + unfold prod_map. @@ -439,9 +436,9 @@ Section greifiers. as J'. simpl in J'. rewrite J'; clear J'. - transitivity (prod_map (λ x0 : Outs (sReifier_ops NotCtxDep (rs !!! sR_idx) + transitivity (prod_map (λ x0 : Outs (sReifier_ops (rs !!! sR_idx) (subEff_opid op)) ♯ X, x0) - (λ st : sReifier_state NotCtxDep (rs !!! sR_idx) ♯ X, + (λ st : sReifier_state (rs !!! sR_idx) ♯ X, (gState_decomp' NotCtxDep sR_idx rs ^-1) (st, rest)) <$> (Some (subEff_outs y, sR_state σ'))). + unfold prod_map. @@ -476,53 +473,55 @@ Section greifiers. Qed. Lemma subReifier_reify_idxI_ctx_dep (r : sReifier CtxDep) - `{!@subReifier sz CtxDep r rs} {X} `{!Cofe X} (op : opid (sReifier_ops CtxDep r)) - (x : Ins (sReifier_ops CtxDep _ op) ♯ X) + `{!@subReifier sz CtxDep r rs} {X} `{!Cofe X} (op : opid (sReifier_ops r)) + (x : Ins (sReifier_ops _ op) ♯ X) (y : laterO X) - (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)) - (s1 s2 : sReifier_state CtxDep r ♯ X) : - sReifier_re CtxDep r op (x, s1, k) ≡ Some (y, s2) ⊢@{iProp} - sReifier_re CtxDep (rs !!! sR_idx) (subEff_opid op) - (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ - Some (y, sR_state s2). + (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) + (s1 s2 : sReifier_state r ♯ X) : + sReifier_re r op (x, s1, k) ≡ Some (y, s2) + ⊢@{iProp} + sReifier_re (rs !!! sR_idx) (subEff_opid op) + (subEff_ins x, sR_state s1, k ◎ (subEff_outs ^-1)) ≡ + Some (y, sR_state s2). Proof. apply uPred.internal_eq_entails=>m. intros H'. - rewrite (sR_re (a := CtxDep)); last first. + rewrite (@sR_re _ CtxDep); last first. - rewrite H'. reflexivity. - reflexivity. Qed. Lemma subReifier_reify_idxI_ctx_indep (r : sReifier NotCtxDep) - `{!@subReifier sz NotCtxDep r rs} {X} `{!Cofe X} (op : opid (sReifier_ops NotCtxDep r)) - (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) - (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) - (s1 s2 : sReifier_state NotCtxDep r ♯ X) : - sReifier_re NotCtxDep r op (x, s1) ≡ Some (y, s2) + `{!@subReifier sz NotCtxDep r rs} {X} `{!Cofe X} (op : opid (sReifier_ops r)) + (x : Ins (sReifier_ops _ op) ♯ X) + (y : Outs (sReifier_ops _ op) ♯ X) + (s1 s2 : sReifier_state r ♯ X) : + sReifier_re r op (x, s1) ≡ Some (y, s2) ⊢@{iProp} - sReifier_re NotCtxDep (rs !!! sR_idx) (subEff_opid op) + sReifier_re (rs !!! sR_idx) (subEff_opid op) (subEff_ins x, sR_state s1) ≡ Some (subEff_outs y, sR_state s2). Proof. apply uPred.internal_eq_entails=>m. - apply (sR_re (a := NotCtxDep)). + apply (@sR_re _ NotCtxDep). Qed. Lemma subReifier_reifyI_ctx_dep (r : sReifier CtxDep) `{!@subReifier sz CtxDep r rs} {X} `{!Cofe X} - (op : opid (sReifier_ops CtxDep r)) - (x : Ins (sReifier_ops CtxDep _ op) ♯ X) (y : laterO X) - (k : (Outs (sReifier_ops CtxDep r op) ♯ X -n> laterO X)) - (σ σ' : sReifier_state CtxDep r ♯ X) (rest : gState_rest CtxDep sR_idx rs ♯ X) : - sReifier_re CtxDep r op (x,σ, k) ≡ Some (y, σ') ⊢@{iProp} - gReifiers_re CtxDep rs (subEff_opid op) - (subEff_ins x, gState_recomp CtxDep rest (sR_state σ), k ◎ (subEff_outs ^-1)) - ≡ Some (y, gState_recomp CtxDep rest (sR_state σ')). + (op : opid (sReifier_ops r)) + (x : Ins (sReifier_ops _ op) ♯ X) (y : laterO X) + (k : (Outs (sReifier_ops r op) ♯ X -n> laterO X)) + (σ σ' : sReifier_state r ♯ X) (rest : gState_rest CtxDep sR_idx rs ♯ X) : + sReifier_re r op (x,σ, k) ≡ Some (y, σ') + ⊢@{iProp} + gReifiers_re CtxDep rs (subEff_opid op) + (subEff_ins x, gState_recomp CtxDep rest (sR_state σ), k ◎ (subEff_outs ^-1)) + ≡ Some (y, gState_recomp CtxDep rest (sR_state σ')). Proof. apply uPred.internal_eq_entails=>m. intros He. - eapply (sR_re (a := CtxDep)) in He. + eapply (@sR_re _ CtxDep) in He. rewrite (gReifiers_re_idx CtxDep)//. rewrite He. simpl. reflexivity. @@ -530,10 +529,10 @@ Section greifiers. Lemma subReifier_reifyI_ctx_indep (r : sReifier NotCtxDep) `{!@subReifier sz NotCtxDep r rs} {X} `{!Cofe X} - (op : opid (sReifier_ops NotCtxDep r)) - (x : Ins (sReifier_ops NotCtxDep _ op) ♯ X) (y : Outs (sReifier_ops NotCtxDep _ op) ♯ X) - (σ σ' : sReifier_state NotCtxDep r ♯ X) (rest : gState_rest NotCtxDep sR_idx rs ♯ X) : - sReifier_re NotCtxDep r op (x,σ) ≡ Some (y, σ') + (op : opid (sReifier_ops r)) + (x : Ins (sReifier_ops _ op) ♯ X) (y : Outs (sReifier_ops _ op) ♯ X) + (σ σ' : sReifier_state r ♯ X) (rest : gState_rest NotCtxDep sR_idx rs ♯ X) : + sReifier_re r op (x,σ) ≡ Some (y, σ') ⊢@{iProp} gReifiers_re NotCtxDep rs (subEff_opid op) (subEff_ins x, gState_recomp NotCtxDep rest (sR_state σ)) @@ -541,15 +540,15 @@ Section greifiers. Proof. apply uPred.internal_eq_entails=>m. intros He. - eapply (sR_re (a := NotCtxDep)) in He. + eapply (@sR_re _ NotCtxDep) in He. pose proof (@gReifiers_re_idx sz NotCtxDep sR_idx rs X _ (subEff_opid op) (subEff_ins x)) as J. simpl in J. simpl. rewrite J//; clear J. - transitivity (prod_map (λ x0 : Outs (sReifier_ops NotCtxDep (rs !!! sR_idx) + transitivity (prod_map (λ x0 : Outs (sReifier_ops (rs !!! sR_idx) (subEff_opid op)) ♯ X, x0) - (λ st : sReifier_state NotCtxDep (rs !!! sR_idx) ♯ X, + (λ st : sReifier_state (rs !!! sR_idx) ♯ X, (gState_decomp' NotCtxDep sR_idx rs ^-1) (st, rest)) <$> (Some (subEff_outs y, sR_state σ'))). @@ -562,3 +561,12 @@ Section greifiers. Qed. End greifiers. + +Arguments gReifiers_cons {_ _}. +Arguments gReifiers_nil {_}. +Arguments gReifiers_ops {_ _}. +Arguments gReifiers_re {_ _}. +Arguments gReifiers_state {_ _}. +Arguments gReifiers_re_idx {_ _}. +Arguments gReifiers_re_idx_type {_ _}. +Arguments gReifiers_re_type {_ _}. diff --git a/theories/gitree/reductions.v b/theories/gitree/reductions.v index cfd14bd..e5833d9 100644 --- a/theories/gitree/reductions.v +++ b/theories/gitree/reductions.v @@ -6,8 +6,8 @@ From gitrees.gitree Require Import core reify. Section sstep. Context {A} `{!Cofe A} {a : is_ctx_dep}. Context (r : sReifier a). - Notation F := (sReifier_ops a r). - Notation stateF := (sReifier_state a r). + Notation F := (sReifier_ops r). + Notation stateF := (sReifier_state r). Notation IT := (IT F A). Notation ITV := (ITV F A). Notation stateO := (stateF ♯ IT). @@ -88,8 +88,8 @@ End sstep. Section istep. Context {A} `{!Cofe A} {a : is_ctx_dep}. Context (r : sReifier a). - Notation F := (sReifier_ops a r). - Notation stateF := (sReifier_state a r). + Notation F := (sReifier_ops r). + Notation stateF := (sReifier_state r). Notation IT := (IT F A). Notation ITV := (ITV F A). Notation stateO := (stateF ♯ IT). @@ -342,8 +342,8 @@ End istep. Section istep_ctx_indep. Context {A} `{!Cofe A}. Context (r : sReifier NotCtxDep). - Notation F := (sReifier_ops NotCtxDep r). - Notation stateF := (sReifier_state NotCtxDep r). + Notation F := (sReifier_ops r). + Notation stateF := (sReifier_state r). Notation IT := (IT F A). Notation ITV := (ITV F A). Notation stateO := (stateF ♯ IT). diff --git a/theories/gitree/reify.v b/theories/gitree/reify.v index 4ba8243..84cc1d8 100644 --- a/theories/gitree/reify.v +++ b/theories/gitree/reify.v @@ -517,3 +517,6 @@ Section reify_props. End reify_props. Arguments reify {_ _ _} _. +Arguments sReifier_ops {_}. +Arguments sReifier_re {_} _ {_ _}. +Arguments sReifier_state {_}. diff --git a/theories/gitree/weakestpre.v b/theories/gitree/weakestpre.v index a3cdb39..45d4d16 100644 --- a/theories/gitree/weakestpre.v +++ b/theories/gitree/weakestpre.v @@ -9,18 +9,18 @@ From gitrees.gitree Require Import core reify greifiers reductions. Definition gReifiers_ucmra {n} (a : is_ctx_dep) (rs : gReifiers a n) (X : ofe) `{!Cofe X} : ucmra := discrete_funUR (λ (i : fin n), - optionUR (exclR (sReifier_state a (rs !!! i) ♯ X))). + optionUR (exclR (sReifier_state (rs !!! i) ♯ X))). (** The resource corresponding to the whole global state *) Definition of_state {n} (a : is_ctx_dep) (rs : gReifiers a n) - (X : ofe) `{!Cofe X} (st : gReifiers_state a rs ♯ X) + (X : ofe) `{!Cofe X} (st : gReifiers_state rs ♯ X) : gReifiers_ucmra a rs X := λ i, Excl' (fstO (gState_decomp a i st)). (** The resource corresponding to a speicific projection out of the global state *) Definition of_idx {n} (a : is_ctx_dep) (rs : gReifiers a n) (X : ofe) `{!Cofe X} (i : fin n) - (st : sReifier_state a (rs !!! i) ♯ X) : gReifiers_ucmra a rs X. + (st : sReifier_state (rs !!! i) ♯ X) : gReifiers_ucmra a rs X. Proof. simple refine (λ j, if (decide (j = i)) then _ else None). simpl. induction e. exact (Excl' st). @@ -28,7 +28,7 @@ Defined. Lemma of_state_recomp_lookup_ne {n} (a : is_ctx_dep) (rs : gReifiers a n) (X : ofe) `{!Cofe X} - i j (σ1 σ2 : sReifier_state a (rs !!! i) ♯ X) rest : + i j (σ1 σ2 : sReifier_state (rs !!! i) ♯ X) rest : i ≠ j → of_state a rs X (gState_recomp a rest σ1) j ≡ of_state a rs X (gState_recomp a rest σ2) j. @@ -61,16 +61,16 @@ Section ucmra. #[export] Instance of_state_proper : Proper ((≡) ==> (≡)) of_state. Proof. apply ne_proper, _. Qed. - Lemma of_state_valid (σ : gReifiers_state a rs ♯ X) : ✓ (of_state σ). + Lemma of_state_valid (σ : gReifiers_state rs ♯ X) : ✓ (of_state σ). Proof. intro; done. Qed. - Lemma of_state_recomp_lookup i (σ : sReifier_state a (rs !!! i) ♯ X) rest : + Lemma of_state_recomp_lookup i (σ : sReifier_state (rs !!! i) ♯ X) rest : of_state (gState_recomp a rest σ) i ≡ Excl' σ. Proof. unfold of_state. rewrite gState_decomp_recomp. done. Qed. - Lemma of_state_decomp_local_update i (σ σ1 σ2 : sReifier_state a (rs !!! i) ♯ X) rest : + Lemma of_state_decomp_local_update i (σ σ1 σ2 : sReifier_state (rs !!! i) ♯ X) rest : (of_state (gState_recomp a rest σ1), of_idx i σ2) ~l~> (of_state (gState_recomp a rest σ), of_idx i σ). Proof. @@ -108,10 +108,10 @@ End ucmra. Section weakestpre. Context {n : nat} (a : is_ctx_dep) (rs : gReifiers a n) {A} `{!Cofe A}. Notation rG := (gReifiers_sReifier a rs). - Notation F := (sReifier_ops a rG). + Notation F := (sReifier_ops rG). Notation IT := (IT F A). Notation ITV := (ITV F A). - Notation stateF := (gReifiers_state a rs). + Notation stateF := (gReifiers_state rs). Notation stateO := (stateF ♯ IT). Notation stateR := (gReifiers_ucmra a rs IT). Let of_state := (of_state a rs IT). @@ -141,10 +141,10 @@ Section weakestpre. Definition has_full_state `{!stateG Σ} (σ : stateO) : iProp Σ := (own stateG_name (◯ (of_state σ)))%I. Definition has_state_idx `{!stateG Σ} - (i : fin n) (σ : sReifier_state a (rs !!! i) ♯ IT) : iProp Σ := + (i : fin n) (σ : sReifier_state (rs !!! i) ♯ IT) : iProp Σ := (own stateG_name (◯ (of_idx i σ)))%I. Definition has_substate {sR : sReifier a} `{!stateG Σ} `{!subReifier sR rs} - (σ : sReifier_state a sR ♯ IT) : iProp Σ := + (σ : sReifier_state sR ♯ IT) : iProp Σ := (own stateG_name (◯ (of_idx sR_idx (sR_state σ))))%I. #[export] Instance state_interp_ne `{!stateG Σ} : NonExpansive state_interp. @@ -163,7 +163,7 @@ Section weakestpre. Qed. Lemma state_interp_has_state_idx_agree (i : fin n) - (σ1 σ2 : sReifier_state a (rs !!! i) ♯ IT) + (σ1 σ2 : sReifier_state (rs !!! i) ♯ IT) (rest : gState_rest a i rs ♯ IT) `{!stateG Σ} : state_interp (gState_recomp a rest σ1) -∗ has_state_idx i σ2 -∗ σ1 ≡ σ2. Proof. @@ -176,7 +176,7 @@ Section weakestpre. Qed. Lemma state_interp_has_state_idx_update (i : fin n) - (σ σ1 σ2 : sReifier_state a (rs !!! i) ♯ IT) + (σ σ1 σ2 : sReifier_state (rs !!! i) ♯ IT) (rest : gState_rest a i rs ♯ IT) `{!stateG Σ} : state_interp (gState_recomp a rest σ1) -∗ has_state_idx i σ2 ==∗ state_interp (gState_recomp a rest σ) ∗ has_state_idx i σ. @@ -402,7 +402,7 @@ Section weakestpre. Opaque gState_recomp. (* We can generalize this based on the stuckness bit *) - Lemma wp_reify_idx E1 E2 s Φ i (lop : opid (sReifier_ops a (rs !!! i))) : + Lemma wp_reify_idx E1 E2 s Φ i (lop : opid (sReifier_ops (rs !!! i))) : let op : opid F := (existT i lop) in forall (x : Ins (F op) ♯ IT) (k : Outs (F op) ♯ IT -n> laterO IT), @@ -450,7 +450,7 @@ Section weakestpre. iRewrite -"Hb". by iFrame. Qed. - Lemma wp_reify E1 s Φ i (lop : opid (sReifier_ops a (rs !!! i))) + Lemma wp_reify E1 s Φ i (lop : opid (sReifier_ops (rs !!! i))) x k σ σ' β : let op : opid F := (existT i lop) in (∀ rest, reify (Vis op x k) (gState_recomp a rest σ) ≡ (gState_recomp a rest σ', Tick β)) → @@ -668,7 +668,7 @@ Section weakestpre_specific. Context {n : nat} {A} `{!Cofe A}. Notation rG a rs := (gReifiers_sReifier (n := n) a rs). - Notation F a rs := (sReifier_ops a (rG a rs)). + Notation F a rs := (sReifier_ops (rG a rs)). Notation IT a rs := (IT (F a rs) A). Notation ITV a rs := (ITV (F a rs) A). Notation stateF a rs := (gReifiers_state a rs). @@ -689,13 +689,13 @@ Section weakestpre_specific. Lemma wp_reify_idx_ctx_dep (rs : gReifiers CtxDep n) `{!@stateG _ CtxDep rs A _ Σ} E1 E2 s Φ i - (lop : opid (sReifier_ops CtxDep (rs !!! i))) : + (lop : opid (sReifier_ops (rs !!! i))) : let op : opid (F CtxDep rs) := (existT i lop) in forall (x : Ins (F CtxDep rs op) ♯ IT CtxDep rs) (k : Outs (F CtxDep rs op) ♯ IT CtxDep rs -n> laterO (IT CtxDep rs)), (|={E1,E2}=> ∃ σ y σ' β, has_state_idx CtxDep rs i σ - ∗ sReifier_re CtxDep (rs !!! i) lop (x, σ, k) ≡ Some (y, σ') + ∗ sReifier_re (rs !!! i) lop (x, σ, k) ≡ Some (y, σ') ∗ y ≡ Next β ∗ ▷ (£ 1 -∗ has_state_idx CtxDep rs i σ' ={E2,E1}=∗ wp CtxDep rs β s E1 Φ)) -∗ wp CtxDep rs (Vis op x k) s E1 Φ. @@ -708,12 +708,12 @@ Section weakestpre_specific. iFrame "Hlst". iIntros (rest). iFrame "H". - iAssert (gReifiers_re CtxDep rs op (x, gState_recomp CtxDep rest σ, _) + iAssert (gReifiers_re rs _ _ op (x, gState_recomp CtxDep rest σ, _) ≡ Some (y, gState_recomp CtxDep rest σ'))%I with "[Hreify]" as "Hgreify". - { rewrite (gReifiers_re_idx CtxDep). + { rewrite (@gReifiers_re_idx _ CtxDep). iAssert (optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) - (sReifier_re CtxDep (rs !!! i) lop (x, σ, k)) + (sReifier_re (rs !!! i) lop (x, σ, k)) ≡ optionO_map (prodO_map idfun (gState_recomp CtxDep rest)) (Some (y, σ')))%I with "[Hreify]" as "H". - iApply (f_equivI with "Hreify"). @@ -726,12 +726,12 @@ Section weakestpre_specific. Lemma wp_reify_idx_ctx_indep (rs : gReifiers NotCtxDep n) `{!@stateG _ NotCtxDep rs A _ Σ} E1 E2 s Φ i - (lop : opid (sReifier_ops NotCtxDep (rs !!! i))) : + (lop : opid (sReifier_ops (rs !!! i))) : let op : opid (F NotCtxDep rs) := (existT i lop) in forall (x : Ins (F NotCtxDep rs op) ♯ IT NotCtxDep rs) (k : Outs (F NotCtxDep rs op) ♯ IT NotCtxDep rs -n> laterO (IT NotCtxDep rs)), (|={E1,E2}=> ∃ σ y σ' β, has_state_idx NotCtxDep rs i σ - ∗ sReifier_re NotCtxDep (rs !!! i) lop (x, σ) ≡ Some (y, σ') + ∗ sReifier_re (rs !!! i) lop (x, σ) ≡ Some (y, σ') ∗ k y ≡ Next β ∗ ▷ (£ 1 -∗ has_state_idx NotCtxDep rs i σ' ={E2,E1}=∗ wp NotCtxDep rs β s E1 Φ)) -∗ wp NotCtxDep rs (Vis op x k) s E1 Φ. @@ -744,7 +744,7 @@ Section weakestpre_specific. iFrame "Hlst". iIntros (rest). iFrame "H". - iAssert (gReifiers_re NotCtxDep rs op (x, gState_recomp NotCtxDep rest σ) + iAssert (@gReifiers_re _ NotCtxDep rs _ _ op (x, gState_recomp NotCtxDep rest σ) ≡ Some (y, gState_recomp NotCtxDep rest σ'))%I with "[Hreify]" as "Hgreify". { pose proof (@gReifiers_re_idx n NotCtxDep i rs (IT NotCtxDep rs)) as J. @@ -752,7 +752,7 @@ Section weakestpre_specific. simpl. rewrite J; clear J. iAssert (optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) - (sReifier_re NotCtxDep (rs !!! i) lop (x, σ)) + (sReifier_re (rs !!! i) lop (x, σ)) ≡ optionO_map (prodO_map idfun (gState_recomp NotCtxDep rest)) (Some (y, σ')))%I with "[Hreify]" as "H". - iApply (f_equivI with "Hreify"). @@ -765,10 +765,10 @@ Section weakestpre_specific. Lemma wp_subreify_ctx_dep' (rs : gReifiers CtxDep n) `{!@stateG _ CtxDep rs A _ Σ} E1 E2 s Φ sR `{!subReifier sR rs} - (op : opid (sReifier_ops CtxDep sR)) (x : Ins (sReifier_ops CtxDep sR op) ♯ (IT CtxDep rs)) + (op : opid (sReifier_ops sR)) (x : Ins (sReifier_ops sR op) ♯ (IT CtxDep rs)) (k : Outs (F CtxDep rs (subEff_opid op)) ♯ IT CtxDep rs -n> laterO (IT CtxDep rs)) : (|={E1,E2}=> ∃ σ y σ' β, has_substate CtxDep rs σ ∗ - sReifier_re CtxDep sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') + sReifier_re sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') ∗ y ≡ Next β ∗ ▷ (£ 1 -∗ has_substate CtxDep rs σ' ={E2,E1}=∗ wp CtxDep rs β s E1 Φ)) -∗ wp CtxDep rs (Vis (subEff_opid op) (subEff_ins x) k) s E1 Φ. @@ -791,10 +791,10 @@ Section weakestpre_specific. Lemma wp_subreify_ctx_indep' (rs : gReifiers NotCtxDep n) `{!@stateG _ NotCtxDep rs A _ Σ} E1 E2 s Φ sR `{!subReifier sR rs} - (op : opid (sReifier_ops NotCtxDep sR)) (x : Ins (sReifier_ops NotCtxDep sR op) ♯ (IT NotCtxDep rs)) + (op : opid (sReifier_ops sR)) (x : Ins (sReifier_ops sR op) ♯ (IT NotCtxDep rs)) (k : Outs (F NotCtxDep rs (subEff_opid op)) ♯ IT NotCtxDep rs -n> laterO (IT NotCtxDep rs)) : (|={E1,E2}=> ∃ σ y σ' β, has_substate NotCtxDep rs σ ∗ - sReifier_re NotCtxDep sR op (x, σ) ≡ Some (y, σ') + sReifier_re sR op (x, σ) ≡ Some (y, σ') ∗ k (subEff_outs y) ≡ Next β ∗ ▷ (£ 1 -∗ has_substate NotCtxDep rs σ' ={E2,E1}=∗ wp NotCtxDep rs β s E1 Φ)) -∗ wp NotCtxDep rs (Vis (subEff_opid op) (subEff_ins x) k) s E1 Φ. @@ -810,11 +810,11 @@ Section weakestpre_specific. Lemma wp_subreify_ctx_dep (rs : gReifiers CtxDep n) `{!@stateG _ CtxDep rs A _ Σ} E1 s Φ sR `{!subReifier sR rs} - (op : opid (sReifier_ops CtxDep sR)) - (x : Ins (sReifier_ops CtxDep sR op) ♯ IT CtxDep rs) (y : laterO (IT CtxDep rs)) + (op : opid (sReifier_ops sR)) + (x : Ins (sReifier_ops sR op) ♯ IT CtxDep rs) (y : laterO (IT CtxDep rs)) (k : Outs (F CtxDep rs (subEff_opid op)) ♯ IT CtxDep rs -n> laterO (IT CtxDep rs)) - (σ σ' : sReifier_state CtxDep sR ♯ IT CtxDep rs) β : - sReifier_re CtxDep sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') → + (σ σ' : sReifier_state sR ♯ IT CtxDep rs) β : + sReifier_re sR op (x, σ, (k ◎ subEff_outs)) ≡ Some (y, σ') → y ≡ Next β → has_substate CtxDep rs σ -∗ ▷ (£ 1 -∗ has_substate CtxDep rs σ' -∗ wp CtxDep rs β s E1 Φ) @@ -844,12 +844,12 @@ Section weakestpre_specific. Lemma wp_subreify_ctx_indep (rs : gReifiers NotCtxDep n) `{!@stateG _ NotCtxDep rs A _ Σ} E1 s Φ sR `{!subReifier sR rs} - (op : opid (sReifier_ops NotCtxDep sR)) - (x : Ins (sReifier_ops NotCtxDep sR op) ♯ IT NotCtxDep rs) - (y : Outs (sReifier_ops NotCtxDep sR op) ♯ IT NotCtxDep rs) + (op : opid (sReifier_ops sR)) + (x : Ins (sReifier_ops sR op) ♯ IT NotCtxDep rs) + (y : Outs (sReifier_ops sR op) ♯ IT NotCtxDep rs) (k : Outs (F NotCtxDep rs (subEff_opid op)) ♯ IT NotCtxDep rs -n> laterO (IT NotCtxDep rs)) - (σ σ' : sReifier_state NotCtxDep sR ♯ IT NotCtxDep rs) β : - sReifier_re NotCtxDep sR op (x, σ) ≡ Some (y, σ') → + (σ σ' : sReifier_state sR ♯ IT NotCtxDep rs) β : + sReifier_re sR op (x, σ) ≡ Some (y, σ') → k (subEff_outs y) ≡ Next β → has_substate NotCtxDep rs σ -∗ ▷ (£ 1 -∗ has_substate NotCtxDep rs σ' -∗ wp NotCtxDep rs β s E1 Φ) @@ -870,7 +870,7 @@ End weakestpre_specific. Section weakestpre_bind. Context {n : nat} (rs : gReifiers NotCtxDep n) {A} `{!Cofe A}. Notation rG := (gReifiers_sReifier NotCtxDep rs). - Notation F := (sReifier_ops NotCtxDep rG). + Notation F := (sReifier_ops rG). Notation IT := (IT F A). Notation ITV := (ITV F A). Notation stateF := (gReifiers_state NotCtxDep rs). @@ -967,111 +967,117 @@ Arguments stateΣ {n _} rs A {_}. Definition notStuck : stuckness := λ e, False. - Notation "'WP@{' re } α @ s ; E {{ Φ } }" := (wp re α s E Φ) +Notation "'WP@{' re } α @ s ; E {{ Φ } }" := + (wp re α s E Φ) (at level 20, α, s, Φ at level 200, only parsing) : bi_scope. - Notation "'WP@{' re } α @ s ; E {{ v , Q } }" := (wp re α s E (λ v, Q)) +Notation "'WP@{' re } α @ s ; E {{ v , Q } }" := + (wp re α s E (λ v, Q)) (at level 20, α, s, Q at level 200, - format "'[hv' 'WP@{' re } α '/' @ s ; E '/' {{ '[' v , '/' Q ']' } } ']'") : bi_scope. + format "'[hv' 'WP@{' re } α '/' @ s ; E '/' {{ '[' v , '/' Q ']' } } ']'") : bi_scope. - Notation "'WP@{' re } α @ s {{ β , Φ } }" := (wp re α s ⊤ (λ β, Φ)) +Notation "'WP@{' re } α @ s {{ β , Φ } }" := + (wp re α s ⊤ (λ β, Φ)) (at level 20, α, Φ at level 200, - format "'WP@{' re } α @ s {{ β , Φ } }") : bi_scope. + format "'WP@{' re } α @ s {{ β , Φ } }") : bi_scope. - Notation "'WP@{' re } α @ s {{ Φ } }" := (wp re α s ⊤ Φ) +Notation "'WP@{' re } α @ s {{ Φ } }" := + (wp re α s ⊤ Φ) (at level 20, α, Φ at level 200, - format "'WP@{' re } α @ s {{ Φ } }") : bi_scope. + format "'WP@{' re } α @ s {{ Φ } }") : bi_scope. - Notation "'WP@{' re } α {{ β , Φ } }" := (wp re α notStuck ⊤ (λ β, Φ)) +Notation "'WP@{' re } α {{ β , Φ } }" := + (wp re α notStuck ⊤ (λ β, Φ)) (at level 20, α, Φ at level 200, - format "'WP@{' re } α {{ β , Φ } }") : bi_scope. + format "'WP@{' re } α {{ β , Φ } }") : bi_scope. - Notation "'WP@{' re } α {{ Φ } }" := (wp re α notStuck ⊤ Φ) +Notation "'WP@{' re } α {{ Φ } }" := + (wp re α notStuck ⊤ Φ) (at level 20, α, Φ at level 200, format "'WP@{' re } α {{ Φ } }") : bi_scope. - Lemma wp_adequacy cr Σ `{!invGpreS Σ} n a (rs : gReifiers a n) - {A} `{!Cofe A} `{!statePreG rs A Σ} - (α : IT _ A) σ βv σ' s k (ψ : (ITV (gReifiers_ops a rs) A) → Prop) : - ssteps (gReifiers_sReifier a rs) α σ (IT_of_V βv) σ' k → - (∀ `{H1 : !invGS Σ} `{H2: !stateG rs A Σ}, - ∃ Φ, NonExpansive Φ ∧ (∀ βv, Φ βv ⊢ ⌜ψ βv⌝) - ∧ (£ cr ∗ has_full_state σ ⊢ WP@{rs} α @ s {{ Φ }})%I) → - ψ βv. - Proof. - intros Hst Hprf. - cut (⊢ ⌜ψ βv⌝ : iProp Σ)%I. - { intros HH. eapply uPred.pure_soundness; eauto. } - eapply (step_fupdN_soundness_lc _ 0 (cr + 3*k)). - intros Hinv. iIntros "[Hcr Hlc]". - iMod (new_state_interp a rs σ) as (sg) "[Hs Hs2]". - destruct (Hprf Hinv sg) as (Φ & HΦ & HΦψ & Hprf'). - iPoseProof (Hprf' with "[$Hcr $Hs2]") as "Hic". - iPoseProof (wp_ssteps with "[$Hs $Hic]") as "Hphi". - { eassumption. } - iMod ("Hphi" with "Hlc") as "[Hst H]". - rewrite wp_val_inv; eauto. - iMod "H" as "H". - rewrite HΦψ. iFrame "H". - by iApply fupd_mask_intro_discard. - Qed. +Lemma wp_adequacy cr Σ `{!invGpreS Σ} n a (rs : gReifiers a n) + {A} `{!Cofe A} `{!statePreG rs A Σ} + (α : IT _ A) σ βv σ' s k (ψ : (ITV (gReifiers_ops rs) A) → Prop) : + ssteps (gReifiers_sReifier a rs) α σ (IT_of_V βv) σ' k → + (∀ `{H1 : !invGS Σ} `{H2: !stateG rs A Σ}, + ∃ Φ, NonExpansive Φ ∧ (∀ βv, Φ βv ⊢ ⌜ψ βv⌝) + ∧ (£ cr ∗ has_full_state σ ⊢ WP@{rs} α @ s {{ Φ }})%I) → + ψ βv. +Proof. + intros Hst Hprf. + cut (⊢ ⌜ψ βv⌝ : iProp Σ)%I. + { intros HH. eapply uPred.pure_soundness; eauto. } + eapply (step_fupdN_soundness_lc _ 0 (cr + 3*k)). + intros Hinv. iIntros "[Hcr Hlc]". + iMod (new_state_interp a rs σ) as (sg) "[Hs Hs2]". + destruct (Hprf Hinv sg) as (Φ & HΦ & HΦψ & Hprf'). + iPoseProof (Hprf' with "[$Hcr $Hs2]") as "Hic". + iPoseProof (wp_ssteps with "[$Hs $Hic]") as "Hphi". + { eassumption. } + iMod ("Hphi" with "Hlc") as "[Hst H]". + rewrite wp_val_inv; eauto. + iMod "H" as "H". + rewrite HΦψ. iFrame "H". + by iApply fupd_mask_intro_discard. +Qed. - Lemma wp_safety cr Σ `{!invGpreS Σ} n a (rs : gReifiers a n) - {A} `{!Cofe A} `{!statePreG rs A Σ} s k - (α β : IT (gReifiers_ops a rs) A) (σ σ' : gReifiers_state a rs ♯ IT (gReifiers_ops a rs) A) : - (∀ Σ P Q, @disjunction_property Σ P Q) → - ssteps (gReifiers_sReifier a rs) α σ β σ' k → - IT_to_V β ≡ None → - (∀ `{H1 : !invGS_gen HasLc Σ} `{H2: !stateG rs A Σ}, - ∃ Φ, NonExpansive Φ ∧ (£ cr ∗ has_full_state σ ⊢ WP@{rs} α @ s {{ Φ }})%I) → - ((∃ β1 σ1, sstep (gReifiers_sReifier a rs) β σ' β1 σ1) - ∨ (∃ e, β ≡ Err e ∧ s e)). - Proof. - Opaque istep. - intros Hdisj Hstep Hbv Hwp. - cut (⊢@{iProp Σ} (∃ β1 σ1, istep (gReifiers_sReifier a rs) β σ' β1 σ1) - ∨ (∃ e, β ≡ Err e ∧ ⌜s e⌝))%I. - { intros [Hprf | Hprf]%Hdisj. - - left. - apply (istep_safe_sstep _ (Σ:=Σ)). - { apply Hdisj. } - done. - - right. - destruct (IT_dont_confuse β) - as [[e Ha] | [[m Ha] | [ [g Ha] | [[α' Ha]|[op [i [ko Ha]]]] ]]]. - + exists e. split; eauto. - eapply uPred.pure_soundness. - iPoseProof (Hprf) as "H". - iDestruct "H" as (e') "[He %Hs]". rewrite Ha. - iPoseProof (Err_inj' with "He") as "%He". - iPureIntro. rewrite He//. - + exfalso. eapply uPred.pure_soundness. - iPoseProof (Hprf) as "H". - iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. - iApply (IT_ret_err_ne with "Ha"). - + exfalso. eapply uPred.pure_soundness. - iPoseProof (Hprf) as "H". - iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. - iApply (IT_fun_err_ne with "Ha"). - + exfalso. eapply uPred.pure_soundness. - iPoseProof (Hprf) as "H". - iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. - iApply (IT_tick_err_ne with "Ha"). - + exfalso. eapply uPred.pure_soundness. - iPoseProof (Hprf) as "H". - iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. - iApply (IT_vis_err_ne with "Ha"). } - eapply (step_fupdN_soundness_lc _ 0 (cr + (3*k+2))). - intros Hinv. iIntros "[Hcr Hlc]". - iMod (new_state_interp a rs σ) as (sg) "[Hs Hs2]". - destruct (Hwp Hinv sg) as (Φ & HΦ & Hprf'). - iPoseProof (Hprf' with "[$Hs2 $Hcr]") as "Hic". - iPoseProof (wp_ssteps_isafe with "[$Hs $Hic]") as "H". - { eassumption. } - iMod ("H" with "Hlc") as "[H | H]". - { iDestruct "H" as (βv) "%Hbeta". - exfalso. rewrite Hbeta in Hbv. - inversion Hbv. } - iFrame "H". - by iApply fupd_mask_intro_discard. - Qed. +Lemma wp_safety cr Σ `{!invGpreS Σ} n a (rs : gReifiers a n) + {A} `{!Cofe A} `{!statePreG rs A Σ} s k + (α β : IT (gReifiers_ops rs) A) (σ σ' : gReifiers_state rs ♯ IT (gReifiers_ops rs) A) : + (∀ Σ P Q, @disjunction_property Σ P Q) → + ssteps (gReifiers_sReifier a rs) α σ β σ' k → + IT_to_V β ≡ None → + (∀ `{H1 : !invGS_gen HasLc Σ} `{H2: !stateG rs A Σ}, + ∃ Φ, NonExpansive Φ ∧ (£ cr ∗ has_full_state σ ⊢ WP@{rs} α @ s {{ Φ }})%I) → + ((∃ β1 σ1, sstep (gReifiers_sReifier a rs) β σ' β1 σ1) + ∨ (∃ e, β ≡ Err e ∧ s e)). +Proof. + Opaque istep. + intros Hdisj Hstep Hbv Hwp. + cut (⊢@{iProp Σ} (∃ β1 σ1, istep (gReifiers_sReifier a rs) β σ' β1 σ1) + ∨ (∃ e, β ≡ Err e ∧ ⌜s e⌝))%I. + { intros [Hprf | Hprf]%Hdisj. + - left. + apply (istep_safe_sstep _ (Σ:=Σ)). + { apply Hdisj. } + done. + - right. + destruct (IT_dont_confuse β) + as [[e Ha] | [[m Ha] | [ [g Ha] | [[α' Ha]|[op [i [ko Ha]]]] ]]]. + + exists e. split; eauto. + eapply uPred.pure_soundness. + iPoseProof (Hprf) as "H". + iDestruct "H" as (e') "[He %Hs]". rewrite Ha. + iPoseProof (Err_inj' with "He") as "%He". + iPureIntro. rewrite He//. + + exfalso. eapply uPred.pure_soundness. + iPoseProof (Hprf) as "H". + iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. + iApply (IT_ret_err_ne with "Ha"). + + exfalso. eapply uPred.pure_soundness. + iPoseProof (Hprf) as "H". + iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. + iApply (IT_fun_err_ne with "Ha"). + + exfalso. eapply uPred.pure_soundness. + iPoseProof (Hprf) as "H". + iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. + iApply (IT_tick_err_ne with "Ha"). + + exfalso. eapply uPred.pure_soundness. + iPoseProof (Hprf) as "H". + iDestruct "H" as (e') "[Ha Hs]". rewrite Ha. + iApply (IT_vis_err_ne with "Ha"). } + eapply (step_fupdN_soundness_lc _ 0 (cr + (3*k+2))). + intros Hinv. iIntros "[Hcr Hlc]". + iMod (new_state_interp a rs σ) as (sg) "[Hs Hs2]". + destruct (Hwp Hinv sg) as (Φ & HΦ & Hprf'). + iPoseProof (Hprf' with "[$Hs2 $Hcr]") as "Hic". + iPoseProof (wp_ssteps_isafe with "[$Hs $Hic]") as "H". + { eassumption. } + iMod ("H" with "Hlc") as "[H | H]". + { iDestruct "H" as (βv) "%Hbeta". + exfalso. rewrite Hbeta in Hbv. + inversion Hbv. } + iFrame "H". + by iApply fupd_mask_intro_discard. +Qed. diff --git a/theories/lang_generic.v b/theories/lang_generic.v index a4a3663..79b5a69 100644 --- a/theories/lang_generic.v +++ b/theories/lang_generic.v @@ -1,11 +1,10 @@ -From gitrees Require Import prelude. -From gitrees Require Import gitree. +From gitrees Require Import prelude gitree utils.finite_sets. Require Import List. Import ListNotations. Require Import Binding.Lib Binding.Set. -Section interp. +Section ctx_interp. Local Open Scope type. Context {E: opsInterp}. Context {R} `{!Cofe R}. @@ -26,6 +25,9 @@ Section interp. solve_proper. Qed. + Program Definition ı_scope : interp_scope Empty_set + := λne (x : ∅), match x with end. + Definition interp_scope_split {S1 S2 : Set} : interp_scope (sum S1 S2) -n> interp_scope S1 * interp_scope S2. Proof. @@ -63,7 +65,7 @@ Section interp. Program Definition ren_scope {S S'} (δ : S [→] S') (env : interp_scope S') : interp_scope S := λne x, env (δ x). -End interp. +End ctx_interp. (* Common definitions and lemmas for Kripke logical relations *) Section kripke_logrel. @@ -73,7 +75,7 @@ Section kripke_logrel. Variable rs : gReifiers a sz. Context {R} `{!Cofe R}. - Notation F := (gReifiers_ops a rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ}. @@ -133,7 +135,7 @@ Section kripke_logrel_ctx_indep. Variable rs : gReifiers NotCtxDep sz. Context {R} `{!Cofe R}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ}. @@ -165,3 +167,157 @@ Section kripke_logrel_ctx_indep. End kripke_logrel_ctx_indep. Arguments expr_pred_bind {_ _ _ _ _ _ _ _ _ _} f {_ _}. + +Section tm_interp. + Context {sz : nat} {a : is_ctx_dep}. + Variable rs : gReifiers a sz. + Context {R} `{!Cofe R}. + + Notation F := (gReifiers_ops rs). + Notation IT := (IT F R). + Notation ITV := (ITV F R). + Context `{!invGS Σ, !stateG rs R Σ}. + Notation iProp := (iProp Σ). + + Context {A : ofe}. + Variable (P : A -n> iProp). + + Variable (ty : Set). + Variable (interp_ty : ty → (ITV -n> iProp)). + Variable (kripke : IT → (ITV -n> iProp) → iProp). + + Definition ssubst_valid1 {S : Set} + (Γ : S -> ty) + (ss : interp_scope S) : iProp := + (∀ x, □ kripke (ss x) (interp_ty (Γ x)))%I. + + Global Instance ssubst_valid_pers `{∀ τ β, Persistent (interp_ty τ β)} + {S : Set} (Γ : S → ty) ss : Persistent (ssubst_valid1 Γ ss). + Proof. apply _. Qed. + +End tm_interp. + +Section tm_interp_fin. + Context {sz : nat} {a : is_ctx_dep}. + Variable rs : gReifiers a sz. + Context {R} `{!Cofe R}. + + Notation F := (gReifiers_ops rs). + Notation IT := (IT F R). + Notation ITV := (ITV F R). + Context `{!invGS Σ, !stateG rs R Σ}. + Notation iProp := (iProp Σ). + + Context {A : ofe}. + Variable (P : A -n> iProp). + + Variable (ty : Set). + Variable (interp_ty : ty → (ITV -n> iProp)). + Variable (kripke : IT → (ITV -n> iProp) → iProp). + + Program Definition ssubst_valid_fin1 {S : Set} `{!EqDecision S} `{!Finite S} + (Ω : S → ty) (ss : interp_scope S) : iProp + := ([∗ set] x ∈ (fin_to_set S), + (kripke (ss x) (interp_ty (Ω x)))%I). + + Context (Q : iProp). + + Definition valid_fin1 {S : Set} `{!EqDecision S} `{!Finite S} (Ω : S → ty) + (α : interp_scope S -n> IT) (τ : ty) : iProp := + ∀ ss, Q + -∗ (ssubst_valid_fin1 Ω ss) + -∗ kripke (α ss) (interp_ty τ). + + Lemma ssubst_valid_fin_empty1 (αs : interp_scope ∅) : + ⊢ ssubst_valid_fin1 □ αs. + Proof. + iStartProof. + unfold ssubst_valid_fin1. + rewrite fin_to_set_empty. + by iApply big_sepS_empty. + Qed. + + Lemma ssubst_valid_fin_app1 + {S1 S2 : Set} `{!EqDecision S1} `{!Finite S1} + `{!EqDecision S2} `{!Finite S2} + `{!EqDecision (S1 + S2)} `{!Finite (S1 + S2)} + (Ω1 : S1 → ty) (Ω2 : S2 → ty) + (αs : interp_scope (sum S1 S2)) : + (ssubst_valid_fin1 (sum_map' Ω1 Ω2) αs) ⊢ + (ssubst_valid_fin1 Ω1 (interp_scope_split αs).1) + ∗ (ssubst_valid_fin1 Ω2 (interp_scope_split αs).2). + Proof. + iIntros "H". + rewrite /ssubst_valid_fin1 fin_to_set_sum big_sepS_union; first last. + { + apply elem_of_disjoint. + intros [x | x]. + - rewrite !elem_of_list_to_set. + intros _ H2. + apply elem_of_list_fmap_2 in H2. + destruct H2 as [y [H2 H2']]; inversion H2. + - rewrite !elem_of_list_to_set. + intros H1 _. + apply elem_of_list_fmap_2 in H1. + destruct H1 as [y [H1 H1']]; inversion H1. + } + iDestruct "H" as "(H1 & H2)". + iSplitL "H1". + - rewrite big_opS_list_to_set; first last. + { + apply NoDup_fmap. + - intros ??; by inversion 1. + - apply NoDup_elements. + } + rewrite big_sepL_fmap /=. + rewrite big_sepS_elements. + iFrame "H1". + - rewrite big_opS_list_to_set; first last. + { + apply NoDup_fmap. + - intros ??; by inversion 1. + - apply NoDup_elements. + } + rewrite big_sepL_fmap /=. + rewrite big_sepS_elements. + iFrame "H2". + Qed. + + Lemma ssubst_valid_fin_cons1 {S : Set} `{!EqDecision S} `{!Finite S} + (Ω : S → ty) (αs : interp_scope S) τ t : + ssubst_valid_fin1 Ω αs ∗ kripke t (interp_ty τ) ⊢ ssubst_valid_fin1 (Ω ▹ τ) (extend_scope αs t). + Proof. + iIntros "(H & G)". + rewrite /ssubst_valid_fin1. + rewrite fin_to_set_inc /=. + rewrite big_sepS_union; first last. + { + apply elem_of_disjoint. + intros [| x]. + - rewrite !elem_of_list_to_set. + intros _ H2. + apply elem_of_list_fmap_2 in H2. + destruct H2 as [y [H2 H2']]; inversion H2. + - rewrite !elem_of_list_to_set. + intros H1 _. + apply elem_of_singleton_1 in H1. + inversion H1. + } + iSplitL "G". + - rewrite big_opS_singleton. + iFrame "G". + - erewrite big_opS_set_map. + + iFrame "H". + + intros ?? H; by inversion H. + Qed. + + Lemma ssubst_valid_fin_lookup1 {S : Set} `{!EqDecision S} `{!Finite S} + (Ω : S → ty) (αs : interp_scope S) x : + ssubst_valid_fin1 Ω αs ⊢ kripke (αs x) (interp_ty (Ω x)). + Proof. + iIntros "H". + iDestruct (big_sepS_elem_of_acc _ _ x with "H") as "($ & _)"; + first apply elem_of_fin_to_set. + Qed. + +End tm_interp_fin. diff --git a/theories/lib/factorial.v b/theories/lib/factorial.v index 259395e..e30fc24 100644 --- a/theories/lib/factorial.v +++ b/theories/lib/factorial.v @@ -4,9 +4,10 @@ From gitrees.effects Require Import store. From gitrees.lib Require Import while. Section fact. - Definition rs : gReifiers NotCtxDep 2 := - gReifiers_cons NotCtxDep reify_io (gReifiers_cons NotCtxDep reify_store (gReifiers_nil NotCtxDep)). - Notation F := (gReifiers_ops NotCtxDep rs). + Context (n' : nat) (r : gReifiers NotCtxDep n'). + Definition rs : gReifiers NotCtxDep (S (S n')) := + (gReifiers_cons reify_io (gReifiers_cons reify_store r)). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R, !SubOfe unitO R}. Notation IT := (IT F R). diff --git a/theories/lib/iter.v b/theories/lib/iter.v index e19dc18..c43e367 100644 --- a/theories/lib/iter.v +++ b/theories/lib/iter.v @@ -62,7 +62,7 @@ Section iter_wp. Variable (rs : gReifiers NotCtxDep sz). Context {R} `{!Cofe R}. Context `{!SubOfe natO R}. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Notation IT := (IT F R). Notation ITV := (ITV F R). Context `{!invGS Σ, !stateG rs R Σ}. diff --git a/theories/prelude.v b/theories/prelude.v index 84cd443..9592851 100644 --- a/theories/prelude.v +++ b/theories/prelude.v @@ -7,6 +7,9 @@ From iris.si_logic Require Import bi siprop. From iris.proofmode Require Import classes tactics modality_instances coq_tactics reduction. +Definition sum_map' {A B C : Set} (f : A → C) (g : B → C) : sum A B → C := + λ x, match x with | inl x' => f x' | inr x' => g x' end. + Program Definition idfun {A : ofe} : A -n> A := λne x, x. (** OFEs stuff *) diff --git a/theories/program_logic.v b/theories/program_logic.v index d4d514f..304248d 100644 --- a/theories/program_logic.v +++ b/theories/program_logic.v @@ -4,7 +4,7 @@ From gitrees Require Import gitree. Section program_logic. Context {sz : nat} {a : is_ctx_dep}. Variable rs : gReifiers a sz. - Notation F := (gReifiers_ops a rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Notation IT := (IT F R). Notation ITV := (ITV F R). @@ -27,7 +27,7 @@ End program_logic. Section program_logic_ctx_indep. Context {sz : nat}. Variable rs : gReifiers NotCtxDep sz. - Notation F := (gReifiers_ops NotCtxDep rs). + Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Notation IT := (IT F R). Notation ITV := (ITV F R). diff --git a/theories/utils/finite_sets.v b/theories/utils/finite_sets.v new file mode 100644 index 0000000..463e388 --- /dev/null +++ b/theories/utils/finite_sets.v @@ -0,0 +1,221 @@ +From stdpp Require Export finite gmap. +Require Export Binding.Resolver Binding.Lib Binding.Set Binding.Auto Binding.Env. + +Lemma fin_to_set_sum {S1 S2 : Set} `{!EqDecision S1} `{!EqDecision S2} + `{!Finite S1} `{!Finite S2} `{HE : EqDecision (S1 + S2)} `{HF : @Finite (S1 + S2) HE} + : fin_to_set (S1 + S2) + = (set_map inl (fin_to_set S1 : gset S1)) + ∪ (set_map inr (fin_to_set S2 : gset S2)) + :> @gset (S1 + S2) HE (@finite_countable _ HE HF). +Proof. + apply set_eq. + intros [x|x]; simpl; split; intros _. + - apply elem_of_union; left. + apply elem_of_map_2. + apply elem_of_fin_to_set. + - apply elem_of_fin_to_set. + - apply elem_of_union; right. + apply elem_of_map_2. + apply elem_of_fin_to_set. + - apply elem_of_fin_to_set. +Qed. + +Lemma fin_to_set_empty `{HE : EqDecision ∅} `{HF : @Finite ∅ HE} : + fin_to_set ∅ = empty :> @gset ∅ HE (@finite_countable _ HE HF). +Proof. + apply set_eq; intros []. +Qed. + +Lemma fin_to_set_inc {S : Set} `{!EqDecision S} `{!Finite S} + `{HE : EqDecision (inc S)} `{HF : @Finite (inc S) HE} + : fin_to_set (inc S) = ({[VZ]} : gset (inc S)) ∪ (set_map VS (fin_to_set S : gset S)) + :> @gset (inc S) HE (@finite_countable _ HE HF). +Proof. + apply set_eq. + intros [| x]. + - split. + + intros _; apply elem_of_union; left. + by apply elem_of_singleton. + + intros _; apply elem_of_fin_to_set. + - split. + + intros _; apply elem_of_union; right. + apply elem_of_map_2, elem_of_fin_to_set. + + intros H. + apply elem_of_fin_to_set. +Qed. + +Section Sum. + (* Kenny Loggins starts playing in a background *) + + Lemma EqDecisionLeft {S1 S2 : Set} {H : EqDecision (S1 + S2)} : EqDecision S1. + Proof. + intros x y. + destruct (decide (inl x = inl y)) as [G | G]; + [left; by inversion G | right; intros C; by subst]. + Qed. + + Lemma EqDecisionRight {S1 S2 : Set} {H : EqDecision (S1 + S2)} : EqDecision S2. + Proof. + intros x y. + destruct (decide (inr x = inr y)) as [G | G]; + [left; by inversion G | right; intros C; by subst]. + Qed. + + Lemma FiniteLeft {S1 S2 : Set} `{EqDecision S1} + `{EqDecision (S1 + S2)} `{Finite (S1 + S2)} + : Finite S1. + Proof. + unshelve econstructor. + - apply (foldr (λ x acc, match x with + | inl x => x :: acc + | inr _ => acc + end) [] (enum (S1 + S2))). + - set (l := enum (S1 + S2)). + assert (NoDup l) as K; [apply (NoDup_enum (S1 + S2)) |]. + clearbody l. + induction l as [| a l IH]; [constructor |]. + destruct a as [a | a]; simpl. + + constructor. + * intros C. + assert (inl a ∈ l) as C'. + { + clear -C. + induction l as [| b l IH]; [inversion C |]. + destruct b as [b | b]; simpl. + - rewrite foldr_cons in C. + rewrite elem_of_cons in C. + destruct C as [-> | C]. + + apply elem_of_cons. + by left. + + right. + apply IH. + apply C. + - apply elem_of_cons. + right. + rewrite foldr_cons in C. + apply IH. + apply C. + } + by inversion K. + * apply IH. + by inversion K. + + apply IH. + by inversion K. + - intros x. + set (l := enum (S1 + S2)). + assert (inl x ∈ l) as K; [apply elem_of_enum |]. + clearbody l. + induction l as [| a l IH]; [inversion K |]. + destruct a as [a | a]; simpl. + + rewrite elem_of_cons in K. + destruct K as [K | K]. + * inversion K; subst. + apply elem_of_cons; by left. + * apply elem_of_cons; right; by apply IH. + + rewrite elem_of_cons in K. + destruct K as [K | K]; [inversion K |]. + by apply IH. + Qed. + + Lemma FiniteRight {S1 S2 : Set} `{EqDecision S2} + `{EqDecision (S1 + S2)} `{H : Finite (S1 + S2)} + : Finite S2. + Proof. + unshelve econstructor. + - apply (foldr (λ x acc, match x with + | inl _ => acc + | inr x => x :: acc + end) [] (enum (S1 + S2))). + - set (l := enum (S1 + S2)). + assert (NoDup l) as K; [apply NoDup_enum |]. + clearbody l. + induction l as [| a l IH]; [constructor |]. + destruct a as [a | a]; simpl. + + apply IH. + by inversion K. + + constructor. + * intros C. + assert (inr a ∈ l) as C'. + { + clear -C. + induction l as [| b l IH]; [inversion C |]. + destruct b as [b | b]; simpl. + - apply elem_of_cons. + right. + rewrite foldr_cons in C. + apply IH, C. + - rewrite foldr_cons in C. + rewrite elem_of_cons in C. + destruct C as [-> | C]. + + apply elem_of_cons. + by left. + + right. + apply IH, C. + } + by inversion K. + * apply IH. + by inversion K. + - intros x. + set (l := enum (S1 + S2)). + assert (inr x ∈ l) as K; [apply elem_of_enum |]. + clearbody l. + induction l as [| a l IH]; [inversion K |]. + destruct a as [a | a]; simpl. + + rewrite elem_of_cons in K. + destruct K as [K | K]; [inversion K |]. + by apply IH. + + rewrite elem_of_cons in K. + destruct K as [K | K]. + * inversion K; subst. + apply elem_of_cons; by left. + * apply elem_of_cons; right; by apply IH. + Qed. + +End Sum. + +Section Inc. + Context (S : Set). + + Global Instance EqDecisionIncN {HS : EqDecision S} (n : nat) : EqDecision (Init.Nat.iter n inc S). + Proof using S. + induction n; simpl. + - apply _. + - intros [|x] [|y]. + + by left. + + by right. + + by right. + + destruct (decide (x = y)) as [-> |]; [by left |]. + right; by inversion 1. + Qed. + + Global Instance EqDecisionInc {HS : EqDecision S} : EqDecision (inc S). + Proof using S. + assert (inc S = Init.Nat.iter 1 inc S) as ->; [done |]. + by apply EqDecisionIncN. + Qed. + + Global Instance FiniteIncN {HS : EqDecision S} (HF : Finite S) (n : nat) {HS' : EqDecision (Init.Nat.iter n inc S)} : Finite (Init.Nat.iter n inc S). + Proof using S. + induction n. + - apply (@surjective_finite S HS HF _ _ id). + apply _. + - simpl. + unshelve eapply (@surjective_finite (option (Init.Nat.iter n inc S))); simpl in *. + + intros [x |]. + * apply (VS x). + * apply VZ. + + apply _. + + intros [| x]; simpl. + * exists None; reflexivity. + * exists (Some x); reflexivity. + Qed. + + Global Instance FiniteInc {HS : EqDecision S} (HF : Finite S) (HE : EqDecision (inc S)) : Finite (inc S). + Proof using S. + assert (J : @Finite (Init.Nat.iter 1 inc S) HE). + { apply FiniteIncN, HF. } + simpl in J. + apply J. + Qed. + +End Inc. From 0487b73567509d981763ffe66b23b989bcf5250d Mon Sep 17 00:00:00 2001 From: Kaptch Date: Fri, 2 Feb 2024 12:47:13 +0100 Subject: [PATCH 7/7] minor simplification --- theories/lib/factorial.v | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/theories/lib/factorial.v b/theories/lib/factorial.v index e30fc24..3d8b324 100644 --- a/theories/lib/factorial.v +++ b/theories/lib/factorial.v @@ -4,9 +4,9 @@ From gitrees.effects Require Import store. From gitrees.lib Require Import while. Section fact. - Context (n' : nat) (r : gReifiers NotCtxDep n'). - Definition rs : gReifiers NotCtxDep (S (S n')) := - (gReifiers_cons reify_io (gReifiers_cons reify_store r)). + Context (n' : nat) (rs : gReifiers NotCtxDep n'). + Context `{!subReifier reify_store rs}. + Context `{!subReifier reify_io rs}. Notation F := (gReifiers_ops rs). Context {R} `{!Cofe R}. Context `{!SubOfe natO R, !SubOfe unitO R}. @@ -103,11 +103,11 @@ Section fact. simpl. rewrite get_ret_ret. iApply (wp_alloc with "Hctx"). { solve_proper. } - fold rs. iNext. iNext. + iNext. iNext. iIntros (acc) "Hacc". simpl. iApply (wp_alloc with "Hctx"). { solve_proper. } - fold rs. iNext. iNext. + iNext. iNext. iIntros (ℓ) "Hl". simpl. iApply wp_seq. { solve_proper. }