Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Abstract contexts structures
+
+
Our whole development utilizes well-scoped and well-typed terms. In this style, the most
+common binder representation is by using lists as contexts and well-typed DeBruijn indices
+as variables. This is nice and good, but sometimes we would rather have had another,
+isomorphic, representation. Indeed in an intensional type theory like Coq, sometimes
+crossing an isomorphism is pretty painful.
+
As such, we here define what it means for an alternate definition of contexts to
+"behave like lists and DeBruijn indices".
+
This is a very natural extension to the theory of substitution by Marcelo Fiore and
+Dimitri Szamozvancev, and in fact a variant of the "Nameless, Painless" approach by
+Nicolas Pouillard.
+
+
Theory
+
Categorically, it very simple. Contexts are represented by a set C, a distinguished
+element ∅ and a binary operation - +▶ -. Additionally it has a map representing
+variable 𝐯 : C → 𝐀 where 𝐀 is a sufficiently well-behaved category, typically a
+presheaf category. We then ask that
+
+
𝐯 ∅ ≈ ⊥, where ⊥ is the initial object in 𝐀 and
+
𝐯 (Γ +▶ Δ) ≈ 𝐯 Γ + 𝐯 Δ where + is the coproduct in 𝐀.
+
+
Our category of contexts is then the full image of 𝐯, which has the structure of a
+commutative monoid. Then, given a family X : C → 𝐀, it is easy to define
+assignments as:
+
+Γ =[X]> Δ ≔ 𝐀[ 𝐯 Γ , X Δ ]
+
+
And renamings as:
+
+Γ ⊆ Δ ≔ Γ =[ 𝐯 ]> Δ
+ ≔ 𝐀[ 𝐯 Γ , 𝐯 Δ ]
+
+
Assuming 𝐀 is (co-)powered over Set, the substitution tensor product and substitution
+internal hom in C → 𝐀 are given by:
+
+( X ⊗ Y ) Γ ≔ ∫^Δ X Δ × (Δ =[ Y ]> Γ)
+⟦ X , Y ⟧ Γ ≔ ∫_Δ (Γ =[ X ]> Δ) → Y Δ
+
+
More generally, given a category 𝐁 (co-)powered over Set we can define the the
+following functors, generalizing the substitution tensor and hom to heretogeneous
+settings:
+
+( - ⊗ - ) : (C → 𝐁) → (C → 𝐀) → (C → 𝐁)
+( X ⊗ Y ) Γ ≔ ∫^Δ X Δ × (Δ =[ Y ]> Γ)
+
+⟦ - , - ⟧ : (C → 𝐀) → (C → 𝐁) → (C → 𝐁)
+⟦ X , Y ⟧ Γ ≔ ∫_Δ (Γ =[ X ]> Δ) → Y Δ
+
+
+
+
Concretely
+
We here apply the above theory to the case where 𝐀 is the family category T → Set
+for some set T. Taking T to the set of types of some object language, families
+C → 𝐀, ie C → T → Set, model nicely well-scoped and well-typed families. To
+capture non-typed syntactic categories, indexed only by a scope, we develop the special
+case where 𝐁 is just Set.
+
We then instanciate this abstract structure with the usual lists and DeBruijn indices,
+but also with two useful instances: the direct sum, where the notion of variables 𝐯
+is the pointwise coproduct and the subset, where the notion of variables is preserved.
+
In further work we wish to tacle this in more generality, notable treating the case for
+the idiomatic presentation of well-scoped untyped syntax where 𝐀 is Set, C
+is ℕ and 𝐯 is Fin. Currently, we do treat untyped syntax, but using the
+non-idiomatic "unityped" presentation.
+
The first part of the context structure: the distinguished empty context, the binary
+concatenation and the family of variables.
+
Classcontext (TC : Type) := {
+ c_emp : C ;
+ c_cat : C -> C -> C ;
+ c_var : C -> T -> Type ;
+}.
+#[global] Notation"∅" := c_emp : ctx_scope.
+#[global] Notation"Γ +▶ Δ" := (c_cat Γ Δ) : ctx_scope.
+#[global] Notation"Γ ∋ t" := (c_var Γ%ctx t).
Next we need to formalize the isomorphisms stating that variables map the empty context
+to the initial family and the concatenation to the coproduct family. To do this we will
+not write the usual forward map and backward map that compose to the identity, but we
+will rather formalize injections with inhabited fibers. The reason for this equivalent
+presentation is that it will ease dependent pattern matching by providing the isomorphisms
+as "views" (see "The view from the left" by Conor McBride).
+
First lets define the two injection maps.
+
Classcontext_cat_wkn (TC : Type) {CC : context T C} := {
+ r_cat_l {Γ Δ} [t] : Γ ∋ t -> Γ +▶ Δ ∋ t ;
+ r_cat_r {Γ Δ} [t] : Δ ∋ t -> Γ +▶ Δ ∋ t ;
+}.
Then, assuming such a structure, we define two inductive families characterizing the
+fibers.
+
Sectionwith_param.
+Context `{C : context_cat_wkn X T}.
+
+Variantc_emp_viewt : ∅ ∋ t -> Type := .
Variantc_cat_viewΓΔt : Γ +▶ Δ ∋ t -> Type :=
+ | Vcat_l (i : Γ ∋ t) : c_cat_view Γ Δ t (r_cat_l i)
+ | Vcat_r (j : Δ ∋ t) : c_cat_view Γ Δ t (r_cat_r j)
+ .
Endwith_param.
Finally we give the rest of the laws: functions witnessing the inhabitedness of the fibers
+and the injectivity of the two injection maps. This last statement is split into 3 as we
+have separated the embedding for coproducts into two separate maps.
+
Classcontext_lawsTC {CC : context T C} := {
+ c_var_cat :: context_cat_wkn T C ;
+ c_view_emp {t} i : c_emp_view t i ;
+ c_view_cat {Γ Δ t} i : c_cat_view Γ Δ t i ;
+ r_cat_l_inj {Γ Δ t} : injective (@r_cat_l _ _ _ _ Γ Δ t) ;
+ r_cat_r_inj {Γ Δ t} : injective (@r_cat_r _ _ _ _ Γ Δ t) ;
+ r_cat_disj {Γ Δ t} (i : Γ ∋ t) (j : Δ ∋ t) : ¬ (r_cat_l i = r_cat_r j) ;
+} .
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Adequacy (Prop. 7)
+
+
We prove in this module that the composition of strategy is adequate.
+The proof essentially proceeds by showing that "evaluating and observing",
+i.e., reduce, is a solution of the same equations as is the composition
+of strategies.
+
This argument assumes we can rely on the unicity of such a solution (Prop. 6):
+we prove this fact by proving that these equations are eventually guarded in
+OGS/CompGuarded.v.
+
We consider a language abstractly captured as a machine satisfying an
+appropriate axiomatization.
+
Context {TC} {CC : context T C} {CL : context_laws T C}.
+ Context {val} {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+ Context {conf} {CM : subst_module val conf} {CML : subst_module_laws val conf}.
+ Context {obs : obs_struct T C} {M : machine val conf obs} {ML : machine_laws val conf obs}.
+ Context {VV : var_assumptions val}.
We define reduce, called "zip-then-eval-then-observe" (z-e-obs) in the paper.
Equipped with eventually guarded equations, we are ready to prove the adequacy.
+
First a technical lemma, which unfolds one step of composition, then reduce to a simpler
+more explicit form.
+
Lemmacompo_reduce_simpl {Δ} (x : reduce_t Δ) :
+ (compo_body x >>= fun_r =>
+ match r with
+ | inl y => reduce' _ y
+ | inr o => Ret' (o : (fun_ => obs∙ _) _)
+ end)
+ ≊
+ (eval x.(red_act).(ms_conf) >>= fun_n =>
+ match c_view_cat (nf_var n) with
+ | Vcat_l i => Ret' (i ⋅ nf_obs n)
+ | Vcat_r j => reduce' _ (RedT (m_strat_resp x.(red_pas) (j ⋅ nf_obs n))
+ (x.(red_act).(ms_env) ;⁻ nf_args n))
+ end).
+ Proof.
+ do2 (etransitivity; [ nowapply fmap_bind_com | ]).
+ apply (subst_eq (RX := fun_ => eq)); eauto.
+ intros ? [ ? i [ ? ? ] ] ? <-.
+ unfold m_strat_wrap; cbn.
+ nowdestruct (c_view_cat i).
+ Qed.
Next we derive a variant of "evaluation respects substitution", working with
+partial assignments [ a_id , e ] instead of full assignments.
+
Definitioneval_split_sub {ΓΔ} (c : conf (Δ +▶ Γ)) (e : Γ =[val]> Δ)
+ : delay (nf obs val Δ)
+ := eval c >>= fun'T1_0n =>
+ match c_view_cat (nf_var n) with
+ | Vcat_l i => Ret' (i ⋅ nf_obs n ⦇ nf_args n ⊛ [ v_var, e ] ⦈)
+ | Vcat_r j => eval (e _ j ⊙ nf_obs n ⦗ nf_args n ⊛ [ v_var, e ] ⦘)
+ end .
+
+ Lemmaeval_split {ΓΔ} (c : conf (Δ +▶ Γ)) (e : Γ =[val]> Δ) :
+ eval (c ₜ⊛ [ a_id , e ]) ≋ eval_split_sub c e .
+ Proof.
+ unfold eval_split_sub.
+ rewrite (eval_sub c ([ v_var , e ])).
+ apply (subst_eq (RX := fun_ => eq)); eauto.
+ intros [] [ ? i [ o a ] ] ? <-; unfold emb; cbn.
+ destruct (c_view_cat i).
+ + unfold nf_args, fill_args, cut_r.
+ rewrite app_sub, v_sub_var.
+ cbn; rewrite c_view_cat_simpl_l.
+ nowapply (eval_nf_ret (i ⋅ o ⦇ a ⊛ [a_id, e] ⦈)).
+ + rewrite app_sub, v_sub_var.
+ cbn; nowrewrite c_view_cat_simpl_r.
+ Qed.
Note the use of iter_evg_uniq: the proof of adequacy is proved by unicity
+of the fixed point, which is made possible by equivalently viewing the fixpoint
+combinator used to define the composition of strategy as a fixpoint of eventually
+guarded equations. We here prove the generalization of adequacy, namely that
+reduce is a fixed point of the composition equation.
+
Lemmaadequacy_gen {Δa} (c : m_strat_act Δ a) (e : m_strat_pas Δ a) :
+ reduce (RedT c e) ≊ (c ∥g e).
+ Proof.
+ refine (iter_evg_uniq (fun'T1_0u => compo_body u) (fun'T1_0u => reduce u) _ _ T1_0 _).
+ clear a c e; intros [] [ ? [ u v ] ].
+ etransitivity; [ | symmetry; apply compo_reduce_simpl ].
+ unfold reduce at1, evalₒ at1; rewrite eval_split.
+ etransitivity; [ apply bind_fmap_com | ].
+ apply (subst_eq (RX := fun_ => eq)); eauto.
+ intros [] [ ? i [ o a ] ] ? <-; unfold emb; cbn.
+ destruct (c_view_cat i).
+ - rewrite c_view_cat_simpl_l.
+ apply it_eq_unstep; cbn; do2econstructor.
+ - rewrite c_view_cat_simpl_r.
+ unfold reduce; cbn; unfold nf_args, cut_r, fill_args; cbn.
+ assert (H : (bicollapse v red_pas cut_ty j ⊙ o ⦗ a ⊛ [a_id, bicollapse v red_pas] ⦘)
+ = ((ₐ↓ red_pas cut_ty j ᵥ⊛ r_emb r_cat3_1) ⊙ o ⦗ r_cat_rr ᵣ⊛ a_id ⦘
+ ₜ⊛ [a_id, [bicollapse v red_pas, a ⊛ [a_id, bicollapse v red_pas]]])).
+ rewrite app_sub, <- v_sub_sub.
+
+ (* AAAAA setoid rewriting!!!! *)
+ etransitivity; cycle1.
+ unshelveerewrite v_sub_proper.
+ 5: { rewrite (a_ren_r_simpl _ r_cat3_1 _).
+ nowrewrite r_cat3_1_simpl; eauto.
+ }
+ 3,4: reflexivity.
+
+ change (?r ᵣ⊛ a_id)%asgn with (r_emb r); rewrite a_ren_r_simpl.
+ unfold r_cat_rr; rewrite a_ren_l_comp.
+ rewrite2 a_cat_proj_r.
+
+ pose proof (H := collapse_fix_pas v red_pas _ j).
+ cbnin H; nowrewrite H.
+ exact _.
+
+ pose (xx := bicollapse v red_pas cut_ty j ⊙ o ⦗ a ⊛ [a_id, bicollapse v red_pas] ⦘).
+ change (bicollapse v _ _ _ ⊙ o ⦗ _ ⦘) with xx in H |- *.
+ nowrewrite H.
+ Qed.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
+
+
From OGS.Ctx Require Export Abstract Family Assignment Renaming.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Assignments
+
+
In this file we define assignments for a given abstract context structure. We also define
+several combinators based on them: the internal substitution hom, filled operations,
+copairing, etc.
+
+
Definition of assignments (Def. 5)
+
We distinguish assignments, mapping variables in a context to terms, from
+substitutions, applying an assignment to a term. Assignments are intrinsically
+typed, mapping variables of a context Γ to open terms with variables in Δ.
+
Definitionassignment (F : Fam₁ T C) (ΓΔ : C) := forallx, Γ ∋ x -> F Δ x.
+
+Notation"Γ =[ F ]> Δ" := (assignment F Γ%ctx Δ%ctx).
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Equivalence (asgn_eq Γ Δ)
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Equivalence (asgn_eq Γ Δ)
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Reflexive (asgn_eq Γ Δ)
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Symmetric (asgn_eq Γ Δ)
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Transitive (asgn_eq Γ Δ)
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Reflexive (asgn_eq Γ Δ)
nowintros u ? i.
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Symmetric (asgn_eq Γ Δ)
intros u h ? ? i; symmetry; nowapply (H _ i).
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ, Δ: C
Transitive (asgn_eq Γ Δ)
intros u v w h1 h2 ? i; transitivity (v _ i); [ nowapply h1 | nowapply h2 ].
+Qed.
Internal hom functors. The monoidal product for substitution is a bit cumbersome as
+it is expressed as a coend, that is, a quotient. Following the formalization by
+Dima Szamozvancev, we instead prefer to express everything in terms of its adjoint,
+the internal hom.
+
For example the monoidal multiplication map will not be typed X ⊗ X ⇒ X but
+X ⇒ ⟦ X , X ⟧.
+
It is an end, that is, a subset, which are far more easy to work with as they can
+simply be encoded as a record pairing an element and a proof. The proof part is not
+written here but encoded later on.
+
The real one is ⟦-,-⟧₁, but in fact we can define an analogue to this bifunctor
+for any functor category ctx → C, which we do here for Fam₀ and Fam₂ (unsorted and
+bisorted, scoped families). This will be helpful to define what it means to
+substitute other kinds of syntactic objects.
+
See Ctx/Abstract.v for more details on the theory. This is called
+the _power object_ (Def. 6) in the paper.
+
Definitioninternal_hom₀ : Fam₁ T C -> Fam₀ T C -> Fam₀ T C
+ := funFGΓ => forallΔ, Γ =[F]> Δ -> G Δ.
+Definitioninternal_hom₁ : Fam₁ T C -> Fam₁ T C -> Fam₁ T C
+ := funFGΓx => forallΔ, Γ =[F]> Δ -> G Δ x.
+Definitioninternal_hom₂ : Fam₁ T C -> Fam₂ T C -> Fam₂ T C
+ := funFGΓxy => forallΔ, Γ =[F]> Δ -> G Δ x y.
+
+Notation"⟦ F , G ⟧₀" := (internal_hom₀ F G).
+Notation"⟦ F , G ⟧₁" := (internal_hom₁ F G).
+Notation"⟦ F , G ⟧₂" := (internal_hom₂ F G).
Constructing a sorted family of operations from O with arguments taken from the family F.
+
Recordfilled_op (O : Oper T C) (F : Fam₁ T C) (Γ : C) (t : T) :=
+ OFill { fill_op : O t ; fill_args : o_dom fill_op =[F]> Γ }.
+Notation"S # F" := (filled_op S F).
We can forget the arguments and get back a bare operation.
+
Definitionforget_args {O : Oper T C} {F} : (O # F) ⇒₁ ⦉ O ⦊
+ := fun__ => fill_op.
The empty assignment.
+
Definitiona_empty {FΓ} : ∅ =[F]> Γ
+ := fun_i => match c_view_emp i withend.
The copairing of assignments.
+
Definitiona_cat {FΓ1Γ2Δ} (u : Γ1 =[F]> Δ) (v : Γ2 =[F]> Δ) : (Γ1 +▶ Γ2) =[F]> Δ
+ := funti => match c_view_cat i with
+ | Vcat_l i => u _ i
+ | Vcat_r j => v _ j
+ end.
A kind relaxed of pointwise mapping.
+
Definitiona_map {FG : Fam₁ T C} {ΓΔ1Δ2} (u : Γ =[F]> Δ1) (f : forallx, F Δ1 x -> G Δ2 x)
+ : Γ =[G]> Δ2
+ := fun_i => f _ (u _ i) .
Copairing respects pointwise equality.
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C x, y: Γ1 =[ F ]> Δ Hu: x ≡ₐ y x0, y0: Γ2 =[ F ]> Δ Hv: x0 ≡ₐ y0 a: T a0: Γ1 +▶ Γ2 ∋ a
match c_view_cat a0 with
+| Vcat_l i => x a i
+| Vcat_r j => x0 a j
+end =
+match c_view_cat a0 with
+| Vcat_l i => y a i
+| Vcat_r j => y0 a j
+end
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Eventual guardedness of the composition (Prop. 6)
+
+
We want to prove adequacy by unicity of solutions to the set of equations defining the
+composition. To do so, we rely on the notion of guarded iteration introduced in
+ITree/Guarded.v. Through this file, we prove that the composition is
+indeed eventually guarded. The proof relies crucially on the "focused redex" assumption
+(Def. 28, here denoted by the law eval_app_not_var).
+
We consider a language abstractly captured as a machine satisfying an
+appropriate axiomatization.
+
Context `{CC : context T C} {CL : context_laws T C}.
+ Context {val} {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+ Context {conf} {CM : subst_module val conf} {CML : subst_module_laws val conf}.
+ Context {obs : obs_struct T C} {M : machine val conf obs}.
+ Context {ML : machine_laws val conf obs} {VV : var_assumptions val}.
+
+ Notationogs_ctx := (ogs_ctx (C:=C)).
A central elements in the proof lies in ensuring that there exists a unique solution
+to compo_body. Since the composition of strategies is defined as such a solution,
+adequacy can be established by proving that evaluating and observing the substituted
+term is also a solution of compo_body.
+
By iter_evg_uniq, we know that eventually guarded equations admit a unique fixed point:
+we hence start by proving this eventual guardedness, captured in lemma
+compo_body_guarded.
+
This proof will among other be using an induction on the age of a variable, ie it's
+height in the OGS position Φ. We provide this definition and utilities.
+
Equationsvar_heightΦp {x} : ↓[p]Φ ∋ x -> nat :=
+ var_height ∅ₓ Pas i with c_view_emp i := { | ! } ;
+ var_height ∅ₓ Act i with c_view_emp i := { | ! } ;
+ var_height (Φ ▶ₓ Γ) Pas i := var_height Φ Act i ;
+ var_height (Φ ▶ₓ Γ) Act i with c_view_cat i := {
+ | Vcat_l j := var_height Φ Pas j ;
+ | Vcat_r j := 1 + c_length Φ } .
+ #[global] Arguments var_height {Φ p x} i.
+
+ Lemmavar_height_pos {Φpx} (i : ↓[p]Φ ∋ x) : 0 < var_height i .
+ Proof.
+ funelim (var_height i).
+ - apply H.
+ - rewrite <- Heqcall; apply H.
+ - rewrite <- Heqcall; apply Nat.lt_0_succ.
+ Qed.
+
+ Equationslt_bound : polarity -> relation nat :=
+ lt_bound Act := lt ;
+ lt_bound Pas := le .
+
+ Lemmavar_height_bound {Φpx} (i : ↓[p]Φ ∋ x)
+ : lt_bound p (var_height i) (1 + c_length Φ).
+ Proof.
+ funelim (var_height i); cbn.
+ - nowapply Nat.lt_succ_r, Nat.le_le_succ_r, Nat.le_le_succ_r.
+ - rewrite Heq; nowapply Nat.lt_succ_r.
+ - rewrite Heq; apply Nat.lt_succ_diag_r.
+ Qed.
+
+ Equationsvar_height' {ΔΦpx} : Δ +▶ ↓[p]Φ ∋ x -> nat :=
+ var_height' i with c_view_cat i := {
+ | Vcat_l j := O ;
+ | Vcat_r j := var_height j } .
This is the main theorem about height: if we lookup a variable in an OGS environment and
+obtain another variable, then this new variable is strictly lower than the first one.
We are now ready to prove eventual guardedness. In fact the main lemma will have a slightly
+different statement, considering a particular kind of pair of OGS states: the ones we
+obtain when restarting after an interaction. As the statement is quite long we factor
+it in a definition first.
+
Definitioncompo_body_guarded_aux_stmt {ΔΦ} (u : ogs_env Δ Act Φ) (v : ogs_env Δ Pas Φ)
+ {x} (o : obs x) (γ : dom o =[val]> (Δ +▶ ↓⁺ Φ)) (j : ↓⁺ Φ ∋ x) : Prop
+ := ev_guarded
+ (fun'T1_0 => compo_body)
+ (compo_body (RedT (MS ((ₐ↓v _ j ᵥ⊛ᵣ r_cat3_1) ⊙ o⦗r_cat_rr ᵣ⊛ a_id⦘)
+ (v ;⁺))
+ (u ;⁻ γ))) .
Now the main proof.
+
The reference ogs_env was not found in the current
+environment.
+Proof.
First we setup an induction on the accessibility of the current observation in the relation
+given by the "no-infinite redex" property.
+
revert Φ u v γ j.
+pose (o' := (x ,' o)); change x with (projT1 o'); change o with (projT2 o').
+generalize o'; clear x o o'; intro o.
+pose (wf := eval_app_not_var o); inductionwf.
+destruct x as [ x o ]; cbn [ projT1 projT2 ] in *.
+clear H; rename H0 into IHred; intros Φ u v γ j.
Next we setup an induction on the height of the current variable we have restarted on.
+
pose (h := lt_wf (var_height j)); remember (var_height j) as n.
+revert Φ u v γ j Heqn; induction h.
+clear H; rename x0 into n, H0 into IHvar; intros Φ u v γ j Heqn.
Then we case split on whether ₐ↓ v x j is a variable or not.
+
unfold compo_body_guarded_aux_stmt.
+pose (vv := ₐ↓ v x j); fold vv.
+ destruct (is_var_dec vv).
Next, we are evaluating a normal form, so we know by hypothesis eval_nf_ret that it
+will reduce to the same normal form. We need some trickery to rewrite by this bisimilarity.
+
pose proof (Heval :=
+ eval_nf_ret ((r_cat3_1 x (is_var_get i)) ⋅ o ⦇ r_cat_rr ᵣ⊛ a_id ⦈)).
+ unfold comp_eq in Heval; apply it_eq_step in Heval.
+ change (eval _) with
+ (eval (a_id (r_cat3_1 x (is_var_get i)) ⊙ o ⦗ r_cat_rr ᵣ⊛ a_id ⦘))
+ in Heval.
+ remember (eval (a_id (r_cat3_1 x (is_var_get i)) ⊙ o ⦗ r_cat_rr ᵣ⊛ a_id ⦘))
+ as tt; clear Heqtt.
+ remember (r_cat3_1 x (is_var_get i) ⋅ o ⦇ r_cat_rr ᵣ⊛ a_id ⦈) as ttn.
+ cbnin Heval; dependent destruction Heval.
+ rewrite Heqttn in r_rel; clear Heqttn.
+ dependent elimination r_rel.
+ unfold observe in x; rewrite <- x; clear x; cbn.
+ unfold m_strat_wrap, r_cat3_1; cbn.
+ remember (ₐ↓v _ j) as vv.
Now, we case split to see whether this variable was a final variable or one given by
+the opponent.
+
destruct i; cbn; destruct (c_view_cat i).
In case of a final variable the composition is ended, hence guarded, hence eventually
+guarded.
+
+rewrite c_view_cat_simpl_l; nowdo2econstructor.
In the other case, there is an interaction. Since we have looked-up a variable and
+obtained another variable, we know it is strictly older and use our induction hypothesis
+on the height of the current variable.
Case 2: the looked-up value is not a variable. In this case we look at one step of the
+evaluation. If there is a redex we are happy. Else, our resumed configuration was
+still a normal form and we can exhibit a proof of the bad instanciation relation, which
+enables us to call our other induction hypothesis.
Now the actual proof of eventual guardedness is just about unfolding a bit of the beginning
+of the interactions until we attain a resume and can apply our main lemma.
+
Lemmacompo_body_guarded {Δ} : eqn_ev_guarded (fun'T1_0 => compo_body (Δ := Δ)).
+ Proof.
+ intros [] [ Γ [ c u ] v ]; unfold m_strat_pas in v.
+ unfold ev_guarded; cbn.
+ pose (ot := _observe (eval c)); change (_observe (eval c)) with ot.
+ destruct ot; trynowdo2econstructor.
+ destruct r as [ x i [ o a ] ]; unfold m_strat_wrap; cbn.
+ destruct (c_view_cat i); trynowdo2econstructor.
+ refine (GNext _); apply compo_body_guarded_aux.
+ Qed.
From the previous proof we can directly apply the strong fixed point construction on
+eventually guarded equations and obtain a composition operator without any Tau node
+at interaction points (Prop. 6).
+
Definitioncompo_ev_guarded {Δa} (u : m_strat_act Δ a) (v : m_strat_pas Δ a)
+ : delay (obs∙ Δ)
+ := iter_ev_guarded _ compo_body_guarded T1_0 (RedT u v).
+
+Endwith_param.
+#[global] Notation"u ∥g v" := (compo_ev_guarded u v) (at level40).
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Congruence (Prop. 4)
+
+
We prove in this module that weak bisimilarity is a congruence for composition. The proof
+makes a slight technical side step: we prove the composition to be equivalent to an
+alternate definition, dully named compo_alt.
+
Indeed, congruence is a very easy result, demanding basically no assumption. What is
+actually hard, is to manage weak bisimilarity proofs, which in contrast to strong
+bisimilarity can be hard to tame: instead of synchronizing, at every step they can eat
+any number of Tau node on either side, forcing us to do complex inductions in the
+middle of our bisimilarity proofs.
+
Because of this, since our main definition compo is the specific case of composition
+of two instances of the machine strategy in the OGS game, we know much more than what we
+actually care about. Hence we define the much more general compo_alt composing
+any two abstract strategy for the OGS game and prove that one congruent w.r.t. weak
+bisimilarity. We then connect these two alternative definitions with a strong bisimilarity,
+much more structured, hence easy to prove.
+
We consider a language abstractly captured as a machine satisfying an
+appropriate axiomatization.
+
Context {TC} {CC : context T C} {CL : context_laws T C}.
+ Context {val} {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+ Context {conf} {CM : subst_module val conf} {CML : subst_module_laws val conf}.
+ Context {obs : obs_struct T C} {M : machine val conf obs}.
+ Context {ML : machine_laws val conf obs} {VV : var_assumptions val}.
We start off by defining this new, opaque composition.
+
The reference C was not found in the current
+environment.
+#[global] Arguments AltT {Δ Φ} u v : rename.
+#[global] Arguments alt_ctx {Δ}.
+#[global] Arguments alt_act {Δ}.
+#[global] Arguments alt_pas {Δ}.
+
+ Definitioncompo_alt_body {Δ}
+ : compo_alt_t Δ -> delay (compo_alt_t Δ + obs∙ Δ)
+ := cofix _compo_body x :=
+ go match x.(alt_act).(_observe) with
+ | RetF r => RetD (inr r)
+ | TauF t => TauD (_compo_body (AltT t x.(alt_pas)))
+ | VisF e k => RetD (inl (AltT (x.(alt_pas) e) k))
+ end .
+
+ Definitioncompo_alt {Δa} (u : ogs_act Δ a) (v : ogs_pas Δ a) : delay (obs∙ Δ)
+ := iter_delay compo_alt_body (AltT u v).
Now we define our bisimulation candidates, for weak and strong bisimilarity.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Covering for free contexts
+
+
In Ctx/Ctx.v we have instanciated the relevant part of the abstract interface
+for concrete contexts. Here we tackle the rest of the structure. Basically, we need to
+provide case-splitting for deciding if a variable i : (Γ +▶ Δ) ∋ x is in Γ or in
+Δ. We could do this directly but we prove a slight generalization, going through the
+definition of coverings, a ternary relationship on concrete contexts witnessing that a
+context is covered by the two others (respecting order).
+
+
A covering
+
This inductive family provides proofs that some context zs is obtained by "zipping"
+together two other contexts xs and ys.
A covering gives us two injective renamings, left embedding and right embedding.
+
Equationsr_cover_l {xsyszs} : xs ⊎ ys ≡ zs -> xs ⊆ zs :=
+ r_cover_l (CLeft c) _ top := top ;
+ r_cover_l (CLeft c) _ (pop i) := pop (r_cover_l c _ i) ;
+ r_cover_l (CRight c) _ i := pop (r_cover_l c _ i) .
+
+Equationsr_cover_r {xsyszs} : xs ⊎ ys ≡ zs -> ys ⊆ zs :=
+ r_cover_r (CLeft c) _ i := pop (r_cover_r c _ i) ;
+ r_cover_r (CRight c) _ top := top ;
+ r_cover_r (CRight c) _ (pop i) := pop (r_cover_r c _ i) .
Injectivity.
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ∋ x H: r_cover_l p x i = r_cover_l p x j
i = j
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ∋ x H: r_cover_l p x i = r_cover_l p x j
i = j
+
X: Type F: Fam₁ X (ctx X) x: X i, j: ∅ₓ ∋ x H: r_cover_l CNil x i = r_cover_l CNil x j
i = j
X: Type F: Fam₁ X (ctx X) x0: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ▶ₓ x0 ∋ x H: r_cover_l (CLeft p) x i = r_cover_l (CLeft p) x j IHp: forallij : xs ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
i = j
X: Type F: Fam₁ X (ctx X) x0: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ∋ x H: r_cover_l (CRight p) x i =
+r_cover_l (CRight p) x j IHp: forallij : xs ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
i = j
+
X: Type F: Fam₁ X (ctx X) x: X i, j: ∅ₓ ∋ x H: r_cover_l CNil x i = r_cover_l CNil x j
i = j
dependent elimination i.
+
X: Type F: Fam₁ X (ctx X) x0: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ▶ₓ x0 ∋ x H: r_cover_l (CLeft p) x i = r_cover_l (CLeft p) x j IHp: forallij : xs ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
i = j
X: Type F: Fam₁ X (ctx X) y0: X Γ1, ys, zs: ctx X p: Γ1 ⊎ ys ≡ zs x: X v, v0: var x Γ1 H: r_cover_l (CLeft p) x (pop v) =
+r_cover_l (CLeft p) x (pop v0) IHp: forallij : Γ1 ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
pop v = pop v0
+
X: Type F: Fam₁ X (ctx X) y0: X Γ1, ys, zs: ctx X p: Γ1 ⊎ ys ≡ zs x: X v, v0: var x Γ1 H: NoConfusion (r_cover_l (CLeft p) x (pop v))
+ (r_cover_l (CLeft p) x (pop v0)) IHp: forallij : Γ1 ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
pop v = pop v0
+f_equal; nowapply IHp.
+
X: Type F: Fam₁ X (ctx X) x0: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ∋ x H: r_cover_l (CRight p) x i =
+r_cover_l (CRight p) x j IHp: forallij : xs ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
i = j
X: Type F: Fam₁ X (ctx X) x0: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs x: X i, j: xs ∋ x H: NoConfusion (r_cover_l (CRight p) x i)
+ (r_cover_l (CRight p) x j) IHp: forallij : xs ∋ x, r_cover_l p x i = r_cover_l p x j -> i = j
i = j
+nowapply IHp.
+Qed.
+
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs y: X i, j: ys ∋ y H: r_cover_r p y i = r_cover_r p y j
i = j
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs y: X i, j: ys ∋ y H: r_cover_r p y i = r_cover_r p y j
i = j
+
X: Type F: Fam₁ X (ctx X) y: X i, j: ∅ₓ ∋ y H: r_cover_r CNil y i = r_cover_r CNil y j
i = j
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs y: X i, j: ys ∋ y H: r_cover_r (CLeft p) y i = r_cover_r (CLeft p) y j IHp: forallij : ys ∋ y, r_cover_r p y i = r_cover_r p y j -> i = j
i = j
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs y: X i, j: ys ▶ₓ x ∋ y H: r_cover_r (CRight p) y i =
+r_cover_r (CRight p) y j IHp: forallij : ys ∋ y, r_cover_r p y i = r_cover_r p y j -> i = j
i = j
+
X: Type F: Fam₁ X (ctx X) y: X i, j: ∅ₓ ∋ y H: r_cover_r CNil y i = r_cover_r CNil y j
i = j
dependent elimination i.
+
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs y: X i, j: ys ∋ y H: r_cover_r (CLeft p) y i = r_cover_r (CLeft p) y j IHp: forallij : ys ∋ y, r_cover_r p y i = r_cover_r p y j -> i = j
i = j
apply noConfusion_inv in H; nowapply IHp.
+
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs y: X i, j: ys ▶ₓ x ∋ y H: r_cover_r (CRight p) y i =
+r_cover_r (CRight p) y j IHp: forallij : ys ∋ y, r_cover_r p y i = r_cover_r p y j -> i = j
i = j
X: Type F: Fam₁ X (ctx X) y1: X xs, Γ1, zs: ctx X p: xs ⊎ Γ1 ≡ zs y: X v, v0: var y Γ1 H: r_cover_r (CRight p) y (pop v) =
+r_cover_r (CRight p) y (pop v0) IHp: forallij : Γ1 ∋ y, r_cover_r p y i = r_cover_r p y j -> i = j
pop v = pop v0
+
X: Type F: Fam₁ X (ctx X) y1: X xs, Γ1, zs: ctx X p: xs ⊎ Γ1 ≡ zs y: X v, v0: var y Γ1 H: NoConfusion (r_cover_r (CRight p) y (pop v))
+ (r_cover_r (CRight p) y (pop v0)) IHp: forallij : Γ1 ∋ y, r_cover_r p y i = r_cover_r p y j -> i = j
pop v = pop v0
+f_equal; nowapply IHp.
+Qed.
The two embeddings have disjoint images.
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs a: X i: xs ∋ a j: ys ∋ a H: r_cover_l p a i = r_cover_r p a j
T0
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs a: X i: xs ∋ a j: ys ∋ a H: r_cover_l p a i = r_cover_r p a j
T0
+
X: Type F: Fam₁ X (ctx X) a: X i, j: ∅ₓ ∋ a H: r_cover_l CNil a i = r_cover_r CNil a j
T0
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs a: X i: xs ▶ₓ x ∋ a j: ys ∋ a H: r_cover_l (CLeft p) a i = r_cover_r (CLeft p) a j IHp: forall (i : xs ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs a: X i: xs ∋ a j: ys ▶ₓ x ∋ a H: r_cover_l (CRight p) a i =
+r_cover_r (CRight p) a j IHp: forall (i : xs ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
+
X: Type F: Fam₁ X (ctx X) a: X i, j: ∅ₓ ∋ a H: r_cover_l CNil a i = r_cover_r CNil a j
T0
inversion i.
+
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs a: X i: xs ▶ₓ x ∋ a j: ys ∋ a H: r_cover_l (CLeft p) a i = r_cover_r (CLeft p) a j IHp: forall (i : xs ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) a: X Γ, ys, zs: ctx X p: Γ ⊎ ys ≡ zs j: ys ∋ a H: r_cover_l (CLeft p) a top =
+r_cover_r (CLeft p) a j IHp: forall (i : Γ ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) y: X Γ0, ys, zs: ctx X p: Γ0 ⊎ ys ≡ zs a: X v: var a Γ0 j: ys ∋ a H: r_cover_l (CLeft p) a (pop v) =
+r_cover_r (CLeft p) a j IHp: forall (i : Γ0 ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
+
X: Type F: Fam₁ X (ctx X) a: X Γ, ys, zs: ctx X p: Γ ⊎ ys ≡ zs j: ys ∋ a H: r_cover_l (CLeft p) a top =
+r_cover_r (CLeft p) a j IHp: forall (i : Γ ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
inversion H.
+
X: Type F: Fam₁ X (ctx X) y: X Γ0, ys, zs: ctx X p: Γ0 ⊎ ys ≡ zs a: X v: var a Γ0 j: ys ∋ a H: r_cover_l (CLeft p) a (pop v) =
+r_cover_r (CLeft p) a j IHp: forall (i : Γ0 ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) y: X Γ0, ys, zs: ctx X p: Γ0 ⊎ ys ≡ zs a: X v: var a Γ0 j: ys ∋ a H: r_cover_l p a v = r_cover_r p a j IHp: forall (i : Γ0 ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
+exact (IHp v j H).
+
X: Type F: Fam₁ X (ctx X) x: X xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs a: X i: xs ∋ a j: ys ▶ₓ x ∋ a H: r_cover_l (CRight p) a i =
+r_cover_r (CRight p) a j IHp: forall (i : xs ∋ a) (j : ys ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) a: X xs, Γ, zs: ctx X p: xs ⊎ Γ ≡ zs i: xs ∋ a H: r_cover_l (CRight p) a i =
+r_cover_r (CRight p) a top IHp: forall (i : xs ∋ a) (j : Γ ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) y: X xs, Γ0, zs: ctx X p: xs ⊎ Γ0 ≡ zs a: X i: xs ∋ a v: var a Γ0 H: r_cover_l (CRight p) a i =
+r_cover_r (CRight p) a (pop v) IHp: forall (i : xs ∋ a) (j : Γ0 ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
+
X: Type F: Fam₁ X (ctx X) a: X xs, Γ, zs: ctx X p: xs ⊎ Γ ≡ zs i: xs ∋ a H: r_cover_l (CRight p) a i =
+r_cover_r (CRight p) a top IHp: forall (i : xs ∋ a) (j : Γ ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
inversion H.
+
X: Type F: Fam₁ X (ctx X) y: X xs, Γ0, zs: ctx X p: xs ⊎ Γ0 ≡ zs a: X i: xs ∋ a v: var a Γ0 H: r_cover_l (CRight p) a i =
+r_cover_r (CRight p) a (pop v) IHp: forall (i : xs ∋ a) (j : Γ0 ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
X: Type F: Fam₁ X (ctx X) y: X xs, Γ0, zs: ctx X p: xs ⊎ Γ0 ≡ zs a: X i: xs ∋ a v: var a Γ0 H: r_cover_l p a i = r_cover_r p a v IHp: forall (i : xs ∋ a) (j : Γ0 ∋ a), ¬ (r_cover_l p a i = r_cover_r p a j)
T0
+exact (IHp i v H).
+Qed.
Now we can start proving that the copairing of the two embeddings has non-empty fibers,
+ie, we can case split.
+
Variantcover_view {xsyszs} (p : xs ⊎ ys ≡ zs) [z] : zs ∋ z -> Type :=
+ | Vcov_l (i : xs ∋ z) : cover_view p (r_cover_l p _ i)
+ | Vcov_r (j : ys ∋ z) : cover_view p (r_cover_r p _ j)
+ .
+#[global] Arguments Vcov_l {xs ys zs p z}.
+#[global] Arguments Vcov_r {xs ys zs p z}.
Indeed!
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs z: X i: zs ∋ z
cover_view p i
+
X: Type F: Fam₁ X (ctx X) xs, ys, zs: ctx X p: xs ⊎ ys ≡ zs z: X i: zs ∋ z
cover_view p i
+
X: Type F: Fam₁ X (ctx X) Γ: ctx X z: X IHzs: forall (i : Γ ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ Γ), cover_view p i xs, ys: ctx X p: xs ⊎ ys ≡ (Γ ▶ₓ z)
cover_view p top
X: Type F: Fam₁ X (ctx X) Γ0: ctx X y, z: X v: var z Γ0 IHzs: forall (i : Γ0 ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ Γ0), cover_view p i xs, ys: ctx X p: xs ⊎ ys ≡ (Γ0 ▶ₓ y)
cover_view p (pop v)
+
X: Type F: Fam₁ X (ctx X) Γ: ctx X z: X IHzs: forall (i : Γ ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ Γ), cover_view p i xs, ys: ctx X p: xs ⊎ ys ≡ (Γ ▶ₓ z)
X: Type F: Fam₁ X (ctx X) Γ0: ctx X y, z: X v: var z Γ0 IHzs: forall (i : Γ0 ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ Γ0), cover_view p i xs, ys: ctx X p: xs ⊎ ys ≡ (Γ0 ▶ₓ y)
cover_view p (pop v)
X: Type F: Fam₁ X (ctx X) zs: ctx X x, z: X v: var z zs IHzs: forall (i : zs ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ zs), cover_view p i xs0, ys0: ctx X c: xs0 ⊎ ys0 ≡ zs
cover_view (CLeft c) (pop v)
X: Type F: Fam₁ X (ctx X) zs0: ctx X x0, z: X v: var z zs0 IHzs: forall (i : zs0 ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ zs0), cover_view p i xs1, ys1: ctx X c0: xs1 ⊎ ys1 ≡ zs0
cover_view (CRight c0) (pop v)
+
X: Type F: Fam₁ X (ctx X) zs: ctx X x, z: X v: var z zs IHzs: forall (i : zs ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ zs), cover_view p i xs0, ys0: ctx X c: xs0 ⊎ ys0 ≡ zs
X: Type F: Fam₁ X (ctx X) zs0: ctx X x0, z: X v: var z zs0 IHzs: forall (i : zs0 ∋ z) (xsys : ctx X)
+ (p : xs ⊎ ys ≡ zs0), cover_view p i xs1, ys1: ctx X c0: xs1 ⊎ ys1 ≡ zs0
X: Type F: Fam₁ X (ctx X) Γ1, Γ2, Γ3, Δ: ctx X H: Γ1 ⊎ Γ2 ≡ Γ3 x, y: Γ1 =[ F ]> Δ H1: x ≡ₐ y x0, y0: Γ2 =[ F ]> Δ H2: x0 ≡ₐ y0 a: X i: Γ3 ∋ a
match view_cover H i with
+| Vcov_l i => x a i
+| Vcov_r j => x0 a j
+end =
+match view_cover H i with
+| Vcov_l i => y a i
+| Vcov_r j => y0 a j
+end
+destruct (view_cover H i); [ nowapply H1 | nowapply H2 ].
+Qed.
+Endwith_param.
+
+#[global] Notation"a ⊎ b ≡ c" := (cover a b c).
+#[global] Notation"a ▶ₐ v" := (a_append a v) : asgn_scope.
+#[global] Notation"[ u , H , v ]" := (a_cover H u v) : asgn_scope.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Free context structure
+
+
Here we instanciate our abstract notion of context structure with the most common example:
+lists and DeBruijn indices. More pedantically, this is the free context structure over
+a given set X.
+
+
Contexts
+
Contexts are simply lists, with the purely aesthetic choice of representing cons as
+coming from the right. Note on paper: we write here "Γ ▶ x" instead of "Γ , x".
+
Inductivectx (X : Type) : Type :=
+| cnil : ctx X
+| ccon : ctx X -> X -> ctx X.
We wish to manipulate intrinsically typed terms. We hence need a tightly typed notion of
+position in the context: rather than a loose index, var x Γ is a proof of membership
+of x to Γ: a well-scoped and well-typed DeBruijn index. See Ex. 5.
+
Inductivevar {X} (x : X) : ctx X -> Type :=
+| top {Γ} : var x (Γ ▶ₓ x)
+| pop {Γ y} : var x Γ -> var x (Γ ▶ₓ y).
+Definitioncvar {X} Γx := @var X x Γ.
A couple basic functions: length, concatenation and pointwise function application.
+
Equationsc_length {X} (Γ : ctx X) : nat :=
+ c_length ∅ₓ := Datatypes.O ;
+ c_length (Γ ▶ₓ _) := Datatypes.S (c_length Γ) .
+
+Equationsccat {X} : ctx X -> ctx X -> ctx X :=
+ ccat Γ ∅ₓ := Γ ;
+ ccat Γ (Δ ▶ₓ x) := ccat Γ Δ ▶ₓ x .
+
+Equationsc_map {XY} : (X -> Y) -> ctx X -> ctx Y :=
+ c_map f ∅ₓ := cnil ;
+ c_map f (Γ ▶ₓ x) := c_map f Γ ▶ₓ f x .
Implementation of the revelant part of the abstract context interface.
This is the end for now, but we have not yet instanciated the rest of the abstract
+context structure. This is a bit more work, it takes place in another file:
+Ctx/Covering.v
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Delay Monad
+
+
Instead of defining the delay monad from scratch, we construct it as a special case of
+itree, namely ones with an empty signature ∅ₑ (disallowing Vis nodes) and
+indexed over the singleton type T1.
+
The delay monad (Def. 9) is formalized as an itree over an empty signature: in the
+absence of events, Tau and Ret are the only possible transitions.
+
Relevant definitions can be found:
+
+
for the underlying itree datatype, in ITree/ITree.v,
Embedding delay into itrees over arbitrary signatures.
+
Definitionemb_delay {I} {E : event I I} {Xi} : delay X -> itree E (X @ i) i :=
+ cofix _emb_delay x :=
+ go (match x.(_observe) with
+ | RetF r => RetF (Fib r)
+ | TauF t => TauF (_emb_delay t)
+ | VisF e k => match e withend
+ end).
+
+#[global] Notation"'RetD' x" := (RetF (x : (fun_ : T1 => _) T1_0)) (at level40).
+#[global] Notation"'TauD' t" := (TauF (t : itree ∅ₑ (fun_ : T1 => _) T1_0)) (at level40).
Specialization of the operations to the delay monad. Most of them are a bit cumbersome
+since our definition of itree is indexed, but delay should not. As such there
+is a bit of a dance around the singleton type T1 which is used as index. Too bad,
+since Coq does not have typed conversion we don't get definitional eta-law for the unit
+type...
+
Definitionret_delay {X} : X -> delay X := funx => Ret' x .
+
+Definitiontau_delay {X} : delay X -> delay X :=
+ funt => go (TauF (t : itree ∅ₑ (fun_ : T1 => X) T1_0)) .
Binding a delay computation in the context of an itree.
+
Definitionbind_delay {I} {E : event I I} {XYi}
+ : delay X -> (X -> itree E Y i) -> itree E Y i
+ := funxf => bind (emb_delay x) (fun_ '(Fib x) => f x) .
Simpler definition of bind when the conclusion is again in delay.
+See Not. 4.
+
Definitionbind_delay' {XY}
+ : delay X -> (X -> delay Y) -> delay Y
+ := funxf => bind x (fun'T1_0y => f y).
Functorial action on maps.
+
Definitionfmap_delay {XY} (f : X -> Y) : delay X -> delay Y :=
+ fmap (fun_ => f) T1_0 .
Iteration operator.
+
Definitioniter_delay {XY} : (X -> delay (X + Y)) -> X -> delay Y :=
+ funf => iter (fun'T1_0 => f) T1_0 .
Alternative to bind_delay, written from scratch.
+
Definitionsubst_delay {I} {E : event I I} {XYi} (f : X -> itree E Y i)
+ : delay X -> itree E Y i
+ := cofix _subst_delay x := go match x.(_observe) with
+ | RetF r => (f r).(_observe)
+ | TauF t => TauF (_subst_delay t)
+ | VisF e k => match e withend
+ end .
+
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0) (subst_delay y y0) t1, t2: itree ∅ₑ (fun_ : T1 => X) T1_0 H: it_eq (fun_ : T1 => RX) t1 t2
it_eq_bt E RY R i (subst_delay x t1)
+ (subst_delay y t2)
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0) (subst_delay y y0) t1, t2: itree ∅ₑ (fun_ : T1 => X) T1_0 H: it_eqF ∅ₑ (fun_ : T1 => RX)
+ (it_eq (fun_ : T1 => RX)) T1_0 (_observe t1)
+ (_observe t2)
it_eqF E RY (it_eq_t E RY R) i
+ match _observe t1 with
+ | RetF r => _observe (x r)
+ | TauF t => TauF (subst_delay x t)
+ | VisF e _ => match e return (itree' E Y i) with
+ end
+ end
+ match _observe t2 with
+ | RetF r => _observe (y r)
+ | TauF t => TauF (subst_delay y t)
+ | VisF e _ => match e return (itree' E Y i) with
+ end
+ end
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0) (subst_delay y y0) t2: itree ∅ₑ (fun_ : T1 => X) T1_0 ot1: itree' ∅ₑ (fun_ : T1 => X) T1_0 H: it_eqF ∅ₑ (fun_ : T1 => RX)
+ (it_eq (fun_ : T1 => RX)) T1_0 ot1
+ (_observe t2)
it_eqF E RY (it_eq_t E RY R) i
+ match ot1 with
+ | RetF r => _observe (x r)
+ | TauF t => TauF (subst_delay x t)
+ | VisF e _ => match e return (itree' E Y i) with
+ end
+ end
+ match _observe t2 with
+ | RetF r => _observe (y r)
+ | TauF t => TauF (subst_delay y t)
+ | VisF e _ => match e return (itree' E Y i) with
+ end
+ end
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0) (subst_delay y y0) ot1, ot2: itree' ∅ₑ (fun_ : T1 => X) T1_0 H: it_eqF ∅ₑ (fun_ : T1 => RX)
+ (it_eq (fun_ : T1 => RX)) T1_0 ot1 ot2
it_eqF E RY (it_eq_t E RY R) i
+ match ot1 with
+ | RetF r => _observe (x r)
+ | TauF t => TauF (subst_delay x t)
+ | VisF e _ => match e return (itree' E Y i) with
+ end
+ end
+ match ot2 with
+ | RetF r => _observe (y r)
+ | TauF t => TauF (subst_delay y t)
+ | VisF e _ => match e return (itree' E Y i) with
+ end
+ end
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0) (subst_delay y y0) r1, r2: (fun_ : T1 => X) T1_0 r_rel: (fun_ : T1 => RX) T1_0 r1 r2
it_eqF E RY (it_eq_t E RY R) i (_observe (x r1))
+ (_observe (y r2))
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0)
+ (subst_delay y y0) t1, t2: itree ∅ₑ (fun_ : T1 => X) T1_0 t_rel: it_eq (fun_ : T1 => RX) t1 t2
it_eqF E RY (it_eq_t E RY R) i
+ (TauF (subst_delay x t1))
+ (TauF (subst_delay y t2))
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0)
+ (subst_delay y y0) q: e_qry T1_0 k1, k2: forallx : e_rsp q,
+itree ∅ₑ (fun_ : T1 => X) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (fun_ : T1 => RX) (k1 r) (k2 r)
it_eqF E RY (it_eq_t E RY R) i
+ match q return (itree' E Y i) with
+ endmatch q return (itree' E Y i) with
+ end
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0)
+ (subst_delay y y0) r1, r2: (fun_ : T1 => X) T1_0 r_rel: (fun_ : T1 => RX) T1_0 r1 r2
it_eqF E RY (it_eq_t E RY R) i
+ (_observe (x r1)) (_observe (y r2))
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0)
+ (subst_delay y y0) r1, r2: (fun_ : T1 => X) T1_0 r_rel: (fun_ : T1 => RX) T1_0 r1 r2
it_eqF E RY (gfp (it_eq_map E RY)) i
+ (_observe (x r1)) (_observe (y r2))
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0)
+ (subst_delay y y0) t1, t2: itree ∅ₑ (fun_ : T1 => X) T1_0 t_rel: it_eq (fun_ : T1 => RX) t1 t2
it_eqF E RY (it_eq_t E RY R) i
+ (TauF (subst_delay x t1))
+ (TauF (subst_delay y t2))
econstructor; nowapply CIH.
+
I: Type E: event I I X: Type RX: relation X Y: psh I RY: relᵢ Y Y i: I x, y: X -> itree E Y i Hf: (RX ==> it_eq RY (i:=i))%signature x y R: relᵢ (itree E Y) (itree E Y) CIH: forallx0y0 : itree ∅ₑ (fun_ : T1 => X) T1_0,
+it_eq (fun_ : T1 => RX) x0 y0 ->
+it_eq_t E RY R i (subst_delay x x0)
+ (subst_delay y y0) q: e_qry T1_0 k1, k2: forallx : e_rsp q,
+itree ∅ₑ (fun_ : T1 => X) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (fun_ : T1 => RX) (k1 r) (k2 r)
it_eqF E RY (it_eq_t E RY R) i
+ match q return (itree' E Y i) with
+ endmatch q return (itree' E Y i) with
+ end
destruct q.
+Qed.
+
+
+
diff --git a/docs/DirectSum.html b/docs/DirectSum.html
new file mode 100644
index 0000000..ff858d1
--- /dev/null
+++ b/docs/DirectSum.html
@@ -0,0 +1,136 @@
+
+
+
+
+
+
+Direct sum of context structure
+
+
+
+
+
+
+
+
+
+
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Direct sum of context structure
+
+
A random calculus you may find in the wild will most likely not fit into the rigid
+well-scoped and well-typed format with concrete contexts and DeBruijn indices. Sometimes
+this is because the designers have found useful to formalize this calculus with two
+contexts, storing two kinds of variables for different use. The usual trick is to change
+the set of types to a coproduct and store everything in one big context, eg using
+ctx (S + T) instead of ctx S × ctx T. This does work, but it might make some things
+more cumbersome.
+
Using our abstraction over context structures, we can readily express the case of two
+separate contexts: the direct sum.
Now the interesting part: variables! The new set of types is either a type from the
+first context or from the second: T₁ + T₂. We then compute the set of variables
+by case splitting on this coproduct: either a variable from the left or from the right.
forall (ΓΔ : dsum) (t : T1 + T2),
+injective
+ match
+ t as s
+ return (c_var2 Γ s -> c_var2 (c_cat2 Γ Δ) s)
+ with
+ | inl a => funi : fst Γ ∋ a => r_cat_l i
+ | inr b => funi : snd Γ ∋ b => r_cat_l i
+ end
forall (ΓΔ : dsum) (t : T1 + T2),
+injective
+ match
+ t as s
+ return (c_var2 Δ s -> c_var2 (c_cat2 Γ Δ) s)
+ with
+ | inl a => funi : fst Δ ∋ a => r_cat_r i
+ | inr b => funi : snd Δ ∋ b => r_cat_r i
+ end
forall (ΓΔ : dsum) (t : T1 + T2),
+injective
+ match
+ t as s
+ return (c_var2 Γ s -> c_var2 (c_cat2 Γ Δ) s)
+ with
+ | inl a => funi : fst Γ ∋ a => r_cat_l i
+ | inr b => funi : snd Γ ∋ b => r_cat_l i
+ end
forall (ΓΔ : dsum) (t : T1 + T2),
+injective
+ match
+ t as s
+ return (c_var2 Δ s -> c_var2 (c_cat2 Γ Δ) s)
+ with
+ | inl a => funi : fst Δ ∋ a => r_cat_r i
+ | inr b => funi : snd Δ ∋ b => r_cat_r i
+ end
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Interaction Trees: (weak) bisimilarity
+
+
As is usual with coinductive types in Coq, eq is not the right notion of equivalence.
+We define through this file strong and weak bisimilarity of itrees, and their
+generalization lifting value relations.
+
These coinductive relations are implemented using Damien Pous's coinduction library.
+Indeed, our previous coinductive definitions like itree where implemented using Coq's
+native coinductive types, but their manipulation is a bit brittle due to the syntactic
+guardedness criterion. For the case of bisimilarity---which is also a coinductive types---
+we have better tools. Indeed bisimilarity is a Prop-valued relation, and since Coq's
+Prop universe feature impredicativity, the set of Prop-valued relation enjoy the
+structure of a complete lattice. This enables us to derive greatest fixpoints ourselves,
+using an off-the-shelf fixpoint construction on complete lattices.
+
The version of coinduction we use constructs greatest fixpoints using the "companion"
+construction, enjoying good properties w.r.t. up-to techniques: we will be able to
+discharge bisimilarity proof by providing less than a full-blown bisimulation relation.
+Since the time of writing, the library has been upgraded to an even more practical
+construction based on tower-induction, but we have not yet ported our code to this
+upgraded API.
+
+
Strong bisimilarity
+
Strong bisimilarity is very useful as it is the natural notion of extensional
+equality for coinductive types. Here we introduce it_eq RR, a slight generalization
+where the relation on the leaves of the tree does not need to be equality on the type
+of the leaves, but only a proof of the relation RR, which might be heterogeneous. This
+might be described as the relational lifting of itree arising from strong bisimilarity.
+
We will write ≅ for strong bisimilarity, aka it_eq (eqᵢ _).
+
First, we define a monotone endofunction it_eq_map over indexed relations between
+trees. Strong bisimilarity is then defined the greatest fixpoint of it_eq_map RR,
+for a fixed value relation RR.
Equationsit_eqF_mon {IEX1X2RXY1Y2}
+ : Proper (leq ==> leq) (@it_eqF I E X1 X2 RX Y1 Y2) :=
+ it_eqF_mon _ _ H1 _ _ _ (EqRet r_rel) := EqRet r_rel ;
+ it_eqF_mon _ _ H1 _ _ _ (EqTau t_rel) := EqTau (H1 _ _ _ t_rel) ;
+ it_eqF_mon _ _ H1 _ _ _ (EqVis k_rel) := EqVis (funr => H1 _ _ _ (k_rel r)) .
+#[global] Existing Instanceit_eqF_mon.
+
+Definitionit_eq_map {I} E {X1X2} RX : mon (relᵢ (@itree I E X1) (@itree I E X2)) := {|
+ body RR i x y := it_eqF E RX RR i (observe x) (observe y) ;
+ Hbody _ _ H _ _ _ r := it_eqF_mon _ _ H _ _ _ r ;
+|}.
Now the definition of the bisimilarity itself as greatest fixed point. See Def. 10.
+
Definitionit_eq {IEX1X2} RX [i] := gfp (@it_eq_map I E X1 X2 RX) i.
+#[global] Notationit_eq_t E RX := (t (it_eq_map E RX)).
+#[global] Notationit_eq_bt E RX := (bt (it_eq_map E RX)).
+#[global] Notationit_eq_T E RX := (T (it_eq_map E RX)).
+#[global] Notation"a ≊ b" := (it_eq (eqᵢ _) a b) (at level20).
+
Basic properties
+
Definitionit_eq' {IEX1X2} RX [i]
+ := @it_eqF I E X1 X2 RX (itree E X1) (itree E X2) (it_eq RX) i.
+
+Definitionit_eq_step {IEX1X2RX} : it_eq RX <= @it_eq_map I E X1 X2 RX (it_eq RX)
+ := funixy => proj1 (gfp_fp (it_eq_map E RX) i x y) .
+
+Definitionit_eq_unstep {IEX1X2RX} : @it_eq_map I E X1 X2 RX (it_eq RX) <= it_eq RX
+ := funixy => proj2 (gfp_fp (it_eq_map E RX) i x y) .
+
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
Proper (leq ==> leq) (it_eq_bt E RX)
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
Proper (leq ==> leq) (it_eq_bt E RX)
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R1, R2: relᵢ (itree E X1) (itree E X2) H: R1 <= R2 i: I x: itree E X1 i y: itree E X2 i
it_eq_bt E RX R1 i x y <= it_eq_bt E RX R2 i x y
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R1, R2: relᵢ (itree E X1) (itree E X2) H: R1 <= R2 i: I x: itree E X1 i y: itree E X2 i
it_eq_t E RX R1 <= it_eq_t E RX R2
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R1, R2: relᵢ (itree E X1) (itree E X2) H: R1 <= R2 i: I x: itree E X1 i y: itree E X2 i
it_eq_t E RX R2 <= it_eq_t E RX R2
reflexivity.
+Qed.
Justifying strong bisimulations up-to reflexivity, symmetry, and transitivity.
+
Sectionit_eq_facts.
+Context {I} {E : event I I} {X : psh I} {RX : relᵢ X X}.
Reversal, symmetry.
+
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX Y1, Y2: psh I RR: relᵢ Y1 Y2
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX Y1, Y2, Y3: psh I R1: relᵢ Y1 Y2 R2: relᵢ Y2 Y3
(it_eqF E RX R1 ⨟ it_eqF E RX R2) <=
+it_eqF E RX (R1 ⨟ R2)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX Y1, Y2, Y3: psh I R1: relᵢ Y1 Y2 R2: relᵢ Y2 Y3
(it_eqF E RX R1 ⨟ it_eqF E RX R2) <=
+it_eqF E RX (R1 ⨟ R2)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX Y1, Y2, Y3: psh I R1: relᵢ Y1 Y2 R2: relᵢ Y2 Y3 a: I a0: itreeF E X Y1 a a1: itreeF E X Y3 a x: itreeF E X Y2 a u: it_eqF E RX R1 a a0 x v: it_eqF E RX R2 a x a1
it_eqF E RX (R1 ⨟ R2) a a0 a1
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX Y1, Y2, Y3: psh I R1: relᵢ Y1 Y2 R2: relᵢ Y2 Y3 a: I r3, r1, r2: X a r_rel: RX a r1 r2 r_rel0: RX a r2 r3
it_eqF E RX (R1 ⨟ R2) a (RetF r1) (RetF r3)
+econstructor; transitivity r2; auto.
+Qed.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
squareᵢ <= it_eq_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
squareᵢ <= it_eq_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX a: relᵢ (itree E X) (itree E X) a0: I a1, a2, x: itree E X a0 H: it_eqF E RX a a0 (observe a1) (observe x) /\
+it_eqF E RX a a0 (observe x) (observe a2)
it_eqF E RX (a ⨟ a) a0 (observe a1) (observe a2)
+apply it_eqF_tra; eauto.
+Qed.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX RR: relᵢ (itree E X) (itree E X)
Transitiveᵢ (it_eq_t E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX RR: relᵢ (itree E X) (itree E X)
We can now package the previous properties as equivalences.
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
Equivalenceᵢ (it_eq_t E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
Equivalenceᵢ (it_eq_t E RX RR)
econstructor; typeclasses eauto.Qed.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
Equivalenceᵢ (it_eq_bt E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
Equivalenceᵢ (it_eq_bt E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
eqᵢ (itree E X) <= it_eq_bt E RX RR
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
converseᵢ (it_eq_bt E RX RR) <= it_eq_bt E RX RR
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
squareᵢ (it_eq_bt E RX RR) <= it_eq_bt E RX RR
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
eqᵢ (itree E X) <= it_eq_bt E RX RR
apply (fbt_bt it_eq_up2rfl).
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
converseᵢ (it_eq_bt E RX RR) <= it_eq_bt E RX RR
apply (fbt_bt it_eq_up2sym).
+
I: Type E: event I I X: psh I RX: relᵢ X X Equivalenceᵢ0: Equivalenceᵢ RX RR: relᵢ (itree E X) (itree E X)
squareᵢ (it_eq_bt E RX RR) <= it_eq_bt E RX RR
apply (fbt_bt it_eq_up2tra).
+Qed.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X RR: relᵢ (itree E X) (itree E X)
Subrelationᵢ (it_eq RX) (it_eq_bt E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X RR: relᵢ (itree E X) (itree E X)
Subrelationᵢ (it_eq RX) (it_eq_bt E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X RR: relᵢ (itree E X) (itree E X) i: I x, y: itree E X i r: it_eq RX x y
it_eq_bt E RX RR i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X RR: relᵢ (itree E X) (itree E X) i: I x, y: itree E X i r: it_eq_map E RX (gfp (it_eq_map E RX)) i x y
it_eq_bt E RX RR i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X RR: relᵢ (itree E X) (itree E X) i: I x, y: itree E X i r: it_eqF E RX (gfp (it_eq_map E RX)) i (observe x)
+ (observe y)
it_eqF E RX (it_eq_t E RX RR) i (observe x)
+ (observe y)
+
I: Type E: event I I X: psh I RR: relᵢ (itree E X) (itree E X) i: I RX: relᵢ X X t1, t2: itree E X i t_rel: gfp (it_eq_map E RX) i t1 t2 x0: itree E X i
it_eqF E RX (it_eq_t E RX RR) i (TauF t1) (TauF t2)
I: Type E: event I I X: psh I RR: relᵢ (itree E X) (itree E X) i: I RX: relᵢ X X q: e_qry i k1, k2: forallx : e_rsp q, itree E X (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RX) (e_nxt r) (k1 r) (k2 r) x0: itree E X i
it_eqF E RX (it_eq_t E RX RR) i
+ (VisF q k1) (VisF q k2)
+
I: Type E: event I I X: psh I RR: relᵢ (itree E X) (itree E X) i: I RX: relᵢ X X t1, t2: itree E X i t_rel: gfp (it_eq_map E RX) i t1 t2 x0: itree E X i
it_eqF E RX (it_eq_t E RX RR) i (TauF t1) (TauF t2)
econstructor; apply (gfp_t (it_eq_map E RX)); auto.
+
I: Type E: event I I X: psh I RR: relᵢ (itree E X) (itree E X) i: I RX: relᵢ X X q: e_qry i k1, k2: forallx : e_rsp q, itree E X (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RX) (e_nxt r) (k1 r) (k2 r) x0: itree E X i
it_eqF E RX (it_eq_t E RX RR) i (VisF q k1)
+ (VisF q k2)
Similarly to strong bisimilarity, we define weak bisimilarity as the greatest fixpoint
+of a monotone endofunction. A characteristic of weak bisimilarity is that it can
+"skip over" a finite number of Tau nodes on either side. As such, the endofunction
+will allow "eating" a number of Tau nodes before a synchronization step.
+
Note that itree encodes deterministic LTSs, hence we do not need to worry about
+allowing to strip off Tau nodes after the synchronization as well.
+
We will start by defining an inductive "eating" relation, such that intuitively
+it_eat X Y := ∃ n, X = Tau^n(Y).
+
Sectionit_eat.
+Context {I : Type} {E : event I I} {R : psh I}.
+
+Inductiveit_eati : itree' E R i -> itree' E R i -> Prop :=
+ | EatRefl {t} : it_eat i t t
+ | EatStep {t1 t2} : it_eat _ (observe t1) t2 -> it_eat i (TauF t1) t2
+ .
Let's prove some easy properties.
+
I: Type E: event I I R: psh I
Transitiveᵢ it_eat
+
I: Type E: event I I R: psh I
Transitiveᵢ it_eat
+intros i x y z r1 r2; dependent induction r1; auto.
+Defined.
+
+Equationseat_cmp : (revᵢ it_eat ⨟ it_eat) <= (it_eat ∨ᵢ revᵢ it_eat) :=
+ eat_cmp i' x' y' (ex_intro _ z' (conj p' q')) := _eat_cmp p' q'
+ where _eat_cmp {i x y z}
+ : it_eat i x y -> it_eat i x z -> (it_eat i y z \/ it_eat i z y) :=
+ _eat_cmp (EatRefl) q := or_introl q ;
+ _eat_cmp p (EatRefl) := or_intror p ;
+ _eat_cmp (EatStep p) (EatStep q) := _eat_cmp p q .
+
+Definitionit_eat' : relᵢ (itree E R) (itree E R) :=
+ funiuv => it_eat i u.(_observe) v.(_observe).
+
+
I: Type E: event I I R: psh I i: I x, y: itree E R i H: it_eat i (TauF x) (TauF y)
it_eat i (_observe x) (_observe y)
+
I: Type E: event I I R: psh I i: I x, y: itree E R i H: it_eat i (TauF x) (TauF y)
it_eat i (_observe x) (_observe y)
+
I: Type E: event I I R: psh I i: I x, y: itree E R i H: it_eat i (observe x) (TauF y) IHit_eat: forally0x0 : itree E R i,
+observe x = TauF x0 ->
+TauF y = TauF y0 ->
+it_eat i (_observe x0) (_observe y0)
#[global] Notation"a ≈ b" := (it_wbisim (eqᵢ _) _ a b) (at level20).
+
Properties
+
Definitionit_wbisim_step {IEX1X2RX} :
+ it_wbisim RX <= @it_wbisim_map I E X1 X2 RX (it_wbisim RX) :=
+ funixy => proj1 (gfp_fp (it_wbisim_map E RX) i x y) .
+
+Definitionit_wbisim_unstep {IEX1X2RX} :
+ @it_wbisim_map I E X1 X2 RX (it_wbisim RX) <= it_wbisim RX :=
+ funixy => proj2 (gfp_fp (it_wbisim_map E RX) i x y) .
Weak bisimilarity up to synchronization.
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
it_eq_map E RX <= it_wbisim_t E RX
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
it_eq_map E RX <= it_wbisim_t E RX
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree E X1 i y: itree E X2 i H: (it_eq_map E RX ° it_wbisim_map E RX) R i x y
bT (it_wbisim_map E RX) (it_eq_map E RX) R i x y
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x0, t1: itree E X1 i t2: itree E X2 i t_rel: it_wbisimF E RX R i (observe t1) (observe t2)
it_wbisim_T E RX (it_eq_map E RX) R i t1 t2
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x0: itree E X1 i y: itree E X2 i q: e_qry i k1: forallx : e_rsp q, itree E X1 (e_nxt x) k2: forallx : e_rsp q, itree E X2 (e_nxt x) k_rel: forallr : e_rsp q,
+it_wbisimF E RX R (e_nxt r)
+ (observe (k1 r)) (observe (k2 r))
forallr : e_rsp q,
+it_wbisim_T E RX (it_eq_map E RX) R
+ (e_nxt r) (k1 r) (k2 r)
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x0, t1: itree E X1 i t2: itree E X2 i t_rel: it_wbisimF E RX R i (observe t1) (observe t2)
it_wbisim_T E RX (it_eq_map E RX) R i t1 t2
apply (b_T (it_wbisim_map E RX)), t_rel.
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x0: itree E X1 i y: itree E X2 i q: e_qry i k1: forallx : e_rsp q, itree E X1 (e_nxt x) k2: forallx : e_rsp q, itree E X2 (e_nxt x) k_rel: forallr : e_rsp q,
+it_wbisimF E RX R (e_nxt r)
+ (observe (k1 r)) (observe (k2 r))
forallr : e_rsp q,
+it_wbisim_T E RX (it_eq_map E RX) R (e_nxt r) (k1 r)
+ (k2 r)
Transitivity will be quite more involved to prove than for strong bisimilarity. In order
+to prove it, we will need quite a bit of lemmata for moving synchronization points around.
+
First a helper for go/_observe ("in"/"out") maps of the final coalgebra.
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 i: I x: itree E X1 i y: itree E X2 i
it_wbisim RX i x y ->
+it_wbisim RX i {| _observe := _observe x |}
+ {| _observe := _observe y |}
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 i: I x: itree E X1 i y: itree E X2 i
it_wbisim RX i x y ->
+it_wbisim RX i {| _observe := _observe x |}
+ {| _observe := _observe y |}
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 i: I x: itree E X1 i y: itree E X2 i H: it_wbisim RX i x y
it_wbisim RX i {| _observe := _observe x |}
+ {| _observe := _observe y |}
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 i: I x: itree E X1 i y: itree E X2 i H: it_wbisim_map E RX (it_wbisim RX) i x y
it_wbisim RX i {| _observe := _observe x |}
+ {| _observe := _observe y |}
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 i: I x: itree E X1 i y: itree E X2 i H: it_wbisim_map E RX (it_wbisim RX) i x y
it_wbisim_map E RX (it_wbisim RX) i
+ {| _observe := _observe x |}
+ {| _observe := _observe y |}
+exact H.
+Qed.
Strong bisimilarity implies weak bisimilarity.
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
it_eq RX <= it_wbisim RX
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
it_eq RX <= it_wbisim RX
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
forall (a : I) (a0 : itree E X1 a) (a1 : itree E X2 a),
+Basics.impl (it_eq RX a0 a1)
+ (gfp (it_wbisim_map E RX) a a0 a1)
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2
forall (a : I) (a0 : itree E X1 a) (a1 : itree E X2 a),
+it_eq RX a0 a1 -> gfp (it_wbisim_map E RX) a a0 a1
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) CIH: forall (a : I) (a0 : itree E X1 a) (a1 : itree E X2 a),
+it_eq RX a0 a1 -> it_wbisim_t E RX R a a0 a1 i: I x: itree E X1 i y: itree E X2 i H: it_eq RX x y
it_wbisim_bt E RX R i x y
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) CIH: forall (a : I) (a0 : itree E X1 a) (a1 : itree E X2 a),
+it_eq RX a0 a1 -> it_wbisim_t E RX R a a0 a1 i: I x: itree E X1 i y: itree E X2 i H: it_eqF E RX (it_eq RX) i (observe x) (observe y)
it_wbisimF E RX (it_wbisim_t E RX R) i (observe x)
+ (observe y)
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree' E X1 i y: itree E X2 i
it_wbisimF E RX R i x (observe y) ->
+it_wbisimF E RX R i x (TauF y)
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree' E X1 i y: itree E X2 i
it_wbisimF E RX R i x (observe y) ->
+it_wbisimF E RX R i x (TauF y)
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree' E X1 i y: itree E X2 i x1: itree' E X1 i x2: itree' E X2 i r1: it_eat i x x1 r2: it_eat i (observe y) x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX R i x (TauF y)
exact (WBisim r1 (EatStep r2) rr).
+Qed.
+
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree E X1 i y: itree' E X2 i
it_wbisimF E RX R i (observe x) y ->
+it_wbisimF E RX R i (TauF x) y
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree E X1 i y: itree' E X2 i
it_wbisimF E RX R i (observe x) y ->
+it_wbisimF E RX R i (TauF x) y
+
I: Type E: event I I X1, X2: psh I RX: relᵢ X1 X2 R: relᵢ (itree E X1) (itree E X2) i: I x: itree E X1 i y: itree' E X2 i x1: itree' E X1 i x2: itree' E X2 i r1: it_eat i (observe x) x1 r2: it_eat i y x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX R i (TauF x) y
exact (WBisim (EatStep r1) r2 rr).
+Qed.
Removing a Tau left or right.
+
Equationswbisim_step_l {ixy} :
+ it_wbisim' E RX i x (TauF y) ->
+ it_wbisim' E RX i x (observe y)
+ :=
+ wbisim_step_l (WBisim p (EatRefl) (EqTau r))
+ with it_wbisim_step _ _ _ r :=
+ { | WBisim w1 w2 s := WBisim (eat_trans _ _ _ _ p (EatStep w1)) w2 s } ;
+ wbisim_step_l (WBisim p (EatStep q) v) := WBisim p q v.
+
+Equationswbisim_step_r {ixy} :
+ it_wbisim' E RX i (TauF x) y ->
+ it_wbisim' E RX i (observe x) y
+ :=
+ wbisim_step_r (WBisim (EatRefl) q (EqTau r))
+ with it_wbisim_step _ _ _ r :=
+ { | WBisim w1 w2 s := WBisim w1 (eat_trans _ _ _ _ q (EatStep w2)) s } ;
+ wbisim_step_r (WBisim (EatStep p) q v) := WBisim p q v.
Pulling a Tau synchronization point up.
+
Equationswbisim_tau_up_r {ixyz}
+ (u : it_eat i x (TauF y))
+ (v : it_eqF E RX (it_wbisim RX) i (TauF y) z) :
+ it_eqF E RX (it_wbisim RX) i x z
+ :=
+ wbisim_tau_up_r (EatRefl) q := q ;
+ wbisim_tau_up_r (EatStep p) (EqTau q) with it_wbisim_step _ _ _ q := {
+ | WBisim w1 w2 s :=
+ EqTau (it_wbisim_unstep _ _ _ (WBisim (eat_trans _ _ _ _ p (EatStep w1)) w2 s))
+ }.
+
+Equationswbisim_tau_up_l {ixyz}
+ (u : it_eqF E RX (it_wbisim RX) i x (TauF y))
+ (v : it_eat i z (TauF y)) :
+ it_eqF E RX (it_wbisim RX) i x z
+ :=
+ wbisim_tau_up_l p (EatRefl) := p ;
+ wbisim_tau_up_l (EqTau p) (EatStep q) with it_wbisim_step _ _ _ p := {
+ | WBisim w1 w2 s :=
+ EqTau (it_wbisim_unstep _ _ _ (WBisim w1 (eat_trans _ _ _ _ q (EatStep w2)) s))
+ }.
Pushing a Ret or Vis synchronization down.
+
Equationswbisim_ret_down_l {ixyr} :
+ it_wbisim' E RX i x y ->
+ it_eat i y (RetF r) ->
+ (it_eat ⨟ it_eqF E RX (it_wbisim RX)) i x (RetF r)
+ :=
+ wbisim_ret_down_l (WBisim p (EatRefl) w) (EatRefl) := p ⨟⨟ w ;
+ wbisim_ret_down_l w (EatStep q) := wbisim_ret_down_l
+ (wbisim_step_l w) q.
+
+Equationswbisim_ret_down_r {ixyr} :
+ it_eat i x (RetF r) ->
+ it_wbisim' E RX i x y ->
+ (it_eqF E RX (it_wbisim RX) ⨟ revᵢ it_eat) i (RetF r) y
+ :=
+ wbisim_ret_down_r (EatRefl) (WBisim (EatRefl) q w) := w ⨟⨟ q ;
+ wbisim_ret_down_r (EatStep p) w := wbisim_ret_down_r p
+ (wbisim_step_r w).
+
+Equationswbisim_vis_down_l {ixyek} :
+ it_wbisim' E RX i x y ->
+ it_eat i y (VisF e k) ->
+ (it_eat ⨟ it_eqF E RX (it_wbisim RX)) i x (VisF e k)
+ :=
+ wbisim_vis_down_l (WBisim p (EatRefl) w) (EatRefl) := p ⨟⨟ w ;
+ wbisim_vis_down_l w (EatStep q) := wbisim_vis_down_l
+ (wbisim_step_l w) q.
+
+Equationswbisim_vis_down_r {ixyek} :
+ it_eat i x (VisF e k) ->
+ it_wbisim' E RX i x y ->
+ (it_eqF E RX (it_wbisim RX) ⨟ revᵢ it_eat) i (VisF e k) y
+ :=
+ wbisim_vis_down_r (EatRefl) (WBisim (EatRefl) q w) := w ⨟⨟ q ;
+ wbisim_vis_down_r (EatStep p) w := wbisim_vis_down_r p
+ (wbisim_step_r w) .
+Endwbisim_facts_het.
We are now ready to prove the useful properties.
+
Sectionwbisim_facts_hom.
+Context {I : Type} {E : event I I} {X : psh I} {RX : relᵢ X X}.
Registering that strong bisimilarity is a subrelation.
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX
const (eqᵢ (itree E X)) <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX
const (eqᵢ (itree E X)) <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX
const (eqᵢ (itree E X)) ° it_wbisim_map E RX <=
+it_wbisim_map E RX ° const (eqᵢ (itree E X))
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX a: relᵢ (itree E X) (itree E X) a0: I a1, a2: itree E X a0 H: (const (eqᵢ (itree E X)) ° it_wbisim_map E RX) a
+ a0 a1 a2
(it_wbisim_map E RX ° const (eqᵢ (itree E X))) a a0 a1
+ a2
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX a: relᵢ (itree E X) (itree E X) a0: I a1, a2: itree E X a0 H: (const (eqᵢ (itree E X)) ° it_wbisim_map E RX) a
+ a0 a1 a2
(it_wbisim_map E RX ° const (eqᵢ (itree E X))) a a0 a2
+ a2
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX a: relᵢ (itree E X) (itree E X) a0: I a1, a2: itree E X a0 H: (const (eqᵢ (itree E X)) ° it_wbisim_map E RX) a
+ a0 a1 a2
it_eqF E RX (const (eqᵢ (itree E X)) a) a0
+ (observe a2) (observe a2)
+nowapply it_eqF_rfl.
+Qed.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX RR: relᵢ (itree E X) (itree E X)
Reflexiveᵢ (it_wbisim_t E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X Reflexiveᵢ0: Reflexiveᵢ RX RR: relᵢ (itree E X) (itree E X)
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX
converseᵢ <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX
converseᵢ <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX a: relᵢ (itree E X) (itree E X) a0: I a1, a2: itree E X a0 x1, x2: itree' E X a0 r1: it_eat a0 (observe a2) x1 r2: it_eat a0 (observe a1) x2 rr: it_eqF E RX a a0 x1 x2
(it_wbisim_map E RX ° converseᵢ) a a0 a1 a2
+
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX a: relᵢ (itree E X) (itree E X) a0: I a1, a2: itree E X a0 x1, x2: itree' E X a0 r1: it_eat a0 (observe a2) x1 r2: it_eat a0 (observe a1) x2 rr: it_eqF E RX a a0 x1 x2
it_eqF E RX (converseᵢ a) a0 x2 x1
+nowapply it_eqF_sym.
+Qed.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX RR: relᵢ (itree E X) (itree E X)
Symmetricᵢ (it_wbisim_t E RX RR)
+
I: Type E: event I I X: psh I RX: relᵢ X X Symmetricᵢ0: Symmetricᵢ RX RR: relᵢ (itree E X) (itree E X)
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
(it_wbisim' E RX ⨟ it_wbisim' E RX) <=
+it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
(it_wbisim' E RX ⨟ it_wbisim' E RX) <=
+it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, z, x1, x2: itree' E X i u1: it_eat i x x1 u2: it_eat i z x2 uS: it_eqF E RX (it_wbisim RX) i x1 x2 y1, y2: itree' E X i v1: it_eat i z y1 v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: it_eat i x2 y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i x2 y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: it_eat i x2 y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 r: X i y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (RetF r) y2 w: it_eat i x2 (RetF r)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 t: itree E X i y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (TauF t) y2 w: it_eat i x2 (TauF t)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 q: e_qry i k: forallr : e_rsp q, itree E X (e_nxt r) y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (VisF q k) y2 w: it_eat i x2 (VisF q k)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 r: X i y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (RetF r) y2 w: it_eat i x2 (RetF r)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i r: X i y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (RetF r) y2 z: itree' E X i w1: it_eat i x z ww: it_eqF E RX (it_wbisim RX) i z (RetF r)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i r, r2: X i v2: it_eat i y (RetF r2) r_rel: RX i r r2 r1: X i w1: it_eat i x (RetF r1) r_rel0: RX i r1 r
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+refine (WBisim w1 v2 (EqRet _)); nowtransitivity r.
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 t: itree E X i y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (TauF t) y2 w: it_eat i x2 (TauF t)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 q: e_qry i k: forallr : e_rsp q, itree E X (e_nxt r) y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (VisF q k) y2 w: it_eat i x2 (VisF q k)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i q: e_qry i k: forallr : e_rsp q, itree E X (e_nxt r) y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i (VisF q k) y2 z: itree' E X i w1: it_eat i x z ww: it_eqF E RX (it_wbisim RX) i z (VisF q k)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i q: e_qry i k, k2: forallx : e_rsp q, itree E X (e_nxt x) v2: it_eat i y (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim RX (e_nxt r) (k r) (k2 r) k1: forallx : e_rsp q, itree E X (e_nxt x) w1: it_eat i x (VisF q k1) k_rel0: forallr : e_rsp q,
+it_wbisim RX (e_nxt r) (k1 r) (k r)
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1, x2: itree' E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 x2 y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i x2 y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i r: X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (RetF r) y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i (RetF r) y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i t: itree E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (TauF t) y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i (TauF t) y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i q: e_qry i k: forallr : e_rsp q, itree E X (e_nxt r) u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (VisF q k) y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i (VisF q k) y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i r: X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (RetF r) y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i (RetF r) y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i r: X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (RetF r) y1, y2, z: itree' E X i ww: it_eqF E RX (it_wbisim RX) i (RetF r) z w1: revᵢ it_eat i z y
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y: itree' E X i r1, r: X i u1: it_eat i x (RetF r1) r_rel: RX i r1 r y1, y2: itree' E X i r2: X i r_rel0: RX i r r2 w1: revᵢ it_eat i (RetF r2) y
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+refine (WBisim u1 w1 (EqRet _)); nowtransitivity r.
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i t: itree E X i u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (TauF t) y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i (TauF t) y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
exact (WBisim u1 v2 (it_eqF_tra _ _ _ (wbisim_tau_up_l uS w ⨟⨟ vS))).
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i q: e_qry i k: forallr : e_rsp q, itree E X (e_nxt r) u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (VisF q k) y1, y2: itree' E X i v2: it_eat i y y2 vS: it_eqF E RX (it_wbisim RX) i y1 y2 w: revᵢ it_eat i (VisF q k) y1
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y, x1: itree' E X i q: e_qry i k: forallr : e_rsp q, itree E X (e_nxt r) u1: it_eat i x x1 uS: it_eqF E RX (it_wbisim RX) i x1 (VisF q k) y1, y2, z: itree' E X i ww: it_eqF E RX (it_wbisim RX) i (VisF q k) z w1: revᵢ it_eat i z y
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX i: I x, y: itree' E X i q: e_qry i k1, k: forallr : e_rsp q, itree E X (e_nxt r) u1: it_eat i x (VisF q k1) k_rel: forallr : e_rsp q,
+it_wbisim RX (e_nxt r) (k1 r) (k r) y1, y2: itree' E X i k2: forallx : e_rsp q, itree E X (e_nxt x) k_rel0: forallr : e_rsp q,
+it_wbisim RX (e_nxt r) (k r) (k2 r) w1: revᵢ it_eat i (VisF q k2) y
it_wbisimF E RX (it_wbisim RX ⨟ it_wbisim RX) i x y
We have proven that strong bisimilarity entails weak bisimilarity, but now we prove the
+much more powerful fact that we can prove weak bisimilarity up-to strong bisimilarity.
+That is, we will be allowed to close any weak bisimulation candidate by strong bisimilarity.
+Let us first define a helper relation taking a relation to its saturation by strong
+bisimilarity.
+
Varianteq_clo (R : relᵢ (itree E X) (itree E X)) i (xy : itree E X i) : Prop :=
+ | EqClo {a b} : it_eq RX x a -> it_eq RX b y -> R i a b -> eq_clo R i x y
+ .
+#[global] Arguments EqClo {R i x y a b}.
This helper is monotone...
+
Definitioneq_clo_map : mon (relᵢ (itree E X) (itree E X)) :=
+ {| body R := eq_clo R ;
+ Hbody _ _ H _ _ _ '(EqClo p q r) := EqClo p q (H _ _ _ r) |}.
... and below the companion of weak bisimilarity, hence justifying the up-to.
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
eq_clo_map <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
eq_clo_map <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eq RX a c v: it_eq RX d b x1, x2: itree' E X i r1: it_eat i (observe c) x1 r2: it_eat i (observe d) x2 rr: it_eqF E RX R i x1 x2
bT (it_wbisim_map E RX) eq_clo_map R i a b
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eqF E RX (it_eq RX) i (observe a) (observe c) v: it_eqF E RX (it_eq RX) i (observe d) (observe b) x1, x2: itree' E X i r1: it_eat i (observe c) x1 r2: it_eat i (observe d) x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i
+ (observe a) (observe b)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I b, c, d: itree E X i oa: itree' E X i u: it_eqF E RX (it_eq RX) i oa (observe c) v: it_eqF E RX (it_eq RX) i (observe d) (observe b) x1, x2: itree' E X i r1: it_eat i (observe c) x1 r2: it_eat i (observe d) x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ (observe b)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I c, d: itree E X i oa: itree' E X i u: it_eqF E RX (it_eq RX) i oa (observe c) ob: itree' E X i v: it_eqF E RX (it_eq RX) i (observe d) ob x1, x2: itree' E X i r1: it_eat i (observe c) x1 r2: it_eat i (observe d) x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I d: itree E X i oa, oc: itree' E X i u: it_eqF E RX (it_eq RX) i oa oc ob: itree' E X i v: it_eqF E RX (it_eq RX) i (observe d) ob x1, x2: itree' E X i r1: it_eat i oc x1 r2: it_eat i (observe d) x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I oa, oc: itree' E X i u: it_eqF E RX (it_eq RX) i oa oc ob, od: itree' E X i v: it_eqF E RX (it_eq RX) i od ob x1, x2: itree' E X i r1: it_eat i oc x1 r2: it_eat i od x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t, oa, ob, od, x2: itree' E X i u: it_eqF E RX (it_eq RX) i oa t r2: it_eat i od x2 v: it_eqF E RX (it_eq RX) i od ob rr: it_eqF E RX R i t x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t1: itree E X i t2: itree' E X i r1: it_eat i (observe t1) t2 IHr1: foralloaobodx2 : itree' E X i,
+it_eqF E RX (it_eq RX) i oa (observe t1) ->
+it_eat i od x2 ->
+it_eqF E RX (it_eq RX) i od ob ->
+it_eqF E RX R i t2 x2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob oa, ob, od, x2: itree' E X i u: it_eqF E RX (it_eq RX) i oa (TauF t1) r2: it_eat i od x2 v: it_eqF E RX (it_eq RX) i od ob rr: it_eqF E RX R i t2 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t, oa, ob, od, x2: itree' E X i u: it_eqF E RX (it_eq RX) i oa t r2: it_eat i od x2 v: it_eqF E RX (it_eq RX) i od ob rr: it_eqF E RX R i t x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t, t0, oa, ob: itree' E X i u: it_eqF E RX (it_eq RX) i oa t v: it_eqF E RX (it_eq RX) i t0 ob rr: it_eqF E RX R i t t0
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t: itree' E X i t1: itree E X i t2: itree' E X i r2: it_eat i (observe t1) t2 IHr2: foralloaob : itree' E X i,
+it_eqF E RX (it_eq RX) i oa t ->
+it_eqF E RX (it_eq RX) i (observe t1) ob ->
+it_eqF E RX R i t t2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob oa, ob: itree' E X i u: it_eqF E RX (it_eq RX) i oa t v: it_eqF E RX (it_eq RX) i (TauF t1) ob rr: it_eqF E RX R i t t2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t, t0, oa, ob: itree' E X i u: it_eqF E RX (it_eq RX) i oa t v: it_eqF E RX (it_eq RX) i t0 ob rr: it_eqF E RX R i t t0
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I r3, r5, r4: X i r_rel0: RX i r3 r4 r1: X i r_rel1: RX i r1 r5 r_rel: RX i r4 r1
RX i r3 r5
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t3, t5, t4: itree E X i t_rel0: it_eq RX t3 t4 t1: itree E X i t_rel1: it_eq RX t1 t5 t_rel: R i t4 t1
it_wbisim_T E RX eq_clo_map R i t3 t5
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I q: e_qry i k3, k5, k4: forallx : e_rsp q, itree E X (e_nxt x) k_rel0: forallr : e_rsp q, it_eq RX (k3 r) (k4 r) k1: forallx : e_rsp q, itree E X (e_nxt x) k_rel1: forallr : e_rsp q, it_eq RX (k1 r) (k5 r) k_rel: forallr : e_rsp q, R (e_nxt r) (k4 r) (k1 r)
forallr : e_rsp q,
+it_wbisim_T E RX eq_clo_map R (e_nxt r) (k3 r) (k5 r)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I r3, r5, r4: X i r_rel0: RX i r3 r4 r1: X i r_rel1: RX i r1 r5 r_rel: RX i r4 r1
RX i r3 r5
transitivity r4; auto; transitivity r1; auto.
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t3, t5, t4: itree E X i t_rel0: it_eq RX t3 t4 t1: itree E X i t_rel1: it_eq RX t1 t5 t_rel: R i t4 t1
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I q: e_qry i k3, k5, k4: forallx : e_rsp q, itree E X (e_nxt x) k_rel0: forallr : e_rsp q, it_eq RX (k3 r) (k4 r) k1: forallx : e_rsp q, itree E X (e_nxt x) k_rel1: forallr : e_rsp q, it_eq RX (k1 r) (k5 r) k_rel: forallr : e_rsp q, R (e_nxt r) (k4 r) (k1 r)
forallr : e_rsp q,
+it_wbisim_T E RX eq_clo_map R (e_nxt r) (k3 r) (k5 r)
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t: itree' E X i t1: itree E X i t2: itree' E X i r2: it_eat i (observe t1) t2 IHr2: foralloaob : itree' E X i,
+it_eqF E RX (it_eq RX) i oa t ->
+it_eqF E RX (it_eq RX) i (observe t1) ob ->
+it_eqF E RX R i t t2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob oa, ob: itree' E X i u: it_eqF E RX (it_eq RX) i oa t v: it_eqF E RX (it_eq RX) i (TauF t1) ob rr: it_eqF E RX R i t t2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t: itree' E X i t3: itree E X i t2: itree' E X i r2: it_eat i (observe t3) t2 IHr2: foralloaob : itree' E X i,
+it_eqF E RX (it_eq RX) i oa t ->
+it_eqF E RX (it_eq RX) i (observe t3) ob ->
+it_eqF E RX R i t t2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob oa: itree' E X i t4: itree E X i u: it_eqF E RX (it_eq RX) i oa t t_rel: it_eq RX t3 t4 rr: it_eqF E RX R i t t2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ (TauF t4)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t: itree' E X i t3: itree E X i t2: itree' E X i r2: it_eat i (observe t3) t2 IHr2: foralloaob : itree' E X i,
+it_eqF E RX (it_eq RX) i oa t ->
+it_eqF E RX (it_eq RX) i (observe t3) ob ->
+it_eqF E RX R i t t2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob oa: itree' E X i t4: itree E X i u: it_eqF E RX (it_eq RX) i oa t t_rel: it_eq_map E RX (it_eq RX) i t3 t4 rr: it_eqF E RX R i t t2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ (TauF t4)
+apply wbisim_unstep_l, IHr2; auto.
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t1: itree E X i t2: itree' E X i r1: it_eat i (observe t1) t2 IHr1: foralloaobodx2 : itree' E X i,
+it_eqF E RX (it_eq RX) i oa (observe t1) ->
+it_eat i od x2 ->
+it_eqF E RX (it_eq RX) i od ob ->
+it_eqF E RX R i t2 x2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob oa, ob, od, x2: itree' E X i u: it_eqF E RX (it_eq RX) i oa (TauF t1) r2: it_eat i od x2 v: it_eqF E RX (it_eq RX) i od ob rr: it_eqF E RX R i t2 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i oa
+ ob
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t4: itree E X i t2: itree' E X i r1: it_eat i (observe t4) t2 IHr1: foralloaobodx2 : itree' E X i,
+it_eqF E RX (it_eq RX) i oa (observe t4) ->
+it_eat i od x2 ->
+it_eqF E RX (it_eq RX) i od ob ->
+it_eqF E RX R i t2 x2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob t3: itree E X i ob, od, x2: itree' E X i t_rel: it_eq RX t3 t4 r2: it_eat i od x2 v: it_eqF E RX (it_eq RX) i od ob rr: it_eqF E RX R i t2 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i
+ (TauF t3) ob
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I t4: itree E X i t2: itree' E X i r1: it_eat i (observe t4) t2 IHr1: foralloaobodx2 : itree' E X i,
+it_eqF E RX (it_eq RX) i oa (observe t4) ->
+it_eat i od x2 ->
+it_eqF E RX (it_eq RX) i od ob ->
+it_eqF E RX R i t2 x2 ->
+it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R)
+ i oa ob t3: itree E X i ob, od, x2: itree' E X i t_rel: it_eq_map E RX (it_eq RX) i t3 t4 r2: it_eat i od x2 v: it_eqF E RX (it_eq RX) i od ob rr: it_eqF E RX R i t2 x2
it_wbisimF E RX (it_wbisim_T E RX eq_clo_map R) i
+ (TauF t3) ob
+apply wbisim_unstep_r, (IHr1 (observe t3) ob od x2); auto.
+Qed.
We now do the same for up-to eating.
+
Varianteat_clo (R : relᵢ (itree E X) (itree E X)) i (xy : itree E X i) : Prop :=
+ | EatClo {a b} : it_eat' i x a -> it_eat' i y b -> R i a b -> eat_clo R i x y
+ .
+#[global] Arguments EatClo {R i x y a b}.
+
+Definitioneat_clo_map : mon (relᵢ (itree E X) (itree E X)) :=
+ {| body R := eat_clo R ;
+ Hbody _ _ H _ _ _ '(EatClo p q r) := EatClo p q (H _ _ _ r) |}.
+
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX
eat_clo_map <= it_wbisim_t E RX
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat' i a c v: it_eat' i b d x1, x2: itree' E X i r1: it_eat i (observe c) x1 r2: it_eat i (observe d) x2 rr: it_eqF E RX R i x1 x2
(it_wbisim_map E RX ° eat_clo_map) R i a b
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_wbisimF E RX (eat_clo R) i (_observe a)
+ (_observe b)
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_eat i (_observe a) ?x1
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_eat i (_observe b) ?x2
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_eqF E RX (eat_clo R) i ?x1?x2
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_eat i (_observe a) ?x1
etransitivity; [ exact u | exact r1 ].
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_eat i (_observe b) ?x2
etransitivity; [ exact v | exact r2 ].
+
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2 rr: it_eqF E RX R i x1 x2
it_eqF E RX (eat_clo R) i x1 x2
I: Type E: event I I X: psh I RX: relᵢ X X Transitiveᵢ0: Transitiveᵢ RX R: relᵢ (itree E X) (itree E X) i: I a, b, c, d: itree E X i u: it_eat i (_observe a) (_observe c) v: it_eat i (_observe b) (_observe d) x1, x2: itree' E X i r1: it_eat i (_observe c) x1 r2: it_eat i (_observe d) x2
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Game event
+
+
In order to specify the set of visible moves of an indexed interaction tree, we use the
+notion of indexed polynomial functor. See "Indexed Containers" by Thorsten Altenkirch,
+Neil Ghani, Peter Hancock, Conor McBride, and Peter Morris.
+
In our interaction-oriented nomenclature, an indexed container consists of a family of
+queries e_qry, a family of responses e_rsp and a function assigning the next
+index to each response e_nxt.
+
Recordevent (IJ : Type) : Type := Event {
+ e_qry : I -> Type ;
+ e_rsp : foralli, e_qry i -> Type ;
+ e_nxt : foralli (q : e_qry i), e_rsp i q -> J
+}.
Finally we can interpret an event as a functor from families over J to families
+over I.
+
Definitione_interp {IJ} (E : event I J) (X : psh J) : psh I :=
+ funi => { q : E.(e_qry) i & forall (r : E.(e_rsp) q), X (E.(e_nxt) r) } .
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Scoped and sorted families
+
+
In this file we bootstrap the abstract theory of contexts, variables and substitution
+by defining a special case of type-families: families indexed by a set of contexts C
+and possibly a number of times by a set of types T.
+
DefinitionFam₀ (TC : Type) := C -> Type .
+DefinitionFam₁ (TC : Type) := C -> T -> Type .
+DefinitionFam₂ (TC : Type) := C -> T -> T -> Type .
Additionally, to represent binders we use the following alternative to sorted families
+Fam₁. In this definition, the set of binders is indexed by the set of types T but
+not by the set of contexts C. Instead, we have a projection into C. This is important
+for reasons of bidirectional typing information flow. We could go with Fam₁ instead, but
+this would entail quite a bit of equality dance when manipulating it.
+
RecordOperTC := {
+ o_op : T -> Type ;
+ o_dom : forallx, o_op x -> C ;
+}.
Now a first combinator: the formal cut of two families. This takes two 1-indexed families
+and returns the 0-indexed family consisting pairs of matching elements.
+
Recordf_cut {TC} (FG : Fam₁ T C) (Γ : C) :=
+ Cut { cut_ty : T ; cut_l : F Γ cut_ty ; cut_r : G Γ cut_ty }.
We embed binder descriptions into sorted families, by dropping the context.
+
Definitionbare_op {TC} (O : Oper T C) : Fam₁ T C := fun_x => O x.
We define maps of these families, their identity and composition.
+
Definitionarr_fam₀ (FG : Fam₀ T C) := forallΓ, F Γ -> G Γ.
+Definitionarr_fam₁ (FG : Fam₁ T C) := forallΓx, F Γ x -> G Γ x.
+Definitionarr_fam₂ (FG : Fam₂ T C) := forallΓxy, F Γ x y -> G Γ x y.
+
+Notation"F ⇒₀ G" := (arr_fam₀ F G).
+Notation"F ⇒₁ G" := (arr_fam₁ F G).
+Notation"F ⇒₂ G" := (arr_fam₂ F G).
+
+Definitionf_id₀ {F : Fam₀ T C} : F ⇒₀ F := fun_a => a.
+Definitionf_id₁ {F : Fam₁ T C} : F ⇒₁ F := fun__a => a.
+Definitionf_id₂ {F : Fam₂ T C} : F ⇒₂ F := fun___a => a.
+
+Definitionf_comp₀ {FGH : Fam₀ T C} (u : G ⇒₀ H) (v : F ⇒₀ G) : F ⇒₀ H
+ := fun_x => u _ (v _ x).
+Definitionf_comp₁ {FGH : Fam₁ T C} (u : G ⇒₁ H) (v : F ⇒₁ G) : F ⇒₁ H
+ := fun__x => u _ _ (v _ _ x).
+Definitionf_comp₂ {FGH : Fam₂ T C} (u : G ⇒₂ H) (v : F ⇒₂ G) : F ⇒₂ H
+ := fun___x => u _ _ _ (v _ _ _ x).
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
An half games (Def 5.1) is composed of "moves" as an indexed family of types,
+and a "next" map computing the next index following a move.
+
Recordhalf_game (IJ : Type) := {
+ g_move : I -> Type ;
+ g_next : foralli, g_move i -> J
+}.
Action (h_actv) and reaction (h_pasv) functors (Def 5.8)
+
Definitionh_actv {IJ} (H : half_game I J) (X : psh J) : psh I :=
+ funi => { m : H.(g_move) i & X (H.(g_next) m) } .
+
+Definitionh_actvR {IJ} (H : half_game I J) {XY : psh J} (R : relᵢ X Y)
+ : relᵢ (h_actv H X) (h_actv H Y) :=
+ funiuv => existsp : projT1 u = projT1 v , R _ (rew p in projT2 u) (projT2 v) .
+
+Definitionh_pasv {IJ} (H : half_game I J) (X : psh J) : psh I :=
+ funi => forall (m : H.(g_move) i), X (H.(g_next) m) .
+
+Definitionh_pasvR {IJ} (H : half_game I J) {XY : psh J} (R : relᵢ X Y)
+ : relᵢ (h_pasv H X) (h_pasv H Y) := funiuv => forallm, R _ (u m) (v m) .
A game (Def 5.4) is composed of two compatible half games.
+
Recordgame (IJ : Type) : Type := {
+ g_client : half_game I J ;
+ g_server : half_game J I
+}.
Given a game, we can construct an event. See ITree/Event.v
+
Definitione_of_g {IJ} (G : game I J) : event I I :=
+ {| e_qry := funi => G.(g_client).(g_move) i ;
+ e_rsp := funiq => G.(g_server).(g_move) (G.(g_client).(g_next) q) ;
+ e_nxt := funiqr => G.(g_server).(g_next) r |} .
+
+
The OGS Game (§ 5.2)
+
First let us define a datatype for polarities, active and passive (called "waiting") in
+the paper.
+
Variantpolarity : Type := Act | Pas .
+DeriveNoConfusionfor polarity.
+
+Equationsp_switch : polarity -> polarity :=
+ p_switch Act := Pas ;
+ p_switch Pas := Act .
+#[global] Notation"p ^" := (p_switch p) (at level5).
+
+Sectionwith_param.
We consider an observation structure, given by a set of types T, a notion of contexts
+C and a operator giving the observations and their domain. See
+Ctx/Family.v and OGS/Obs.v.
+
Context `{CC : context T C} {obs : obs_struct T C}.
Interleaved contexts (Def 5.12) are given by the free context structure over C.
+
Definitionogs_ctx := ctx C.
We define the collapsing functions (Def 5.13).
+
Equationsjoin_pol : polarity -> ogs_ctx -> C :=
+ join_pol Act ∅ₓ := ∅ ;
+ join_pol Act (Ψ ▶ₓ Γ) := join_pol Pas Ψ +▶ Γ ;
+ join_pol Pas ∅ₓ := ∅ ;
+ join_pol Pas (Ψ ▶ₓ Γ) := join_pol Act Ψ .
+
+Notation"↓⁺ Ψ" := (join_pol Act Ψ).
+Notation"↓⁻ Ψ" := (join_pol Pas Ψ).
+Notation"↓[ p ] Ψ" := (join_pol p Ψ).
Finally we define the OGS half-game and game (Def 5.15).
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Guarded Iteration
+
+
The iter combinator, adapted directly from the Interaction Tree library,
+happily builds a computation iter f for any set of equations f.
+The resulting computation satisfies in particular an unfolding equation:
+iter f i ≈ f i >>= case_ (iter f) (ret)
+In general, this solution is however not unique, depriving us from a
+powerful tool to prove the equivalence of two computations.
+
Through this file, we recover uniqueness by introducing an alternate
+combinator iter_guarded restricted to so-called guarded sets of equations:
+they cannot be reduced to immediately returning a new index to iterate over.
+Both combinators agree over guarded equations (w.r.t. to weak bisimilarity), but
+the new one satisfies uniqueness.
+
Then, we refine the result to allow more sets of equations: not all must
+be guarded, but they must always finitely lead to a guarded one.
+
+
Guarded iteration
+
Sectionguarded.
+
+Context {I} {E : event I I}.
A set of equations is an (indexed) family of computations, i.e. body fed to the
+combinator (Def. 30).
+
DefinitioneqnRXY : Type := X ⇒ᵢ itree E (Y +ᵢ R) .
+
+Definitionapply_eqn {RXY} (q : eqn R X Y) : itree E (X +ᵢ R) ⇒ᵢ itree E (Y +ᵢ R) :=
+ fun_t => t >>= fun_r => match r with
+ | inl x => q _ x
+ | inr y => Ret' (inr y)
+ end .
A computation is guarded if it is not reduced to Ret (inl i), i.e.
+if as part of the loop, this specific iteration will be observable w.r.t.
+strong bisimilarity. It therefore may:
+
+
end the iteration
+
produce a silent step
+
produce an external event
+
+
A set of equation is then said to be guarded if all its equations are guarded
+(Def. 33).
+
Variantguarded' {XYi} : itree' E (X +ᵢ Y) i -> Prop :=
+ | GRet y : guarded' (RetF (inr y))
+ | GTau t : guarded' (TauF t)
+ | GVis e k : guarded' (VisF e k)
+ .
+Definitionguarded {XYi} (t : itree E (X +ᵢ Y) i)
+ := guarded' t.(_observe).
+Definitioneqn_guarded {RXY} (e : eqn R X Y) : Prop
+ := foralli (x : X i), guarded (e i x) .
The iter combinator gets away with putting no restriction on the equations
+by aggressively guarding with Tau all recursive calls.
+In contrast, by assuming the equations are guarded, we do not need to introduce
+any spurious guard: they are only legitimate in the case of returning immediately
+a new index, which we can here rule out via elim_guarded.
+
Equationselim_guarded {RXix} : @guarded' R X i (RetF (inl x)) -> T0 := | ! .
+
+Definitioniter_guarded_aux {RX}
+ (e : eqn R X X) (EG : eqn_guarded e)
+ : itree E (X +ᵢ R) ⇒ᵢ itree E R
+ := cofix CIH i t :=
+ t >>= fun_r =>
+ go match r with
+ | inl x => match (e _ x).(_observe) as t return guarded' t -> _ with
+ | RetF (inl x) => fung => ex_falso (elim_guarded g)
+ | RetF (inr r) => fun_ => RetF r
+ | TauF t => fun_ => TauF (CIH _ t)
+ | VisF q k => fun_ => VisF q (funr => CIH _ (k r))
+ end (EG _ x)
+ | inr y => RetF y
+ end .
A better iteration for guarded equations.
+
Definitioniter_guarded {RX}
+ (f : eqn R X X) (EG : eqn_guarded f)
+ : X ⇒ᵢ itree E R
+ := fun_x =>
+ go (match (f _ x).(_observe) as t return guarded' t -> _ with
+ | RetF (inl x) => fung => ex_falso (elim_guarded g)
+ | RetF (inr r) => fun_ => RetF r
+ | TauF t => fun_ => TauF (iter_guarded_aux f EG _ t)
+ | VisF q k => fun_ => VisF q (funr => iter_guarded_aux f EG _ (k r))
+ end (EG _ x)) .
The absence of a spurious guard is reflected in the unfolding equation:
+w.r.t. strong bisimilarity, it directly satisfies what one would expect,
+i.e. we run the current equation, and either terminate or jump to the next
+one.
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I t: itree E (X +ᵢ Y) i
it_eq RY (iter_guarded_aux f EG i t)
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I t: itree E (X +ᵢ Y) i
it_eq RY (iter_guarded_aux f EG i t)
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i
it_eq_bt E RY R i (iter_guarded_aux f EG i t)
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i ot:= _observe t: itree' E (X +ᵢ Y) i
it_eqF E RY (it_eq_t E RY R) i
+ match ot with
+ | RetF (inl x) =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun_ : guarded' (RetF (inr r1)) =>
+ RetF r1
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r0) (k r0))
+ end (EG i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x =>
+ match
+ _observe (f i x) as t0
+ return
+ (guarded' t0 -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) ->
+ itree' E Y i)
+ with
+ | inl x0 =>
+ ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun
+ _ : guarded'
+ (RetF (inr r1)) =>
+ RetF r1
+ end
+ | TauF t0 =>
+ fun_ : guarded' (TauF t0) =>
+ TauF
+ (iter_guarded_aux f EG i t0)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG
+ (e_nxt r0) (k r0))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return
+ (guarded' t -> itree' E Y i)
+ with
+ | RetF r1 =>
+ match
+ r1 as r2
+ return
+ (guarded' (RetF r2) ->
+ itree' E Y i)
+ with
+ | inl x0 =>
+ ex_falso ∘ elim_guarded
+ | inr r2 =>
+ fun
+ _ : guarded'
+ (RetF (inr r2)) =>
+ RetF r2
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k0 =>
+ fun_ : guarded' (VisF q k0) =>
+ VisF q
+ (funr1 : e_rsp q =>
+ iter_guarded_aux f EG
+ (e_nxt r1) (k0 r1))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match ot with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i x: X i
it_eqF E RY (it_eq_t E RY R) i
+ (match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end (EG i x))
+ (match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end (EG i x))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i y: Y i
it_eqF E RY (it_eq_t E RY R) i (RetF y) (RetF y)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t, t0: itree E (X +ᵢ Y) i
it_eqF E RY (it_eq_t E RY R) i
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun
+ _ : guarded' (RetF (inr r1))
+ => RetF r1
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r0)
+ (k r0))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r)
it_eqF E RY (it_eq_t E RY R) i
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r1 =>
+ match
+ r1 as r2
+ return
+ (guarded' (RetF r2) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r2 =>
+ fun
+ _ : guarded' (RetF (inr r2))
+ => RetF r2
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr1 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r1)
+ (k r1))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i x: X i y: Y i
it_eqF E RY (it_eq_t E RY R) i (RetF y) (RetF y)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i x: X i t0: itree E (X +ᵢ Y) i
it_eqF E RY (it_eq_t E RY R) i
+ (TauF (iter_guarded_aux f EG i t0))
+ (TauF (iter_guarded_aux f EG i t0))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r)
it_eqF E RY (it_eq_t E RY R) i
+ (VisF e
+ (funr : e_rsp e =>
+ iter_guarded_aux f EG (e_nxt r) (k r)))
+ (VisF e
+ (funr : e_rsp e =>
+ iter_guarded_aux f EG (e_nxt r) (k r)))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i y: Y i
it_eqF E RY (it_eq_t E RY R) i (RetF y) (RetF y)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t, t0: itree E (X +ᵢ Y) i
it_eqF E RY (it_eq_t E RY R) i
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun
+ _ : guarded' (RetF (inr r1))
+ => RetF r1
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r0)
+ (k r0))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r)
it_eqF E RY (it_eq_t E RY R) i
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r1 =>
+ match
+ r1 as r2
+ return
+ (guarded' (RetF r2) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r2 =>
+ fun
+ _ : guarded' (RetF (inr r2))
+ => RetF r2
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr1 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r1)
+ (k r1))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t, t0: itree E (X +ᵢ Y) i
it_eq_t E RY R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun_ : guarded' (RetF (inr r1))
+ => RetF r1
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r0)
+ (k r0))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r)
forallr : e_rsp q,
+it_eq_t E RY R (e_nxt r)
+ ((fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r1 =>
+ match
+ r1 as r2
+ return
+ (guarded' (RetF r2) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r2 =>
+ fun_ : guarded' (RetF (inr r2))
+ => RetF r2
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr1 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r1)
+ (k r1))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ ((fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t, t0: itree E (X +ᵢ Y) i
it_eq_t E RY R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun_ : guarded' (RetF (inr r1))
+ => RetF r1
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r0)
+ (k r0))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) r: e_rsp q
it_eq_t E RY R (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) ->
+ itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r1 =>
+ fun_ : guarded' (RetF (inr r1))
+ => RetF r1
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr0 : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r0)
+ (k r0))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t, t0: itree E (X +ᵢ Y) i
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_eq_t E RY R i
+ {|
+ _observe :=
+ match x1 with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) =>
+ RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}
+ match x2 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E (X +ᵢ Y) i),
+it_eq_t E RY R i (iter_guarded_aux f EG i t0)
+(t0 >>=
+(fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+match r with
+| inl x => iter_guarded f EG i0 x
+| inr y => Ret' y
+end)) i: I t: itree E (X +ᵢ Y) i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) r: e_rsp q
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_eq_t E RY R i
+ {|
+ _observe :=
+ match x1 with
+ | inl x =>
+ match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return
+ (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x0 => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) =>
+ RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end (EG i x)
+ | inr y => RetF y
+ end
+ |}
+ match x2 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+all: intros ? ? x2 ->; destruct x2; auto.
+Qed.
Guarded iteration is a fixed point of the equation w.r.t. strong bisimilarity.
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I x: X i
it_eq RY (iter_guarded f EG i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I x: X i
it_eq RY (iter_guarded f EG i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I x: X i
it_eqF E RY (it_eq RY) i
+ (match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end (EG i x))
+ match _observe (f i x) with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I x: X i p:= EG i x: guarded' (_observe (f i x))
it_eqF E RY (it_eq RY) i
+ (match
+ _observe (f i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end p)
+ match _observe (f i x) with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Reflexiveᵢ RY f: eqn Y X X EG: eqn_guarded f i: I x: X i ot:= _observe (f i x): itree' E (X +ᵢ Y) i p:= EG i x: guarded' ot
it_eqF E RY (it_eq RY) i
+ (match
+ ot as t return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF (iter_guarded_aux f EG i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux f EG (e_nxt r) (k r))
+ end p)
+ match ot with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+destruct p; econstructor; [ nowapply H | | intro r ]; apply iter_guarded_aux_unfold.
+Qed.
The payoff: iter_guarded does not only deliver a solution to the equations,
+but it also is the unique such.
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end))
forall (i : I) (x : X i),
+it_eq RY (g i x) (iter_guarded f EG i x)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end))
forall (i : I) (x : X i),
+it_eq RY (g i x) (iter_guarded f EG i x)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i
it_eq_bt E RY R i (g i x) (iter_guarded f EG i x)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i
it_eq_bt E RY R i (g i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i
it_eqF E RY (it_eq_t E RY R) i
+ match _observe (f i x) with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+ match _observe (f i x) with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i a:= _observe (f i x): itree' E (X +ᵢ Y) i
it_eqF E RY (it_eq_t E RY R) i
+ match a with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+ match a with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i a:= _observe (f i x): itree' E (X +ᵢ Y) i Ha:= EG i x: guarded' a
it_eqF E RY (it_eq_t E RY R) i
+ match a with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+ match a with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i t: itree E (X +ᵢ Y) i
it_eq_t E RY R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< t)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) r: e_rsp e
it_eq_t E RY R (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i t: itree E (X +ᵢ Y) i
it_eq_t E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) r: e_rsp e
it_eq_t E RY R (e_nxt r)
+ (k r >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (k r >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i t: itree E (X +ᵢ Y) i
(it_eq_t E RY ° it_eq_t E RY) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) r: e_rsp e
(it_eq_t E RY ° it_eq_t E RY) R (e_nxt r)
+ (k r >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (k r >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i t: itree E (X +ᵢ Y) i
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_eq_t E RY R i
+ match x1 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ match x2 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) r: e_rsp e
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_eq_t E RY R i
+ match x1 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ match x2 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+all: intros ? ? ? <-; destruct x1; auto.
+Qed.
Furthermore, w.r.t. weak bisimilarity, we compute the same solution as what iter does.
+
I: Type E: event I I X, Y: psh I i: I s, t: itree' E (X +ᵢ Y) i EQ: it_eq' (eqᵢ (X +ᵢ Y)) s t g: guarded' s
guarded' t
+
I: Type E: event I I X, Y: psh I i: I s, t: itree' E (X +ᵢ Y) i EQ: it_eq' (eqᵢ (X +ᵢ Y)) s t g: guarded' s
guarded' t
+destruct EQ as [ [] | | ]; [ dependent elimination g | rewrite <- r_rel | | ]; econstructor.
+Qed.
+
+Definitionguarded_cong {XY} {i} (st : itree E (X +ᵢ Y) i)
+ (EQ : s ≊ t) (g : guarded s) : guarded t
+ := guarded_cong' s.(_observe) t.(_observe) (it_eq_step _ _ _ EQ) g .
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f i: I x: X i
iter_guarded f EG i x ≈ iter f i x
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f i: I x: X i
iter_guarded f EG i x ≈ iter f i x
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i (iter_guarded f EG i x)
+ (iter f i x)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
(eq_clo_map ° it_wbisim_bt E (eqᵢ Y)) R i
+ (iter_guarded f EG i x) (iter f i x)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
iter_guarded f EG i x ≊ ?a
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
?b ≊ iter f i x
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i ?a?b
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
?b ≊ iter f i x
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end)) ?b
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i g: guarded (f i x)
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i g: guarded t
it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i g: guarded' (_observe t)
it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_guarded f EG i x) (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i ot: itree' E (X +ᵢ Y) i Heqot: ot = _observe t g: guarded' ot
it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i y: Y i Heqot: RetF (inr y) = _observe t
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (RetF y) (RetF y)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t
it_wbisim_t E (eqᵢ Y) R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t r: e_rsp e
it_wbisim_t E (eqᵢ Y) R
+ (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t0)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t r: e_rsp e
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R
+ (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_wbisim_t E (eqᵢ Y) R i
+ match x1 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t r: e_rsp e
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_wbisim_t E (eqᵢ Y) R i
+ match x1 with
+ | inl x => iter_guarded f EG i x
+ | inr y => Ret' y
+ end
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t i0: I x0: X i0
it_wbisim_t E (eqᵢ Y) R i0
+ (iter_guarded f EG i0 x0)
+ (Tau' (iter f i0 x0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t r: e_rsp e i0: I x0: X i0
it_wbisim_t E (eqᵢ Y) R i0
+ (iter_guarded f EG i0 x0)
+ (Tau' (iter f i0 x0))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t, t0: itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t i0: I x0: X i0
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R i0
+ (iter_guarded f EG i0 x0)
+ (Tau' (iter f i0 x0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_guarded f EG i x)
+ (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) Heqot: VisF e k = _observe t r: e_rsp e i0: I x0: X i0
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R i0
+ (iter_guarded f EG i0 x0)
+ (Tau' (iter f i0 x0))
I: Type E: event I I R, X: psh I i: I t: itree E (R +ᵢ X) i p, q: guarded t
p = q
+apply guarded'_irrelevant.
+Qed.
+
+Endguarded.
+
+
Eventually guarded iteration
+
For our purpose, proving the soundness of the ogs, guardedness is a bit too
+restrictive: while not all equations are guarded, they all inductively lead
+to a guarded one. We capture this intuition via the notion of "eventually
+guarded" set of equations, and establish similar result to ones in the
+guarded case (these properties are collectively refered to as Prop. 5 in the
+paper).
+
Sectioneventually_guarded.
+
+Context {I} {E : event I I}.
A set of equations is eventually guarded if they all admit an inductive path following
+its non-guarded indirections that leads to a guarded equation (Def. 34).
+
Unset Elimination Schemes.
+Inductiveev_guarded' {XY} (e : eqn Y X X) {i} : itree' E (X +ᵢ Y) i -> Prop :=
+ | GHead t : guarded' t -> ev_guarded' e t
+ | GNext x : ev_guarded' e (e i x).(_observe) -> ev_guarded' e (RetF (inl x))
+ .
+#[global] Arguments GHead {X Y e i t}.
+#[global] Arguments GNext {X Y e i x}.
+
+Schemeev_guarded'_ind := Induction for ev_guarded' SortProp.
+Set Elimination Schemes.
+
+Definitionev_guarded {XY} (e : eqn Y X X) {i} (t : itree E (X +ᵢ Y) i)
+ := ev_guarded' e t.(_observe).
+Definitioneqn_ev_guarded {XY} (e : eqn Y X X) : Type
+ := foralli (x : X i), ev_guarded e (e i x) .
+
+Equationselim_ev_guarded' {XYeix} (g : @ev_guarded' X Y e i (RetF (inl x)))
+ : ev_guarded' e (e i x).(_observe) :=
+ elim_ev_guarded' (GNext g) := g .
Given eventually guarded equations e, we can build a set of guarded equations
+eqn_evg_unroll_guarded e: we simply map all indices to their reachable guarded
+equations. Iteration over eventually guarded equations is hence defined as guarded
+iteration over their guarded counterpart.
+
Fixpointevg_unroll' {XY} (e : eqn Y X X) {i}
+ (t : itree' E (X +ᵢ Y) i) (g : ev_guarded' e t) { structg } : itree' E (X +ᵢ Y) i
+ := match t as t' return ev_guarded' e t' -> _ with
+ | RetF (inl x) => fung => evg_unroll' e (e _ x).(_observe) (elim_ev_guarded' g)
+ | RetF (inr y) => fun_ => RetF (inr y)
+ | TauF t => fun_ => TauF t
+ | VisF q k => fun_ => VisF q k
+ end g .
+
+Definitionevg_unroll {XY} (e : eqn Y X X) {i}
+ (t : itree E (X +ᵢ Y) i) (g : ev_guarded e t) : itree E (X +ᵢ Y) i
+ := go (evg_unroll' e t.(_observe) g) .
+
+Definitioneqn_evg_unroll {XY} (e : eqn Y X X) (H : eqn_ev_guarded e) : eqn Y X X
+ := fun_x => evg_unroll e _ (H _ x) .
+
+
I: Type E: event I I X, Y: psh I e: eqn Y X X i: I t: itree' E (X +ᵢ Y) i g: ev_guarded' e t
guarded' (evg_unroll' e t g)
+
I: Type E: event I I X, Y: psh I e: eqn Y X X i: I t: itree' E (X +ᵢ Y) i g: ev_guarded' e t
guarded' (evg_unroll' e t g)
+
I: Type E: event I I X, Y: psh I e: eqn Y X X i: I t: itree' E (X +ᵢ Y) i g: guarded' t
guarded' (evg_unroll' e t (GHead g))
+
I: Type E: event I I X, Y: psh I e: eqn Y X X i: I x: X i g: guarded' (RetF (inl x))
guarded' (evg_unroll' e (RetF (inl x)) (GHead g))
+dependent elimination g.
+Qed.
+
+Definitionevg_unroll_guarded {XY} (e : eqn Y X X) {i}
+ (t : itree E (X +ᵢ Y) i) (EG : ev_guarded e t) : guarded (evg_unroll e t EG)
+ := evg_unroll_guarded' e t.(_observe) EG.
+
+Definitioneqn_evg_unroll_guarded {XY}
+ (e : eqn Y X X) (EG : eqn_ev_guarded e) : eqn_guarded (eqn_evg_unroll e EG)
+ := fun_x => evg_unroll_guarded e _ (EG _ x) .
Definitioniter_ev_guarded {RX}
+ (e : eqn R X X) (EG : eqn_ev_guarded e) : X ⇒ᵢ itree E R
+ := fun_x => iter_guarded _ (eqn_evg_unroll_guarded e EG) _ x .
Once again, we establish the expected unfolding lemma, for eventually guarded iteration.
+
I: Type E: event I I X, Y: psh I e: eqn Y X X i: I x: X i g: ev_guarded' e (RetF (inl x))
evg_unroll' e (RetF (inl x)) g =
+evg_unroll' e (_observe (e i x)) (elim_ev_guarded' g)
+
I: Type E: event I I X, Y: psh I e: eqn Y X X i: I x: X i g: ev_guarded' e (RetF (inl x))
evg_unroll' e (RetF (inl x)) g =
+evg_unroll' e (_observe (e i x)) (elim_ev_guarded' g)
I: Type E: event I I R, X: psh I e: eqn X R R i: I t: itree' E (R +ᵢ X) i p, q: ev_guarded' e t
p = q
+
I: Type E: event I I R, X: psh I e: eqn X R R i: I t: itree' E (R +ᵢ X) i p, q: ev_guarded' e t
p = q
+
I: Type E: event I I R, X: psh I e: eqn X R R i: I t: itree' E (R +ᵢ X) i g: guarded' t q: ev_guarded' e t
GHead g = q
I: Type E: event I I R, X: psh I e: eqn X R R i: I x: R i p: ev_guarded' e (_observe (e i x)) q: ev_guarded' e (RetF (inl x)) IHp: forallq : ev_guarded' e (_observe (e i x)), p = q
GNext p = q
+
I: Type E: event I I R, X: psh I e: eqn X R R i: I t: itree' E (R +ᵢ X) i g: guarded' t q: ev_guarded' e t
GHead g = q
destruct t as [ [] | | ]; [ dependent elimination g | | | ];
+ dependent elimination q; f_equal; apply guarded'_irrelevant.
+
I: Type E: event I I R, X: psh I e: eqn X R R i: I x: R i p: ev_guarded' e (_observe (e i x)) q: ev_guarded' e (RetF (inl x)) IHp: forallq : ev_guarded' e (_observe (e i x)), p = q
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2)
it_eq RY (iter_ev_guarded f EG i x1)
+ (iter_ev_guarded f EG i x2)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2)
it_eq RY (iter_ev_guarded f EG i x1)
+ (iter_ev_guarded f EG i x2)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2)
it_eq RY
+ (iter_guarded (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i x1)
+ (iter_ev_guarded f EG i x2)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2)
it_eq RY
+ (eqn_evg_unroll f EG i x1 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x =>
+ iter_guarded (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i x
+ | inr y => Ret' y
+ end)) (iter_ev_guarded f EG i x2)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2)
it_eq RY
+ (eqn_evg_unroll f EG i x1 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) (iter_ev_guarded f EG i x2)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2) q: ev_guarded f (f i x1)
it_eq RY
+ (evg_unroll f (f i x1) q >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) (iter_ev_guarded f EG i x2)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i p: _observe (f i x1) = RetF (inl x2) q: ev_guarded f (f i x1)
it_eqF E RY (it_eq RY) i
+ match evg_unroll' f (_observe (f i x1)) q with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end (observe (iter_ev_guarded f EG i x2))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i q: ev_guarded' f (RetF (inl x2))
it_eqF E RY (it_eq RY) i
+ match evg_unroll' f (RetF (inl x2)) q with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end (observe (iter_ev_guarded f EG i x2))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i
it_eqF E RY (it_eq RY) i
+ match
+ evg_unroll' f (_observe (f i x2)) (EG i x2)
+ with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end (observe (iter_ev_guarded f EG i x2))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x1, x2: X i
it_eq_map E RY (it_eq RY) i
+ (evg_unroll f (f i x2) (EG i x2) >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i0 x
+ | inr y => Ret' y
+ end)) (iter_ev_guarded f EG i x2)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i
it_eq RY (iter_ev_guarded f EG i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i
it_eq RY (iter_ev_guarded f EG i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ match _observe (f i x) with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i g: ev_guarded' f (_observe (f i x))
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ match _observe (f i x) with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i i0: itree' E (X +ᵢ Y) i Heqi0: i0 = _observe (f i x) g: ev_guarded' f i0
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ match i0 with
+ | RetF r =>
+ _observe
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x, x0: X i Heqi0: RetF (inl x0) = _observe (f i x) g: ev_guarded' f (RetF (inl x0))
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (_observe (iter_ev_guarded f EG i x0))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i y: Y i Heqi0: RetF (inr y) = _observe (f i x) g: ev_guarded' f (RetF (inr y))
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (_observe (Ret' y))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i t: itree E (X +ᵢ Y) i Heqi0: TauF t = _observe (f i x) g: ev_guarded' f (TauF t)
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) Heqi0: VisF q k = _observe (f i x) g: ev_guarded' f (VisF q k)
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i y: Y i Heqi0: RetF (inr y) = _observe (f i x) g: ev_guarded' f (RetF (inr y))
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (_observe (Ret' y))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i t: itree E (X +ᵢ Y) i Heqi0: TauF t = _observe (f i x) g: ev_guarded' f (TauF t)
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) Heqi0: VisF q k = _observe (f i x) g: ev_guarded' f (VisF q k)
it_eqF E RY (it_eq RY) i
+ (observe (iter_ev_guarded f EG i x))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i y: Y i Heqi0: RetF (inr y) = _observe (f i x)
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (_observe (f i x)) (EG i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end
+ (evg_unroll_guarded' f
+ (_observe (f i x))
+ (EG i x))) (RetF y)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i t: itree E (X +ᵢ Y) i Heqi0: TauF t = _observe (f i x)
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (_observe (f i x)) (EG i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end
+ (evg_unroll_guarded' f
+ (_observe (f i x))
+ (EG i x)))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) Heqi0: VisF q k = _observe (f i x)
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (_observe (f i x)) (EG i x) as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end
+ (evg_unroll_guarded' f
+ (_observe (f i x))
+ (EG i x)))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i y: Y i Heqi0: RetF (inr y) = _observe (f i x) g': ev_guarded' f (_observe (f i x))
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (_observe (f i x)) g' as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end (evg_unroll_guarded' f (_observe (f i x)) g'))
+ (RetF y)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i t: itree E (X +ᵢ Y) i Heqi0: TauF t = _observe (f i x) g': ev_guarded' f (_observe (f i x))
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (_observe (f i x)) g' as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end (evg_unroll_guarded' f (_observe (f i x)) g'))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) Heqi0: VisF q k = _observe (f i x) g': ev_guarded' f (_observe (f i x))
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (_observe (f i x)) g' as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end (evg_unroll_guarded' f (_observe (f i x)) g'))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i y: Y i Heqi0: RetF (inr y) = _observe (f i x) g': ev_guarded' f (RetF (inr y))
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (RetF (inr y)) g' as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end (evg_unroll_guarded' f (RetF (inr y)) g'))
+ (RetF y)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i t: itree E (X +ᵢ Y) i Heqi0: TauF t = _observe (f i x) g': ev_guarded' f (TauF t)
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (TauF t) g' as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end (evg_unroll_guarded' f (TauF t) g'))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) Heqi0: VisF q k = _observe (f i x) g': ev_guarded' f (VisF q k)
it_eqF E RY (it_eq RY) i
+ (match
+ evg_unroll' f (VisF q k) g' as t
+ return (guarded' t -> itree' E Y i)
+ with
+ | RetF r0 =>
+ match
+ r0 as r1
+ return (guarded' (RetF r1) -> itree' E Y i)
+ with
+ | inl x => ex_falso ∘ elim_guarded
+ | inr r =>
+ fun_ : guarded' (RetF (inr r)) => RetF r
+ end
+ | TauF t =>
+ fun_ : guarded' (TauF t) =>
+ TauF
+ (iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ | VisF q k =>
+ fun_ : guarded' (VisF q k) =>
+ VisF q
+ (funr : e_rsp q =>
+ iter_guarded_aux
+ (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ end (evg_unroll_guarded' f (VisF q k) g'))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i t: itree E (X +ᵢ Y) i Heqi0: TauF t = _observe (f i x) g: guarded' (TauF t)
it_eq RY
+ (iter_guarded_aux (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG) i t)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i q: e_qry i k: forallr : e_rsp q, itree E (X +ᵢ Y) (e_nxt r) Heqi0: VisF q k = _observe (f i x) g: guarded' (VisF q k) r: e_rsp q
it_eq RY
+ (iter_guarded_aux (eqn_evg_unroll f EG)
+ (eqn_evg_unroll_guarded f EG)
+ (e_nxt r) (k r))
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+all: apply iter_guarded_aux_unfold.
+Qed.
Uniqueness still holds.
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end))
forall (i : I) (x : X i),
+it_eq RY (g i x) (iter_ev_guarded f EG i x)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end))
forall (i : I) (x : X i),
+it_eq RY (g i x) (iter_ev_guarded f EG i x)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_ev_guarded f EG i x) i: I x: X i
it_eq_bt E RY R i (g i x) (iter_ev_guarded f EG i x)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_ev_guarded f EG i x) i: I x: X i
it_eq_bt E RY R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_ev_guarded f EG i x) i: I x: X i Ha: ev_guarded' f (_observe (f i x))
it_eq_bt E RY R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_ev_guarded f EG i x) i: I x: X i t: itree E (X +ᵢ Y) i Ha: ev_guarded' f (_observe t)
it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i), it_eq_t E RY R i (g i x) (iter_ev_guarded f EG i x) i: I x: X i t: itree E (X +ᵢ Y) i ot: itree' E (X +ᵢ Y) i Heqot: ot = _observe t Ha: ev_guarded' f ot
it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i t: itree' E (X +ᵢ Y) i g0: guarded' t t': itree E (X +ᵢ Y) i Heqot: t = _observe t'
it_eq_bt E RY R i
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x, x0: X i Ha: ev_guarded' f (_observe (f i x0)) IHHa: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) t': itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe t'
it_eq_bt E RY R i
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i t: itree' E (X +ᵢ Y) i g0: guarded' t t': itree E (X +ᵢ Y) i Heqot: t = _observe t'
it_eq_bt E RY R i
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i t0, t': itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t'
it_eq_t E RY R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< t0)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) t': itree E (X +ᵢ Y) i Heqot: VisF e k = _observe t' r: e_rsp e
it_eq_t E RY R (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i t0, t': itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t'
(it_eq_t E RY ° it_eq_t E RY) R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< t0)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t0)
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) t': itree E (X +ᵢ Y) i Heqot: VisF e k = _observe t' r: e_rsp e
(it_eq_t E RY ° it_eq_t E RY) R
+ (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i t0, t': itree E (X +ᵢ Y) i Heqot: TauF t0 = _observe t'
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_eq_t E RY R i
+ match x1 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ match x2 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) t': itree E (X +ᵢ Y) i Heqot: VisF e k = _observe t' r: e_rsp e
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_eq_t E RY R i
+ match x1 with
+ | inl x => g i x
+ | inr y => Ret' y
+ end
+ match x2 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+all: intros ? ? x2 ->; destruct x2; auto.
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x, x0: X i Ha: ev_guarded' f (_observe (f i x0)) IHHa: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) t': itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe t'
it_eq_bt E RY R i
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t' >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x, x0: X i Ha: ev_guarded' f (_observe (f i x0)) IHHa: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) t': itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe t'
it_eqF E RY (it_eq_t E RY R) i
+ (_observe (g i x0))
+ (_observe (iter_ev_guarded f EG i x0))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x, x0: X i Ha: ev_guarded' f (_observe (f i x0)) IHHa: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) t': itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe t'
it_eq_map E RY (it_eq_t E RY R) i
+ (g i x0) (iter_ev_guarded f EG i x0)
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) H: Equivalenceᵢ RY f: eqn Y X X g: X ⇒ᵢ itree E Y EG: eqn_ev_guarded f EQ: forall (i : I) (x : X i),
+it_eq RY (g i x)
+ (f i x >>=
+ (fun (i0 : I) (r : (X +ᵢ Y) i0) =>
+ match r with
+ | inl x0 => g i0 x0
+ | inr y => Ret' y
+ end)) R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_eq_t E RY R i (g i x)
+ (iter_ev_guarded f EG i x) i: I x, x0: X i Ha: ev_guarded' f (_observe (f i x0)) IHHa: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_eq_bt E RY R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) t': itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe t'
it_eq_map E RY (it_eq_t E RY R) i
+ (f i x0 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => g i x
+ | inr y => Ret' y
+ end))
+ (f i x0 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+apply IHHa; auto.
+Qed.
And w.r.t. weak bisimilarity, we still perform the same computation.
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i
iter_ev_guarded f EG i x ≈ iter f i x
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f i: I x: X i
iter_ev_guarded f EG i x ≈ iter f i x
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i (iter_ev_guarded f EG i x)
+ (iter f i x)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
(eq_clo_map ° it_wbisim_bt E (eqᵢ Y)) R i
+ (iter_ev_guarded f EG i x) (iter f i x)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
iter_ev_guarded f EG i x ≊ ?a
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
?b ≊ iter f i x
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i ?a?b
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
?b ≊ iter f i x
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) ?b
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i g: ev_guarded f (f i x)
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i g: ev_guarded f t
it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i g: ev_guarded' f (_observe t)
it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t: itree E (X +ᵢ Y) i ot: itree' E (X +ᵢ Y) i Heqot: ot = _observe t g: ev_guarded' f ot
it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t: itree' E (X +ᵢ Y) i g: guarded' t u: itree E (X +ᵢ Y) i Heqot: t = _observe u
it_wbisim_bt E (eqᵢ Y) R i
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t: itree' E (X +ᵢ Y) i g: guarded' t u: itree E (X +ᵢ Y) i Heqot: t = _observe u
it_wbisim_bt E (eqᵢ Y) R i
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i y: Y i u: itree E (X +ᵢ Y) i Heqot: RetF (inr y) = _observe u
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (RetF y) (RetF y)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t))
+ (TauF
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ match r0 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r))
+ (VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u
it_wbisim_t E (eqᵢ Y) R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u r: e_rsp e
it_wbisim_t E (eqᵢ Y) R
+ (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R i
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< t)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u r: e_rsp e
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R
+ (e_nxt r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end) =<< k r)
+ ((fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_wbisim_t E (eqᵢ Y) R i
+ match x1 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u r: e_rsp e
forall (i : I) (x1x2 : (X +ᵢ Y) i),
+eqᵢ (X +ᵢ Y) i x1 x2 ->
+it_wbisim_t E (eqᵢ Y) R i
+ match x1 with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u i0: I x0: X i0
it_wbisim_t E (eqᵢ Y) R i0
+ (iter_ev_guarded f EG i0 x0) (Tau' (iter f i0 x0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u r: e_rsp e i0: I x0: X i0
it_wbisim_t E (eqᵢ Y) R i0
+ (iter_ev_guarded f EG i0 x0)
+ (Tau' (iter f i0 x0))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x: X i t, u: itree E (X +ᵢ Y) i Heqot: TauF t = _observe u i0: I x0: X i0
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R i0
+ (iter_ev_guarded f EG i0 x0) (Tau' (iter f i0 x0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x: X i e: e_qry i k: forallr : e_rsp e, itree E (X +ᵢ Y) (e_nxt r) u: itree E (X +ᵢ Y) i Heqot: VisF e k = _observe u r: e_rsp e i0: I x0: X i0
(it_wbisim_t E (eqᵢ Y) ° it_wbisim_t E (eqᵢ Y)) R i0
+ (iter_ev_guarded f EG i0 x0)
+ (Tau' (iter f i0 x0))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (u >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisimF E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) R) i
+ (_observe (iter_ev_guarded f EG i x0))
+ (TauF (iter f i x0))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i (iter_ev_guarded f EG i x0)
+ (Tau' (iter f i x0))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
?f <= it_wbisim_t E (eqᵢ Y)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
(?f ° it_wbisim_bt E (eqᵢ Y)) R i
+ (iter_ev_guarded f EG i x0)
+ (Tau' (iter f i x0))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
(eat_clo_map ° it_wbisim_bt E (eqᵢ Y)) R i
+ (iter_ev_guarded f EG i x0) (Tau' (iter f i x0))
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i (iter_ev_guarded f EG i x0)
+ (iter f i x0)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
?f <= it_wbisim_t E (eqᵢ Y)
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
(?f ° it_wbisim_bt E (eqᵢ Y)) R i
+ (iter_ev_guarded f EG i x0)
+ (iter f i x0)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
(eq_clo_map ° it_wbisim_bt E (eqᵢ Y)) R i
+ (iter_ev_guarded f EG i x0) (iter f i x0)
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
iter_ev_guarded f EG i x0 ≊ ?a
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
?b ≊ iter f i x0
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i ?a?b
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
?b ≊ iter f i x0
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i
+ (iter_ev_guarded f EG i x)
+ (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (t >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x0 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end)) ?b
+
I: Type E: event I I X, Y: psh I f: eqn Y X X EG: eqn_ev_guarded f R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (x : X i),
+it_wbisim_t E (eqᵢ Y) R i (iter_ev_guarded f EG i x) (iter f i x) i: I x, x0: X i g: ev_guarded' f (_observe (f i x0)) IHg: forallt : itree E (X +ᵢ Y) i,
+_observe (f i x0) = _observe t ->
+it_wbisim_bt E (eqᵢ Y) R i
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+match r with
+| inl x => iter_ev_guarded f EG i x
+| inr y => Ret' y
+end))
+(t >>=
+(fun (i : I) (r : (X +ᵢ Y) i) =>
+{|
+_observe :=
+match r with
+| inl x => TauF (iter f i x)
+| inr y => RetF y
+end
+|})) u: itree E (X +ᵢ Y) i Heqot: RetF (inl x0) = _observe u
it_wbisim_bt E (eqᵢ Y) R i
+ (f i x0 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter_ev_guarded f EG i x
+ | inr y => Ret' y
+ end))
+ (f i x0 >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Interaction Trees: Definition
+
+
We implement a subset of the Interaction Tree library in an indexed setting.
+These indices provide dynamic control over the set of available external events
+during the computation. In particular, the interface is composed of some event I I,
+i.e.:
+
+
a family of available query at each index;
+
a domain of answers associated to each query;
+
a transition function assigning the new index associated with the answer
+
+
+
Indexed Interaction Trees
+
Sectionitree.
+Context {I : Type}.
+Context (E : event I I).
+Context (R : psh I).
+
+VariantitreeF (REC : psh I) (i : I) : Type :=
+ | RetF (r : R i) : itreeF REC i
+ | TauF (t : REC i) : itreeF REC i
+ | VisF (q : E.(e_qry) i) (k : forall (r : E.(e_rsp) q), REC (E.(e_nxt) r)) : itreeF REC i
+ .
+DeriveNoConfusionfor itreeF.
+
+CoInductiveitree (i : I) : Type := go { _observe : itreeF itree i }.
+
+Enditree.
Notationitree' E R := (itreeF E R (itree E R)).
The following function defines the coalgebra structure.
+
Definitionobserve {IERi} (t : @itree I E R i) : itree' E R i := t.(_observe).
+
+NotationRet' x := (go (RetF x)).
+NotationTau' t := (go (TauF t)).
+NotationVis' e k := (go (VisF e k)).
+
+
+
diff --git a/docs/Machine.html b/docs/Machine.html
new file mode 100644
index 0000000..b84cfdd
--- /dev/null
+++ b/docs/Machine.html
@@ -0,0 +1,81 @@
+
+
+
+
+
+
+Evaluation structure and language machine
+
+
+
+
+
+
+
+
+
+
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Evaluation structure and language machine
+
+
In this file we pull everything together and define evaluation structures and
+the abstract characterization of a language machine, together with all the associated
+laws.
+
Note that we have here a minor mismatch between the formalization and the paper:
+we have realized a posteriori a more elegant way to decompose the abstract
+machine as exposed in the paper.
+
Here, instead of only requiring an embedding of normal forms into
+configurations (refold, from Def. 11), we require an observation
+application function oapp describing how to build a configuration
+from a value, an observation, and an assignment. Since normal forms are triples
+of a variable, an observation and an assignment, the sole difference is in the
+first component: instead of just a variable, oapp takes any value. Both
+are easily interdefineable. While refold is more compact to introduce,
+it is slightly more cumbersome to work with, which is why we axiomatize oapp
+directly.
+
Patching the development to follow more closely the paper's presentation would be slightly
+technical due to the technicality of the mechanized proofs involved, but essentially
+straightforward.
+
We here define what is called an evaluation structure in the paper (Def. 11). It tells
+us how to evaluate a configuration to a normal form and how to stitch one back together
+from the data sent by Opponent.
+
Classmachine (val : Fam₁ T C) (conf : Fam₀ T C) (obs : obs_struct T C) := {
+ eval : conf ⇒₀ (delay ∘ nf obs val) ;
+ oapp [Γ x] (v : val Γ x) (o : obs x) : dom o =[val]> Γ -> conf Γ ;
+}.
Next we define "evaluating then observing" (Def. 14).
oapp commutes with substitutions. This is the equivalent of the second
+equation at the end of Def. 13, in terms of oapp instead of refold.
+
app_sub {Γ1 Γ2 x} (v : val Γ1 x) (o : obs x) (a : dom o =[val]> Γ1) (b : Γ1 =[val]> Γ2)
+ : (v ⊙ o⦗a⦘) ₜ⊛ b = (v ᵥ⊛ b) ⊙ o⦗a ⊛ b⦘ ;
"evaluation respects substitution". This is the core hypothesis of the OGS
+soundness, stating essentially "substituting, then evaluating" is equivalent to
+"evaluating, then substituting, then evaluating once more". It is the equvalent
+of the first equation at the end of Def. 13.
+
eval_sub {Γ Δ} (c : conf Γ) (a : Γ =[val]> Δ)
+ : eval (c ₜ⊛ a) ≋ bind_delay' (eval c) (funn => eval (emb n ₜ⊛ a)) ;
Evaluating the embedding of a normal form is equivalent to returning the normal
+form. This is part of the evaluation structure (Def. 11) in the paper.
+
eval_nf_ret {Γ} (u : nf obs val Γ)
+ : eval (emb u) ≋ ret_delay u ;
At last the mystery hypothesis, stating that the machine has focused redexes (Def. 28).
+It is necessary for establishing that the composition can be defined by eventually
+guarded iteration.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
+
+
From OGS Require Import Prelude.
+From OGS.Utils Require Import Rel.
+From OGS.Ctx Require Import All Ctx Covering Subst.
+From OGS.ITree Require Import Event ITree Eq Delay.
+From OGS.OGS Require Import Obs Game Machine.
+
+Sectionwith_param.
+ Context `{CC : context T C} {CL : context_laws T C} (obs : obs_struct T C).
+ Context {val} {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+ Context {conf} {CM : subst_module val conf} {CML : subst_module_laws val conf}.
+ Context (M : machine val conf obs).
+
+ (* NF bisimulation game *)
+ Definitionnfb_g : game C (C × C)
+ := {| g_client := {| g_move Γ := obs∙ Γ ;
+ g_next Γ o := (Γ , dom o.(cut_r)) |} ;
+ g_server := {| g_move '(Γ , Δ) := obs∙ Δ ;
+ g_next '(Γ , Δ) o := Γ +▶ dom o.(cut_r) |} |}.
+
+ Definitionnfb_e : event C C
+ := e_of_g nfb_g.
+
+ (* active NF bisimulation *)
+ Definitionnfb_act := itree nfb_e ∅ᵢ.
+ (* single-var passive NF bisimulation *)
+ Definitionnfb_pas₁ (Γ : C) (x : T) := forallo : obs x, nfb_act (Γ +▶ dom o) .
+ (* full passive NF bisimulation *)
+ Definitionnfb_pas '(Γ , Δ) := Δ =[nfb_pas₁]> Γ .
+
+ (* active NF bisimulation with cut-off *)
+ Definitionnfb_finΔΓ := itree nfb_e (fun_ => obs∙ Δ) Γ.
+ (* single-var passive NF bisimulation with cut-off *)
+ Definitionnfb_fin_pas₁ (ΔΓ : C) (x : T) := forallo : obs x, nfb_fin Δ (Γ +▶ dom o) .
+
+ (* renaming active NF bisim *)
+ Definitionnfb_ren : forallΓ1Γ2, Γ1 ⊆ Γ2 -> nfb_act Γ1 -> nfb_act Γ2 :=
+ cofix _nfb_ren _ _ ρ t :=
+ go match t.(_observe) with
+ | RetF r => match r withend
+ | TauF t => TauF (_nfb_ren _ _ ρ t)
+ | VisF q k =>
+ VisF (ρ _ q.(cut_l) ⋅ q.(cut_r) : nfb_e.(e_qry) _)
+ (funr => _nfb_ren _ _ (r_shift (m_dom r) ρ) (k r))
+ end .
+
+ Definitionnfb_var : c_var ⇒₁ nfb_pas₁
+ := cofix _nfb_var Γ x i o :=
+ Vis' (r_cat_l i ⋅ o : nfb_e.(e_qry) _)
+ (funr => _nfb_var _ _ (r_cat_r r.(cut_l)) r.(cut_r)) .
+
+ (* renaming active NF bisim with cut-off *)
+ Program Definitionnfb_fin_ren {Δ} : forallΓ1Γ2, Γ1 ⊆ Γ2 -> nfb_fin Δ Γ1 -> nfb_fin Δ Γ2 :=
+ cofix _nfb_fin_ren _ _ ρ t :=
+ go match t.(_observe) with
+ | RetF r => RetF r
+ | TauF t => TauF (_nfb_fin_ren _ _ ρ t)
+ | VisF q k =>
+ VisF (ρ _ q.(cut_l) ⋅ q.(cut_r) : nfb_e.(e_qry) _)
+ (funr => _nfb_fin_ren _ _ (r_shift (m_dom r) ρ) (k r))
+ end .
+
+ (* tail-cutting of NF bisim *)
+ Program Definitionnfb_stop : forallΔΓ, nfb_act (Δ +▶ Γ) -> nfb_fin Δ Γ :=
+ cofix _nfb_stop Δ _ t :=
+ go match t.(_observe) with
+ | RetF r => match r withend
+ | TauF t => TauF (_nfb_stop _ _ t)
+ | VisF q k => ltac:(exact
+ match c_view_cat q.(cut_l) with
+ | Vcat_l i => RetF (i ⋅ q.(cut_r))
+ | Vcat_r j => VisF (j ⋅ q.(cut_r) : nfb_e.(e_qry) _)
+ (funr => _nfb_stop _ _ (nfb_ren _ _ r_assoc_r (k r)))
+ end)
+ end.
+
+ (* embed active NF-bisim with cut-off to active OGS strategy (generalized) *)
+ Definitionnfb_to_ogs_aux {Δ}
+ : forallΦ (ks : ↓⁻Φ =[nfb_fin_pas₁ Δ]> ↓⁺Φ), nfb_fin Δ ↓⁺Φ -> ogs_act (obs:=obs) Δ Φ.
+ cofix _nfb_to_ogs; intros Φ ks t.
+ apply go; destruct (t.(_observe)).
+ + apply RetF; exact r.
+ + apply TauF; exact (_nfb_to_ogs _ ks t0).
+ + unshelveeapply VisF; [ exact q | ].
+ intro r.
+ pose (ks' := [ ks , fun_jo => k (j ⋅ o) ]%asgn).
+ refine (_nfb_to_ogs (Φ ▶ₓ _ ▶ₓ _)%ctx _ (ks' _ r.(cut_l) r.(cut_r))).
+ intros ? i o.
+ refine (nfb_fin_ren _ _ [ r_cat_l ᵣ⊛ r_cat_l , r_cat_r ]%asgn (ks' _ i o)).
+ Defined.
+
+ (* embed active NF-bisim with cut-off to active OGS strategy (initial state) *)
+ Definitionnfb_to_ogs {ΔΓ} : nfb_fin Δ Γ -> ogs_act Δ (∅ ▶ₓ Γ)
+ := funu => nfb_to_ogs_aux (∅ ▶ₓ Γ) ! (nfb_fin_ren _ _ r_cat_r u) .
+
+ Definitionapp_no_arg {Γx} (v : val Γ x) (o : obs x)
+ : conf (Γ +▶ dom o)
+ := v ᵥ⊛ᵣ r_cat_l ⊙ o ⦗ r_emb r_cat_r ⦘.
+
+ (* language machine to NF-bisim *)
+ Definitionm_nfb_act : conf ⇒₀ nfb_act
+ := cofix _m_nfb Γ c
+ := subst_delay
+ (funn => Vis' (nf_to_obs _ n : nfb_e.(e_qry) _)
+ (funo => _m_nfb _ (app_no_arg (nf_args n _ o.(cut_l)) o.(cut_r))))
+ (eval c).
+
+ Definitionm_nfb_pas : val ⇒₁ nfb_pas₁
+ := fun__vo => m_nfb_act _ (app_no_arg v o).
+Endwith_param.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Observation Structure (§ 4.4)
+
+
The messages that player and opponent exchange in the OGS
+arise from splitting normal forms into the head variable
+on which it is stuck, an observation, and an assignment.
+An observation structure axiomatizes theses observations
+of the language. They consist of:
+
+
a family o_obs indexed by the types of the language and,
+
a projection from observations to contexts of the language.
+
+
In fact as these make sense quite generally, we have already defined
+them along with generic combinators on sorted families in
+Ctx/Family.v. Much of the content of this file will
+be to wrap these generic definitions with more suitable names and
+provide specific utilities on top.
+
As explained above, observation structures are just operators. We
+rename their domain o_dom to just dom.
+
#[global] Notationobs_struct T C := (Oper T C).
+#[global] Notationdom := o_dom.
+
System Message: WARNING/2 (theories/OGS/Obs.v, line 3); backlink
+Duplicate explicit target name: "ctx/family.v".
+
Pointed observations consist of a pair of a variable and a matching observation.
+We construct them using the generic formal cut combinator on families defined
+in Ctx/Family.v.
+
Definitionpointed_obs `{CC : context T C} (O : Oper T C) : Fam₀ T C
+ := c_var ∥ₛ ⦉O⦊.
+#[global] Notation"O ∙" := (pointed_obs O).
+#[global] Notationm_var o := (o.(cut_l)).
+#[global] Notationm_obs o := (o.(cut_r)).
+#[global] Notationm_dom o := (o_dom o.(cut_r)).
Next we define normal forms, as triplets of a variable, an observation
+and an assignment filling the domain of the observation. For this we again
+use the formal cut combinator and the "observation filling" combinator,
+defined in Ctx/Assignment.v.
+
Definitionnf `{CC : context T C} (O : obs_struct T C) (X : Fam₁ T C) : Fam₀ T C
+ := c_var ∥ₛ (O # X).
We now define two utilities for projecting normal forms to pointed observations.
+
Definitionnf_to_obs `{CC : context T C} {O} {X : Fam₁ T C} : forallΓ, nf O X Γ -> O∙ Γ
+ := f_cut_map f_id₁ forget_args.
+#[global] Coercionnf_to_obs : nf >-> pointed_obs.
+
+Definitionthen_to_obs `{CC : context T C} {O} {X : Fam₁ T C} {Γ}
+ : delay (nf O X Γ) -> delay (O∙ Γ)
+ := fmap_delay (nf_to_obs Γ).
Next, we define shortcuts for projecting normal forms into their various components.
+
Sectionwith_param.
+Context `{CC : context T C} {O : obs_struct T C} {X : Fam₁ T C}.
+
+ Definitionnf_ty {Γ} (n : nf O X Γ) : T
+ := n.(cut_ty).
+ Definitionnf_var {Γ} (n : nf O X Γ) : Γ ∋ nf_ty n
+ := n.(cut_l).
+ Definitionnf_obs {Γ} (n : nf O X Γ) : O (nf_ty n)
+ := n.(cut_r).(fill_op).
+ Definitionnf_dom {Γ} (n : nf O X Γ) : C
+ := dom n.(cut_r).(fill_op).
+ Definitionnf_args {Γ} (n : nf O X Γ) : nf_dom n =[X]> Γ
+ := n.(cut_r).(fill_args).
As normal forms contain an assignment, Coq's equality is not the right notion of
+equivalence: we wish to consider two normal forms as equivalent even if their assignment
+are merely pointwise equal.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Coq Prelude
+
+
This file contains a bunch of setup, global options and imports.
+
Our notations have evolved quite organically. It would need refactoring, for now we
+just disable warnings related to it.
+
#[global] Set Warnings"-notation-overridden".
+#[global] Set Warnings"-parsing".
Enable negative records with definitional eta-equivalence. A requirement for well-behaved
+coinductive types.
+
#[global] Set Primitive Projections.
We use a lot of implicit variable declaration around typeclasses.
+
#[global] Generalizable All Variables.
Import the Equations library by Matthieu Sozeau, for Agda-like dependent pattern-matching.
We import basic definition for the use of the universe of strict proposition SProp.
+
Require Export Coq.Logic.StrictProp.
A bunch of notations and definitions for basic datatypes.
+
#[global] Notation"f ∘ g" := (funx => f (g x))
+ (at level40, left associativity) : function_scope.
+
+#[global] Notation"a ,' b" := (existT _ a b) (at level30).
+
+(*
+Definition uncurry {A B} {C : A -> B -> Type} (f : forall a b, C a b) (i : A * B)
+ := f (fst i) (snd i).
+
+Definition curry {A B} {C : A -> B -> Type} (f : forall i, C (fst i) (snd i)) a b
+ := f (a , b).
+
+#[global] Notation curry' := (fun f a b => f (a ,' b)).
+#[global] Notation uncurry' := (fun f x => f (projT1 x) (projT2 x)).
+ *)
+
+#[global] Notation"A × B" := (prod A%type B%type) (at level20).
Finite types.
+
VariantT0 : Type := .
+DeriveNoConfusionfor T0.
+
+VariantT1 : Type := T1_0.
+DeriveNoConfusionfor T1.
Absurdity elimination.
+
Definitionex_falso {A : Type} (bot : T0) : A := match bot withend.
Decidable predicates.
+
#[global] Notation"¬ P" := (P -> T0) (at level5).
+Variantdecidable (X : Type) : Type :=
+| Yes : X -> decidable X
+| No : ¬X -> decidable X
+.
+DeriveNoConfusion NoConfusionHom for decidable.
Subset types based on strict propositions.
+
RecordsigS {X : Type} (P : X -> SProp) := { sub_elt : X ; sub_prf : P sub_elt }.
+Arguments sub_elt {X P}.
+Arguments sub_prf {X P}.
+
+
X: Type P: X -> SProp a, b: sigS P H: sub_elt a = sub_elt b
a = b
+
X: Type P: X -> SProp e: X p: P e b: sigS P H: e = sub_elt b
{| sub_elt := e; sub_prf := p |} = b
+
X: Type P: X -> SProp e: X b: sigS P H: e = sub_elt b p: P (sub_elt b)
{| sub_elt := sub_elt b; sub_prf := p |} = b
+change p with b.(sub_prf); reflexivity.
+Qed.
Misc.
+
DefinitionsubstS {X : SProp} (P : X -> Type) (ab : X) : P a -> P b := funp => p .
+
+
A, B, C: Type x: A y: B f: A -> B -> C
f_equal2 f eq_refl eq_refl = eq_refl
+apply YesUIP.
+Qed.
+
+Definitioninjective {AB : Type} (f : A -> B) := forallxy : A, f x = f y -> x = y .
+
+
+
diff --git a/docs/Properties.html b/docs/Properties.html
new file mode 100644
index 0000000..248b783
--- /dev/null
+++ b/docs/Properties.html
@@ -0,0 +1,3016 @@
+
+
+
+
+
+
+Interaction Trees: Theory of the structure
+
+
+
+
+
+
+
+
+
+
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Interaction Trees: Theory of the structure
+
+
We collect in these file the main lemmas capturing the applicative,
+monadic, and (unguarded) iteration structure of itrees.
+
SectionwithE.
+Context {I} {E : event I I}.
+
+
I: Type E: event I I X: psh I i: I t: itree E X i
Tau' t ≈ t
+
I: Type E: event I I X: psh I i: I t: itree E X i
Tau' t ≈ t
+
I: Type E: event I I X: psh I i: I t: itree E X i
it_eat i (observe (Tau' t)) ?x1
I: Type E: event I I X: psh I i: I t: itree E X i
it_eat i (observe t) ?x2
I: Type E: event I I X: psh I i: I t: itree E X i
it_eqF E (eqᵢ X) (gfp (it_wbisim_map E (eqᵢ X))) i
+ ?x1?x2
+
I: Type E: event I I X: psh I i: I t: itree E X i
it_eat i (observe (Tau' t)) ?x1
apply EatStep, EatRefl.
+
I: Type E: event I I X: psh I i: I t: itree E X i
it_eat i (observe t) ?x2
apply EatRefl.
+
I: Type E: event I I X: psh I i: I t: itree E X i
it_eqF E (eqᵢ X) (gfp (it_wbisim_map E (eqᵢ X))) i
+ (observe t) (observe t)
destruct (observe t); noweconstructor.
+Qed.
fmap respects strong bisimilarity.
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x)
Proper
+ (dpointwise_relation (funi : I => RX i ==> RY i) ==>
+ dpointwise_relation
+ (funi : I => it_eq RX (i:=i) ==> it_eq RY (i:=i)))
+ fmap
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x)
Proper
+ (dpointwise_relation (funi : I => RX i ==> RY i) ==>
+ dpointwise_relation
+ (funi : I => it_eq RX (i:=i) ==> it_eq RY (i:=i)))
+ fmap
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g
dpointwise_relation
+ (funi : I =>
+ (it_eq RX (i:=i) ==> it_eq RY (i:=i))%signature)
+ (fmap f) (fmap g)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g
forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+gfp (it_eq_map E RY) a (f <$> x) (g <$> y)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y -> it_eq_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i h: gfp (it_eq_map E RX) i x y
it_eq_bt E RY R i (f <$> x) (g <$> y)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y -> it_eq_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i h: it_eq_map E RX (gfp (it_eq_map E RX)) i x y
it_eq_bt E RY R i (f <$> x) (g <$> y)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y -> it_eq_t E RY R a (f <$> x) (g <$> y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2
it_eqF E RY (it_eq_t E RY R) i (RetF (f i r1))
+ (RetF (g i r2))
+econstructor; nowapply H1.
+Qed.
fmap respects weak bisimilarity.
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x)
Proper
+ (dpointwise_relation (funi : I => RX i ==> RY i) ==>
+ dpointwise_relation
+ (funi : I => it_wbisim RX i ==> it_wbisim RY i))
+ fmap
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x)
Proper
+ (dpointwise_relation (funi : I => RX i ==> RY i) ==>
+ dpointwise_relation
+ (funi : I => it_wbisim RX i ==> it_wbisim RY i))
+ fmap
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g
dpointwise_relation
+ (funi : I =>
+ (it_wbisim RX i ==> it_wbisim RY i)%signature)
+ (fmap f) (fmap g)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g
forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+gfp (it_wbisim_map E RY) a (f <$> x) (g <$> y)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y -> it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i h: gfp (it_wbisim_map E RX) i x y
it_wbisim_bt E RY R i (f <$> x) (g <$> y)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: dpointwise_relation
+ (funi : I => (RX i ==> RY i)%signature) f g R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y -> it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i h: it_wbisim_map E RX (gfp (it_wbisim_map E RX)) i x
+ y
it_wbisim_bt E RY R i (f <$> x) (g <$> y)
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y -> it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i h: it_wbisimF E RX (gfp (it_wbisim_map E RX)) i
+ (_observe x) (_observe y)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y -> it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i x1, x2: itree' E X i r1: it_eat i (_observe x) x1 r2: it_eat i (_observe y) x2 rr: it_eqF E RX (gfp (it_wbisim_map E RX)) i x1 x2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y -> it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
itree' E Y i
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
itree' E Y i
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_eat i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end?x1
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_eat i
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end?x2
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
itree' E Y i
exact (RetF (f i r0)).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
itree' E Y i
exact (RetF (g i r3)).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_eat i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end (RetF (f i r0))
I: Type E: event I I X, Y: psh I f: foralla : I, X a -> Y a i: I r0: X i ot: itree' E X i r1: it_eat i ot (RetF r0)
it_eat i
+ match ot with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end (RetF (f i r0))
+
I: Type E: event I I X: psh I i: I t1: itree E X i r0: X i r1: it_eat i (observe t1) (RetF r0) IHr1: forall (Y : psh I)
+ (f : foralla : I, X a -> Y a)
+ (r1 : X i),
+RetF r0 = RetF r1 ->
+it_eat i
+ match observe t1 with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k =>
+ VisF e (funr : e_rsp e => f <$> k r)
+ end (RetF (f i r1)) Y: psh I f: foralla : I, X a -> Y a
it_eat i (observe (f <$> t1)) (RetF (f i r0))
+exact (IHr1 Y f r0 eq_refl).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_eat i
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end (RetF (g i r3))
I: Type E: event I I X, Y: psh I g: foralla : I, X a -> Y a i: I r3: X i ot: itree' E X i r2: it_eat i ot (RetF r3)
it_eat i
+ match ot with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end (RetF (g i r3))
+
I: Type E: event I I X: psh I i: I t1: itree E X i r3: X i r2: it_eat i (observe t1) (RetF r3) IHr2: forall (Y : psh I)
+ (g : foralla : I, X a -> Y a)
+ (r4 : X i),
+RetF r3 = RetF r4 ->
+it_eat i
+ match observe t1 with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k =>
+ VisF e (funr : e_rsp e => g <$> k r)
+ end (RetF (g i r4)) Y: psh I g: foralla : I, X a -> Y a
it_eat i (observe (g <$> t1)) (RetF (g i r3))
+exact (IHr2 Y g r3 eq_refl).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
it_eqF E RY (it_wbisim_t E RY R) i
+ (RetF (f i r0)) (RetF (g i r3))
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i r0: X i r1: it_eat i (_observe x) (RetF r0) r3: X i r2: it_eat i (_observe y) (RetF r3) r_rel: RX i r0 r3
RY i (f i r0) (g i r3)
+nowapply H1.
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
itree' E Y i
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
itree' E Y i
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_eat i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end?x1
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_eat i
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end?x2
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
itree' E Y i
exact (TauF (f <$> t1)).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
itree' E Y i
exact (TauF (g <$> t2)).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_eat i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end (TauF (f <$> t1))
I: Type E: event I I X, Y: psh I f: foralla : I, X a -> Y a i: I t1: itree E X i ot: itree' E X i r1: it_eat i ot (TauF t1)
it_eat i
+ match ot with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end (TauF (f <$> t1))
+
I: Type E: event I I X: psh I i: I t0, t1: itree E X i r1: it_eat i (observe t0) (TauF t1) IHr1: forall (Y : psh I)
+ (f : foralla : I, X a -> Y a)
+ (t2 : itree E X i),
+TauF t1 = TauF t2 ->
+it_eat i
+ match observe t0 with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k =>
+ VisF e (funr : e_rsp e => f <$> k r)
+ end (TauF (f <$> t2)) Y: psh I f: foralla : I, X a -> Y a
it_eat i (observe (f <$> t0)) (TauF (f <$> t1))
+exact (IHr1 Y f t1 eq_refl).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_eat i
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end (TauF (g <$> t2))
I: Type E: event I I X, Y: psh I g: foralla : I, X a -> Y a i: I t2: itree E X i ot: itree' E X i r2: it_eat i ot (TauF t2)
it_eat i
+ match ot with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end (TauF (g <$> t2))
+
I: Type E: event I I X: psh I i: I t1, t2: itree E X i r2: it_eat i (observe t1) (TauF t2) IHr2: forall (Y : psh I)
+ (g : foralla : I, X a -> Y a)
+ (t3 : itree E X i),
+TauF t2 = TauF t3 ->
+it_eat i
+ match observe t1 with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k =>
+ VisF e (funr : e_rsp e => g <$> k r)
+ end (TauF (g <$> t3)) Y: psh I g: foralla : I, X a -> Y a
it_eat i (observe (g <$> t1)) (TauF (g <$> t2))
+exact (IHr2 Y g t2 eq_refl).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_eqF E RY (it_wbisim_t E RY R) i
+ (TauF (f <$> t1)) (TauF (g <$> t2))
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y, t1: itree E X i r1: it_eat i (_observe x) (TauF t1) t2: itree E X i r2: it_eat i (_observe y) (TauF t2) t_rel: gfp (it_wbisim_map E RX) i t1 t2
it_wbisim_t E RY R i (f <$> t1) (g <$> t2)
+nowapply CIH.
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
itree' E Y i
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
itree' E Y i
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end?x1
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end?x2
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
itree' E Y i
exact (VisF q (funr => f <$> k1 r)).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
itree' E Y i
exact (VisF q (funr => g <$> k2 r)).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe x with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end (VisF q (funr : e_rsp q => f <$> k1 r))
I: Type E: event I I X, Y: psh I f: foralla : I, X a -> Y a i: I q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) ot: itree' E X i r1: it_eat i ot (VisF q k1)
it_eat i
+ match ot with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => f <$> k r)
+ end (VisF q (funr : e_rsp q => f <$> k1 r))
+
I: Type E: event I I X: psh I i: I t1: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (observe t1) (VisF q k1) IHr1: forall (Y : psh I)
+ (f : foralla : I, X a -> Y a)
+ (q0 : e_qry i)
+ (k2 : forallx : e_rsp q0,
+ itree E X (e_nxt x)),
+VisF q k1 = VisF q0 k2 ->
+it_eat i
+ match observe t1 with
+ | RetF r => RetF (f i r)
+ | TauF t => TauF (f <$> t)
+ | VisF e k =>
+ VisF e (funr : e_rsp e => f <$> k r)
+ end
+ (VisF q0 (funr : e_rsp q0 => f <$> k2 r)) Y: psh I f: foralla : I, X a -> Y a
it_eat i (observe (f <$> t1))
+ (VisF q (funr : e_rsp q => f <$> k1 r))
+exact (IHr1 Y f q k1 eq_refl).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe y with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end (VisF q (funr : e_rsp q => g <$> k2 r))
I: Type E: event I I X, Y: psh I g: foralla : I, X a -> Y a i: I q: e_qry i k2: forallx : e_rsp q, itree E X (e_nxt x) ot: itree' E X i r2: it_eat i ot (VisF q k2)
it_eat i
+ match ot with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end (VisF q (funr : e_rsp q => g <$> k2 r))
+
I: Type E: event I I X: psh I i: I t1: itree E X i q: e_qry i k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (observe t1) (VisF q k2) IHr2: forall (Y : psh I)
+ (g : foralla : I, X a -> Y a)
+ (q0 : e_qry i)
+ (k3 : forallx : e_rsp q0,
+ itree E X (e_nxt x)),
+VisF q k2 = VisF q0 k3 ->
+it_eat i
+ match observe t1 with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k =>
+ VisF e (funr : e_rsp e => g <$> k r)
+ end
+ (VisF q0 (funr : e_rsp q0 => g <$> k3 r)) Y: psh I g: foralla : I, X a -> Y a
it_eat i (observe (g <$> t1))
+ (VisF q (funr : e_rsp q => g <$> k2 r))
+exact (IHr2 Y g q k2 eq_refl).
+
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
it_eqF E RY (it_wbisim_t E RY R) i
+ (VisF q (funr : e_rsp q => f <$> k1 r))
+ (VisF q (funr : e_rsp q => g <$> k2 r))
I: Type E: event I I X: psh I RX: forallx : I, relation (X x) Y: psh I RY: forallx : I, relation (Y x) f, g: foralla : I, X a -> Y a H1: foralla : I,
+(RX a ==> RY a)%signature (f a) (g a) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_wbisim_map E RX) a x y ->
+it_wbisim_t E RY R a (f <$> x) (g <$> y) i: I x, y: itree E X i q: e_qry i k1: forallx : e_rsp q, itree E X (e_nxt x) r1: it_eat i (_observe x) (VisF q k1) k2: forallx : e_rsp q, itree E X (e_nxt x) r2: it_eat i (_observe y) (VisF q k2) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k1 r) (k2 r)
forallr : e_rsp q,
+it_wbisim_t E RY R (e_nxt r) (f <$> k1 r) (g <$> k2 r)
+intro r; nowapply CIH.
+Qed.
subst respects strong bisimilarity.
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y
Proper
+ (dpointwise_relation
+ (funi : I => RX i ==> it_eq RY (i:=i)) ==>
+ dpointwise_relation
+ (funi : I => it_eq RX (i:=i) ==> it_eq RY (i:=i)))
+ subst
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y
Proper
+ (dpointwise_relation
+ (funi : I => RX i ==> it_eq RY (i:=i)) ==>
+ dpointwise_relation
+ (funi : I => it_eq RX (i:=i) ==> it_eq RY (i:=i)))
+ subst
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y
forallxy : foralla : I, X a -> itree E Y a,
+(forall (a : I) (x0y0 : X a),
+ RX a x0 y0 ->
+ gfp (it_eq_map E RY) a (x a x0) (y a y0)) ->
+forall (a : I) (x0y0 : itree E X a),
+gfp (it_eq_map E RX) a x0 y0 ->
+gfp (it_eq_map E RY) a (x =<< x0) (y =<< y0)
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y -> it_eq_t E RY R a (f =<< x) (g =<< y) i: I x, y: itree E X i h2: gfp (it_eq_map E RX) i x y
it_eq_bt E RY R i (f =<< x) (g =<< y)
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y -> it_eq_t E RY R a (f =<< x) (g =<< y) i: I x, y: itree E X i h2: it_eq_map E RX (gfp (it_eq_map E RX)) i x y
it_eq_bt E RY R i (f =<< x) (g =<< y)
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y -> it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2
it_eqF E RY (it_eq_t E RY R) i (_observe (f i r1))
+ (_observe (g i r2))
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 h3:= h1 i r1 r2 r_rel: gfp (it_eq_map E RY) i (f i r1) (g i r2)
it_eqF E RY (it_eq_t E RY R) i
+ (_observe (f i r1)) (_observe (g i r2))
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 h3: it_eq_map E RY (gfp (it_eq_map E RY)) i
+ (f i r1) (g i r2)
it_eqF E RY (it_eq_t E RY R) i
+ (_observe (f i r1)) (_observe (g i r2))
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 h3: it_eqF E RY (gfp (it_eq_map E RY)) i
+ (_observe (f i r1)) (_observe (g i r2))
it_eqF E RY (it_eq_t E RY R) i
+ (_observe (f i r1)) (_observe (g i r2))
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 t1, t2: itree E Y i t_rel: gfp (it_eq_map E RY) i t1 t2
it_eq_t E RY R i t1 t2
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 q: e_qry i k1, k2: forallx : e_rsp q, itree E Y (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RY) (e_nxt r) (k1 r) (k2 r)
forallr : e_rsp q,
+it_eq_t E RY R (e_nxt r) (k1 r) (k2 r)
+
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 t1, t2: itree E Y i t_rel: gfp (it_eq_map E RY) i t1 t2
it_eq_t E RY R i t1 t2
I: Type E: event I I X, Y: psh I RX: relᵢ X X RY: relᵢ Y Y f, g: foralla : I, X a -> itree E Y a h1: forall (a : I) (xy : X a),
+RX a x y ->
+gfp (it_eq_map E RY) a (f a x) (g a y) R: relᵢ (itree E Y) (itree E Y) CIH: forall (a : I) (xy : itree E X a),
+gfp (it_eq_map E RX) a x y ->
+it_eq_t E RY R a (f =<< x) (g =<< y) i: I x0: itree E X i r1, r2: X i r_rel: RX i r1 r2 q: e_qry i k1, k2: forallx : e_rsp q, itree E Y (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RY) (e_nxt r) (k1 r) (k2 r) r: e_rsp q
it_eq_t E RY R (e_nxt r) (k1 r) (k2 r)
+all: nowapply (gfp_t (it_eq_map E RY)).
+Qed.
A slight generalization of the monad law t ≊ t >>= η.
+
I: Type E: event I I X, Y: psh I RX: relᵢ Y Y H: Reflexiveᵢ RX f: X ⇒ᵢ Y i: I t: itree E X i
it_eq RX (f <$> t)
+ (t >>= (fun (i : I) (x : X i) => Ret' (f i x)))
+
I: Type E: event I I X, Y: psh I RX: relᵢ Y Y H: Reflexiveᵢ RX f: X ⇒ᵢ Y i: I t: itree E X i
it_eq RX (f <$> t)
+ (t >>= (fun (i : I) (x : X i) => Ret' (f i x)))
+
I: Type E: event I I X, Y: psh I RX: relᵢ Y Y H: Reflexiveᵢ RX f: X ⇒ᵢ Y R: relᵢ (itree E Y) (itree E Y) CIH: forall (i : I) (t0 : itree E X i),
+it_eq_t E RX R i (f <$> t0)
+(t0 >>= (fun (i0 : I) (x : X i0) => Ret' (f i0 x))) i: I t: itree E X i
it_eq_bt E RX R i (f <$> t)
+ (t >>= (fun (i : I) (x : X i) => Ret' (f i x)))
+cbn; destruct (_observe t); auto.
+Qed.
+
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E Y i: I
Proper (it_eat' i ==> it_eat' i) (subst f i)
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E Y i: I
Proper (it_eat' i ==> it_eat' i) (subst f i)
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E Y i: I x, y: itree E X i H: it_eat i (_observe x) (_observe y)
it_eat i
+ match _observe x with
+ | RetF r => _observe (f i r)
+ | TauF t => TauF (f =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => f =<< k r)
+ end
+ match _observe y with
+ | RetF r => _observe (f i r)
+ | TauF t => TauF (f =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => f =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E Y i: I x, y: itree E X i _x: itree' E X i H: it_eat i _x (_observe y)
it_eat i
+ match _x with
+ | RetF r => _observe (f i r)
+ | TauF t => TauF (f =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => f =<< k r)
+ end
+ match _observe y with
+ | RetF r => _observe (f i r)
+ | TauF t => TauF (f =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => f =<< k r)
+ end
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E Y i: I x, y: itree E X i _x, _y: itree' E X i H: it_eat i _x _y
it_eat i
+ match _x with
+ | RetF r => _observe (f i r)
+ | TauF t => TauF (f =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => f =<< k r)
+ end
+ match _y with
+ | RetF r => _observe (f i r)
+ | TauF t => TauF (f =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => f =<< k r)
+ end
+dependent induction H; noweconstructor.
+Qed.
Composition law of bind.
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ itree E Z i: I x: itree E X i
it_eq RZ ((x >>= f) >>= g) (x >>= (f >=> g))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ itree E Z i: I x: itree E X i
it_eq RZ ((x >>= f) >>= g) (x >>= (f >=> g))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ itree E Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i ((x >>= f) >>= g)
+ (x >>= (f >=> g)) i: I x: itree E X i
it_eq_bt E RZ R i ((x >>= f) >>= g) (x >>= (f >=> g))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ itree E Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i ((x >>= f) >>= g)
+ (x >>= (f >=> g)) i: I x: itree E X i r: X i
it_eqF E RZ (it_eq_t E RZ R) i
+ match _observe (f i r) with
+ | RetF r => _observe (g i r)
+ | TauF t => TauF (g =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => g =<< k r)
+ end
+ match _observe (f i r) with
+ | RetF r => _observe (g i r)
+ | TauF t => TauF (g =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => g =<< k r)
+ end
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ itree E Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i ((x >>= f) >>= g)
+ (x >>= (f >=> g)) i: I x: itree E X i r: X i r0: Y i
it_eqF E RZ (it_eq_t E RZ R) i
+ (_observe (g i r0))
+ (_observe (g i r0))
+destruct ((g i r0).(_observe)); eauto.
+Qed.
Composition law of fmap.
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ Z i: I x: itree E X i
it_eq RZ (g <$> (f <$> x))
+ ((funi : I => g i ∘ f i) <$> x)
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ Z i: I x: itree E X i
it_eq RZ (g <$> (f <$> x))
+ ((funi : I => g i ∘ f i) <$> x)
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i (g <$> (f <$> x))
+ ((funi0 : I => g i0 ∘ f i0) <$> x) i: I x: itree E X i
it_eq_bt E RZ R i (g <$> (f <$> x))
+ ((funi : I => g i ∘ f i) <$> x)
+cbn; destruct (x.(_observe)); eauto.
+Qed.
bind and fmap interaction.
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ itree E Z i: I x: itree E X i
it_eq RZ ((f <$> x) >>= g)
+ (x >>= (funi : I => g i ∘ f i))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ itree E Z i: I x: itree E X i
it_eq RZ ((f <$> x) >>= g)
+ (x >>= (funi : I => g i ∘ f i))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ itree E Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i ((f <$> x) >>= g)
+ (x >>= (funi0 : I => g i0 ∘ f i0)) i: I x: itree E X i
it_eq_bt E RZ R i ((f <$> x) >>= g)
+ (x >>= (funi : I => g i ∘ f i))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ Y g: Y ⇒ᵢ itree E Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i ((f <$> x) >>= g)
+ (x >>= (funi0 : I => g i0 ∘ f i0)) i: I x: itree E X i r: X i
it_eqF E RZ (it_eq_t E RZ R) i
+ (_observe (g i (f i r)))
+ (_observe (g i (f i r)))
+destruct ((g i (f i r)).(_observe)); eauto.
+Qed.
+
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ Z i: I x: itree E X i
it_eq RZ (g <$> (x >>= f))
+ (x >>= (funi : I => fmap g i ∘ f i))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ Z i: I x: itree E X i
it_eq RZ (g <$> (x >>= f))
+ (x >>= (funi : I => fmap g i ∘ f i))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i (g <$> (x >>= f))
+ (x >>= (funi0 : I => fmap g i0 ∘ f i0)) i: I x: itree E X i
it_eq_bt E RZ R i (g <$> (x >>= f))
+ (x >>= (funi : I => fmap g i ∘ f i))
+
I: Type E: event I I X, Y, Z: psh I RZ: forallx : I, relation (Z x) Reflexiveᵢ0: Reflexiveᵢ RZ f: X ⇒ᵢ itree E Y g: Y ⇒ᵢ Z R: relᵢ (itree E Z) (itree E Z) CIH: forall (i : I) (x : itree E X i),
+it_eq_t E RZ R i (g <$> (x >>= f))
+ (x >>= (funi0 : I => fmap g i0 ∘ f i0)) i: I x: itree E X i r: X i
it_eqF E RZ (it_eq_t E RZ R) i
+ match _observe (f i r) with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
+ match _observe (f i r) with
+ | RetF r => RetF (g i r)
+ | TauF t => TauF (g <$> t)
+ | VisF e k => VisF e (funr : e_rsp e => g <$> k r)
+ end
+destruct ((f i r).(_observe)); eauto.
+Qed.
Rewording the composition law of bind.
+
I: Type E: event I I X, Y, Z: psh I f: Y ⇒ᵢ itree E Z g: X ⇒ᵢ itree E Y i: I x: itree E X i
((x >>= g) >>= f) ≊ (x >>= (g >=> f))
+
I: Type E: event I I X, Y, Z: psh I f: Y ⇒ᵢ itree E Z g: X ⇒ᵢ itree E Y i: I x: itree E X i
((x >>= g) >>= f) ≊ (x >>= (g >=> f))
apply bind_bind_com.Qed.
Proof of the up-to bind principle: we may close bisimulation candidates by prefixing
+related elements by bisimilar computations.
+
VariantbindR {X1X2Y1Y2} (RX : relᵢ X1 X2)
+ (R : relᵢ (itree E X1) (itree E X2))
+ (S : relᵢ (itree E Y1) (itree E Y2)) :
+ relᵢ (itree E Y1) (itree E Y2) :=
+ | BindR {i t1 t2 k1 k2}
+ (u : R i t1 t2)
+ (v : forallix1x2, RX i x1 x2 -> S i (k1 i x1) (k2 i x2))
+ : bindR RX R S i (t1 >>= k1) (t2 >>= k2).
+Arguments BindR {X1 X2 Y1 Y2 RX R S i t1 t2 k1 k2}.
+Hint Constructors bindR : core.
Up-to bind is valid for strong bisimilarity.
+
Program DefinitionbindR_eq_map {X1X2Y1Y2} (RX : relᵢ X1 X2) : mon (relᵢ (itree E Y1) (itree E Y2)) :=
+ {| body S := bindR RX (@it_eq _ E _ _ RX) S ;
+ Hbody _ _ H _ _ _ '(BindR u v) := BindR u (funi__r => H _ _ _ (v i _ _ r)) |}.
+
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2
bindR_eq_map RX <= it_eq_t E RY
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2
bindR_eq_map RX <= it_eq_t E RY
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x u: it_eq RX t1 t2 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eq_map E RY R i (k1 i x1) (k2 i x2)
bT (it_eq_map E RY) (bindR_eq_map RX) R i0 (t1 >>= k1)
+ (t2 >>= k2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x u: it_eq_map E RX (gfp (it_eq_map E RX)) i0 t1 t2 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eq_map E RY R i (k1 i x1) (k2 i x2)
bT (it_eq_map E RY) (bindR_eq_map RX) R i0 (t1 >>= k1)
+ (t2 >>= k2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x u: it_eqF E RX (gfp (it_eq_map E RX)) i0
+ (_observe t1) (_observe t2) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r1: X1 i0 r2: X2 i0 r_rel: RX i0 r1 r2 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ (_observe (k1 i0 r1)) (_observe (k2 i0 r2))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 t3: itree E X2 i0 t_rel: gfp (it_eq_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ (TauF (k1 =<< t0)) (TauF (k2 =<< t3))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) k3: forallx : e_rsp q, itree E X2 (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RX) (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ (VisF q (funr : e_rsp q => k1 =<< k0 r))
+ (VisF q (funr : e_rsp q => k2 =<< k3 r))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r1: X1 i0 r2: X2 i0 r_rel: RX i0 r1 r2 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ (_observe (k1 i0 r1))
+ (_observe (k2 i0 r2))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 t3: itree E X2 i0 t_rel: gfp (it_eq_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ (TauF (k1 =<< t0)) (TauF (k2 =<< t3))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 t3: itree E X2 i0 t_rel: gfp (it_eq_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
(bindR_eq_map RX ° it_eq_T E RY (bindR_eq_map RX)) R
+ i0 (k1 =<< t0) (k2 =<< t3)
+refine (BindR t_rel (funi__r => b_T (it_eq_map E RY) _ _ _ _ _ (v i _ _ r))).
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) k3: forallx : e_rsp q, itree E X2 (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RX) (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_eq_T E RY (bindR_eq_map RX) R) i0
+ (VisF q (funr : e_rsp q => k1 =<< k0 r))
+ (VisF q (funr : e_rsp q => k2 =<< k3 r))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x0: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) k3: forallx : e_rsp q, itree E X2 (e_nxt x) k_rel: forallr : e_rsp q,
+gfp (it_eq_map E RX) (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_eqF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) r: e_rsp q
(bindR_eq_map RX ° it_eq_T E RY (bindR_eq_map RX)) R
+ (e_nxt r) (k1 =<< k0 r)
+ (k2 =<< k3 r)
+refine (BindR (k_rel r) (funi__r => b_T (it_eq_map E RY) _ _ _ _ _ (v i _ _ r))).
+Qed.
Up-to bind is valid for weak bisimilarity.
+
Program DefinitionbindR_w_map {X1X2Y1Y2} (RX : relᵢ X1 X2) : mon (relᵢ (itree E Y1) (itree E Y2)) :=
+ {| body S := bindR RX (@it_wbisim _ E _ _ RX) S ;
+ Hbody _ _ H _ _ _ '(BindR u v) := BindR u (funi__r => H _ _ _ (v i _ _ r)) |}.
+
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2
bindR_w_map RX <= it_wbisim_t E RY
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2
bindR_w_map RX <= it_wbisim_t E RY
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x u: it_wbisim RX i0 t1 t2 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisim_map E RY R i (k1 i x1) (k2 i x2)
bT (it_wbisim_map E RY) (bindR_w_map RX) R i0
+ (t1 >>= k1) (t2 >>= k2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x u: it_wbisim_map E RX (gfp (it_wbisim_map E RX)) i0
+ t1 t2 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisim_map E RY R i (k1 i x1) (k2 i x2)
bT (it_wbisim_map E RY) (bindR_w_map RX) R i0
+ (t1 >>= k1) (t2 >>= k2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x u: it_wbisimF E RX (gfp (it_wbisim_map E RX)) i0
+ (_observe t1) (_observe t2) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) x1: itree' E Y1 i0 x2: itree' E Y2 i0 s1: it_eat i0 (observe (k1 i0 r0)) x1 s2: it_eat i0 (observe (k2 i0 r3)) x2 ss: it_eqF E RY R i0 x1 x2
it_eat i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) x1: itree' E Y1 i0 x2: itree' E Y2 i0 s1: it_eat i0 (observe (k1 i0 r0)) x1 s2: it_eat i0 (observe (k2 i0 r3)) x2 ss: it_eqF E RY R i0 x1 x2
it_eat i0
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) x1: itree' E Y1 i0 x2: itree' E Y2 i0 s1: it_eat i0 (observe (k1 i0 r0)) x1 s2: it_eat i0 (observe (k2 i0 r3)) x2 ss: it_eqF E RY R i0 x1 x2
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) x1: itree' E Y1 i0 x2: itree' E Y2 i0 s1: it_eat i0 (observe (k1 i0 r0)) x1 s2: it_eat i0 (observe (k2 i0 r3)) x2 ss: it_eqF E RY R i0 x1 x2
it_eat i0
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) x1: itree' E Y1 i0 x2: itree' E Y2 i0 s1: it_eat i0 (observe (k1 i0 r0)) x1 s2: it_eat i0 (observe (k2 i0 r3)) x2 ss: it_eqF E RY R i0 x1 x2
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ x1 ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x r0: X1 i0 r1: it_eat i0 (_observe t1) (RetF r0) r3: X2 i0 r2: it_eat i0 (_observe t2) (RetF r3) r_rel: RX i0 r0 r3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) x1: itree' E Y1 i0 x2: itree' E Y2 i0 s1: it_eat i0 (observe (k1 i0 r0)) x1 s2: it_eat i0 (observe (k2 i0 r3)) x2 ss: it_eqF E RY R i0 x1 x2
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ x1 x2
+nowapply (it_eqF_mon _ _ (id_T _ _ R)).
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eat i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eat i0
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eat i0
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ (_observe (k1 =<< Tau' t0))
+ ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ (_observe (k1 =<< Tau' t0))
+ (_observe (k2 =<< Tau' t3))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisim_T E RY (bindR_w_map RX) R i0
+ (k1 =<< t0) (k2 =<< t3)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x t0: itree E X1 i0 r1: it_eat i0 (_observe t1) (TauF t0) t3: itree E X2 i0 r2: it_eat i0 (_observe t2) (TauF t3) t_rel: gfp (it_wbisim_map E RX) i0 t0 t3 v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
(bindR_w_map RX ° it_wbisim_T E RY (bindR_w_map RX)) R
+ i0 (k1 =<< t0) (k2 =<< t3)
+econstructor; [ apply t_rel | intros; nowapply (b_T (it_wbisim_map E RY)), v ].
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_wbisimF E RY (it_wbisim_T E RY (bindR_w_map RX) R)
+ i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eat i0
+ match _observe t1 with
+ | RetF r => _observe (k1 i0 r)
+ | TauF t => TauF (k1 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k1 =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eat i0
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eat i0
+ match _observe t2 with
+ | RetF r => _observe (k2 i0 r)
+ | TauF t => TauF (k2 =<< t)
+ | VisF e k => VisF e (funr : e_rsp e => k2 =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ (_observe (k1 =<< Vis' q k0))
+ ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2))
it_eqF E RY (it_wbisim_T E RY (bindR_w_map RX) R) i0
+ (_observe (k1 =<< Vis' q k0))
+ (_observe (k2 =<< Vis' q k3))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) r: e_rsp q
it_wbisim_T E RY (bindR_w_map RX) R
+ (e_nxt r) (k1 =<< k0 r)
+ (k2 =<< k3 r)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 R: relᵢ (itree E Y1) (itree E Y2) i: I x: itree E Y1 i y: itree E Y2 i i0: I t1: itree E X1 i0 t2: itree E X2 i0 k1: forallx : I, X1 x -> itree E Y1 x k2: forallx : I, X2 x -> itree E Y2 x q: e_qry i0 k0: forallx : e_rsp q, itree E X1 (e_nxt x) r1: it_eat i0 (_observe t1) (VisF q k0) k3: forallx : e_rsp q, itree E X2 (e_nxt x) r2: it_eat i0 (_observe t2) (VisF q k3) k_rel: forallr : e_rsp q,
+gfp (it_wbisim_map E RX)
+ (e_nxt r) (k0 r) (k3 r) v: forall (i : I) (x1 : X1 i) (x2 : X2 i),
+RX i x1 x2 ->
+it_wbisimF E RY R i (observe (k1 i x1))
+ (observe (k2 i x2)) r: e_rsp q
(bindR_w_map RX ° it_wbisim_T E RY (bindR_w_map RX)) R
+ (e_nxt r) (k1 =<< k0 r)
+ (k2 =<< k3 r)
+econstructor; [ apply k_rel | intros; nowapply (b_T (it_wbisim_map E RY)), v ].
+Qed.
Pointwise strongly bisimilar equations have strongly bisimilar iteration.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b)
forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq RY (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b)
forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq RY (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b
it_eq_bt E RY R i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b h1:= H i a b r: it_eq (sumᵣ RX RY) (f i a) (g i b)
it_eq_bt E RY R i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b h1: it_eqF E (sumᵣ RX RY) (it_eq (sumᵣ RX RY)) i
+ (_observe (f i a)) (_observe (g i b))
it_eqF E RY (it_eq_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b r1: (X1 +ᵢ Y1) i r2: (X2 +ᵢ Y2) i r_rel: sumᵣ RX RY i r1 r2
it_eqF E RY (it_eq_t E RY R) i
+ match r1 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ match r2 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i t2: itree E (X2 +ᵢ Y2) i t_rel: it_eq (sumᵣ RX RY) t1 t2
it_eqF E RY (it_eq_t E RY R) i
+ (TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t1))
+ (TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t2))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (sumᵣ RX RY) (k1 r) (k2 r)
it_eqF E RY (it_eq_t E RY R) i
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k1 r))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k2 r))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b r1: (X1 +ᵢ Y1) i r2: (X2 +ᵢ Y2) i r_rel: sumᵣ RX RY i r1 r2
it_eqF E RY (it_eq_t E RY R) i
+ match r1 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ match r2 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
destruct r_rel; eauto.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i t2: itree E (X2 +ᵢ Y2) i t_rel: it_eq (sumᵣ RX RY) t1 t2
it_eqF E RY (it_eq_t E RY R) i
+ (TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t1))
+ (TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t2))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i t2: itree E (X2 +ᵢ Y2) i t_rel: it_eq (sumᵣ RX RY) t1 t2
it_eq_t E RY R i
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t1)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i t2: itree E (X2 +ᵢ Y2) i t_rel: it_eq (sumᵣ RX RY) t1 t2
it_eq_t E RY (it_eq_t E RY R) i
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t1)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i t2: itree E (X2 +ᵢ Y2) i t_rel: it_eq (sumᵣ RX RY) t1 t2
forall (i : I) (x1 : (X1 +ᵢ Y1) i) (x2 : (X2 +ᵢ Y2) i),
+sumᵣ RX RY i x1 x2 ->
+it_eq_t E RY R i
+ {|
+ _observe :=
+ match x1 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}
+intros ? ? ? []; apply (tt_t (it_eq_map E RY)), (b_t (it_eq_map E RY)); cbn; eauto.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (sumᵣ RX RY) (k1 r) (k2 r)
it_eqF E RY (it_eq_t E RY R) i
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k1 r))
+ (VisF q
+ (funr : e_rsp q =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k2 r))
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (sumᵣ RX RY) (k1 r) (k2 r) r0: e_rsp q
it_eq_t E RY R (e_nxt r0)
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k1 r0)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k2 r0)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (sumᵣ RX RY) (k1 r) (k2 r) r0: e_rsp q
it_eq_t E RY (it_eq_t E RY R)
+ (e_nxt r0)
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k1 r0)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k2 r0)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_eq (sumᵣ RX RY) (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_eq_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) k_rel: forallr : e_rsp q,
+it_eq (sumᵣ RX RY) (k1 r) (k2 r) r0: e_rsp q
forall (i : I) (x1 : (X1 +ᵢ Y1) i) (x2 : (X2 +ᵢ Y2) i),
+sumᵣ RX RY i x1 x2 ->
+it_eq_t E RY R i
+ {|
+ _observe :=
+ match x1 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}
+intros ? ? ? []; apply (bt_t (it_eq_map E RY)); cbn; eauto.
+Qed.
+
+#[global] Instanceiter_proper_strong {XY} {RX : relᵢ X X} {RY : relᵢ Y Y} :
+ Proper (dpointwise_relation (funi => RX i ==> it_eq (sumᵣ RX RY) (i:=i))
+ ==> dpointwise_relation (funi => RX i ==> it_eq RY (i:=i)))
+ (@iter I E X Y)
+ := iter_cong_strong.
Pointwise weakly bisimilar equations have weakly bisimilar iteration.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b)
forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim RY i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b)
forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim RY i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b
it_wbisim_bt E RY R i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b H1:= H i a b r: it_wbisim (sumᵣ RX RY) i (f i a) (g i b)
it_wbisim_bt E RY R i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b H1: it_wbisim_map E (sumᵣ RX RY)
+ (it_wbisim (sumᵣ RX RY)) i (f i a) (g i b)
it_wbisim_bt E RY R i (iter f i a) (iter g i b)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b H1: it_wbisimF E (sumᵣ RX RY)
+ (it_wbisim (sumᵣ RX RY)) i (_observe (f i a))
+ (_observe (g i b))
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b -> it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b r0: (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (RetF r0) r3: (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (RetF r3) r_rel: sumᵣ RX RY i r0 r3
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b r0: (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (RetF r0) r3: (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (RetF r3) r_rel: sumᵣ RX RY i r0 r3
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 H2:= H i x1 x2 H0: it_wbisim (sumᵣ RX RY) i (f i x1) (g i x2)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 H2: it_wbisim_map E (sumᵣ RX RY)
+ (it_wbisim (sumᵣ RX RY)) i
+ (f i x1) (g i x2)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_eat i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Ret' (inl x1)))
+ ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Ret' (inl x1)))
+ (_observe
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< Ret' (inl x2)))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b x1: X1 i r1: it_eat i (_observe (f i a)) (RetF (inl x1)) x2: X2 i r2: it_eat i (_observe (g i b)) (RetF (inl x2)) H0: RX i x1 x2 x0: itree' E (X1 +ᵢ Y1) i x3: itree' E (X2 +ᵢ Y2) i s1: it_eat i (observe (f i x1)) x0 s2: it_eat i (observe (g i x2)) x3 ss: it_eqF E (sumᵣ RX RY) (it_wbisim (sumᵣ RX RY)) i
+ x0 x3
it_wbisim_t E RY R i (iter f i x1) (iter g i x2)
+nowapply CIH.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_eat i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Ret' (inr y1)))
+ ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b y1: Y1 i r1: it_eat i (_observe (f i a)) (RetF (inr y1)) y2: Y2 i r2: it_eat i (_observe (g i b)) (RetF (inr y2)) H0: RY i y1 y2
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Ret' (inr y1)))
+ (_observe
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< Ret' (inr y2)))
+nowcbn; econstructor.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_eat i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Tau' t1))
+ ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Tau' t1))
+ (_observe
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< Tau' t2))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_wbisim_t E RY R i
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t1)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
it_wbisim_t E RY (it_wbisim_t E RY R) i
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t1)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t2)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b t1: itree E (X1 +ᵢ Y1) i r1: it_eat i (_observe (f i a)) (TauF t1) t2: itree E (X2 +ᵢ Y2) i r2: it_eat i (_observe (g i b)) (TauF t2) t_rel: it_wbisim (sumᵣ RX RY) i t1 t2
forall (i : I) (x1 : (X1 +ᵢ Y1) i) (x2 : (X2 +ᵢ Y2) i),
+sumᵣ RX RY i x1 x2 ->
+it_wbisim_t E RY R i
+ {|
+ _observe :=
+ match x1 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}
+intros ? ? ? []; apply (tt_t (it_wbisim_map E RY)), (b_t (it_wbisim_map E RY)); cbn; eauto.
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_wbisimF E RY (it_wbisim_t E RY R) i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe (f i a) with
+ | RetF (inl x) => TauF (iter f i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x1
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_eqF E RY (it_wbisim_t E RY R) i ?x1?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_eat i
+ match _observe (g i b) with
+ | RetF (inl x) => TauF (iter g i x)
+ | RetF (inr y) => RetF y
+ | TauF t =>
+ TauF
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< t)
+ | VisF e k =>
+ VisF e
+ (funr : e_rsp e =>
+ (fun (i : I) (r0 : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r0 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k r)
+ end?x2
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Vis' q k1))
+ ?x2
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r)
it_eqF E RY (it_wbisim_t E RY R) i
+ (_observe
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< Vis' q k1))
+ (_observe
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< Vis' q k2))
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r) r0: e_rsp q
it_wbisim_t E RY R (e_nxt r0)
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k1 r0)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k2 r0)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r) r0: e_rsp q
it_wbisim_t E RY (it_wbisim_t E RY R)
+ (e_nxt r0)
+ ((fun (i : I) (r : (X1 +ᵢ Y1) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}) =<< k1 r0)
+ ((fun (i : I) (r : (X2 +ᵢ Y2) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}) =<< k2 r0)
+
I: Type E: event I I X1, X2, Y1, Y2: psh I RX: relᵢ X1 X2 RY: relᵢ Y1 Y2 f: forallx : I, X1 x -> itree E (X1 +ᵢ Y1) x g: forallx : I, X2 x -> itree E (X2 +ᵢ Y2) x H: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim (sumᵣ RX RY) i (f i a) (g i b) R: relᵢ (itree E Y1) (itree E Y2) CIH: forall (i : I) (a : X1 i) (b : X2 i),
+RX i a b ->
+it_wbisim_t E RY R i (iter f i a) (iter g i b) i: I a: X1 i b: X2 i r: RX i a b q: e_qry i k1: forallx : e_rsp q, itree E (X1 +ᵢ Y1) (e_nxt x) r1: it_eat i (_observe (f i a)) (VisF q k1) k2: forallx : e_rsp q, itree E (X2 +ᵢ Y2) (e_nxt x) r2: it_eat i (_observe (g i b)) (VisF q k2) k_rel: forallr : e_rsp q,
+it_wbisim (sumᵣ RX RY) (e_nxt r) (k1 r) (k2 r) r0: e_rsp q
forall (i : I) (x1 : (X1 +ᵢ Y1) i) (x2 : (X2 +ᵢ Y2) i),
+sumᵣ RX RY i x1 x2 ->
+it_wbisim_t E RY R i
+ {|
+ _observe :=
+ match x1 with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}
+ {|
+ _observe :=
+ match x2 with
+ | inl x => TauF (iter g i x)
+ | inr y => RetF y
+ end
+ |}
+intros ? ? ? []; apply (tt_t (it_wbisim_map E RY)), (b_t (it_wbisim_map E RY)); cbn; eauto.
+Qed.
+
+#[global] Instanceiter_weq {XY} {RX : relᵢ X X} {RY : relᵢ Y Y} :
+ Proper (dpointwise_relation (funi => RX i ==> it_wbisim (sumᵣ RX RY) i)
+ ==> dpointwise_relation (funi => RX i ==> it_wbisim RY i))
+ (@iter I E X Y)
+ := iter_cong_weak.
Iteration is a strong fixed point of the folowing guarded equation.
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) Reflexiveᵢ0: Reflexiveᵢ RY f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
it_eq RY (iter f i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) Reflexiveᵢ0: Reflexiveᵢ RY f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
it_eq RY (iter f i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) Reflexiveᵢ0: Reflexiveᵢ RY f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
it_eq_map E RY (gfp (it_eq_map E RY)) i
+ (iter f i x)
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+
I: Type E: event I I X, Y: psh I RY: forallx : I, relation (Y x) Reflexiveᵢ0: Reflexiveᵢ RY f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i r: (X +ᵢ Y) i
it_eqF E RY (gfp (it_eq_map E RY)) i
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+destruct r; eauto.
+Qed.
Iteration is a weak fixed point of the equation (Prop. 3).
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
iter f i x
+≈ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter f i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
iter f i x
+≈ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter f i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
(f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+≈ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter f i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
it_wbisim_t E (eqᵢ Y) (it_wbisim_t E (eqᵢ Y) bot) i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter f i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i
bindR_w_map (eqᵢ (X +ᵢ Y)) (it_wbisim_t E (eqᵢ Y) bot)
+ i
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ {|
+ _observe :=
+ match r with
+ | inl x => TauF (iter f i x)
+ | inr y => RetF y
+ end
+ |}))
+ (f i x >>=
+ (fun (i : I) (r : (X +ᵢ Y) i) =>
+ match r with
+ | inl x => iter f i x
+ | inr y => Ret' y
+ end))
+
I: Type E: event I I X, Y: psh I f: X ⇒ᵢ itree E (X +ᵢ Y) i: I x: X i i0: I x0: X i0
it_wbisim_t E (eqᵢ Y) bot i0 (Tau' (iter f i0 x0))
+ (iter f i0 x0)
+exact (skip_tau _).
+Qed.
+EndwithE.
Misc. utilities.
+
Variantis_tau' {I} {E : event I I} {Xi} : itree' E X i -> Prop :=
+ | IsTau {t : itree E X i} : is_tau' (TauF t) .
+Definitionis_tau {I} {E : event I I} {Xi} (t : itree E X i) : Prop := is_tau' t.(_observe).
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Type families
+
+
This file is dully-named Psh.v but in fact is actually about
+type families, that is maps I → Type for some I : Type.
+
Notationpsh I := (I -> Type).
Pointwise arrows of families.
+
Definitionarrᵢ {I} (XY : psh I) : Type := foralli, X i -> Y i.
+#[global] Infix"⇒ᵢ" := (arrᵢ) (at level20) : indexed_scope.
A bunch of combinators.
+
Definitionvoidᵢ {I : Type} : I -> Type := fun_ => T0.
+#[global] Notation"∅ᵢ" := (voidᵢ) : indexed_scope.
+
+Definitionsumᵢ {I} (XY : psh I) : psh I := funi => (X i + Y i)%type.
+#[global] Infix"+ᵢ" := (sumᵢ) (at level20) : indexed_scope.
+
+Definitionprodᵢ {I} (XY : psh I) : psh I := funi => (X i * Y i)%type.
+#[global] Infix"×ᵢ" := (prodᵢ) (at level20) : indexed_scope.
Fibers of a function are a particularly import kind of type family.
+
Inductivefiber {AB} (f : A -> B) : psh B := Fib a : fiber f (f a).
+Arguments Fib {A B f}.
+DeriveNoConfusionfor fiber.
+
+Equationsfib_extr {AB} {f : A -> B} {b : B} : fiber f b -> A :=
+ fib_extr (Fib a) := a .
+Equationsfib_coh {AB} {f : A -> B} {b : B} : forallp : fiber f b, f (fib_extr p) = b :=
+ fib_coh (Fib _) := eq_refl .
+Definitionfib_constr {AB} {f : A -> B} a : forallb (p : f a = b), fiber f b :=
+ eq_rect (f a) (fiber f) (Fib a).
These next two functions form an isomorphism (fiber f ⇒ᵢ X) ≅ (∀ a → X (f a))
+
Equationsfib_into {AB} {f : A -> B} X (h : foralla, X (f a)) : fiber f ⇒ᵢ X :=
+ fib_into _ h _ (Fib a) := h a .
+Definitionfib_from {AB} {f : A -> B} X (h : fiber f ⇒ᵢ X) a : X (f a) :=
+ h _ (Fib a).
Using the fiber construction, we can define a "sparse" type family which will
+be equal to some set at one point of the index and empty otherwise. This will
+enable us to have an isomorphism (X @ i ⇒ᵢ Y) ≅ (X → Y i).
+
See the functional pearl by Conor McBride "Kleisli arrows of outrageous fortune"
+for background on this construction and its use.
Built with Alectryon, running . Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
An abstract, certified account of operational game semantics
+
+
This is the companion artifact to the ESOP paper. The main contributions of
+this library are:
+
+
an independent implementation of an indexed counterpart to the
+InteractionTree Coq library with support for guarded and eventually guarded
+recursion
+
an abstract OGS model of an axiomatised language proven sound w.r.t.
+substitution equivalence
+
several instantiations of this abstract result to concrete
+languages: simply-typed lambda-calculus with recursive functions,
+untyped lambda calculus, Downen and Ariola's polarized system D and
+system L.
To simply typecheck the Coq proofs, we provide a Docker image preloaded with
+all the dependencies. The instructions to run it are given below. If instead
+you want to manually install the OGS library on your own system, follow the
+instructions from the section "Local Installation" at the end of the file.
+
First, ensure that your docker daemon is running.
+
We recommend building the docker image from source, by executing the following
+command from the root of the repository. Note that this requires network access
+and will download around 1.5GiB from hub.docker.com.
+
+makedocker-build
+
+
Alternatively, if you prefer, you can download the precompiled image
+docker_coq-ogs.tar.gz from the
+Zenodo archive of this artifact,
+and load it into docker with the following command.
+
+dockerimageload-ipath/to/docker_coq-ogs.tar.gz
+
+
After the image is built or loaded from the file, verify that it is indeed
+listed by the following command.
+
+dockerimagelscoq-ogs:latest
+
+
This image contains the full code artifact in the directory /home/coq/ogs. The
+default command for the image typechecks the whole repository. Run it with a
+tty to see the progress with the following command. This should take around
+3-5min and conclude by displaying information about several soudness theorems.
+
+dockerrun--ttycoq-ogs:latest
+
+
+
+
Step-by-Step Reproduction
+
The above "Getting Started" section already describes how to typecheck the
+whole repository, validating our claims of certification from the paper. The
+concluding output arises from a special file which we have included,
+Checks.v, which imports core theorems, displays their type and list
+of arguments, as well as the assumptions (axioms) they depend on.
+
If you wish to further inspect the repository, the following command will
+start an interactive shell inside the container.
+
+dockerrun--tty--interactivecoq-ogs:latest
+
+
The following section details in more precision the content of each file
+and their relationship to the paper. We have furthermore tried to thoroughly
+comment most parts of the development. If you prefer to navigate the code
+in a web browser, an HTML rendering of the whole code together with the proof
+state during intermediate steps is provided at the following URL:
Ctx/ directory: general metatheory of substitution. This material has been
+largely left untold in the paper, and as explained in the artifact report, we
+introduce a novel gadget for abstracting over scope and variable
+representations.
+
Family.v: Definition of scoped
+and sorted families (Def. 4).
+
Abstract.v: Definition of
+scope structures. The comments make this file a good entry-point to the
+understand the constructions from this directory.
+
Assignment.v: Generic
+definition of assignments (Def. 5 and 6).
+
Renaming.v: Generic
+definition of renamings as variable assignments, together with their
+important equational laws.
+
Ctx.v: Definition of concrete
+contexts (lists) and dependently-typed DeBruijn indices.
+
Covering.v: Instanciation of
+the scope structure for concrete contexts from
+Ctx.v.
Eq.v: Strong and weak bisimilarity
+over interaction trees.
+
Structure.v: Combinators
+(definitions) for the monadic structure and unguarded iteration (Def. 31).
+
Properties.v: General
+properties of the combinators (Prop. 3).
+
Guarded.v: (Eventually)
+guarded equations and iterations over them, together with their unicity
+property (Def. 30, 33 and 34 and Prop. 5).
+
Delay.v: Definition of the
+delay monad (as a special case of interaction trees over the empty event)
+(Def. 9 and 10).
+
+
+
OGS/ directory: construction of a sound OGS model for an abstract language
+machine.
+
Obs.v: Axiomatization of
+observation structure (Def. 12) and normal forms (part of Def. 13).
+
Machine.v: Axiomatization of
+evaluation structures (Def. 11), language machines (Def. 13) and focused
+redexes (Def. 28).
+
Game.v: Abstract games (Def. 16
+and 18) and OGS game definition (Def. 21--23).
+
Strategy.v: Machine strategy
+(Def. 24--26) and composition.
+
CompGuarded.v: Proof of
+eventual guardedness of the equation defining the composition of strategies
+(Prop. 6).
+
Adequacy.v: Proof of
+adequacy of composition (Prop. 7).
+
Congruence.v: Proof of
+congruence of composition (Prop. 4).
+
Soundness.v: Proof of
+soundness of the OGS (Thm. 8).
+
+
+
Examples/ directory: concrete language machines instanciating the generic
+construction and soundness theorem.
+
STLC_CBV.v: Simply typed,
+call-by-value, lambda calculus, with a unit type and recursive functions.
+This example is the most commented, with a complete walk-through of the
+instanciation.
+
ULC_CBV.v: Pure, untyped,
+call-by-value, lambda calculus. This example demonstrate that the
+intrinsically typed and scoped approach still handles untyped calculi, by
+treating them as "unityped".
+
SystemD.v: Mu-mu-tilde
+calculus variant System D from Downen and Ariola, polarized. This is
+our "flagship" example, as it is a very expressive calculus. We have
+dropped existential and universal type quantifier as our framework only
+captures simple types. We have added a slightly ad-hoc construction to
+enable general recursion, making the calculus non-normalizing.
+
SystemL_CBV.v:
+mu-mu-tilde calculus variant System L, in call by value
+(i.e., lambda-bar-mu-mu-tilde-Q calculus from Herbelin and Curien).
+
+
+
Checks.v: Interactively display
+information about the most important theorems.
+
+
+
+
+
Axioms
+
The whole development relies only on axiom K, a conventional and sound axiom
+from Coq's standard library (more precisely, Eq_rect_eq.eq_rect_eq).
+It is used by Equations to perform some dependent pattern matching.
+
This fact can be verified from the output of typechecking Checks.v. It can
+also be double checked in an interactive mode.
+
+
For the abstract result of soundness of the OGS by running
+Print Assumptions ogs_correction. at the end of
+Soundness.v.
+
For any particular example, for instance by running
+Print Assumptions stlc_ciu_correct. at the end of
+STLC_CBV.v.
+
+
+
+
Local Installation
+
The most convenient way to experiment and develop with this library is to
+install Coq locally and use some IDE such as emacs. To do so, first ensure you
+have the source code for the project or download it with the following command.
To install the Coq dependencies, first ensure you have a working opam
+installation. This should usually be obtained from your systems package
+distribution, see https://opam.ocaml.org/doc/Install.html for further
+information.
+
Check if you have added the coq-released package repository with the command
+opam repo. If it does not appear, add it with the following command.
Then, from the root of the repository, install the dependencies with
+the following command
+
+opaminstall--deps-only.
+
+
We stress that the development has been only checked to compile against these
+specific dependencies. In particular, it does not compiled at the moment
+against latest version of coq-coinduction due to major changes in the API.
+
Finally, typecheck the code with the following command, again from the root of
+the repository. This should take around 3-5min.
+
+dunebuild
+
+
+
+
Generating the Alectryon documentation
+
To build the html documentation, first install Alectryon:
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Indexed relations
+
+
A few standard constructions to work with indexed relations.
+
From OGS Require Import Prelude.
+From OGS.Utils Require Import Psh.
+
+From Coq.Classes Require Export RelationClasses.
+From Coinduction Require Export lattice.
+
+#[export] Existing InstanceCompleteLattice_dfun.
+
+#[global] Notation"∀ₕ x , R" := (forall_relation (funx => R%signature))
+ (x binder, at level60).
+Notationrelᵢ A B := (foralli, A i -> B i -> Prop).
+
+#[global] NotationReflexiveᵢ R := (foralli, Reflexive (R i)).
+#[global] NotationSymmetricᵢ R := (foralli, Symmetric (R i)).
+#[global] NotationTransitiveᵢ R := (foralli, Transitive (R i)).
+#[global] NotationEquivalenceᵢ R := (foralli, Equivalence (R i)).
+#[global] NotationSubrelationᵢ R S := (foralli, subrelation (R i) (S i)).
+#[global] NotationPreOrderᵢ R := (foralli, PreOrder (R i)).
+
+Definitioneqᵢ {I} (X : psh I) : relᵢ X X := fun_xy => x = y.
+Arguments eqᵢ _ _ _ /.
+
+
I: Type X: psh I
Reflexiveᵢ (eqᵢ X)
+
I: Type X: psh I
Reflexiveᵢ (eqᵢ X)
intros ? ?; reflexivity.Qed.
+
+
I: Type X: psh I
Symmetricᵢ (eqᵢ X)
+
I: Type X: psh I
Symmetricᵢ (eqᵢ X)
intros ? ? ? ?; nowsymmetry.Qed.
+
+
I: Type X: psh I
Transitiveᵢ (eqᵢ X)
+
I: Type X: psh I
Transitiveᵢ (eqᵢ X)
intros i x y z a b; nowtransitivity y.Qed.
+
+Variantsumᵣ {I} {X1X2Y1Y2 : psh I} (R : relᵢ X1 X2) (S : relᵢ Y1 Y2) : relᵢ (X1 +ᵢ Y1) (X2 +ᵢ Y2) :=
+ | RLeft {i x1 x2} : R i x1 x2 -> sumᵣ R S i (inl x1) (inl x2)
+ | RRight {i y1 y2} : S i y1 y2 -> sumᵣ R S i (inr y1) (inr y2)
+.
+#[global] Hint Constructors sumᵣ : core.
+
+
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S
Equivalenceᵢ (sumᵣ R S)
+
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S
Equivalenceᵢ (sumᵣ R S)
+
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I
Reflexive (sumᵣ R S i)
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I
Symmetric (sumᵣ R S i)
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I
Transitive (sumᵣ R S i)
+
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I
Reflexive (sumᵣ R S i)
intros []; eauto.
+
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I
Symmetric (sumᵣ R S i)
intros ? ? []; econstructor; symmetry; eauto.
+
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I
Transitive (sumᵣ R S i)
I: Type X, Y: psh I R: relᵢ X X H: Equivalenceᵢ R S: relᵢ Y Y H0: Equivalenceᵢ S i: I x, y, z: (X +ᵢ Y) i H1: sumᵣ R S i x y H2: sumᵣ R S i y z
sumᵣ R S i x z
dependent destruction H1;
+ dependent destruction H2; econstructor; etransitivity; eauto.
+Qed.
+
+Definitionseqᵢ {I} {XYZ : psh I} (R0 : relᵢ X Y) (R1 : relᵢ Y Z) : relᵢ X Z :=
+ funixz => existsy, R0 i x y /\ R1 i y z.
+#[global] Infix"⨟" := (seqᵢ) (at level120).
+#[global] Notation"u ⨟⨟ v" := (ex_intro _ _ (conj u v)) (at level70).
+#[global] Hint Unfold seqᵢ : core.
+
+
I: Type X, Y, Z: psh I
Proper (leq ==> leq ==> leq) seqᵢ
+
I: Type X, Y, Z: psh I
Proper (leq ==> leq ==> leq) seqᵢ
I: Type X, Y, Z: psh I x, y: relᵢ X Y H1: x <= y x0, y0: relᵢ Y Z H2: x0 <= y0 a: I a0: X a a1: Z a z: Y a H: x a a0 z H0: x0 a z a1
(y ⨟ y0) a a0 a1
I: Type X, Y, Z: psh I x, y: relᵢ X Y H1: x <= y x0, y0: relᵢ Y Z H2: x0 <= y0 a: I a0: X a a1: Z a z: Y a H: x a a0 z H0: x0 a z a1
y a a0 z /\ y0 a z a1
I: Type X, Y, Z: psh I x, y: relᵢ X Y H1: x <= y x0, y0: relᵢ Y Z H2: x0 <= y0 a: I a0: X a a1: Z a z: Y a H: x a a0 z H0: x0 a z a1
y a a0 z
I: Type X, Y, Z: psh I x, y: relᵢ X Y H1: x <= y x0, y0: relᵢ Y Z H2: x0 <= y0 a: I a0: X a a1: Z a z: Y a H: x a a0 z H0: x0 a z a1
y0 a z a1
I: Type X, Y, Z: psh I x, y: relᵢ X Y H1: x <= y x0, y0: relᵢ Y Z H2: x0 <= y0 a: I a0: X a a1: Z a z: Y a H: x a a0 z H0: x0 a z a1
y0 a z a1
nowapply H2.Qed.
+
+Definitionsquareᵢ {I} {X : psh I} : mon (relᵢ X X) :=
+ {| body R := R ⨟ R ; Hbody _ _ H := seq_mon _ _ H _ _ H |}.
+
+Definitionrevᵢ {I} {XY : psh I} (R : relᵢ X Y) : relᵢ Y X := funixy => R i y x.
+#[global] Hint Unfold revᵢ : core.
+
+
I: Type X, Y: psh I
Proper (leq ==> leq) revᵢ
+
I: Type X, Y: psh I
Proper (leq ==> leq) revᵢ
firstorder.Qed.
+
+Definitionconverseᵢ {I} {X : psh I} : mon (relᵢ X X) :=
+ {| body := revᵢ ; Hbody := rev_mon |}.
+
+Definitionorᵢ {I} {XY : psh I} (RS : relᵢ X Y) : relᵢ X Y := funixy => R i x y \/ S i x y.
+#[global] Infix"∨ᵢ" := (orᵢ) (at level70).
+
I: Type X, Y: psh I
Proper (leq ==> leq ==> leq) orᵢ
+
I: Type X, Y: psh I
Proper (leq ==> leq ==> leq) orᵢ
firstorder.Qed.
+
+
I: Type X: psh I R: relᵢ X X
eqᵢ X <= R -> Reflexiveᵢ R
+
I: Type X: psh I R: relᵢ X X
eqᵢ X <= R -> Reflexiveᵢ R
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R i: I x: X i
R i x x
nowapply H.Qed.
+
+
I: Type X: psh I R: relᵢ X X H: Reflexiveᵢ R
eqᵢ X <= R
+
I: Type X: psh I R: relᵢ X X H: Reflexiveᵢ R
eqᵢ X <= R
intros ? ? ? ->; nowreflexivity.Qed.
+
+
I: Type X: psh I R: relᵢ X X
converseᵢ R <= R -> Symmetricᵢ R
+
I: Type X: psh I R: relᵢ X X
converseᵢ R <= R -> Symmetricᵢ R
I: Type X: psh I R: relᵢ X X H: converseᵢ R <= R i: I x, y: X i H0: R i x y
R i y x
nowapply H.Qed.
+
+
I: Type X: psh I R: relᵢ X X H: Symmetricᵢ R
converseᵢ R <= R
+
I: Type X: psh I R: relᵢ X X H: Symmetricᵢ R
converseᵢ R <= R
intros ? ? ? ?; nowsymmetry.Qed.
+
+
I: Type X: psh I R: relᵢ X X
squareᵢ R <= R -> Transitiveᵢ R
+
I: Type X: psh I R: relᵢ X X
squareᵢ R <= R -> Transitiveᵢ R
I: Type X: psh I R: relᵢ X X H: squareᵢ R <= R i: I x, y, z: X i r: R i x y s: R i y z
R i x z
apply H; nowexistsy.Qed.
+
+
I: Type X: psh I R: relᵢ X X H: Transitiveᵢ R
squareᵢ R <= R
+
I: Type X: psh I R: relᵢ X X H: Transitiveᵢ R
squareᵢ R <= R
I: Type X: psh I R: relᵢ X X H: Transitiveᵢ R a: I a0, a1, y: X a H0: R a a0 y H1: R a y a1
R a a0 a1
nowtransitivity y.Qed.
+
+
I: Type X: psh I R: relᵢ X X
eqᵢ X <= R ->
+converseᵢ R <= R -> squareᵢ R <= R -> Equivalenceᵢ R
+
I: Type X: psh I R: relᵢ X X
eqᵢ X <= R ->
+converseᵢ R <= R -> squareᵢ R <= R -> Equivalenceᵢ R
+
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R H0: converseᵢ R <= R H1: squareᵢ R <= R i: I
Reflexive (R i)
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R H0: converseᵢ R <= R H1: squareᵢ R <= R i: I
Symmetric (R i)
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R H0: converseᵢ R <= R H1: squareᵢ R <= R i: I
Transitive (R i)
+
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R H0: converseᵢ R <= R H1: squareᵢ R <= R i: I
Symmetric (R i)
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R H0: converseᵢ R <= R H1: squareᵢ R <= R i: I
Transitive (R i)
+
I: Type X: psh I R: relᵢ X X H: eqᵢ X <= R H0: converseᵢ R <= R H1: squareᵢ R <= R i: I
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Renamings
+
+
As explained in the abstract theory, renamings are a particular kind
+of assignements, where variables are mapped to variables.
+
In this file we define renamings for a given abstract context structure and provide
+all their properties that we will use throughout the development.
+
+
Definition
+
Context inclusion, or renaming is defined as an assignment of variables in Γ to
+variables in Δ.
+
Notation"Γ ⊆ Δ" := (@assignment T C CC c_var Γ%ctx Δ%ctx).
The identity inclusion, whose renaming is the identity.
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3: C x, y: Γ1 ⊆ Γ2 H1: x ≡ₐ y x0, y0: Γ2 =[ F ]> Γ3 H2: x0 ≡ₐ y0 a: T a0: Γ1 ∋ a
x0 a (x a a0) = y0 a (y a a0)
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3: C x, y: Γ1 ⊆ Γ2 H1: x ≡ₐ y x0, y0: Γ2 =[ F ]> Γ3 H2: x0 ≡ₐ y0 a: T a0: Γ1 ∋ a
y0 a (x a a0) = y0 a (y a a0)
+f_equal; apply H1.
+Qed.
+
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2: C a: Γ1 =[ F ]> Γ2
r_id ᵣ⊛ a ≡ₐ a
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2: C a: Γ1 =[ F ]> Γ2
r_id ᵣ⊛ a ≡ₐ a
reflexivity.Qed.
+
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Γ4: C u: Γ1 ⊆ Γ2 v: Γ2 ⊆ Γ3 w: Γ3 =[ F ]> Γ4
(u ᵣ⊛ v) ᵣ⊛ w ≡ₐ u ᵣ⊛ (v ᵣ⊛ w)
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Γ4: C u: Γ1 ⊆ Γ2 v: Γ2 ⊆ Γ3 w: Γ3 =[ F ]> Γ4
(u ᵣ⊛ v) ᵣ⊛ w ≡ₐ u ᵣ⊛ (v ᵣ⊛ w)
reflexivity.Qed.
The fiber of the inclusion c_var (Γ +▶ Δ) ↪ c_var Γ + c_var Δ is a subsingleton.
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2: C x: T i: Γ1 +▶ Γ2 ∋ x a, b: c_cat_view Γ1 Γ2 x i
a = b
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2: C x: T i: Γ1 +▶ Γ2 ∋ x a, b: c_cat_view Γ1 Γ2 x i
a = b
+
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C i, i0: Γ1 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_l i0) x1: r_cat_l i = r_cat_l i0 x: Vcat_l i ~= b0
Vcat_l i0 = b0
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C j: Γ2 ∋ x0 i0: Γ1 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_l i0) x1: r_cat_r j = r_cat_l i0 x: Vcat_r j ~= b0
Vcat_l i0 = b0
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C i: Γ1 ∋ x0 j: Γ2 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_r j) x1: r_cat_l i = r_cat_r j x: Vcat_l i ~= b0
Vcat_r j = b0
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C j0, j: Γ2 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_r j) x1: r_cat_r j0 = r_cat_r j x: Vcat_r j0 ~= b0
Vcat_r j = b0
+
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C i, i0: Γ1 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_l i0) x1: r_cat_l i = r_cat_l i0 x: Vcat_l i ~= b0
Vcat_l i0 = b0
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C i, i0: Γ1 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_l i0) x1: i = i0 x: Vcat_l i ~= b0
Vcat_l i0 = b0
+nowrewrite x1 in x |-; rewrite x.
+
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C j: Γ2 ∋ x0 i0: Γ1 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_l i0) x1: r_cat_r j = r_cat_l i0 x: Vcat_r j ~= b0
Vcat_l i0 = b0
symmetry in x1; destruct (r_cat_disj _ _ x1).
+
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C i: Γ1 ∋ x0 j: Γ2 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_r j) x1: r_cat_l i = r_cat_r j x: Vcat_l i ~= b0
Vcat_r j = b0
destruct (r_cat_disj _ _ x1).
+
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C j0, j: Γ2 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_r j) x1: r_cat_r j0 = r_cat_r j x: Vcat_r j0 ~= b0
Vcat_r j = b0
T, C: Type CC: context T C Γ1, Γ2: C x0: T CL: context_laws T C j0, j: Γ2 ∋ x0 b0: c_cat_view Γ1 Γ2 x0 (r_cat_r j) x1: j0 = j x: Vcat_r j0 ~= b0
Vcat_r j = b0
+nowrewrite x1 in x |-; rewrite x.
+Qed.
Simplifications of the embedding.
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2: C x: T i: Γ1 ∋ x
c_view_cat (r_cat_l i) =
+(Vcat_l i : c_cat_view Γ1 Γ2 x (r_cat_l i))
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2: C x: T i: Γ1 ∋ x
c_view_cat (r_cat_l i) =
+(Vcat_l i : c_cat_view Γ1 Γ2 x (r_cat_l i))
apply view_cat_irr.Qed.
+
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2: C x: T i: Γ2 ∋ x
c_view_cat (r_cat_r i) =
+(Vcat_r i : c_cat_view Γ1 Γ2 x (r_cat_r i))
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2: C x: T i: Γ2 ∋ x
c_view_cat (r_cat_r i) =
+(Vcat_r i : c_cat_view Γ1 Γ2 x (r_cat_r i))
apply view_cat_irr.Qed.
Simplifying copairing.
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ
r_cat_l ᵣ⊛ [u, v] ≡ₐ u
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: (Γ1 +▶ Γ2) =[ F ]> Δ H1: u ≡ₐ r_cat_l ᵣ⊛ w H2: v ≡ₐ r_cat_r ᵣ⊛ w
[u, v] ≡ₐ w
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: (Γ1 +▶ Γ2) =[ F ]> Δ H1: u ≡ₐ r_cat_l ᵣ⊛ w H2: v ≡ₐ r_cat_r ᵣ⊛ w
[u, v] ≡ₐ w
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: (Γ1 +▶ Γ2) =[ F ]> Δ H1: u ≡ₐ r_cat_l ᵣ⊛ w H2: v ≡ₐ r_cat_r ᵣ⊛ w a: T i: Γ1 +▶ Γ2 ∋ a
match c_view_cat i with
+| Vcat_l i => u a i
+| Vcat_r j => v a j
+end = w a i
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ
r_cat3_1 ᵣ⊛ [u, [v, w]] ≡ₐ [u, v]
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ
r_cat3_1 ᵣ⊛ [u, [v, w]] ≡ₐ [u, v]
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l i) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = u a i
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T j: Γ2 ∋ a
match c_view_cat (r_cat_r (r_cat_l j)) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = v a j
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l i) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = u a i
nowrewrite c_view_cat_simpl_l.
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T j: Γ2 ∋ a
match c_view_cat (r_cat_r (r_cat_l j)) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = v a j
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ
r_cat3_2 ᵣ⊛ [u, [v, w]] ≡ₐ [u, w]
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ
r_cat3_2 ᵣ⊛ [u, [v, w]] ≡ₐ [u, w]
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l i) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = u a i
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r (r_cat_r j)) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = w a j
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l i) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = u a i
nowrewrite c_view_cat_simpl_l.
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r (r_cat_r j)) with
+| Vcat_l i => u a i
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => v a i
+ | Vcat_r j0 => w a j0
+ end
+end = w a j
nowrewrite2 c_view_cat_simpl_r.
+Qed.
+
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ
r_cat3_3 ᵣ⊛ [[u, v], w] ≡ₐ [v, w]
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ
r_cat3_3 ᵣ⊛ [[u, v], w] ≡ₐ [v, w]
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T i: Γ2 ∋ a
match c_view_cat (r_cat_l (r_cat_r i)) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => u a i0
+ | Vcat_r j => v a j
+ end
+| Vcat_r j => w a j
+end = v a i
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r j) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => u a i0
+ | Vcat_r j => v a j
+ end
+| Vcat_r j => w a j
+end = w a j
+
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T i: Γ2 ∋ a
match c_view_cat (r_cat_l (r_cat_r i)) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => u a i0
+ | Vcat_r j => v a j
+ end
+| Vcat_r j => w a j
+end = v a i
T, C: Type CC: context T C CL: context_laws T C F: Fam₁ T C Γ1, Γ2, Γ3, Δ: C u: Γ1 =[ F ]> Δ v: Γ2 =[ F ]> Δ w: Γ3 =[ F ]> Δ a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r j) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => u a i0
+ | Vcat_r j => v a j
+ end
+| Vcat_r j => w a j
+end = w a j
nowrewrite c_view_cat_simpl_r.
+Qed.
These last two are pretty interesting, they are the proofs witnessing the associativity
+isomorphism Γ₁ +▶ (Γ₂ +▶ Γ₃) ≈ (Γ₁ +▶ Γ₂) +▶ Γ₃. Here isomorphism means isomorphism
+of the set of variables.
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C
r_assoc_l ᵣ⊛ r_assoc_r ≡ₐ r_id
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C
r_assoc_l ᵣ⊛ r_assoc_r ≡ₐ r_id
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l (r_cat_l i)) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_l i
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ2 +▶ Γ3 ∋ a
match
+ c_view_cat
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j => r_cat_r j
+ end
+with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_r j
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l (r_cat_l i)) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_l i
nowrewrite2 c_view_cat_simpl_l.
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ2 +▶ Γ3 ∋ a
match
+ c_view_cat
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j => r_cat_r j
+ end
+with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_r j
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ2 ∋ a
match c_view_cat (r_cat_l (r_cat_r i)) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_r (r_cat_l i)
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r j) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_r (r_cat_r j)
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ2 ∋ a
match c_view_cat (r_cat_l (r_cat_r i)) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_r (r_cat_l i)
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r j) with
+| Vcat_l i =>
+ match c_view_cat i with
+ | Vcat_l i0 => r_cat_l i0
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+| Vcat_r j => r_cat_r (r_cat_r j)
+end = r_cat_r (r_cat_r j)
nowrewrite c_view_cat_simpl_r.
+Qed.
+
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C
r_assoc_r ᵣ⊛ r_assoc_l ≡ₐ r_id
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C
r_assoc_r ᵣ⊛ r_assoc_l ≡ₐ r_id
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ1 +▶ Γ2 ∋ a
match
+ c_view_cat
+ match c_view_cat i with
+ | Vcat_l i => r_cat_l i
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_l i
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ3 ∋ a
match c_view_cat (r_cat_r (r_cat_r j)) with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_r j
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ1 +▶ Γ2 ∋ a
match
+ c_view_cat
+ match c_view_cat i with
+ | Vcat_l i => r_cat_l i
+ | Vcat_r j => r_cat_r (r_cat_l j)
+ end
+with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_l i
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l i) with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_l (r_cat_l i)
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ2 ∋ a
match c_view_cat (r_cat_r (r_cat_l j)) with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_l (r_cat_r j)
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T i: Γ1 ∋ a
match c_view_cat (r_cat_l i) with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_l (r_cat_l i)
nowrewrite c_view_cat_simpl_l.
+
T, C: Type CC: context T C CL: context_laws T C Γ1, Γ2, Γ3: C a: T j: Γ2 ∋ a
match c_view_cat (r_cat_r (r_cat_l j)) with
+| Vcat_l i => r_cat_l (r_cat_l i)
+| Vcat_r j =>
+ match c_view_cat j with
+ | Vcat_l i => r_cat_l (r_cat_r i)
+ | Vcat_r j0 => r_cat_r j0
+ end
+end = r_cat_l (r_cat_r j)
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
A Minimal Example: Call-by-Value Simply Typed Lambda Calculus
+
+
We demonstrate how to instantiate our framework to define the OGS associated
+to the CBV λ-calculus. With the instance comes the proof that bisimilarity of
+the OGS entails substitution equivalence, which coincides with
+CIU-equivalence.
+
+
Note
+
This example is designed to be minimal, hiding by nature some
+difficulties. In particular it has no positive types, which simplifies the
+development a lot.
+
+
+
Syntax
+
+
Types
+
As discussed in the paper, our framework applies not really to a language but
+more to an abstract machine. In order to ease this instanciation, we will
+focus directly on a CBV machine and define evaluation contexts early on.
+Working with intrinsically typed syntax, we need to give types to these
+contexts: we will type them by the "formal negation" of the type of their
+hole. In order to do so we first define the usual types ty0 of STLC with
+functions and a ground type.
Typing contexts are now simply defined as lists of tagged types. This is
+perhaps the slightly surprising part: contexts will now contain "continuation
+variables". Rest assured terms will make no use of them. These are solely
+needed for evaluation contexts: as they are only typed with their hole, we are
+missing the type of their outside. We fix this problem by naming the outside
+and make evaluation contexts end with a continuation variable.
+
+
Note
+
Our choices make a bit more sense if we realize that what we are
+writing is exactly the subset of λμ-calculus that is the image of the CBV
+translation from λ-calculus.
+
+
+
The reference ctx was not found in the current
+environment.
+
+
Terms and Values and ...
+
We now have all that is needed to define terms, which we define mutually with
+values. As discussed above, they are are indexed by a list of tagged types (of
+which they only care about the non-negated elements) and by an untagged type.
+
The only fanciness is the recursive lambda abstraction, which we include to
+safeguard ourselves from accidentally using the fact that the language is
+strongly normalizing.
+
Inductiveterm (Γ : t_ctx) : ty0 -> Type :=
+| Val {a} : val Γ a -> term Γ a
+| App {a b} : term Γ (a → b) -> term Γ a -> term Γ b
+with val (Γ : t_ctx) : ty0 -> Type :=
+| Var {a} : Γ ∋ + a -> val Γ a
+| TT : val Γ ι
+| Lam {a b} : term (Γ ▶ₓ (a → b) ▶ₓ a) b -> val Γ (a → b)
+.
Evaluation contexts follow. As discussed, there is a "covariable" case, to
+"end" the context. The other cases are usual: evaluating the argument of an
+application and evaluating the function.
+
Inductiveev_ctx (Γ : t_ctx) : ty0 -> Type :=
+| K0 {a} : Γ ∋ ¬ a -> ev_ctx Γ a
+| K1 {a b} : term Γ (a → b) -> ev_ctx Γ b -> ev_ctx Γ a
+| K2 {a b} : val Γ a -> ev_ctx Γ b -> ev_ctx Γ (a → b)
+.
Next are the machine states. They consist of an explicit cut and represent
+a term in the process of being executed in some context. Notice how states
+they are only indexed by a typing context: these are proper "diverging
+programs".
+
Definitionstate := term ∥ₛ ev_ctx.
Finally the last of syntactic categories: "machine values". This category is
+typed with a tagged type and encompasses both values (for non-negated types)
+and evaluation contexts (for negated types). The primary use-case for this
+category is to denote "things by which we can substitute (tagged) variables".
+
In fact when working with intrinsically-typed syntax, substitution is modelled
+as a monoidal multiplication (see [AACMM21] and [FS22]). We will prove later
+that val_mΓ is indeed a monoid relative to hasΓ and that contexts and
+assignments form a category.
+
Equationsval_m : t_ctx -> ty -> Type :=
+ val_m Γ (+ a) := val Γ a ;
+ val_m Γ (¬ a) := ev_ctx Γ a .
Together with machine values we define a smart constructor for "machine var",
+embedding tagged variables into these generalized values. It conveniently
+serves as the identity assignment, a fact we use to give it this mysterious
+point-free style type, which desugars to foralla,Γ∋a->val_mΓa.
+
Equationsa_id {Γ} : Γ =[val_m]> Γ :=
+ a_id (+ _) i := Var i ;
+ a_id (¬ _) i := K0 i .
+
+
+
Substitution and Renaming
+
In order to define substitution we first need to dance the intrinsically-typed
+dance and define renamings, from which we will derive weakenings and then only
+define substitution.
+
+
Renaming
+
Lets write intrinsically-typed parallel renaming for all of our syntactic
+categories! If you have never seen such intrinsically-typed definitions you
+might be surprised by the absence of error-prone de-bruijn index manipulation.
+Enjoy this beautiful syntax traversal!
+
Equationst_rename {ΓΔ} : Γ ⊆ Δ -> term Γ ⇒ᵢ term Δ :=
+ t_rename f _ (Val v) := Val (v_rename f _ v) ;
+ t_rename f _ (App t1 t2) := App (t_rename f _ t1) (t_rename f _ t2) ;
+with v_rename {Γ Δ} : Γ ⊆ Δ -> val Γ ⇒ᵢ val Δ :=
+ v_rename f _ (Var i) := Var (f _ i) ;
+ v_rename f _ (TT) := TT ;
+ v_rename f _ (Lam t) := Lam (t_rename (r_shift2 f) _ t) .
+
+Equationse_rename {ΓΔ} : Γ ⊆ Δ -> ev_ctx Γ ⇒ᵢ ev_ctx Δ :=
+ e_rename f _ (K0 i) := K0 (f _ i) ;
+ e_rename f _ (K1 t π) := K1 (t_rename f _ t) (e_rename f _ π) ;
+ e_rename f _ (K2 v π) := K2 (v_rename f _ v) (e_rename f _ π) .
+
+Equationss_rename {ΓΔ} : Γ ⊆ Δ -> state Γ -> state Δ :=
+ s_rename f (Cut t π) := Cut (t_rename f _ t) (e_rename f _ π).
+
+Equationsm_rename {ΓΔ} : Γ ⊆ Δ -> val_m Γ ⇒ᵢ val_m Δ :=
+ m_rename f (+ _) v := v_rename f _ v ;
+ m_rename f (¬ _) π := e_rename f _ π .
+
+#[global] Arguments s_rename _ _ _ /.
+#[global] Arguments m_rename _ _ _ /.
We can recast m_rename as an operator on assigments, more precisely as
+renaming an assigment on the left.
As discussed above, we can now obtain our precious weakenings. Here are the
+three we will need.
+
Definitiont_shift {Γa} := @t_rename Γ (Γ ▶ₓ a) r_pop.
+Definitionm_shift2 {Γab} := @m_rename Γ (Γ ▶ₓ a ▶ₓ b) (r_pop ᵣ⊛ r_pop).
+Definitiona_shift2 {ΓΔab} (f : Γ =[val_m]> Δ)
+ : (Γ ▶ₓ a ▶ₓ b) =[val_m]> (Δ ▶ₓ a ▶ₓ b)
+ := [ [ a_map f m_shift2 ,ₓ a_id a (pop top) ] ,ₓ a_id b top ].
+
+
Substitutions
+
With weakenings in place we are now equipped to define substitutions. This
+goes pretty much like renaming. We have abstained from defining generic syntax
+traversal tools like Allais et al's "Kits" to keep our example minimal... And
+incidentally showcase the intrinsically-typed style with Matthieu Sozeau's
+Equations.
+
Equationst_subst {ΓΔ} : Γ =[val_m]> Δ -> term Γ ⇒ᵢ term Δ :=
+ t_subst f _ (Val v) := Val (v_subst f _ v) ;
+ t_subst f _ (App t1 t2) := App (t_subst f _ t1) (t_subst f _ t2) ;
+with v_subst {Γ Δ} : Γ =[val_m]> Δ -> val Γ ⇒ᵢ val Δ :=
+ v_subst f _ (Var i) := f _ i ;
+ v_subst f _ (TT) := TT ;
+ v_subst f _ (Lam t) := Lam (t_subst (a_shift2 f) _ t) .
+
+Equationse_subst {ΓΔ} : Γ =[val_m]> Δ -> ev_ctx Γ ⇒ᵢ ev_ctx Δ :=
+ e_subst f _ (K0 i) := f _ i ;
+ e_subst f _ (K1 t π) := K1 (t_subst f _ t) (e_subst f _ π) ;
+ e_subst f _ (K2 v π) := K2 (v_subst f _ v) (e_subst f _ π) .
These two last substitutions, for states and generalized values, are the ones
+we will give to the abstract interface. For categorical abstraction reasons,
+the abstract interface has the argument reversed for substitutions: first
+the state or value, then the assignment as second argument. To ease instanciation
+we will hence do so too. The type is given with tricky combinators but expands to:
+
+
m_subst Γ a : val_m Γ a -> forall Δ, Γ =[val_m]> Δ -> val_m Δ a
+
Equationsm_subst : val_m ⇒₁ ⟦ val_m , val_m ⟧₁ :=
+ m_subst _ (+ _) v _ f := v `ᵥ⊛ f ;
+ m_subst _ (¬ _) π _ f := π `ₑ⊛ f .
+
+Definitions_subst : state ⇒₀ ⟦ val_m , state ⟧₀ :=
+ fun_s_f => (s.(cut_l) `ₜ⊛ f) ⋅ (s.(cut_r) `ₑ⊛ f).
+#[global] Arguments s_subst _ _ /.
+
We can now instanciate the substitution monoid and module structures.
Finally we define a more usual substitution function which only substitutes
+the top two variables instead of doing parallel substitution.
+
Definitionassign2 {Γab} v1v2
+ : (Γ ▶ₓ a ▶ₓ b) =[val_m]> Γ
+ := [ [ a_id ,ₓ v1 ] ,ₓ v2 ] .
+Definitiont_subst2 {Γabc} (t : term _ c) v1v2 := t `ₜ⊛ @assign2 Γ a b v1 v2.
+Notation"t /[ v1 , v2 ]" := (t_subst2 t v1 v2) (at level50, left associativity).
+
+
+
An Evaluator
+
As motivated earlier, the evaluator will be a defined as a state-machine, the
+core definition being a state-transition function. To stick to the intrinsic
+style, we wish that this state-machine stops only at evidently normal forms,
+instead of stoping at states which happen to be in normal form. Perhaps more
+simply said, we want to type the transition function as stateΓ→(state
+Γ+nfΓ), where returning in the right component means stoping and outputing
+a normal form.
+
In order to do this we first need to define normal forms! But instead of
+defining an inductive definition of normal forms as is customary, we will take
+an other route more tailored to OGS based on ultimate patterns.
+
+
Patterns and Observations
+
As discussed in the paper, a central aspect of OGS is to circumvent the
+problem of naive trace semantics for higher-order languages by mean of
+a notion of "abstract values", or more commonly ultimate patterns, which
+define the "shareable part" of a value. For clarity reasons, instead of
+patterns we here take the dual point of view of "observations", which can be
+seen as patterns at the dual type (dual in the sense of swapping the tag). For
+our basic λ-calculus, all types are negative -- that is, unsheareble -- hence
+the observations are pretty simple:
+
+
Observing a continuation of type ¬a means returning a (hidden) value to
+it. In terms of pattern this corresponds to the "fresh variable" pattern
+Varx.
+
+
Observing a function of type a→b means giving it a hidden value and
+a hidden continuation. In terms of patterns this corresponds to the
+"application" co-pattern K2(Varx)(K0y).
+
Variantobs : ty -> Type :=
+| ORet {a} : obs (¬ a)
+| OApp {a b} : obs (a → b)
+.
+
+
As observations correspond to a subset of (machine) values where all the
+variables are "fresh", these variables have no counter-part in de-bruijn
+notation (there is no meaningful notion of freshness). As such we have not
+indexed obs by any typing context, but to complete the definition we now
+need to define a projection into typing contexts denoting the "domain",
+"support" or more accurately "set of nameless fresh variables" of an
+observation.
+
Equationsobs_dom {a} : obs a -> t_ctx :=
+ obs_dom (@ORet a) := ∅ ▶ₓ + a ;
+ obs_dom (@OApp a b) := ∅ ▶ₓ + a ▶ₓ ¬ b .
Given a value, an observation on its type and a value for each fresh variable
+of the observation, we can "refold" everything and form a new state which will
+indeed observe this value.
+
Equationsobs_app {Γa} (v : val_m Γ a) (p : obs a) (γ : obs_dom p =[val_m]> Γ) : state Γ :=
+ obs_app v OApp γ := Val v ⋅ K2 (γ _ (pop top)) (γ _ top) ;
+ obs_app π ORet γ := Val (γ _ top) ⋅ (π : ev_ctx _ _) .
+
+
Normal forms
+
Normal forms for CBV λ-calculus can be characterized as either a value v or
+a stuck application in evaluation context E[xv] (see "eager-normal forms"
+in [L05]). Now it doesn't take much squinting to see that in our setting,
+this corresponds respectively to states of the form ⟨v|K0x⟩ and ⟨Var
+x|K2vπ⟩. Squinting a bit more, we can reap the benefits of our unified
+treatment of terms and contexts and see that both of these cases work in the
+same way: normal states are states given by a variable facing an observation
+whose fresh variables have been substituted with values.
+
Having already defined observation and their set of fresh variables, an
+observation stuffed with values in typing context Γ can be represented
+simply by a formal substitution of an observation o and an assigment
+obs_domo=[val]>Γ. This "split" definition of normal forms is the one we
+will take.
+
Definitionnf := c_var ∥ₛ (obs_op # val_m).
+
+
The CBV Machine
+
Everything is now in place to define our state transition function. The
+reduction rules should come to no surprise:
+
+
⟨t1t2|π⟩→⟨t2|t1⋅1π⟩
+
⟨v|x⟩ normal
+
⟨v|t⋅1π⟩→⟨t|v⋅2π⟩
+
⟨x|v⋅2π⟩ normal
+
⟨λfx.t|v⋅2π⟩→⟨t[f↦λfx.t;x↦v]|π⟩
+
+
Rules 1,3,5 step to a new configuration, while cases 2,4 are stuck on normal
+forms.
+
Equationseval_step {Γ : t_ctx} : state Γ -> sum (state Γ) (nf Γ) :=
+ eval_step (Cut (App t1 t2) (π)) := inl (Cut t2 (K1 t1 π)) ;
+ eval_step (Cut (Val v) (K0 i)) := inr (i ⋅ ORet' ⦇ [ ! ,ₓ v ] ⦈) ;
+ eval_step (Cut (Val v) (K1 t π)) := inl (Cut t (K2 v π)) ;
+ eval_step (Cut (Val (Var i)) (K2 v π)) := inr (i ⋅ OApp' ⦇ [ [ ! ,ₓ v ] ,ₓ π ] ⦈) ;
+ eval_step (Cut (Val (Lam t)) (K2 v π)) := inl (Cut (t /[ Lam t , v ]) π) .
Having defined the transition function, we can now iterate it inside the delay
+monad. This constructs a possibly non-terminating computation ending with
+a normal form.
We have now finished all the computational parts of the instanciation, but all
+the proofs are left to be done. Before attacking the OGS-specific hypotheses,
+we will need to prove the usual standard lemmata on renaming and substitution.
+
There will be a stack of lemmata which will all pretty much be simple
+inductions on the syntax, so we start by introducing some helpers for this. In
+fact it is not completely direct to do since terms and values are mutually
+defined: we will need to derive a mutual induction principle.
+
Schemeterm_mut := Induction for term SortProp
+ with val_mut := Induction for val SortProp .
Annoyingly, Coq treats this mutual induction principle as two separate
+induction principles. They both have the exact same premises but differ in
+their conclusion. Thus we define a datatype for these premises, to avoid
+duplicating the proofs. Additionally, evaluation contexts are not defined
+mutually with terms and values, but it doesn't hurt to prove their properties
+simultaneously too, so syn_ind_args is in fact closer to the premises of
+a three-way mutual induction principle between terms, values and evaluation
+contexts.
+
Recordsyn_ind_args (P_t : forallΓA, term Γ A -> Prop)
+ (P_v : forallΓA, val Γ A -> Prop)
+ (P_e : forallΓA, ev_ctx Γ A -> Prop) :=
+ {
+ ind_val {Γ a} v (_ : P_v Γ a v) : P_t Γ a (Val v) ;
+ ind_app {Γ a b} t1 (_ : P_t Γ (a → b) t1) t2 (_ : P_t Γ a t2) : P_t Γ b (App t1 t2) ;
+ ind_var {Γ a} i : P_v Γ a (Var i) ;
+ ind_tt {Γ} : P_v Γ (ι) (TT) ;
+ ind_lamrec {Γ a b} t (_ : P_t _ b t) : P_v Γ (a → b) (Lam t) ;
+ ind_kvar {Γ a} i : P_e Γ a (K0 i) ;
+ ind_kfun {Γ a b} t (_ : P_t Γ (a → b) t) π (_ : P_e Γ b π) : P_e Γ a (K1 t π) ;
+ ind_karg {Γ a b} v (_ : P_v Γ a v) π (_ : P_e Γ b π) : P_e Γ (a → b) (K2 v π)
+ } .
+
+Lemmaterm_ind_mutP_tP_vP_e (H : syn_ind_args P_t P_v P_e) Γat : P_t Γ a t .
+destruct H; nowapply (term_mut P_t P_v).
+Qed.
+
+Lemmaval_ind_mutP_tP_vP_e (H : syn_ind_args P_t P_v P_e) Γav : P_v Γ a v .
+destruct H; nowapply (val_mut P_t P_v).
+Qed.
+
+Lemmactx_ind_mutP_tP_vP_e (H : syn_ind_args P_t P_v P_e) Γaπ : P_e Γ a π .
+induction π.
+-apply (ind_kvar _ _ _ H).
+-apply (ind_kfun _ _ _ H); auto; apply (term_ind_mut _ _ _ H).
+-apply (ind_karg _ _ _ H); auto; apply (val_ind_mut _ _ _ H).
+Qed.
Now equipped we can start with the first lemma: renaming respects pointwise
+equality of assignments. As discussed, we will prove this by mutual induction
+on our three "base" syntactic categories of terms, values and evaluation
+contexts, and then we will also deduce it for the three "derived" notions of
+machine values, states and assigments. Sometimes some of the derived notions
+will be omitted if it is not needed later on.
+
This proof, like all the following ones will follow a simple pattern:
+a simplification; an application of congruence; a fixup for the two-time
+shifted assigment in the case of λ; finally a call to the induction
+hypothesis.
+
+
Note
+
Here is definitely where the generic syntax traversal kit of
+Guillaume Allais et al would shine. Indeed the proof pattern i outlined can
+really be formalized into a generic proof.
Lemma 2: renaming-renaming assocativity. I say "associativity" because it
+definitely looks like associativity if we disregard the subscripts. More
+precisely it could be described as the composition law a right action.
Definitiont_ren_id_l_PΓa (t : term Γ a) : Prop := t ₜ⊛ᵣ r_id = t .
+Definitionv_ren_id_l_PΓa (v : val Γ a) : Prop := v ᵥ⊛ᵣ r_id = v .
+Definitione_ren_id_l_PΓa (π : ev_ctx Γ a) : Prop := π ₑ⊛ᵣ r_id = π.
+Lemmaren_id_l_prf : syn_ind_args t_ren_id_l_P v_ren_id_l_P e_ren_id_l_P .
+ econstructor.
+ all: unfold t_ren_id_l_P, v_ren_id_l_P, e_ren_id_l_P.
+ all: intros; cbn; f_equal; auto.
+ rewrite <- H at2; apply t_ren_eq; auto.
+ intros ? v; nowdo2 (dependent elimination v; eauto).
+Qed.
+
+Lemmat_ren_id_l {Γ} a (t : term Γ a) : t ₜ⊛ᵣ r_id = t .
+ nowapply (term_ind_mut _ _ _ ren_id_l_prf).
+Qed.
+Lemmav_ren_id_l {Γ} a (v : val Γ a) : v ᵥ⊛ᵣ r_id = v.
+ nowapply (val_ind_mut _ _ _ ren_id_l_prf).
+Qed.
+Lemmae_ren_id_l {Γ} a (π : ev_ctx Γ a) : π ₑ⊛ᵣ r_id = π .
+ nowapply (ctx_ind_mut _ _ _ ren_id_l_prf).
+Qed.
+Lemmam_ren_id_l {Γ} a (v : val_m Γ a) : v ₘ⊛ᵣ r_id = v .
+ destruct a; [ nowapply v_ren_id_l | nowapply e_ren_id_l ].
+Qed.
+Lemmas_ren_id_l {Γ} (s : state Γ) : s ₛ⊛ᵣ r_id = s.
+ destruct s; apply (f_equal2 Cut); [ nowapply t_ren_id_l | nowapply e_ren_id_l ].
+Qed.
Lemma 4: right identity law of renaming. This one basically holds
+definitionally, it only needs a case split for some of the derived notions. We
+will also prove a consequence on weakenings: identity law.
+
Lemmam_ren_id_r {ΓΔ} (f : Γ ⊆ Δ) {a} (i : Γ ∋ a) : a_id _ i ₘ⊛ᵣ f = a_id _ (f _ i) .
+ nowdestruct a.
+Qed.
+Lemmaa_ren_id_r {ΓΔ} (f : Γ ⊆ Δ) : a_id ⊛ᵣ f ≡ₐ f ᵣ⊛ a_id .
+ intros ??; nowapply m_ren_id_r.
+Qed.
+Lemmaa_shift2_id {Γxy} : @a_shift2 Γ Γ x y a_id ≡ₐ a_id.
+ intros ? v; do2 (dependent elimination v; auto).
+ exact (m_ren_id_r _ _).
+Qed.
Lemma 5: shifting assigments commutes with left and right renaming.
Having proved all the basic syntactic properties of STLC, we are now ready to
+instanciate our framework!
+
As we only have negative types, we instanciate the interaction specification
+with types and observations. Beware that in more involved cases, the notion of
+"types" we give to the OGS construction does not coincide with the
+"language types": you should only give the negative types, or more intuitively,
+"non-shareable" types.
+
As hinted at the beginning, we instanciate the abstract value notion with our
+"machine values". They form a suitable monoid, which means we get a category
+of assigments, for which we now provide the proof of the laws.
In our generic theorem, there is a finicky lemma that is the counter-part to
+the exclusion of any "infinite chit-chat" that one finds in other accounts of
+OGS and other game semantics. The way we have proved it requires a little bit
+more structure on values. Specifically, we need to show that a_id is
+injective and that its fibers are decidable and invert renamings. These
+technicalities are easily shown by induction on values but help us to
+distinguish conveniently between values which are variables and others.
We now instanciate the machine with stlc_eval as the active step ("compute
+the next observable action") and obs_app as the passive step ("resume from
+a stuck state").
All that is left is to prove our theorem-specific hypotheses. All but another
+technical lemma for the chit-chat problem are again coherence conditions
+between eval and app and the monoidal structure of values and
+configurations.
+
#[global] Instancestlc_machine_law : machine_laws val_m state obs_op.
+econstructor; cbn; intros.
The first one proves that obs_app respects pointwise equality of assigments.
The meat of our abstract proof is this next one. We need to prove that our
+evaluator respects substitution in a suitable sense: evaluating a substituted
+configuration must be the same thing as evaluating the configuration, then
+"substituting" the normal form and continuing the evaluation.
+
While potentially scary, the proof is direct and this actually amount to
+checking that indeed, when unrolling our evaluator, this is what happens.
+
-revert c a; unfold comp_eq, it_eq; coinduction R CIH; intros c e.
+cbn; funelim (eval_step c); cbn.
++destruct (e ¬ a0 i); auto.
+remember (v `ᵥ⊛ e) as v'; clear H v Heqv'.
+ dependent elimination v'; cbn; auto.
+ + econstructor; refold_eval; apply CIH.
+ + remember (e + (a2 → b0) i) as vv; clear H i Heqvv.
+ dependent elimination vv; cbn; auto.
+ + econstructor;
+ refold_eval;
+ change (Lam (t `ₜ⊛ _)) with (Lam t `ᵥ⊛ e);
+ change (?v `ᵥ⊛ ?a) with ((v : val_m _ (+ _)) ᵥ⊛ a).
+ rewrite t_sub2_sub.
+ apply CIH.
+ + econstructor; refold_eval; apply CIH.
Just like the above proof had the flavor of a composition law of module, this
+one has the flavor of an identity law. It states that evaluating a normal form
+is the identity computation.
+
-destruct u as [ a i [ p γ ] ]; cbn.
+dependent elimination p; cbn.
+all: unfold comp_eq; apply it_eq_unstep; cbn; econstructor.
+all: econstructor; intros ? v; do3 (dependent elimination v; auto).
This last proof is the technical condition we hinted at. It is a proof of
+well-foundedness of some relation, and what it amounts to is that if we
+repeatedly instantiate the head variable of a normal form by a value which is
+not a variable, after a finite number of times doing so we will eventually
+reach something that is not a normal form.
+
For our calculus this number is at most 2, the pathological state being
+⟨x|y⟩, which starts by being stuck on y, but when instanciating by
+some non-variable π, ⟨x|π⟩ is still stuck, this time on x. After
+another step it will definitely be unstuck and evaluation will be able to do
+a reduction step.
+
It is slightly tedious to prove but amount again to a "proof by case
+splitting".
+
-intros [ x p ].
+destruct x; dependent elimination p; econstructor.
+*intros [ z p ] H.
+dependent elimination p; dependent elimination H.
+all: dependent elimination v; trynowdestruct (t0 (Vvar _)).
+all: apply it_eq_step in i0; nowinversion i0.
+*intros [ z p ] H.
+dependent elimination p; dependent elimination H; cbnin *.
+dependent elimination v; trynowdestruct (t0 (Vvar _)).
++apply it_eq_step in i0; nowinversion i0.
++remember (a1 _ Ctx.top) as vv; clear a1 Heqvv.
+dependent elimination vv;
+ apply it_eq_step in i0; cbnin i0; dependent elimination i0.
+inversion r_rel.
++econstructor; intros [ z p ] H.
+dependent elimination p; dependent elimination H.
+all: dependent elimination v0; trynowdestruct (t1 (Vvar _)).
+all: apply it_eq_step in i2; nowinversion i2.
+Qed.
At this point we have finished all the hard work! We already enjoy the generic
+correctness theorem but don't know it yet! Lets define some shorthands for
+some generic notions applied to our case, to make it a welcoming nest.
+
The whole semantic is parametrized by typing scope Δ of "final channels".
+Typically this can be instanciated with the singleton [¬ans] for some
+chosen type ans, which will correspond with the outside type of the
+testing-contexts from CIU-equivalence. Usually this answer type is taken among
+the positive (or shareable) types of our language, but in fact using our
+observation machinery we can project the value of any type onto its "shareable
+part". This is why our generic proof abstracts over this answer type and even
+allows several of them at the same time (that is, Δ). In our case, as all
+types in our language are unshareable, the positive part of any value is
+pretty useless: it is always a singleton. Yet our notion of testing still
+distinguishes terminating from non-terminating programs.
+
As discussed in the paper, the "native output" of the generic theorem is
+correctness with respect to an equivalence we call "substitution equivalence".
+We will recover a more standard CIU later on.
Our semantic objects live in what is defined in the generic construction as
+ogs_act, that is active strategies for the OGS game. They come with their
+own notion of equivalence, weak bisimilarity and we get to interpret states
+into semantic objects.
We can now obtain our instance of the correctness result!
+
Theoremstlc_subst_correctΔ {Γ} (xy : state Γ)
+ : OGS⟦x⟧ ≈[Δ]≈ OGS⟦y⟧ -> x ≈S[Δ]≈ y .
+ exact (ogs_correction Δ x y).
+Qed.
+
Recovering CIU-equivalence
+
CIU-equivalence more usually defined as a relation on terms (and not some
+states), and involves an evaluation context. In our formalism it amounts to
+the following definition.
Now from a term we can always construct a state by naming it, that is, placing
+the term opposite of a fresh context variable.
+
Definitionc_init {Γa} (t : term Γ a) : state (Γ ▶ₓ ¬ a)
+ := t_shift _ t ⋅ K0 Ctx.top .
+Notation"T⟦ t ⟧" := (OGS⟦ c_init t ⟧) .
Similarly, from an evaluation context and a substitution, we can form an
+extended substitution. Without surprise these two constructions simplify well
+in terms of substitution.
We can now obtain a correctness theorem with respect to standard
+CIU-equivalence by embedding terms into states. Proving that CIU-equivalence
+entails our substitution equivalence is left to the reader!
+
Theoremstlc_ciu_correctΔ {Γa} (xy : term Γ a)
+ : T⟦ x ⟧ ≈[Δ]≈ T⟦ y ⟧ -> x ≈C[Δ]≈ y .
+ intros H σ k; rewrite2 sub_init.
+ nowapply stlc_subst_correct.
+Qed.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Soundness (Theorem 8)
+
+
Finally, all the pieces are in place to prove that bisimilarity of induced LTS is sound
+w.r.t. substitution equivalence. Having worked hard to establish
+adequacy and congruence
+of weak bisimilarity for composition, very little remains to do here.
+
We consider a language abstractly captured as a machine satisfying an
+appropriate axiomatization.
+
Context {TC} {CC : context T C} {CL : context_laws T C}.
+ Context {val} {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+ Context {conf} {CM : subst_module val conf} {CML : subst_module_laws val conf}.
+ Context {obs : obs_struct T C} {M : machine val conf obs} {ML : machine_laws val conf obs}.
+ Context {VV : var_assumptions val}.
We define substitution equivalence of two language machine configurations (Def. 15).
Our main theorem: bisimilarity of induced OGS machine strategies is sound w.r.t.
+substitution equivalence, by applying barbed equivalence soundness, swapping the naive
+composition with the opaque one and then applying congruence.
+
Theoremogs_correction {Γ} Δ (xy : conf Γ) : x ≈⟦ogs Δ⟧≈ y -> x ≈⟦sub Δ⟧≈ y.
+ Proof.
+ intros H γ; unfold m_conf_eqv in H.
+ nowrewrite2 adequacy, H.
+ Qed.
+Endwith_param.
If you wish to double check these results you can run the following commands at
+this point in the file:
+
ogs_correction not a defined object.
+
The reference ogs_correction was not found
+in the current environment.
The first command will explicit the assumptions of the theorem, which we show how
+to provide with several examples:
The second command will explicit if any axiom has been used to establish the
+result. As stated in the prelude, we exclusively use
+Eq_rect_eq.eq_rect_eq, ie Streicher's axiom K.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
+
+
+
The Machine Strategy (§ 5.3)
+
Having defined the OGS game and axiomatized the
+language machine, we are now ready to construct the machine strategy.
+
We consider a language abstractly captured as a machine.
+
Context `{CC : context T C} {CL : context_laws T C}.
+ Context {val} {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+ Context {conf} {CM : subst_module val conf} {CML : subst_module_laws val conf}.
+ Context {obs : obs_struct T C} {M : machine val conf obs}.
We start off by defining active and passive OGS assignments (Def 5.16). This datastructure
+will hold the memory part of a strategy state, remembering the values which we have
+hidden from Opponent and given as fresh variables.
Next we define the collapsing function on OGS assignments (Def 5.18).
+
Equationscollapse {ΔpΦ} : ogs_env Δ p Φ -> ↓[p^]Φ =[val]> (Δ +▶ ↓[p]Φ) :=
+ collapse ε⁺ := ! ;
+ collapse ε⁻ := ! ;
+ collapse (u ;⁺) := collapse u ⊛ᵣ r_cat3_1 ;
+ collapse (u ;⁻ e) := [ collapse u , e ] .
+ Notation"ₐ↓ γ" := (collapse γ).
We readily define the zipping function (Def 6.10).
+
#[derive(eliminator=no)]
+ Equationsbicollapse {Δ} Φ : ogs_env Δ Act Φ -> ogs_env Δ Pas Φ -> forallp, ↓[p]Φ =[val]> Δ :=
+ bicollapse ∅ₓ (ε⁺) (ε⁻) Act := ! ;
+ bicollapse ∅ₓ (ε⁺) (ε⁻) Pas := ! ;
+ bicollapse (Φ ▶ₓ _) (u ;⁺) (v ;⁻ γ) Act :=
+ [ bicollapse Φ v u Pas , γ ⊛ [ a_id , bicollapse Φ v u Act ] ] ;
+ bicollapse (Φ ▶ₓ _) (v ;⁺) (u ;⁻ e) Pas := bicollapse Φ u v Act .
+ #[global] Arguments bicollapse {Δ Φ} u v {p}.
And the fixpoint property linking collapsing and zipping (Prop 6.13).
+
Lemmacollapse_fix_aux {ΔΦ} (u : ogs_env Δ Act Φ) (v : ogs_env Δ Pas Φ)
+ : ₐ↓u ⊛ [ a_id , bicollapse u v ] ≡ₐ bicollapse u v
+ /\ ₐ↓v ⊛ [ a_id , bicollapse u v ] ≡ₐ bicollapse u v .
+ Proof.
+ induction Φ; dependent destruction u; dependent destruction v.
+ - split; intros ? i; nowdestruct (c_view_emp i).
+ - split; cbn; simp collapse; simp bicollapse.
+ + intros ? i; cbn; rewrite <- v_sub_sub, a_ren_r_simpl, r_cat3_1_simpl.
+ nowapply IHΦ.
+ exact _. (* wtf typeclass?? *)
+ + intros ? i; cbn; destruct (c_view_cat i); eauto.
+ nowapply IHΦ.
+ Qed.
+
+ Lemmacollapse_fix_act {ΔΦ} (u : ogs_env Δ Act Φ) (v : ogs_env Δ Pas Φ)
+ : ₐ↓u ⊛ [ a_id , bicollapse u v ] ≡ₐ bicollapse u v .
+ Proof. nowapply collapse_fix_aux. Qed.
+
+ Lemmacollapse_fix_pas {ΔΦ} (u : ogs_env Δ Act Φ) (v : ogs_env Δ Pas Φ)
+ : ₐ↓v ⊛ [ a_id , bicollapse u v ] ≡ₐ bicollapse u v .
+ Proof. nowapply collapse_fix_aux. Qed.
Here we provide an alternative definition to the collapsing functions, using a more
+precisely typed lookup function. This is more practical to use when reasoning about
+the height of a variable, in the eventual guardedness proof.
+
First we compute actual useful subset of the context used by a particular variable.
+
Equationsctx_domΦp {x} : ↓[p^]Φ ∋ x -> C :=
+ ctx_dom ∅ₓ Act i with c_view_emp i := { | ! } ;
+ ctx_dom ∅ₓ Pas i with c_view_emp i := { | ! } ;
+ ctx_dom (Φ ▶ₓ Γ) Act i := ctx_dom Φ Pas i ;
+ ctx_dom (Φ ▶ₓ Γ) Pas i with c_view_cat i := {
+ | Vcat_l j := ctx_dom Φ Act j ;
+ | Vcat_r j := ↓[Act]Φ } .
+ #[global] Arguments ctx_dom {Φ p x} i.
Next we provide a renaming from this precise domain to the actual current allowed context.
+
Equationsr_ctx_domΦp {x} (i : ↓[p^]Φ ∋ x) : ctx_dom i ⊆ ↓[p]Φ :=
+ r_ctx_dom ∅ₓ Act i with c_view_emp i := { | ! } ;
+ r_ctx_dom ∅ₓ Pas i with c_view_emp i := { | ! } ;
+ r_ctx_dom (Φ ▶ₓ Γ) Act i := r_ctx_dom Φ Pas i ᵣ⊛ r_cat_l ;
+ r_ctx_dom (Φ ▶ₓ Γ) Pas i with c_view_cat i := {
+ | Vcat_l j := r_ctx_dom Φ Act j ;
+ | Vcat_r j := r_id } .
+ #[global] Arguments r_ctx_dom {Φ p x} i.
We can write this more precise lookup function, returning a value in just the necessary
+context.
+
Equationslookup {ΔpΦ} (γ : ogs_env Δ p Φ) [x] (i : ↓[p^]Φ ∋ x)
+ : val (Δ +▶ ctx_dom i) x :=
+ lookup ε⁺ i with c_view_emp i := { | ! } ;
+ lookup ε⁻ i with c_view_emp i := { | ! } ;
+ lookup (γ ;⁺) i := lookup γ i ;
+ lookup (γ ;⁻ e) i with c_view_cat i := {
+ | Vcat_l j := lookup γ j ;
+ | Vcat_r j := e _ j } .
We relate the precise lookup with the previously defined collapse.
We define active and passive states of the machine strategy. Active states consist of
+the pair of a language configuration and an active OGS assignement. Passive states consist
+solely of a passive OGS assignement.
Next we define the action and reaction morphisms of the machine strategy. First
+m_strat_wrap provides the active transition given an already evaluated normal form,
+such that the action morphism m_strat_play only has to evaluate the active part
+of the state. m_strat_resp mostly is a wrapper around oapp, our analogue to
+the embedding from normal forms to language configurations present in the paper.
These action and reaction morphisms define a coalgebra, which we now embed into the
+final coalgebra by looping them coinductively. This constructs the indexed interaction
+tree arising from starting the machine strategy at a given state.
+
Definitionm_strat {Δ} : m_strat_act Δ ⇒ᵢ ogs_act Δ :=
+ cofix _m_strat Φ e :=
+ subst_delay
+ (funr => go match r with
+ | inl m => RetF m
+ | inr (x ,' p) => VisF x (funr : ogs_e.(e_rsp) _ => _m_strat _ (m_strat_resp p r))
+ end)
+ (m_strat_play e).
We also provide a wrapper for the passive version of the above map.
+
Definitionm_stratp {Δ} : m_strat_pas Δ ⇒ᵢ ogs_pas Δ :=
+ fun_xm => m_strat _ (m_strat_resp x m).
We define the notion of equivalence on machine-strategy states given by the weak
+bisimilarity of the induced infinite trees. There is an active version, and a passive
+version, working pointwise.
+
Definitionm_strat_act_eqv {Δ} : relᵢ (m_strat_act Δ) (m_strat_act Δ) :=
+ funixy => m_strat i x ≈ m_strat i y.
+ Notation"x ≈ₐ y" := (m_strat_act_eqv _ x y).
+
+ Definitionm_strat_pas_eqv {Δ} : relᵢ (m_strat_pas Δ) (m_strat_pas Δ) :=
+ funixy => forallm, m_strat_resp x m ≈ₐ m_strat_resp y m .
+ Notation"x ≈ₚ y" := (m_strat_pas_eqv _ x y).
+
+ #[global] Instancem_strat_act_eqv_refl {Δ} : Reflexiveᵢ (@m_strat_act_eqv Δ).
+ Proof. intros ??; unfold m_strat_act_eqv; reflexivity. Qed.
+ #[global] Instancem_strat_pas_eqv_refl {Δ} : Reflexiveᵢ (@m_strat_pas_eqv Δ).
+ Proof. intros ???; reflexivity. Qed.
A technical lemma explaining how the infinite strategy unfolds.
Next we construct the initial states. The active initial state is given by simply a
+configuration from the language machine while a passive initial state is given by
+an assignment into the final context Δ.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Interaction Trees: Structure
+
+
ITrees form an iterative monad. The iter combinator provided here iterates over
+arbitrary sets of equations. The file ITree/Guarded.v provides a
+finer iterator, restricted to notions of guarded equations, enjoying an additional
+uniqueness property.
+
Sectionmonad.
+Context {I} {E : event I I}.
Functorial action on maps.
+
Definitionfmap {XY} (f : X ⇒ᵢ Y) : itree E X ⇒ᵢ itree E Y :=
+ cofix _fmap _ u :=
+ go match u.(_observe) with
+ | RetF r => RetF (f _ r)
+ | TauF t => TauF (_fmap _ t)
+ | VisF e k => VisF e (funr => _fmap _ (k r))
+ end.
Monadic bind.
+
Definitionsubst {XY} (f : X ⇒ᵢ itree E Y) : itree E X ⇒ᵢ itree E Y :=
+ cofix _subst _ u :=
+ go match u.(_observe) with
+ | RetF r => (f _ r).(_observe)
+ | TauF t => TauF (_subst _ t)
+ | VisF e k => VisF e (funr => _subst _ (k r))
+ end.
+
+Definitionbind {XYi} xf := @subst X Y f i x.
+
+Definitionkcomp {XYZ} (f : X ⇒ᵢ itree E Y) (g : Y ⇒ᵢ itree E Z) : X ⇒ᵢ itree E Z :=
+ funix => bind (f i x) g.
Iteration operator (Def. 31).
+
Definitioniter {XY} (f : X ⇒ᵢ itree E (X +ᵢ Y)) : X ⇒ᵢ itree E Y :=
+ cofix _iter _ x :=
+ bind (f _ x) (fun_r => go match r with
+ | inl x => TauF (_iter _ x)
+ | inr y => RetF y
+ end) .
+Endmonad.
+
+#[global] Notation"f <$> x" := (fmap f _ x) (at level30).
+#[global] Notation"x >>= f" := (bind x f) (at level30).
+#[global] Notation"f =<< x" := (subst f _ x) (at level30).
+#[global] Notation"f >=> g" := (kcomp f g) (at level30).
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Sub-context structure
+
+
As motivated in Ctx/DirectSum.v, we sometimes have exotic calculi
+which need specific kinds of contexts. Another such case is when there is a particular
+subset of types, represented by a predicate, and we wish to talk about contexts containing
+only these kind of types. Once again, we could do this using concrete lists, by taking
+the set of types to be elements of the predicate P : T → SProp: ctx (sigS P). The
+problem is once again that DeBruijn indices on this structure is not the right notion of
+variables. The set of contexts of some subset of the types is a subset of the set of contexts
+over that type! In other words, we should rather take sigS (ctx T) (allS P). As such the
+notion of variable should be the same.
+
We assume given a notion of context and a strict predicate on types.
+
Sectionwith_param.
+Context {TC} {CC : context T C} {CL : context_laws T C} {P : T -> SProp}.
We define the strict predicate on context stating that all elements verify P.
+
DefinitionallS (Γ : C) : SProp := forallx, Γ ∋ x -> P x.
+DefinitionctxS : Type := sigS allS.
+Definitioncoe_ctxS : ctxS -> C := sub_elt.
A variable in the refined context is just a variable in the underlying context.
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS
allS (sub_elt Γ +▶ sub_elt Δ)
+
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS x: T i: sub_elt Γ ∋ x
P x
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS x: T j: sub_elt Δ ∋ x
P x
+
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS x: T j: sub_elt Δ ∋ x
P x
+nowapply Δ.(sub_prf).
+Qed.
We now construct the instance.
+
#[global] Instancesubset_context : context T ctxS :=
+ {| c_emp := nilS ;
+ c_cat := catS ;
+ c_var := varS |}.
+
+#[global] Instancesubset_context_cat_wkn : context_cat_wkn T ctxS :=
+ {| r_cat_l Γ Δ t i := @r_cat_l T C _ _ Γ.(sub_elt) Δ.(sub_elt) t i ;
+ r_cat_r Γ Δ t i := @r_cat_r T C _ _ Γ.(sub_elt) Δ.(sub_elt) t i |} .
+
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
context_laws T ctxS
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
context_laws T ctxS
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (t : T) (i : varS nilS t), c_emp_view t i
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T) (i : varS (catS Γ Δ) t),
+c_cat_view Γ Δ t i
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T),
+injective (funi : varS Γ t => r_cat_l i)
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T),
+injective (funi : varS Δ t => r_cat_r i)
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T) (i : varS Γ t)
+ (j : varS Δ t), ¬ (r_cat_l i = r_cat_r j)
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (t : T) (i : varS nilS t), c_emp_view t i
intros ? i; destruct (c_view_emp i).
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T) (i : varS (catS Γ Δ) t),
+c_cat_view Γ Δ t i
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS t: T i: sub_elt Γ ∋ t
c_cat_view Γ Δ t (r_cat_l i)
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS t: T j: sub_elt Δ ∋ t
c_cat_view Γ Δ t (r_cat_r j)
+
T, C: Type CC: context T C CL: context_laws T C P: psh T Γ, Δ: ctxS t: T j: sub_elt Δ ∋ t
c_cat_view Γ Δ t (r_cat_r j)
+refine (Vcat_r _).
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T),
+injective (funi : varS Γ t => r_cat_l i)
intros ????? H; exact (r_cat_l_inj _ _ H).
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T),
+injective (funi : varS Δ t => r_cat_r i)
intros ????? H; exact (r_cat_r_inj _ _ H).
+
T, C: Type CC: context T C CL: context_laws T C P: psh T
forall (ΓΔ : ctxS) (t : T) (i : varS Γ t)
+ (j : varS Δ t), ¬ (r_cat_l i = r_cat_r j)
intros ????? H; exact (r_cat_disj _ _ H).
+Qed.
A couple helpers for manipulating the new variables.
+
Definitions_prf {Γ : ctxS} {x} (i : Γ.(sub_elt) ∋ x) : P x := Γ.(sub_prf) x i .
+
+Definitions_elt_upg {Γ : ctxS} {x} (i : Γ.(sub_elt) ∋ x) : sigS P :=
+ {| sub_prf := Γ.(sub_prf) x i |}.
+
+Definitions_var_upg {Γ : ctxS} {x : T} (i : Γ.(sub_elt) ∋ x)
+ : Γ ∋ (s_elt_upg i).(sub_elt) := i.
+Endwith_param.
In the case of sub-contexts over concrete contexts we provide a wrapper for the
+"append" operation.
+
From OGS.Ctx Require Import Ctx Covering.
+Sectionwith_param.
+Context {T} {P : T -> SProp}.
+
+Program DefinitionconS (Γ : ctxS T (ctx T) P) (x : sigS P) : ctxS T (ctx T) P :=
+ {| sub_elt := Γ.(sub_elt) ▶ₓ x.(sub_elt) |}.
+
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P
allS P (sub_elt Γ ▶ₓ sub_elt x)
+
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T i: cvar (sub_elt Γ ▶ₓ sub_elt x) x0
P x0
+
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T c: ctx T Heqc: NoConfusion c (sub_elt Γ ▶ₓ sub_elt x) i: cvar c x0
P x0
+
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T Γ0: ctx T Heqc: NoConfusion (Γ0 ▶ₓ x0) (sub_elt Γ ▶ₓ sub_elt x)
P x0
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T Γ0: ctx T y: T Heqc: NoConfusion (Γ0 ▶ₓ y) (sub_elt Γ ▶ₓ sub_elt x) i: var x0 Γ0
P x0
+
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T Γ0: ctx T Heqc: NoConfusion (Γ0 ▶ₓ x0) (sub_elt Γ ▶ₓ sub_elt x)
P x0
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T Γ0: ctx T Heqc: NoConfusion (Γ0 ▶ₓ x0) (sub_elt Γ ▶ₓ sub_elt x) H: x0 = sub_elt x
P x0
+rewrite H; nowapply x.(sub_prf).
+
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T Γ0: ctx T y: T Heqc: NoConfusion (Γ0 ▶ₓ y) (sub_elt Γ ▶ₓ sub_elt x) i: var x0 Γ0
P x0
T: Type P: psh T Γ: ctxS T (ctx T) P x: sigS P x0: T Γ0: ctx T y: T Heqc: NoConfusion (Γ0 ▶ₓ y) (sub_elt Γ ▶ₓ sub_elt x) i: var x0 Γ0 H: Γ0 = sub_elt Γ
P x0
+rewrite H in i; nowapply Γ.(sub_prf).
+Qed.
+Endwith_param.
+
+#[global] Notation"Γ ▶ₛ x" := (conS Γ x) : ctx_scope.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Substitution structures (§ 4.1)
+
+
In this file we axiomatize what it means for a family to support substitution.
+
+
Substitution Monoid (Def. 7)
+
The specification of an evaluator will be separated in several steps. First we will ask
+for a family of values, i.e. objects that can be substituted for variables. We formalize
+well-typed well-scoped substitutions in the monoidal style of Fiore et al and Allais et al.
+Here we are generic over a notion of context, treating lists and DeBruijn indices in the
+usual well-typed well-scoped style, but also other similar notions. See
+Ctx/Abstract.v for more background on notions of variables and contexts
+and for the categorical presentation of substitution.
+
Classsubst_monoid `{CC : context T C} (val : Fam₁ T C) := {
+ v_var : c_var ⇒₁ val ;
+ v_sub : val ⇒₁ ⟦ val , val ⟧₁ ;
+}.
By pointwise lifting of substitution we can define composition of assignments.
+
Definitiona_comp `{subst_monoid T C val} {Γ1 Γ2 Γ3}
+ : Γ1 =[val]> Γ2 -> Γ2 =[val]> Γ3 -> Γ1 =[val]> Γ3
+ := funuv_i => u _ i ᵥ⊛ v.
+#[global] Infix"⊛" := a_comp (at level14) : asgn_scope.
The laws for monoids and modules are pretty straightforward. A specificity is that
+assignments are represented by functions from variables to values, as such their
+well-behaved equality is pointwise equality and we require substitution to respect it.
+
Classsubst_monoid_laws `{CC : context T C} (val : Fam₁ T C) {VM : subst_monoid val} := {
+ v_sub_proper :: Proper (∀ₕΓ, ∀ₕ_, eq ==> ∀ₕΔ, asgn_eq Γ Δ ==> eq) v_sub ;
+ v_sub_var {Γ1 Γ2 x} (i : Γ1 ∋ x) (p : Γ1 =[val]> Γ2)
+ : v_var i ᵥ⊛ p = p _ i ;
+ v_var_sub {Γ x} (v : val Γ x)
+ : v ᵥ⊛ a_id = v ;
+ v_sub_sub {Γ1 Γ2 Γ3 x} (v : val Γ1 x) (a : Γ1 =[val]> Γ2) (b : Γ2 =[val]> Γ3)
+ : v ᵥ⊛ (a ⊛ b) = (v ᵥ⊛ a) ᵥ⊛ b ;
+} .
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val H: subst_monoid_laws val Γ1, Γ2: C x: T v: val Γ1 x
Proper (asgn_eq Γ1 Γ2 ==> eq) (v_sub v)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val H: subst_monoid_laws val Γ1, Γ2: C x: T v: val Γ1 x
Proper (asgn_eq Γ1 Γ2 ==> eq) (v_sub v)
nowapply v_sub_proper.Qed.
+
+
Substitution Module (Def. 8)
+
Next, we ask for a module over the monoid of values, to represent the configurations
+of the machine.
+
Classsubst_module `{CC : context T C} (val : Fam₁ T C) (conf : Fam₀ T C) := {
+ c_sub : conf ⇒₀ ⟦ val , conf ⟧₀ ;
+}.
#[global] Notation"c ₜ⊛ a" := (c_sub c a%asgn) (at level30).
Again the laws should not be surprising.
+
Classsubst_module_laws `{CC : context T C} (val : Fam₁ T C) (conf : Fam₀ T C)
+ {VM : subst_monoid val} {CM : subst_module val conf} := {
+ c_sub_proper :: Proper (∀ₕΓ, eq ==> ∀ₕΔ, asgn_eq Γ Δ ==> eq) c_sub ;
+ c_var_sub {Γ} (c : conf Γ) : c ₜ⊛ a_id = c ;
+ c_sub_sub {Γ1 Γ2 Γ3} (c : conf Γ1) (a : Γ1 =[val]> Γ2) (b : Γ2 =[val]> Γ3)
+ : c ₜ⊛ (a ⊛ b) = (c ₜ⊛ a) ₜ⊛ b ;
+} .
+
+
T, C: Type CC: context T C val: Fam₁ T C conf: Fam₀ T C VM: subst_monoid val CM: subst_module val conf H: subst_module_laws val conf Γ1, Γ2: C c: conf Γ1
Proper (asgn_eq Γ1 Γ2 ==> eq) (c_sub c)
+
T, C: Type CC: context T C val: Fam₁ T C conf: Fam₀ T C VM: subst_monoid val CM: subst_module val conf H: subst_module_laws val conf Γ1, Γ2: C c: conf Γ1
Proper (asgn_eq Γ1 Γ2 ==> eq) (c_sub c)
nowapply c_sub_proper.Qed.
Now that we know that our families have a substitution operation and variables, we
+can readily derive a renaming operation. While we could have axiomatized it, together with
+its compatibility with substitution, allowing for possibly more efficient implementations,
+we prefer simplicity and work with this generic implementation in terms of substitution.
+
Sectionrenaming.
+Context `{CC : context T C} {val : Fam₁ T C} {conf : Fam₀ T C}.
+Context {VM : subst_monoid val} {VML : subst_monoid_laws val}.
+Context {CM : subst_module val conf} {CML : subst_module_laws val conf}.
By post-composing with the substitution identity, we can embed renamings into assignments.
Something which we have absolutely pushed under the rug in the paper is a couple more
+mild technical hypotheses on the substitution monoid of values: since in the
+eventual guardedness proof we need to case split on whether
+or not a value is a variable, we need to ask for that to be possible! As such,
+we need the following additional assumptions:
+
+
v_var has decidable fibers
+
v_var is injective
+
the fibers of v_var pull back along renamings
+
+
These assumptions are named "clear-cut" in the paper (Def. 27).
+
We first define the fibers of v_var.
+
Variantis_var `{VM : subst_monoid T C val} {Γ x} : val Γ x -> Type :=
+| Vvar (i : Γ ∋ x) : is_var (v_var i)
+.
+
+Equationsis_var_get `{VM : subst_monoid T C val} {Γ x} {v : val Γ x} : is_var v -> Γ ∋ x :=
+ is_var_get (Vvar i) := i .
Which are obviously stable under renamings.
+
T, C: Type CC: context T C val: Fam₁ T C VM0: subst_monoid val VM: subst_monoid_laws val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T v: val Γ1 x
is_var v -> is_var (v ᵥ⊛ᵣ r)
+
T, C: Type CC: context T C val: Fam₁ T C VM0: subst_monoid val VM: subst_monoid_laws val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T v: val Γ1 x
is_var v -> is_var (v ᵥ⊛ᵣ r)
+
T, C: Type CC: context T C val: Fam₁ T C VM0: subst_monoid val VM: subst_monoid_laws val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T i: Γ1 ∋ x
is_var (a_id i ᵥ⊛ᵣ r)
+
T, C: Type CC: context T C val: Fam₁ T C VM0: subst_monoid val VM: subst_monoid_laws val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T i: Γ1 ∋ x
is_var (r_emb r x i)
+econstructor.
+Qed.
At last we define our last assumptions on variables (Def. 28).
+
Classvar_assumptions `{CC : context T C} (val : Fam₁ T C) {VM : subst_monoid val} := {
+ v_var_inj {Γ x} : injective (@v_var _ _ _ _ _ Γ x) ;
+ is_var_dec {Γ x} (v : val Γ x) : decidable (is_var v) ;
+ is_var_ren {Γ1 Γ2 x} (v : val Γ1 x) (r : Γ1 ⊆ Γ2) : is_var (v ᵥ⊛ᵣ r) -> is_var v ;
+}.
Here we derive a couple helpers around these new assumptions.
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x p, q: is_var v
p = q
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x p, q: is_var v
p = q
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x p, q0: is_var v i: Γ ∋ x w: val Γ x H: w = a_id i q: is_var w
Vvar i = rew [is_var] H in q
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x p, q0: is_var v i, i0: Γ ∋ x H: a_id i0 = a_id i
Vvar i = rew [is_var] H in Vvar i0
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x p, q0: is_var v i, i0: Γ ∋ x H: a_id i0 = a_id i H': i0 = i
Vvar i = rew [is_var] H in Vvar i0
+now dependent elimination H'; dependent elimination H.
+Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T i: Γ ∋ x p: is_var (a_id i)
p = Vvar i
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T i: Γ ∋ x p: is_var (a_id i)
p = Vvar i
apply is_var_irr.Qed.
+
+Variantis_var_ren_view {Γ1Γ2x}
+ (v : val Γ1 x) (r : Γ1 ⊆ Γ2) : is_var (v ᵥ⊛ᵣ r) -> Type :=
+ | Vvren (H : is_var v) : is_var_ren_view v r (ren_is_var r H) .
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C x: T v: val Γ1 x r: Γ1 ⊆ Γ2 H: is_var (v ᵥ⊛ᵣ r)
is_var_ren_view v r H
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C x: T v: val Γ1 x r: Γ1 ⊆ Γ2 H: is_var (v ᵥ⊛ᵣ r)
is_var_ren_view v r H
+rewrite (is_var_irr H (ren_is_var r (is_var_ren v r H))); econstructor.
+Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T v: val Γ1 x H: is_var v
is_var_get (ren_is_var r H) = r x (is_var_get H)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T v: val Γ1 x H: is_var v
is_var_get (ren_is_var r H) = r x (is_var_get H)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T i: Γ1 ∋ x
is_var_get (ren_is_var r (Vvar i)) = r x i
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T i: Γ1 ∋ x
foralli0 : is_var (a_id i ᵥ⊛ᵣ r),
+is_var_get i0 = r x i
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ1, Γ2: C r: Γ1 ⊆ Γ2 x: T i: Γ1 ∋ x H: is_var (a_id (r x i))
is_var_get H = r x i
+nowrewrite (is_var_irr H (Vvar (r _ i))).
+Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x H: is_var v
v = a_id (is_var_get H)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val VA: var_assumptions val Γ: C x: T v: val Γ x H: is_var v
v = a_id (is_var_get H)
nowdestruct H.Qed.
+Endvariables.
Finally we end with a couple derived property on assignments.
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3: C r: Γ1 ⊆ Γ2 a: Γ2 =[ val ]> Γ3
r_emb r ⊛ a ≡ₐ r ᵣ⊛ a
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3: C r: Γ1 ⊆ Γ2 a: Γ2 =[ val ]> Γ3
r_emb r ⊛ a ≡ₐ r ᵣ⊛ a
intros ??; cbn; nowrewrite v_sub_var.Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ: C
r_emb r_id ≡ₐ a_id
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ: C
r_emb r_id ≡ₐ a_id
reflexivity.Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3, Γ4: C a: Γ1 =[ val ]> Γ2 b: Γ2 =[ val ]> Γ3 r: Γ3 ⊆ Γ4
(a ⊛ b) ⊛ᵣ r ≡ₐ a ⊛ (b ⊛ᵣ r)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3, Γ4: C a: Γ1 =[ val ]> Γ2 b: Γ2 =[ val ]> Γ3 r: Γ3 ⊆ Γ4
(a ⊛ b) ⊛ᵣ r ≡ₐ a ⊛ (b ⊛ᵣ r)
intros ??; cbn; nowrewrite <-v_sub_sub.Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3, Γ4: C a: Γ1 =[ val ]> Γ2 r: Γ2 ⊆ Γ3 b: Γ3 =[ val ]> Γ4
(a ⊛ᵣ r) ⊛ b ≡ₐ a ⊛ (r ᵣ⊛ b)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3, Γ4: C a: Γ1 =[ val ]> Γ2 r: Γ2 ⊆ Γ3 b: Γ3 =[ val ]> Γ4
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3, Γ4: C a: Γ1 =[ val ]> Γ2 b: Γ2 =[ val ]> Γ3 c: Γ3 =[ val ]> Γ4
(a ⊛ b) ⊛ c ≡ₐ a ⊛ (b ⊛ c)
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2, Γ3, Γ4: C a: Γ1 =[ val ]> Γ2 b: Γ2 =[ val ]> Γ3 c: Γ3 =[ val ]> Γ4
(a ⊛ b) ⊛ c ≡ₐ a ⊛ (b ⊛ c)
intros ??; cbn; nowrewrite v_sub_sub.Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2: C a: Γ1 =[ val ]> Γ2
a ⊛ a_id ≡ₐ a
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2: C a: Γ1 =[ val ]> Γ2
a ⊛ a_id ≡ₐ a
intros ??; cbn; nowrewrite v_var_sub.Qed.
+
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2: C a: Γ1 =[ val ]> Γ2
a_id ⊛ a ≡ₐ a
+
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2: C a: Γ1 =[ val ]> Γ2
a_id ⊛ a ≡ₐ a
T, C: Type CC: context T C val: Fam₁ T C VM: subst_monoid val VML: subst_monoid_laws val Γ1, Γ2: C a: Γ1 =[ val ]> Γ2 a0: T a1: Γ1 ∋ a0
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
In this file we instanciate our OGS construction with the fully-dual polarized mu-mu-tilde
+calculus 'System D' from P. Downen & Z. Ariola. The presentation may be slightly unusual
+as we go for one-sided sequent. The only real divergence from the original calculus
+is the addition of a restricted form of recursion, making the language non-normalizing.
+
+
Types
+
Type have polarities, basically whether their values are CBV-biased
+or CBN-biased
+
Variantpol : Type := pos | neg .
Syntax of types. We have a fully dual type system with:
+
+
0, positive void (no constructor),
+
⊤, negative unit (no destructor),
+
⊥, negative void (one trivial destructor),
+
1, positive unit (one trivial constructor),
+
A⊗B, positive product (pairs with pattern-matching),
+
A⅋B, negative sum (one binary destructor),
+
A⊕B, positive sum (usual coproduct with injections),
+
A&B, negative product (usual product with projections),
+
↓A, positive shift (thunking),
+
↑A, negative shift ('returners'),
+
⊖A, positive negation (approximately continuations accepting an A),
+
¬A, negative negation (approximately refutations of A).
+
+
We opt for an explicit treatment of polarity, by indexing the family of types.
Notation"0" := (Zer) : ty_scope .
+Notation"1" := (One) : ty_scope .
+Notation"⊤" := (Top) : ty_scope .
+Notation"⊥" := (Bot) : ty_scope .
+Notation"A ⊗ B" := (Tens A B) (at level40) : ty_scope.
+Notation"A ⅋ B" := (Par A B) (at level40) : ty_scope .
+Notation"A ⊕ B" := (Or A B) (at level40) : ty_scope.
+Notation"A & B" := (And A B) (at level40) : ty_scope.
+Notation"↓ A" := (ShiftP A) (at level40) : ty_scope.
+Notation"↑ A" := (ShiftN A) (at level40) : ty_scope.
+Notation"⊖ A" := (NegP A) (at level5) : ty_scope .
+Notation"¬ A" := (NegN A) (at level5) : ty_scope .
As hinted above, we go for one-sided sequents. This enables to have only one context
+instead of two, simplifying the theory of substitution. On the flip-side, as we still
+need two kinds of variables, the 'normal' ones and the co-variables, our contexts will
+not contain bare types but side-annotated types. A variable of type A will thus be
+stored as \`-A in the context while a co-variable of type A will be stored as \`-A.
+
Variantty : Type :=
+| LTy {p} : pre_ty p -> ty
+| RTy {p} : pre_ty p -> ty
+.
Finally we define contexts as backward lists of side-annotated types.
+
The reference Ctx.ctx was not found in the current
+environment.
+
+
Terms
+
We define the well-typed syntax of the language with 3 mutually defined syntactic
+categories: terms, weak head-normal forms and states ('language configurations' in
+the paper).
+
Nothing should be too surprising. A first notable point is that by choosing to have an
+explicit notion of 'side' of a variable, we can have a single construct
+for both mu and mu-tilde. A second notable point is our new slightly exotic RecL and
+RecR constructions. They allow arbitrary recursion at co-terms of positive types and
+at terms of negative types. These polarity restrictions allow us to have minimal
+disruption of the evaluation rules.
+
Inductiveterm : t_ctx -> ty -> Type :=
+| Mu {Γ A} : state (Γ ▶ₓ A†) -> term Γ A
+| RecL {Γ} {A : pre_ty pos} : term (Γ ▶ₓ `-A) `-A -> term Γ `-A
+| RecR {Γ} {A : pre_ty neg} : term (Γ ▶ₓ `+A) `+A -> term Γ `+A
+| Whn {Γ A} : whn Γ A -> term Γ A
+with whn : t_ctx -> ty -> Type :=
+| Var {Γ A} : Γ ∋ A -> whn Γ A
+| ZerL {Γ} : whn Γ `-0
+| TopR {Γ} : whn Γ `+⊤
+| OneR {Γ} : whn Γ `+1
+| OneL {Γ} : state Γ -> whn Γ `-1
+| BotR {Γ} : state Γ -> whn Γ `+⊥
+| BotL {Γ} : whn Γ `-⊥
+| TenR {Γ A B} : whn Γ `+A -> whn Γ `+B -> whn Γ `+(A ⊗ B)
+| TenL {Γ A B} : state (Γ ▶ₓ `+A ▶ₓ `+B) -> whn Γ `-(A ⊗ B)
+| ParR {Γ A B} : state (Γ ▶ₓ `-A ▶ₓ `-B) -> whn Γ `+(A ⅋ B)
+| ParL {Γ A B} : whn Γ `-A -> whn Γ `-B -> whn Γ `-(A ⅋ B)
+| OrR1 {Γ A B} : whn Γ `+A -> whn Γ `+(A ⊕ B)
+| OrR2 {Γ A B} : whn Γ `+B -> whn Γ `+(A ⊕ B)
+| OrL {Γ A B} : state (Γ ▶ₓ `+A) -> state (Γ ▶ₓ `+B) -> whn Γ `-(A ⊕ B)
+| AndR {Γ A B} : state (Γ ▶ₓ `-A) -> state (Γ ▶ₓ `-B) -> whn Γ `+(A & B)
+| AndL1 {Γ A B} : whn Γ `-A -> whn Γ `-(A & B)
+| AndL2 {Γ A B} : whn Γ `-B -> whn Γ `-(A & B)
+| ShiftPR {Γ A} : term Γ `+A -> whn Γ `+(↓ A)
+| ShiftPL {Γ A} : state (Γ ▶ₓ `+A) -> whn Γ `-(↓ A)
+| ShiftNR {Γ A} : state (Γ ▶ₓ `-A) -> whn Γ `+(↑ A)
+| ShiftNL {Γ A} : term Γ `-A -> whn Γ `-(↑ A)
+| NegPR {Γ A} : whn Γ `-A -> whn Γ `+(⊖ A)
+| NegPL {Γ A} : state (Γ ▶ₓ `-A) -> whn Γ `-(⊖ A)
+| NegNR {Γ A} : state (Γ ▶ₓ `+A) -> whn Γ `+(¬ A)
+| NegNL {Γ A} : whn Γ `+A -> whn Γ `-(¬ A)
+with state : t_ctx -> Type :=
+| Cut {Γ} p {A : pre_ty p} : term Γ `+A -> term Γ `-A -> state Γ
+.
+
+DefinitionCut' {ΓA} : term Γ A -> term Γ A† -> state Γ :=
+ match A with
+ | `+A => funt1t2 => Cut _ t1 t2
+ | `-A => funt1t2 => Cut _ t2 t1
+ end .
Values are not exactly weak head-normal forms, but depend on the polarity of the type.
+As positive types have CBV evaluation, their values are weak head-normal forms, but their
+co-values (evaluation contexts) are just any co-term (context) as they are delayed anyways.
+Dually for negative types, values are delayed hence can be any term while co-values must
+be weak head-normal form contexts.
+
Equationsval : t_ctx -> ty -> Type :=
+ val Γ (@LTy pos A) := whn Γ `+A ;
+ val Γ (@RTy pos A) := term Γ `-A ;
+ val Γ (@LTy neg A) := term Γ `+A ;
+ val Γ (@RTy neg A) := whn Γ `-A .
+Arguments val _ _ /.
We provide a 'smart-constructor' for variables, embedding variables in values.
+
Equationsvar : c_var ⇒₁ val :=
+ var _ (@LTy pos _) i := Var i ;
+ var _ (@RTy pos _) i := Whn (Var i) ;
+ var _ (@LTy neg _) i := Whn (Var i) ;
+ var _ (@RTy neg _) i := Var i .
+#[global] Arguments var {Γ} [x] / i.
+
+
Renaming
+
Without surprise parallel renaming goes by a big mutual induction, shifting the renaming
+apropriately while going under binders. Note the use of the internal substitution hom
+to type it.
+
Equationst_rename : term ⇒₁ ⟦ c_var , term ⟧₁ :=
+ t_rename _ _ (Mu c) _ f := Mu (s_rename _ c _ (r_shift1 f)) ;
+ t_rename _ _ (RecL t) _ f := RecL (t_rename _ _ t _ (r_shift1 f)) ;
+ t_rename _ _ (RecR t) _ f := RecR (t_rename _ _ t _ (r_shift1 f)) ;
+ t_rename _ _ (Whn v) _ f := Whn (w_rename _ _ v _ f) ;
+with w_rename : whn ⇒₁ ⟦ c_var , whn ⟧₁ :=
+ w_rename _ _ (Var i) _ f := Var (f _ i) ;
+ w_rename _ _ (ZerL) _ f := ZerL ;
+ w_rename _ _ (TopR) _ f := TopR ;
+ w_rename _ _ (OneR) _ f := OneR ;
+ w_rename _ _ (OneL c) _ f := OneL (s_rename _ c _ f) ;
+ w_rename _ _ (BotR c) _ f := BotR (s_rename _ c _ f) ;
+ w_rename _ _ (BotL) _ f := BotL ;
+ w_rename _ _ (TenR v1 v2) _ f := TenR (w_rename _ _ v1 _ f) (w_rename _ _ v2 _ f) ;
+ w_rename _ _ (TenL c) _ f := TenL (s_rename _ c _ (r_shift2 f)) ;
+ w_rename _ _ (ParR c) _ f := ParR (s_rename _ c _ (r_shift2 f)) ;
+ w_rename _ _ (ParL k1 k2) _ f := ParL (w_rename _ _ k1 _ f) (w_rename _ _ k2 _ f) ;
+ w_rename _ _ (OrR1 v) _ f := OrR1 (w_rename _ _ v _ f) ;
+ w_rename _ _ (OrR2 v) _ f := OrR2 (w_rename _ _ v _ f) ;
+ w_rename _ _ (OrL c1 c2) _ f := OrL (s_rename _ c1 _ (r_shift1 f))
+ (s_rename _ c2 _ (r_shift1 f)) ;
+ w_rename _ _ (AndR c1 c2) _ f := AndR (s_rename _ c1 _ (r_shift1 f))
+ (s_rename _ c2 _ (r_shift1 f)) ;
+ w_rename _ _ (AndL1 k) _ f := AndL1 (w_rename _ _ k _ f) ;
+ w_rename _ _ (AndL2 k) _ f := AndL2 (w_rename _ _ k _ f) ;
+ w_rename _ _ (ShiftPR t) _ f := ShiftPR (t_rename _ _ t _ f) ;
+ w_rename _ _ (ShiftPL c) _ f := ShiftPL (s_rename _ c _ (r_shift1 f)) ;
+ w_rename _ _ (ShiftNR c) _ f := ShiftNR (s_rename _ c _ (r_shift1 f)) ;
+ w_rename _ _ (ShiftNL t) _ f := ShiftNL (t_rename _ _ t _ f) ;
+ w_rename _ _ (NegPR k) _ f := NegPR (w_rename _ _ k _ f) ;
+ w_rename _ _ (NegPL c) _ f := NegPL (s_rename _ c _ (r_shift1 f)) ;
+ w_rename _ _ (NegNR c) _ f := NegNR (s_rename _ c _ (r_shift1 f)) ;
+ w_rename _ _ (NegNL v) _ f := NegNL (w_rename _ _ v _ f) ;
+with s_rename : state ⇒₀ ⟦ c_var , state ⟧₀ :=
+ s_rename _ (Cut _ v k) _ f := Cut _ (t_rename _ _ v _ f) (t_rename _ _ k _ f) .
We extend it to values...
+
Equationsv_rename : val ⇒₁ ⟦ c_var , val ⟧₁ :=
+ v_rename _ (@LTy pos a) := w_rename _ _ ;
+ v_rename _ (@RTy pos a) := t_rename _ _ ;
+ v_rename _ (@LTy neg a) := t_rename _ _ ;
+ v_rename _ (@RTy neg a) := w_rename _ _ .
The following bunch of shifting functions will help us define parallel substitution.
+
Definitiont_shift1 {Γy} : term Γ ⇒ᵢ term (Γ ▶ₓ y) := fun_t => t ₜ⊛ᵣ r_pop.
+Definitionw_shift1 {Γy} : whn Γ ⇒ᵢ whn (Γ ▶ₓ y) := fun_w => w `ᵥ⊛ᵣ r_pop.
+Definitions_shift1 {Γy} : state Γ -> state (Γ ▶ₓ y) := funs => s ₛ⊛ᵣ r_pop.
+Definitionv_shift1 {Γy} : val Γ ⇒ᵢ val (Γ ▶ₓ y) := fun_v => v ᵥ⊛ᵣ r_pop.
+Definitionv_shift2 {Γyz} : val Γ ⇒ᵢ val (Γ ▶ₓ y ▶ₓ z) := fun_v => v ᵥ⊛ᵣ (r_pop ᵣ⊛ r_pop).
+Definitiona_shift1 {ΓΔ} [y] (a : Γ =[val]> Δ) : (Γ ▶ₓ y) =[val]> (Δ ▶ₓ y)
+ := [ fun_i => v_shift1 _ (a _ i) ,ₓ var top ].
+Definitiona_shift2 {ΓΔ} [y z] (a : Γ =[val]> Δ) : (Γ ▶ₓ y ▶ₓ z) =[val]> (Δ ▶ₓ y ▶ₓ z)
+ := [ [ fun_i => v_shift2 _ (a _ i) ,ₓ var (pop top) ] ,ₓ var top ].
We also define two embeddings linking the various syntactical categories.
+
Equationsv_of_wΓA : whn Γ A -> val Γ A :=
+ v_of_w _ (@LTy pos _) v := v ;
+ v_of_w _ (@RTy pos _) u := Whn u ;
+ v_of_w _ (@LTy neg _) u := Whn u ;
+ v_of_w _ (@RTy neg _) k := k .
+Arguments v_of_w {Γ A} v.
+#[global] Coercionv_of_w : whn >-> val.
+
+Equationst_of_vΓA : val Γ A -> term Γ A :=
+ t_of_v _ (@LTy pos _) v := Whn v ;
+ t_of_v _ (@RTy pos _) u := u ;
+ t_of_v _ (@LTy neg _) u := u ;
+ t_of_v _ (@RTy neg _) k := Whn k .
+Arguments t_of_v {Γ A} v.
+#[global] Coerciont_of_v : val >-> term.
+
+
Substitution
+
Having done with renaming, we reapply the same pattern to define parallel substitution.
+Note that substituting a weak head-normal form with values may not yield a weak
+head-normal form, but only a value!
+
Equationst_subst : term ⇒₁ ⟦ val , term ⟧₁ :=
+ t_subst _ _ (Mu c) _ f := Mu (s_subst _ c _ (a_shift1 f)) ;
+ t_subst _ _ (RecL t) _ f := RecL (t_subst _ _ t _ (a_shift1 f)) ;
+ t_subst _ _ (RecR t) _ f := RecR (t_subst _ _ t _ (a_shift1 f)) ;
+ t_subst _ _ (Whn v) _ f := w_subst _ _ v _ f ;
+with w_subst : whn ⇒₁ ⟦ val , val ⟧₁ :=
+ w_subst _ _ (Var i) _ f := f _ i ;
+ w_subst _ _ (ZerL) _ f := Whn ZerL ;
+ w_subst _ _ (TopR) _ f := Whn TopR ;
+ w_subst _ _ (OneR) _ f := OneR ;
+ w_subst _ _ (OneL c) _ f := Whn (OneL (s_subst _ c _ f)) ;
+ w_subst _ _ (BotR c) _ f := Whn (BotR (s_subst _ c _ f)) ;
+ w_subst _ _ (BotL) _ f := BotL ;
+ w_subst _ _ (TenR v1 v2) _ f := TenR (w_subst _ _ v1 _ f) (w_subst _ _ v2 _ f) ;
+ w_subst _ _ (TenL c) _ f := Whn (TenL (s_subst _ c _ (a_shift2 f))) ;
+ w_subst _ _ (ParR c) _ f := Whn (ParR (s_subst _ c _ (a_shift2 f))) ;
+ w_subst _ _ (ParL k1 k2) _ f := ParL (w_subst _ _ k1 _ f) (w_subst _ _ k2 _ f) ;
+ w_subst _ _ (OrR1 v) _ f := OrR1 (w_subst _ _ v _ f) ;
+ w_subst _ _ (OrR2 v) _ f := OrR2 (w_subst _ _ v _ f) ;
+ w_subst _ _ (OrL c1 c2) _ f := Whn (OrL (s_subst _ c1 _ (a_shift1 f))
+ (s_subst _ c2 _ (a_shift1 f))) ;
+ w_subst _ _ (AndR c1 c2) _ f := Whn (AndR (s_subst _ c1 _ (a_shift1 f))
+ (s_subst _ c2 _ (a_shift1 f))) ;
+ w_subst _ _ (AndL1 k) _ f := AndL1 (w_subst _ _ k _ f) ;
+ w_subst _ _ (AndL2 k) _ f := AndL2 (w_subst _ _ k _ f) ;
+ w_subst _ _ (ShiftPR t) _ f := ShiftPR (t_subst _ _ t _ f) ;
+ w_subst _ _ (ShiftPL c) _ f := Whn (ShiftPL (s_subst _ c _ (a_shift1 f))) ;
+ w_subst _ _ (ShiftNR c) _ f := Whn (ShiftNR (s_subst _ c _ (a_shift1 f))) ;
+ w_subst _ _ (ShiftNL t) _ f := ShiftNL (t_subst _ _ t _ f) ;
+ w_subst _ _ (NegPR k) _ f := NegPR (w_subst _ _ k _ f) ;
+ w_subst _ _ (NegPL c) _ f := Whn (NegPL (s_subst _ c _ (a_shift1 f))) ;
+ w_subst _ _ (NegNR c) _ f := Whn (NegNR (s_subst _ c _ (a_shift1 f))) ;
+ w_subst _ _ (NegNL v) _ f := NegNL (w_subst _ _ v _ f) ;
+with s_subst : state ⇒₀ ⟦ val , state ⟧₀ :=
+ s_subst _ (Cut p v k) _ f := Cut p (t_subst _ _ v _ f) (t_subst _ _ k _ f) .
+
+Notation"t `ₜ⊛ a" := (t_subst _ _ t _ a%asgn) (at level30).
+Notation"w `ᵥ⊛ a" := (w_subst _ _ w _ a%asgn) (at level30).
+
+Equationsv_subst : val ⇒₁ ⟦ val , val ⟧₁ :=
+ v_subst _ (@LTy pos a) v _ f := v `ᵥ⊛ f ;
+ v_subst _ (@RTy pos a) t _ f := t `ₜ⊛ f ;
+ v_subst _ (@LTy neg a) t _ f := t `ₜ⊛ f ;
+ v_subst _ (@RTy neg a) k _ f := k `ᵥ⊛ f .
With this in hand we can instanciate the relevant part of substitution monoid and module
+structures for values and states. This will provide us with the missing infix notations.
We now define helpers for substituting the top one or top two variables from a context.
+
Definitionasgn1 {Γa} (v : val Γ a) : (Γ ▶ₓ a) =[val]> Γ := [ var ,ₓ v ] .
+Definitionasgn2 {Γab} (v1 : val Γ a) (v2 : val Γ b) : (Γ ▶ₓ a ▶ₓ b) =[val]> Γ
+ := [ [ var ,ₓ v1 ] ,ₓ v2 ].
+
+Arguments asgn1 {_ _} & _.
+Arguments asgn2 {_ _ _} & _ _.
+
+Notation"₁[ v ]" := (asgn1 v).
+Notation"₂[ v1 , v2 ]" := (asgn2 v1 v2).
+
+
Observations
+
When defining (co-)patterns, we will enforce a form of focalisation, where no negative
+variables are introduced. In this context, 'negative' is a new notion applying to
+side-annotated types, mixing both type polarity and side annotation: a side-annotated
+variable is positive iff it is a positive variable or a negative co-variable.
+
Equationsis_neg : ty -> SProp :=
+ is_neg (@LTy pos a) := sEmpty ;
+ is_neg (@RTy pos a) := sUnit ;
+ is_neg (@LTy neg a) := sUnit ;
+ is_neg (@RTy neg a) := sEmpty .
We define negative types as a subset of types, and negative contexts as a subset of
+contexts. Our generic infrastructure for contexts and variables really shines here as
+the type of variables in a negative context is convertible to the type of variables in
+the underlying context. See Ctx/Subset.v.
+
Definitionneg_ty : Type := sigS is_neg.
+Definitionneg_coe : neg_ty -> ty := sub_elt.
+Global Coercionneg_coe : neg_ty >-> ty.
+
+Definitionneg_ctx : Type := ctxS ty t_ctx is_neg.
+Definitionneg_c_coe : neg_ctx -> ctx ty := sub_elt.
+Global Coercionneg_c_coe : neg_ctx >-> ctx.
We can now define patterns...
+
Inductivepat : ty -> Type :=
+| PVarP (A : pre_ty neg) : pat `+A
+| PVarN (A : pre_ty pos) : pat `-A
+| POne : pat `+1
+| PBot : pat `-⊥
+| PTen {A B} : pat `+A -> pat `+B -> pat `+(A ⊗ B)
+| PPar {A B} : pat `-A -> pat `-B -> pat `-(A ⅋ B)
+| POr1 {A B} : pat `+A -> pat `+(A ⊕ B)
+| POr2 {A B} : pat `+B -> pat `+(A ⊕ B)
+| PAnd1 {A B} : pat `-A -> pat `-(A & B)
+| PAnd2 {A B} : pat `-B -> pat `-(A & B)
+| PShiftP A : pat `+(↓ A)
+| PShiftN A : pat `-(↑ A)
+| PNegP {A} : pat `-A -> pat `+(⊖ A)
+| PNegN {A} : pat `+A -> pat `-(¬ A)
+.
We finally instanciate the observation structure. Note that our generic formalization
+mostly cares about 'observations', that is co-patterns. As such we instanciate observations
+by patterns at the dual type.
+
Definitionobs_op : Oper ty neg_ctx :=
+ {| o_op A := pat A† ; o_dom _ p := p_dom p |} .
Now come a rather tedious set of embeddings between syntactic categories related to
+patterns. We start by embedding patterns into weak head-normal forms.
Now we explain how to split (some) weak head-normal forms into a pattern filled with
+values. I am sorry in advance for your CPU-cycles wasted to typechecking these quite
+hard dependent pattern matchings. We start off by two helpers for refuting impossible
+variables in negative context, which because of the use of SProp give trouble to
+Equations for deriving functional elimination principles if inlined.
+
Definitionelim_var_p {Γ : neg_ctx} {A : pre_ty pos} {X : Type} : Γ ∋ `+A -> X
+ := funi => match s_prf i withend .
+
+Definitionelim_var_n {Γ : neg_ctx} {A : pre_ty neg} {X : Type} : Γ ∋ `-A -> X
+ := funi => match s_prf i withend .
+
+Equationsp_of_w_0p {Γ : neg_ctx} (A : pre_ty pos) : whn Γ `+A -> pat `+A :=
+ p_of_w_0p (0) (Var i) := elim_var_p i ;
+ p_of_w_0p (1) (Var i) := elim_var_p i ;
+ p_of_w_0p (A ⊗ B) (Var i) := elim_var_p i ;
+ p_of_w_0p (A ⊕ B) (Var i) := elim_var_p i ;
+ p_of_w_0p (↓ A) (Var i) := elim_var_p i ;
+ p_of_w_0p (⊖ A) (Var i) := elim_var_p i ;
+ p_of_w_0p (1) (OneR) := POne ;
+ p_of_w_0p (A ⊗ B) (TenR v1 v2) := PTen (p_of_w_0p A v1) (p_of_w_0p B v2) ;
+ p_of_w_0p (A ⊕ B) (OrR1 v) := POr1 (p_of_w_0p A v) ;
+ p_of_w_0p (A ⊕ B) (OrR2 v) := POr2 (p_of_w_0p B v) ;
+ p_of_w_0p (↓ A) (ShiftPR _) := PShiftP A ;
+ p_of_w_0p (⊖ A) (NegPR k) := PNegP (p_of_w_0n A k) ;
+with p_of_w_0n {Γ : neg_ctx} (A : pre_ty neg) : whn Γ `-A -> pat `-A :=
+ p_of_w_0n (⊤) (Var i) := elim_var_n i ;
+ p_of_w_0n (⊥) (Var i) := elim_var_n i ;
+ p_of_w_0n (A ⅋ B) (Var i) := elim_var_n i ;
+ p_of_w_0n (A & B) (Var i) := elim_var_n i ;
+ p_of_w_0n (↑ A) (Var i) := elim_var_n i ;
+ p_of_w_0n (¬ A) (Var i) := elim_var_n i ;
+ p_of_w_0n (⊥) (BotL) := PBot ;
+ p_of_w_0n (A ⅋ B) (ParL k1 k2) := PPar (p_of_w_0n A k1) (p_of_w_0n B k2) ;
+ p_of_w_0n (A & B) (AndL1 k) := PAnd1 (p_of_w_0n A k) ;
+ p_of_w_0n (A & B) (AndL2 k) := PAnd2 (p_of_w_0n B k) ;
+ p_of_w_0n (↑ A) (ShiftNL _) := PShiftN A ;
+ p_of_w_0n (¬ A) (NegNL v) := PNegN (p_of_w_0p A v) .
+
+Equationsp_dom_of_w_0p {Γ : neg_ctx} (A : pre_ty pos) (v : whn Γ `+A)
+ : p_dom (p_of_w_0p A v) =[val]> Γ bystruct A :=
+ p_dom_of_w_0p (0) (Var i) := elim_var_p i ;
+ p_dom_of_w_0p (1) (Var i) := elim_var_p i ;
+ p_dom_of_w_0p (A ⊗ B) (Var i) := elim_var_p i ;
+ p_dom_of_w_0p (A ⊕ B) (Var i) := elim_var_p i ;
+ p_dom_of_w_0p (↓ A) (Var i) := elim_var_p i ;
+ p_dom_of_w_0p (⊖ A) (Var i) := elim_var_p i ;
+ p_dom_of_w_0p (1) (OneR) := a_empty ;
+ p_dom_of_w_0p (A ⊗ B) (TenR v1 v2) := [ p_dom_of_w_0p A v1 , p_dom_of_w_0p B v2 ] ;
+ p_dom_of_w_0p (A ⊕ B) (OrR1 v) := p_dom_of_w_0p A v ;
+ p_dom_of_w_0p (A ⊕ B) (OrR2 v) := p_dom_of_w_0p B v ;
+ p_dom_of_w_0p (↓ A) (ShiftPR x) := a_append a_empty x ;
+ p_dom_of_w_0p (⊖ A) (NegPR k) := p_dom_of_w_0n A k ;
+ with p_dom_of_w_0n {Γ : neg_ctx} (A : pre_ty neg) (k : whn Γ `-A)
+ : p_dom (p_of_w_0n A k) =[val]> Γ bystruct A :=
+ p_dom_of_w_0n (⊤) (Var i) := elim_var_n i ;
+ p_dom_of_w_0n (⊥) (Var i) := elim_var_n i ;
+ p_dom_of_w_0n (A ⅋ B) (Var i) := elim_var_n i ;
+ p_dom_of_w_0n (A & B) (Var i) := elim_var_n i ;
+ p_dom_of_w_0n (↑ A) (Var i) := elim_var_n i ;
+ p_dom_of_w_0n (¬ A) (Var i) := elim_var_n i ;
+ p_dom_of_w_0n (⊥) (BotL) := a_empty ;
+ p_dom_of_w_0n (A ⅋ B) (ParL k1 k2) := [ p_dom_of_w_0n A k1 , p_dom_of_w_0n B k2 ] ;
+ p_dom_of_w_0n (A & B) (AndL1 k) := p_dom_of_w_0n A k ;
+ p_dom_of_w_0n (A & B) (AndL2 k) := p_dom_of_w_0n B k ;
+ p_dom_of_w_0n (↑ A) (ShiftNL x) := a_append a_empty x ;
+ p_dom_of_w_0n (¬ A) (NegNL v) := p_dom_of_w_0p A v .
We can now package up all these auxiliary functions into the following ones, abstracting
+polarity and side-annotation.
+
Equationsp_of_v {Γ : neg_ctx} A : val Γ A -> pat A :=
+ p_of_v (@LTy pos A) v := p_of_w_0p A v ;
+ p_of_v (@RTy pos A) _ := PVarN A ;
+ p_of_v (@LTy neg A) _ := PVarP A ;
+ p_of_v (@RTy neg A) k := p_of_w_0n A k .
+
+Equationsp_dom_of_v {Γ : neg_ctx} A (v : val Γ A) : p_dom (p_of_v A v) =[val]> Γ :=
+ p_dom_of_v (@LTy pos A) v := p_dom_of_w_0p A v ;
+ p_dom_of_v (@RTy pos A) x := [ ! ,ₓ x ] ;
+ p_dom_of_v (@LTy neg A) x := [ ! ,ₓ x ] ;
+ p_dom_of_v (@RTy neg A) k := p_dom_of_w_0n A k .
+
+Definitionv_split_p {Γ : neg_ctx} {A} (v : whn Γ `+A) : (obs_op # val) Γ `-A
+ := (p_of_w_0p A v : o_op obs_op `-A) ⦇ p_dom_of_w_0p A v ⦈.
+
+Definitionv_split_n {Γ : neg_ctx} {A} (v : whn Γ `-A) : (obs_op # val) Γ `+A
+ := (p_of_w_0n A v : o_op obs_op `+_) ⦇ p_dom_of_w_0n A v ⦈.
+
+
Evaluation
+
With patterns and observations now in hand we prepare for the definition of evaluation
+and define a shorthand for normal forms. 'Normal forms' are here understood---as in the
+paper---in our slightly non-standard presentation of triplets of a variable, an
+observation and an assignment.
Now the bulk of evaluation: the step function. Once again we are greatful for
+Equations providing us with a justification for the fact that this complex
+dependent pattern-matching is indeed total.
Finally we define evaluation as the iteration of the step function in the Delay monad,
+and also define application of an observation with arguments to a value.
+
Definitioneval {Γ : neg_ctx} : state Γ -> delay (nf Γ)
+ := iter_delay (func => Ret' (eval_aux c)).
+
+Definitionp_app {ΓA} (v : val Γ A) (m : pat A†) (e : p_dom m =[val]> Γ) : state Γ :=
+ Cut' v (m `ᵥ⊛ e) .
+
+
Metatheory
+
Now comes a rather ugly part: the metatheory of our syntax. Comments will be rather more
+sparse. For a thorough explaination of its structure, see Examples/Lambda/CBLTyped.v.
+We will here be concerned with extensional equality preservation, identity and composition
+laws for renaming and substitution, and also refolding lemmas for splitting and embedding
+patterns. You are encouraged to just skip until line ~1300.
+
Schemeterm_mut := Induction for term SortProp
+ with whn_mut := Induction for whn SortProp
+ with state_mut := Induction for state SortProp.
+
+Recordsyn_ind_args
+ (P_t : forallΓA, term Γ A -> Prop)
+ (P_w : forallΓA, whn Γ A -> Prop)
+ (P_s : forallΓ, state Γ -> Prop) :=
+{
+ ind_mu : forallΓAs (H : P_s _ s), P_t Γ A (Mu s) ;
+ ind_recp : forallΓAt (H : P_t _ _ t), P_t Γ `-A (RecL t) ;
+ ind_recn : forallΓAt (H : P_t _ _ t), P_t Γ `+A (RecR t) ;
+ ind_whn : forallΓAw (H : P_w _ _ w), P_t Γ A (Whn w) ;
+ ind_var : forallΓAh, P_w Γ A (Var h) ;
+ ind_zerl : forallΓ, P_w Γ `-0 ZerL ;
+ ind_topr : forallΓ, P_w Γ `+⊤ TopR ;
+ ind_oner : forallΓ, P_w Γ `+1 OneR ;
+ ind_onel : forallΓs, P_s Γ s -> P_w Γ `-1 (OneL s) ;
+ ind_botr : forallΓs, P_s Γ s -> P_w Γ `+⊥ (BotR s) ;
+ ind_botl : forallΓ, P_w Γ `-⊥ BotL ;
+ ind_tenr : forallΓABw1 (H1 : P_w _ _ w1) w2 (H2 : P_w _ _ w2), P_w Γ `+(A ⊗ B) (TenR w1 w2) ;
+ ind_tenl : forallΓABs (H : P_s _ s), P_w Γ `-(A ⊗ B) (TenL s) ;
+ ind_parr : forallΓABs (H : P_s _ s), P_w Γ `+(A ⅋ B) (ParR s) ;
+ ind_parl : forallΓABw1 (H1 : P_w _ _ w1) w2 (H2 : P_w Γ `-B w2), P_w Γ `-(A ⅋ B) (ParL w1 w2) ;
+ ind_orr1 : forallΓABw (H : P_w _ _ w), P_w Γ `+(A ⊕ B) (OrR1 w) ;
+ ind_orr2 : forallΓABw (H : P_w _ _ w), P_w Γ `+(A ⊕ B) (OrR2 w) ;
+ ind_orl : forallΓABs1 (H1 : P_s _ s1) s2 (H2 : P_s _ s2), P_w Γ `-(A ⊕ B) (OrL s1 s2) ;
+ ind_andr : forallΓABs1 (H1 : P_s _ s1) s2 (H2 : P_s _ s2), P_w Γ `+(A & B) (AndR s1 s2) ;
+ ind_andl1 : forallΓABw (H : P_w _ _ w), P_w Γ `-(A & B) (AndL1 w) ;
+ ind_andl2 : forallΓABw (H : P_w _ _ w), P_w Γ `-(A & B) (AndL2 w) ;
+ ind_shiftpr : forallΓAt (H : P_t _ _ t), P_w Γ `+(↓ A) (ShiftPR t) ;
+ ind_shiftpl : forallΓAs (H : P_s _ s), P_w Γ `-(↓ A) (ShiftPL s) ;
+ ind_shiftnr : forallΓAs (H : P_s _ s), P_w Γ `+(↑ A) (ShiftNR s) ;
+ ind_shiftnl : forallΓAt (H : P_t _ _ t), P_w Γ `-(↑ A) (ShiftNL t) ;
+ ind_negpr : forallΓAw (H : P_w _ _ w), P_w Γ `+(⊖ A) (NegPR w) ;
+ ind_negpl : forallΓAs (H : P_s _ s), P_w Γ `-(⊖ A) (NegPL s) ;
+ ind_negnr : forallΓAs (H : P_s _ s), P_w Γ `+(¬ A) (NegNR s) ;
+ ind_negnl : forallΓAw (H : P_w _ _ w), P_w Γ `-(¬ A) (NegNL w) ;
+ ind_cut : forallΓpAt1 (H1 : P_t _ _ t1) t2 (H2 : P_t _ _ t2), P_s Γ (@Cut _ p A t1 t2)
+} .
+
+Lemmaterm_ind_mutP_tP_wP_s (H : syn_ind_args P_t P_w P_s) ΓAt : P_t Γ A t .
+destruct H; nowapply (term_mut P_t P_w P_s).
+Qed.
+
+Lemmawhn_ind_mutP_tP_wP_s (H : syn_ind_args P_t P_w P_s) ΓAw : P_w Γ A w .
+destruct H; nowapply (whn_mut P_t P_w P_s).
+Qed.
+
+Lemmastate_ind_mutP_tP_wP_s (H : syn_ind_args P_t P_w P_s) Γs : P_s Γ s .
+destruct H; nowapply (state_mut P_t P_w P_s).
+Qed.
+
+Definitiont_ren_proper_PΓA (t : term Γ A) : Prop :=
+ forallΔ (f1f2 : Γ ⊆ Δ), f1 ≡ₐ f2 -> t ₜ⊛ᵣ f1 = t ₜ⊛ᵣ f2 .
+Definitionw_ren_proper_PΓA (v : whn Γ A) : Prop :=
+ forallΔ (f1f2 : Γ ⊆ Δ), f1 ≡ₐ f2 -> v `ᵥ⊛ᵣ f1 = v `ᵥ⊛ᵣ f2 .
+Definitions_ren_proper_PΓ (s : state Γ) : Prop :=
+ forallΔ (f1f2 : Γ ⊆ Δ), f1 ≡ₐ f2 -> s ₛ⊛ᵣ f1 = s ₛ⊛ᵣ f2 .
+Lemmaren_proper_prf : syn_ind_args t_ren_proper_P w_ren_proper_P s_ren_proper_P.
+ econstructor.
+ all: unfold t_ren_proper_P, w_ren_proper_P, s_ren_proper_P.
+ all: intros; cbn; f_equal; eauto.
+ all: first [ apply H | apply H1 | apply H2 ]; auto.
+ all: first [ apply r_shift1_eq | apply r_shift2_eq ]; auto.
+Qed.
+
+#[global] Instancet_ren_eq {ΓatΔ} : Proper (asgn_eq _ _ ==> eq) (t_rename Γ a t Δ).
+ intros f1 f2 H1; nowapply (term_ind_mut _ _ _ ren_proper_prf).
+Qed.
+
+#[global] Instancew_ren_eq {ΓavΔ} : Proper (asgn_eq _ _ ==> eq) (w_rename Γ a v Δ).
+ intros f1 f2 H1; nowapply (whn_ind_mut _ _ _ ren_proper_prf).
+Qed.
+
+#[global] Instances_ren_eq {ΓsΔ} : Proper (asgn_eq _ _ ==> eq) (s_rename Γ s Δ).
+ intros f1 f2 H1; nowapply (state_ind_mut _ _ _ ren_proper_prf).
+Qed.
+
+#[global] Instancev_ren_eq {ΓavΔ} : Proper (asgn_eq _ _ ==> eq) (v_rename Γ a v Δ).
+ destruct a as [ [] | [] ].
+ nowapply w_ren_eq.
+ nowapply t_ren_eq.
+ nowapply t_ren_eq.
+ nowapply w_ren_eq.
+Qed.
+
+#[global] Instancea_ren_eq {Γ1Γ2Γ3}
+ : Proper (asgn_eq _ _ ==> asgn_eq _ _ ==> asgn_eq _ _) (@a_ren Γ1 Γ2 Γ3).
+ intros r1 r2 H1 a1 a2 H2 ? i; cbn; nowrewrite H1, (v_ren_eq _ _ H2).
+Qed.
+
+#[global] Instancea_shift1_eq {ΓΔA} : Proper (asgn_eq _ _ ==> asgn_eq _ _) (@a_shift1 Γ Δ A).
+ intros ? ? H ? h.
+ dependent elimination h; auto; cbn; nowrewrite H.
+Qed.
+
+#[global] Instancea_shift2_eq {ΓΔAB} : Proper (asgn_eq _ _ ==> asgn_eq _ _) (@a_shift2 Γ Δ A B).
+ intros ? ? H ? v.
+ do2 (dependent elimination v; auto).
+ cbn; nowrewrite H.
+Qed.
+
+Definitiont_ren_ren_PΓ1A (t : term Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3),
+ (t ₜ⊛ᵣ f1) ₜ⊛ᵣ f2 = t ₜ⊛ᵣ (f1 ᵣ⊛ f2) .
+Definitionw_ren_ren_PΓ1A (v : whn Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3),
+ (v `ᵥ⊛ᵣ f1) `ᵥ⊛ᵣ f2 = v `ᵥ⊛ᵣ (f1 ᵣ⊛ f2) .
+Definitions_ren_ren_PΓ1 (s : state Γ1) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3),
+ (s ₛ⊛ᵣ f1) ₛ⊛ᵣ f2 = s ₛ⊛ᵣ (f1 ᵣ⊛ f2) .
+
+Lemmaren_ren_prf : syn_ind_args t_ren_ren_P w_ren_ren_P s_ren_ren_P.
+ econstructor.
+ all: unfold t_ren_ren_P, w_ren_ren_P, s_ren_ren_P.
+ all: intros; cbn; f_equal; eauto.
+ all: first [ rewrite r_shift1_comp | rewrite r_shift2_comp ]; eauto.
+Qed.
+
+Lemmat_ren_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3) A (t : term Γ1 A)
+ : (t ₜ⊛ᵣ f1) ₜ⊛ᵣ f2 = t ₜ⊛ᵣ (f1 ᵣ⊛ f2) .
+ nowapply (term_ind_mut _ _ _ ren_ren_prf).
+Qed.
+Lemmaw_ren_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3) A (v : whn Γ1 A)
+ : (v `ᵥ⊛ᵣ f1) `ᵥ⊛ᵣ f2 = v `ᵥ⊛ᵣ (f1 ᵣ⊛ f2) .
+ nowapply (whn_ind_mut _ _ _ ren_ren_prf).
+Qed.
+Lemmas_ren_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3) (s : state Γ1)
+ : (s ₛ⊛ᵣ f1) ₛ⊛ᵣ f2 = s ₛ⊛ᵣ (f1 ᵣ⊛ f2) .
+ nowapply (state_ind_mut _ _ _ ren_ren_prf).
+Qed.
+Lemmav_ren_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 ⊆ Γ3) A (v : val Γ1 A)
+ : (v ᵥ⊛ᵣ f1) ᵥ⊛ᵣ f2 = v ᵥ⊛ᵣ (f1 ᵣ⊛ f2) .
+ destruct A as [ [] | [] ].
+ nowapply w_ren_ren.
+ nowapply t_ren_ren.
+ nowapply t_ren_ren.
+ nowapply w_ren_ren.
+Qed.
+
+Definitiont_ren_id_l_PΓA (t : term Γ A) : Prop := t ₜ⊛ᵣ r_id = t.
+Definitionw_ren_id_l_PΓA (v : whn Γ A) : Prop := v `ᵥ⊛ᵣ r_id = v.
+Definitions_ren_id_l_PΓ (s : state Γ) : Prop := s ₛ⊛ᵣ r_id = s.
+
+Lemmaren_id_l_prf : syn_ind_args t_ren_id_l_P w_ren_id_l_P s_ren_id_l_P.
+ econstructor.
+ all: unfold t_ren_id_l_P, w_ren_id_l_P, s_ren_id_l_P.
+ all: intros; cbn; f_equal; eauto.
+ all: first [ rewrite r_shift1_id | rewrite r_shift2_id ]; eauto.
+Qed.
+
+Lemmat_ren_id_l {Γ} A (t : term Γ A) : t ₜ⊛ᵣ r_id = t.
+ nowapply (term_ind_mut _ _ _ ren_id_l_prf).
+Qed.
+Lemmaw_ren_id_l {Γ} A (v : whn Γ A) : v `ᵥ⊛ᵣ r_id = v.
+ nowapply (whn_ind_mut _ _ _ ren_id_l_prf).
+Qed.
+Lemmas_ren_id_l {Γ} (s : state Γ) : s ₛ⊛ᵣ r_id = s.
+ nowapply (state_ind_mut _ _ _ ren_id_l_prf).
+Qed.
+Lemmav_ren_id_l {Γ} A (v : val Γ A) : v ᵥ⊛ᵣ r_id = v.
+ destruct A as [ [] | [] ].
+ nowapply w_ren_id_l.
+ nowapply t_ren_id_l.
+ nowapply t_ren_id_l.
+ nowapply w_ren_id_l.
+Qed.
+
+Lemmav_ren_id_r {ΓΔ} (f : Γ ⊆ Δ) A (i : Γ ∋ A) : (var i) ᵥ⊛ᵣ f = var (f _ i).
+ nowdestruct A as [ [] | [] ].
+Qed.
+
+Lemmaa_shift1_id {ΓA} : @a_shift1 Γ Γ A var ≡ₐ var.
+ intros [ [] | [] ] i; dependent elimination i; auto.
+Qed.
+
+Lemmaa_shift2_id {ΓAB} : @a_shift2 Γ Γ A B var ≡ₐ var.
+ intros ? v; cbn.
+ do2 (dependent elimination v; cbn; auto).
+ nowdestruct a as [[]|[]].
+Qed.
+
+Arguments var : simpl never.
+Lemmaa_shift1_ren_r {Γ1Γ2Γ3y} (f1 : Γ1 =[ val ]> Γ2) (f2 : Γ2 ⊆ Γ3)
+ : a_shift1 (y:=y) (f1 ⊛ᵣ f2) ≡ₐ a_shift1 f1 ⊛ᵣ r_shift1 f2 .
+ intros ? h; dependent elimination h; cbn.
+ - nowrewrite v_ren_id_r.
+ - nowunfold v_shift1; rewrite2 v_ren_ren.
+Qed.
+
+Lemmaa_shift2_ren_r {Γ1Γ2Γ3yz} (f1 : Γ1 =[ val ]> Γ2) (f2 : Γ2 ⊆ Γ3)
+ : a_shift2 (y:=y) (z:=z) (f1 ⊛ᵣ f2) ≡ₐ a_shift2 f1 ⊛ᵣ r_shift2 f2 .
+ intros ? v; do2 (dependent elimination v; cbn; [ nowrewrite v_ren_id_r | ]).
+ unfold v_shift2; nowrewrite2 v_ren_ren.
+Qed.
+
+Lemmaa_shift1_ren_l {Γ1Γ2Γ3y} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3)
+ : a_shift1 (y:=y) (f1 ᵣ⊛ f2) ≡ₐ r_shift1 f1 ᵣ⊛ a_shift1 f2 .
+ intros ? i; dependent elimination i; auto.
+Qed.
+
+Lemmaa_shift2_ren_l {Γ1Γ2Γ3yz} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3)
+ : a_shift2 (y:=y) (z:=z) (f1 ᵣ⊛ f2) ≡ₐ r_shift2 f1 ᵣ⊛ a_shift2 f2 .
+ intros ? v; do2 (dependent elimination v; auto).
+Qed.
+
+Definitiont_sub_proper_PΓA (t : term Γ A) : Prop :=
+ forallΔ (f1f2 : Γ =[val]> Δ), f1 ≡ₐ f2 -> t `ₜ⊛ f1 = t `ₜ⊛ f2 .
+Definitionw_sub_proper_PΓA (v : whn Γ A) : Prop :=
+ forallΔ (f1f2 : Γ =[val]> Δ), f1 ≡ₐ f2 -> v `ᵥ⊛ f1 = v `ᵥ⊛ f2 .
+Definitions_sub_proper_PΓ (s : state Γ) : Prop :=
+ forallΔ (f1f2 : Γ =[val]> Δ), f1 ≡ₐ f2 -> s ₜ⊛ f1 = s ₜ⊛ f2 .
+
+Lemmasub_proper_prf : syn_ind_args t_sub_proper_P w_sub_proper_P s_sub_proper_P.
+ econstructor.
+ all: unfold t_sub_proper_P, w_sub_proper_P, s_sub_proper_P.
+ all: intros; cbn.
+ all: match goal with
+ | |- Whn _ = Whn _ => do2f_equal
+ | _ => f_equal
+ end .
+ all: first [ apply H | apply H1 | apply H2 ]; auto.
+ all: first [ apply a_shift1_eq | apply a_shift2_eq ]; auto.
+Qed.
+
+#[global] Instancet_sub_eq {ΓatΔ} : Proper (asgn_eq _ _ ==> eq) (t_subst Γ a t Δ).
+ intros f1 f2 H1; nowapply (term_ind_mut _ _ _ sub_proper_prf).
+Qed.
+
+#[global] Instancew_sub_eq {ΓavΔ} : Proper (asgn_eq _ _ ==> eq) (w_subst Γ a v Δ).
+ intros f1 f2 H1; nowapply (whn_ind_mut _ _ _ sub_proper_prf).
+Qed.
+
+#[global] Instances_sub_eq {ΓsΔ} : Proper (asgn_eq _ _ ==> eq) (s_subst Γ s Δ).
+ intros f1 f2 H1; nowapply (state_ind_mut _ _ _ sub_proper_prf).
+Qed.
+
+#[global] Instancev_sub_eq {ΓavΔ} : Proper (asgn_eq _ _ ==> eq) (v_subst Γ a v Δ).
+ destruct a as [[]|[]].
+ - nowapply w_sub_eq.
+ - nowapply t_sub_eq.
+ - nowapply t_sub_eq.
+ - nowapply w_sub_eq.
+Qed.
+
+#[global] Instancea_comp_eq {Γ1Γ2Γ3}
+ : Proper (asgn_eq _ _ ==> asgn_eq _ _ ==> asgn_eq _ _) (@a_comp _ _ _ _ _ Γ1 Γ2 Γ3).
+ intros ? ? H1 ? ? H2 ? ?; cbn; rewrite H1; noweapply v_sub_eq.
+Qed.
+
+Definitiont_ren_sub_PΓ1A (t : term Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3),
+ (t `ₜ⊛ f1) ₜ⊛ᵣ f2 = t `ₜ⊛ (f1 ⊛ᵣ f2) .
+Definitionw_ren_sub_PΓ1A (v : whn Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3),
+ (v `ᵥ⊛ f1) ᵥ⊛ᵣ f2 = v `ᵥ⊛ (f1 ⊛ᵣ f2) .
+Definitions_ren_sub_PΓ1 (s : state Γ1) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3),
+ (s ₜ⊛ f1) ₛ⊛ᵣ f2 = s ₜ⊛ (f1 ⊛ᵣ f2) .
+Lemmaren_sub_prf : syn_ind_args t_ren_sub_P w_ren_sub_P s_ren_sub_P.
+ econstructor.
+ all: unfold t_ren_sub_P, w_ren_sub_P, s_ren_sub_P.
+ all: intros; cbn.
+ 4: destruct A as [ [] | [] ]; cbn.
+ all: match goal with
+ | |- Whn _ = Whn _ => do2f_equal
+ | _ => f_equal
+ end ; eauto.
+ all: first [ rewrite a_shift1_ren_r | rewrite a_shift2_ren_r ]; auto.
+Qed.
+
+Lemmat_ren_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3) A (t : term Γ1 A)
+ : (t `ₜ⊛ f1) ₜ⊛ᵣ f2 = t `ₜ⊛ (f1 ⊛ᵣ f2) .
+ nowapply (term_ind_mut _ _ _ ren_sub_prf).
+Qed.
+Lemmaw_ren_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3) A (v : whn Γ1 A)
+ : (v `ᵥ⊛ f1) ᵥ⊛ᵣ f2 = v `ᵥ⊛ (f1 ⊛ᵣ f2) .
+ nowapply (whn_ind_mut _ _ _ ren_sub_prf).
+Qed.
+Lemmas_ren_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3) (s : state Γ1)
+ : (s ₜ⊛ f1) ₛ⊛ᵣ f2 = s ₜ⊛ (f1 ⊛ᵣ f2) .
+ nowapply (state_ind_mut _ _ _ ren_sub_prf).
+Qed.
+Lemmav_ren_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 ⊆ Γ3) A (v : val Γ1 A)
+ : (v ᵥ⊛ f1) ᵥ⊛ᵣ f2 = v ᵥ⊛ (f1 ⊛ᵣ f2) .
+ destruct A as [ [] | [] ]; cbn [ v_subst ].
+ - nowapply w_ren_sub.
+ - nowapply t_ren_sub.
+ - nowapply t_ren_sub.
+ - nowapply w_ren_sub.
+Qed.
+
+Definitiont_sub_ren_PΓ1A (t : term Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3),
+ (t ₜ⊛ᵣ f1) `ₜ⊛ f2 = t `ₜ⊛ (f1 ᵣ⊛ f2).
+Definitionw_sub_ren_PΓ1A (v : whn Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3),
+ (v `ᵥ⊛ᵣ f1) `ᵥ⊛ f2 = v `ᵥ⊛ (f1 ᵣ⊛ f2).
+Definitions_sub_ren_PΓ1 (s : state Γ1) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3),
+ (s ₛ⊛ᵣ f1) ₜ⊛ f2 = s ₜ⊛ (f1 ᵣ⊛ f2).
+
+Lemmasub_ren_prf : syn_ind_args t_sub_ren_P w_sub_ren_P s_sub_ren_P.
+ econstructor.
+ all: unfold t_sub_ren_P, w_sub_ren_P, s_sub_ren_P.
+ all: intros; cbn.
+ all: match goal with
+ | |- Whn _ = Whn _ => do2f_equal
+ | _ => f_equal
+ end ; eauto.
+ all: first [ rewrite a_shift1_ren_l | rewrite a_shift2_ren_l ]; auto.
+Qed.
+
+Lemmat_sub_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3) A (t : term Γ1 A)
+ : (t ₜ⊛ᵣ f1) `ₜ⊛ f2 = t `ₜ⊛ (f1 ᵣ⊛ f2).
+ nowapply (term_ind_mut _ _ _ sub_ren_prf).
+Qed.
+Lemmaw_sub_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3) A (v : whn Γ1 A)
+ : (v `ᵥ⊛ᵣ f1) `ᵥ⊛ f2 = v `ᵥ⊛ (f1 ᵣ⊛ f2).
+ nowapply (whn_ind_mut _ _ _ sub_ren_prf).
+Qed.
+Lemmas_sub_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3) (s : state Γ1)
+ : (s ₛ⊛ᵣ f1) ₜ⊛ f2 = s ₜ⊛ (f1 ᵣ⊛ f2).
+ nowapply (state_ind_mut _ _ _ sub_ren_prf).
+Qed.
+Lemmav_sub_ren {Γ1Γ2Γ3} (f1 : Γ1 ⊆ Γ2) (f2 : Γ2 =[val]> Γ3) A (v : val Γ1 A)
+ : (v ᵥ⊛ᵣ f1) ᵥ⊛ f2 = v ᵥ⊛ (f1 ᵣ⊛ f2).
+ destruct A as [ [] | [] ].
+ - nowapply w_sub_ren.
+ - nowapply t_sub_ren.
+ - nowapply t_sub_ren.
+ - nowapply w_sub_ren.
+Qed.
+
+Lemmav_sub_id_r {ΓΔ} (f : Γ =[val]> Δ) A (i : Γ ∋ A) : var i ᵥ⊛ f = f A i.
+ destruct A as [ [] | [] ]; auto.
+Qed.
+
+Lemmaa_shift1_comp {Γ1Γ2Γ3A} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3)
+ : @a_shift1 _ _ A (f1 ⊛ f2) ≡ₐ a_shift1 f1 ⊛ a_shift1 f2 .
+ intros x i; dependent elimination i; cbn.
+ - nowrewrite v_sub_id_r.
+ - nowunfold v_shift1; rewrite v_ren_sub, v_sub_ren.
+Qed.
+
+Lemmaa_shift2_comp {Γ1Γ2Γ3AB} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3)
+ : @a_shift2 _ _ A B (f1 ⊛ f2) ≡ₐ a_shift2 f1 ⊛ a_shift2 f2 .
+ intros ? v; do2 (dependent elimination v; cbn; [ nowrewrite v_sub_id_r | ]).
+ nowunfold v_shift2; rewrite v_ren_sub, v_sub_ren.
+Qed.
+
+Definitiont_sub_sub_PΓ1A (t : term Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3),
+ (t `ₜ⊛ f1) `ₜ⊛ f2 = t `ₜ⊛ (f1 ⊛ f2) .
+Definitionw_sub_sub_PΓ1A (v : whn Γ1 A) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3),
+ (v `ᵥ⊛ f1) ᵥ⊛ f2 = v `ᵥ⊛ (f1 ⊛ f2) .
+Definitions_sub_sub_PΓ1 (s : state Γ1) : Prop :=
+ forallΓ2Γ3 (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3),
+ (s ₜ⊛ f1) ₜ⊛ f2 = s ₜ⊛ (f1 ⊛ f2) .
+
+Lemmasub_sub_prf : syn_ind_args t_sub_sub_P w_sub_sub_P s_sub_sub_P.
+ econstructor.
+ all: unfold t_sub_sub_P, w_sub_sub_P, s_sub_sub_P; intros; cbn.
+ 4: destruct A as [ [] | [] ]; cbn.
+ all: match goal with
+ | |- Whn _ = Whn _ => do2f_equal
+ | _ => f_equal
+ end ; eauto.
+ all: first [ rewrite a_shift1_comp | rewrite a_shift2_comp ]; auto.
+Qed.
+
+Lemmat_sub_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3) A (t : term Γ1 A)
+ : (t `ₜ⊛ f1) `ₜ⊛ f2 = t `ₜ⊛ (f1 ⊛ f2) .
+ nowapply (term_ind_mut _ _ _ sub_sub_prf).
+Qed.
+Lemmaw_sub_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3) A (v : whn Γ1 A)
+ : (v `ᵥ⊛ f1) ᵥ⊛ f2 = v `ᵥ⊛ (f1 ⊛ f2) .
+ nowapply (whn_ind_mut _ _ _ sub_sub_prf).
+Qed.
+Lemmas_sub_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3) (s : state Γ1)
+ : (s ₜ⊛ f1) ₜ⊛ f2 = s ₜ⊛ (f1 ⊛ f2) .
+ nowapply (state_ind_mut _ _ _ sub_sub_prf).
+Qed.
+Lemmav_sub_sub {Γ1Γ2Γ3} (f1 : Γ1 =[val]> Γ2) (f2 : Γ2 =[val]> Γ3) A (v : val Γ1 A)
+ : (v ᵥ⊛ f1) ᵥ⊛ f2 = v ᵥ⊛ (f1 ⊛ f2) .
+ destruct A as [ [] | [] ].
+ - nowapply w_sub_sub.
+ - nowapply t_sub_sub.
+ - nowapply t_sub_sub.
+ - nowapply w_sub_sub.
+Qed.
+
+Lemmaa_comp_assoc {Γ1Γ2Γ3Γ4} (u : Γ1 =[val]> Γ2) (v : Γ2 =[val]> Γ3) (w : Γ3 =[val]> Γ4)
+ : (u ⊛ v) ⊛ w ≡ₐ u ⊛ (v ⊛ w).
+ intros ? i; unfold a_comp; nowapply v_sub_sub.
+Qed.
+
+Definitiont_sub_id_l_PΓA (t : term Γ A) : Prop := t `ₜ⊛ var = t.
+Definitionw_sub_id_l_PΓA (v : whn Γ A) : Prop := v `ᵥ⊛ var = v.
+Definitions_sub_id_l_PΓ (s : state Γ) : Prop := s ₜ⊛ var = s.
+
+Lemmasub_id_l_prf : syn_ind_args t_sub_id_l_P w_sub_id_l_P s_sub_id_l_P.
+ econstructor.
+ all: unfold t_sub_id_l_P, w_sub_id_l_P, s_sub_id_l_P; intros; cbn.
+ 4,5: destruct A as [ [] | [] ]; cbn.
+ all: match goal with
+ | |- Whn _ = Whn _ => do2f_equal
+ | _ => f_equal
+ end; eauto.
+ all: first [ rewrite a_shift1_id | rewrite a_shift2_id ]; auto.
+Qed.
+
+Lemmat_sub_id_l {Γ} A (t : term Γ A) : t `ₜ⊛ var = t.
+ nowapply (term_ind_mut _ _ _ sub_id_l_prf).
+Qed.
+Lemmaw_sub_id_l {Γ} A (v : whn Γ A) : v `ᵥ⊛ var = v.
+ nowapply (whn_ind_mut _ _ _ sub_id_l_prf).
+Qed.
+Lemmas_sub_id_l {Γ} (s : state Γ) : s ₜ⊛ var = s.
+ nowapply (state_ind_mut _ _ _ sub_id_l_prf).
+Qed.
+Lemmav_sub_id_l {Γ} A (v : val Γ A) : v ᵥ⊛ var = v.
+ destruct A as [ [] | [] ].
+ - nowapply w_sub_id_l.
+ - nowapply t_sub_id_l.
+ - nowapply t_sub_id_l.
+ - nowapply w_sub_id_l.
+Qed.
+
+Lemmasub1_sub {ΓΔA} (f : Γ =[val]> Δ) (v : val Γ A) :
+ a_shift1 f ⊛ asgn1 (v ᵥ⊛ f) ≡ₐ asgn1 v ⊛ f.
+ intros ? h; dependent elimination h; cbn.
+ - nowrewrite v_sub_id_r.
+ - unfold v_shift1; rewrite v_sub_ren, v_sub_id_r.
+ nowapply v_sub_id_l.
+Qed.
+
+Lemmasub1_ren {ΓΔA} (f : Γ ⊆ Δ) (v : val Γ A) :
+ r_shift1 f ᵣ⊛ asgn1 (v ᵥ⊛ᵣ f) ≡ₐ asgn1 v ⊛ᵣ f.
+ intros ? h; dependent elimination h; auto.
+ cbn; nowrewrite v_ren_id_r.
+Qed.
+
+Lemmav_sub1_sub {ΓΔAB} (f : Γ =[val]> Δ) (v : val Γ A) (w : val (Γ ▶ₓ A) B)
+ : (w ᵥ⊛ a_shift1 f) ᵥ⊛ ₁[ v ᵥ⊛ f ] = (w ᵥ⊛ ₁[ v ]) ᵥ⊛ f .
+ cbn; rewrite2 v_sub_sub.
+ apply v_sub_eq; nowrewrite sub1_sub.
+Qed.
+
+Lemmav_sub1_ren {ΓΔAB} (f : Γ ⊆ Δ) (v : val Γ A) (w : val (Γ ▶ₓ A) B)
+ : (w ᵥ⊛ᵣ r_shift1 f) ᵥ⊛ ₁[ v ᵥ⊛ᵣ f ] = (w ᵥ⊛ ₁[ v ]) ᵥ⊛ᵣ f .
+ cbn; rewrite v_sub_ren, v_ren_sub.
+ apply v_sub_eq; nowrewrite sub1_ren.
+Qed.
+
+Lemmas_sub1_sub {ΓΔA} (f : Γ =[val]> Δ) (v : val Γ A) (s : state (Γ ▶ₓ A))
+ : (s ₜ⊛ a_shift1 f) ₜ⊛ ₁[ v ᵥ⊛ f ] = (s ₜ⊛ ₁[ v ]) ₜ⊛ f .
+ cbn; nowrewrite2 s_sub_sub, sub1_sub.
+Qed.
+
+Lemmas_sub2_sub {ΓΔAB} (f : Γ =[val]> Δ) (s : state (Γ ▶ₓ A ▶ₓ B)) uv
+ : (s ₜ⊛ a_shift2 f) ₜ⊛ ₂[ u ᵥ⊛ f , v ᵥ⊛ f ] = (s ₜ⊛ ₂[ u, v ]) ₜ⊛ f .
+ cbn; rewrite2 s_sub_sub; apply s_sub_eq.
+ intros ? v0; cbn.
+ do2 (dependent elimination v0; cbn; [ nowrewrite v_sub_id_r | ]).
+ unfold v_shift2; rewrite v_sub_ren, v_sub_id_r, <- v_sub_id_l.
+ nowapply v_sub_eq.
+Qed.
+
+Lemmas_sub1_ren {ΓΔA} (f : Γ ⊆ Δ) (v : val Γ A) (s : state (Γ ▶ₓ A))
+ : (s ₛ⊛ᵣ r_shift1 f) ₜ⊛ ₁[ v ᵥ⊛ᵣ f ] = (s ₜ⊛ ₁[ v ]) ₛ⊛ᵣ f .
+ cbn; nowrewrite s_sub_ren, s_ren_sub, sub1_ren.
+Qed.
+
+Lemmat_sub1_sub {ΓΔAB} (f : Γ =[val]> Δ) (v : val Γ A) (t : term (Γ ▶ₓ A) B)
+ : (t `ₜ⊛ a_shift1 f) `ₜ⊛ ₁[ v ᵥ⊛ f ] = (t `ₜ⊛ ₁[ v ]) `ₜ⊛ f .
+ cbn; rewrite2 t_sub_sub.
+ apply t_sub_eq; nowrewrite sub1_sub.
+Qed.
+
+Lemmat_sub1_ren {ΓΔAB} (f : Γ ⊆ Δ) (v : val Γ A) (t : term (Γ ▶ₓ A) B)
+ : (t ₜ⊛ᵣ r_shift1 f) `ₜ⊛ ₁[ v ᵥ⊛ᵣ f ] = (t `ₜ⊛ ₁[ v ]) ₜ⊛ᵣ f .
+ cbn; rewrite t_sub_ren, t_ren_sub.
+ apply t_sub_eq; nowrewrite sub1_ren.
+Qed.
+
+#[global] Instancep_app_eq {ΓA} (v : val Γ A) (m : pat (t_neg A))
+ : Proper (asgn_eq _ _ ==> eq) (p_app v m) .
+ intros u1 u2 H; destruct A as [ [] | []]; cbn; nowrewrite (w_sub_eq u1 u2 H).
+Qed.
+
+Definitionrefold_id_aux_P (Γ : neg_ctx) p : pre_ty p -> Prop :=
+ match p with
+ | pos => funA => forall (v : whn Γ `+A), p_of_w_0p _ v `ᵥ⊛ p_dom_of_w_0p _ v = v
+ | neg => funA => forall (v : whn Γ `-A), p_of_w_0n _ v `ᵥ⊛ p_dom_of_w_0n _ v = v
+ end .
+
+Lemmarefold_id_aux {Γp} A : refold_id_aux_P Γ p A .
+ induction A; intros v.
+ - dependent elimination v; destruct (s_prf c).
+ - dependent elimination v; destruct (s_prf c).
+ - dependent elimination v; [ destruct (s_prf c) | ]; auto.
+ - dependent elimination v; [ destruct (s_prf c) | ]; auto.
+ - dependent elimination v; [ destruct (s_prf c) | ].
+ cbn; f_equal; rewrite w_sub_ren.
+ rewrite <- IHA1; apply w_sub_eq; exact (a_cat_proj_l _ _).
+ rewrite <- IHA2; apply w_sub_eq; exact (a_cat_proj_r _ _).
+ - dependent elimination v; [ destruct (s_prf c) | ].
+ cbn; f_equal; rewrite w_sub_ren.
+ rewrite <- IHA1; apply w_sub_eq; exact (a_cat_proj_l _ _).
+ rewrite <- IHA2; apply w_sub_eq; exact (a_cat_proj_r _ _).
+ - dependent elimination v; [ destruct (s_prf c) | | ];
+ cbn; f_equal; [ apply IHA1 | apply IHA2 ].
+ - dependent elimination v; [ destruct (s_prf c) | | ];
+ cbn; f_equal; [ apply IHA1 | apply IHA2 ].
+ - dependent elimination v; [ destruct (s_prf c) | ]; cbn; f_equal.
+ - dependent elimination v; [ destruct (s_prf c) | ]; cbn; f_equal.
+ - dependent elimination v; [ destruct (s_prf c) | ].
+ cbn; f_equal; apply IHA.
+ - dependent elimination v; [ destruct (s_prf c) | ].
+ cbn; f_equal; apply IHA.
+Qed.
+
+Lemmarefold_id {Γ : neg_ctx} A (v : val Γ A)
+ : p_of_v A v `ᵥ⊛ p_dom_of_v A v = v.
+ destruct A as [ [] A | [] A ]; auto; exact (refold_id_aux A v).
+Qed.
+
+Equationsp_of_w_eq_aux_p {Γ : neg_ctx} (A : pre_ty pos) (p : pat `+A) (e : p_dom p =[val]> Γ)
+ : p_of_w_0p A (p `ᵥ⊛ e) = p
+ bystruct p :=
+ p_of_w_eq_aux_p (1) (POne) e := eq_refl ;
+ p_of_w_eq_aux_p (A ⊗ B) (PTen p1 p2) e := f_equal2 PTen
+ (eq_trans (f_equal _ (w_sub_ren _ _ _ _)) (p_of_w_eq_aux_p A p1 _))
+ (eq_trans (f_equal _ (w_sub_ren _ _ _ _)) (p_of_w_eq_aux_p B p2 _)) ;
+ p_of_w_eq_aux_p (A ⊕ B) (POr1 p) e := f_equal POr1 (p_of_w_eq_aux_p A p e) ;
+ p_of_w_eq_aux_p (A ⊕ B) (POr2 p) e := f_equal POr2 (p_of_w_eq_aux_p B p e) ;
+ p_of_w_eq_aux_p (↓ A) (PShiftP _) e := eq_refl ;
+ p_of_w_eq_aux_p (⊖ A) (PNegP p) e := f_equal PNegP (p_of_w_eq_aux_n A p e) ;
+
+with p_of_w_eq_aux_n {Γ : neg_ctx} (A : pre_ty neg) (p : pat `-A) (e : p_dom p =[val]> Γ)
+ : p_of_w_0n A (p `ᵥ⊛ e) = p
+ bystruct p :=
+ p_of_w_eq_aux_n (⊥) (PBot) e := eq_refl ;
+ p_of_w_eq_aux_n (A ⅋ B) (PPar p1 p2) e := f_equal2 PPar
+ (eq_trans (f_equal _ (w_sub_ren _ _ _ _)) (p_of_w_eq_aux_n A p1 _))
+ (eq_trans (f_equal _ (w_sub_ren _ _ _ _)) (p_of_w_eq_aux_n B p2 _)) ;
+ p_of_w_eq_aux_n (A & B) (PAnd1 p) e := f_equal PAnd1 (p_of_w_eq_aux_n A p e) ;
+ p_of_w_eq_aux_n (A & B) (PAnd2 p) e := f_equal PAnd2 (p_of_w_eq_aux_n B p e) ;
+ p_of_w_eq_aux_n (↑ A) (PShiftN _) e := eq_refl ;
+ p_of_w_eq_aux_n (¬ A) (PNegN p) e := f_equal PNegN (p_of_w_eq_aux_p A p e) .
+
+Definitionp_dom_of_w_eq_P (Γ : neg_ctx) p : pre_ty p -> Prop :=
+ match p with
+ | pos => funA => forall (p : pat `+A) (e : p_dom p =[val]> Γ),
+ rew [funp => p_dom p =[ val ]> Γ] p_of_w_eq_aux_p A p e in p_dom_of_w_0p A (p `ᵥ⊛ e) ≡ₐ e
+ | neg => funA => forall (p : pat `-A) (e : p_dom p =[val]> Γ),
+ rew [funp => p_dom p =[ val ]> Γ] p_of_w_eq_aux_n A p e in p_dom_of_w_0n A (p `ᵥ⊛ e) ≡ₐ e
+ end .
+
+Lemmap_dom_of_v_eq {Γp} A : p_dom_of_w_eq_P Γ p A .
+ induction A; intros p e; dependent elimination p; cbn - [ r_cat_l r_cat_r ].
+ - intros ? h; repeat (dependent elimination h; auto).
+ - intros ? h; repeat (dependent elimination h; auto).
+ - pose (H1 := w_sub_ren r_cat_l e `+A3 (w_of_p p)) .
+ pose (H2 := w_sub_ren r_cat_r e `+B (w_of_p p0)) .
+ pose (x1 := w_of_p p `ᵥ⊛ᵣ r_cat_l `ᵥ⊛ e).
+ pose (x2 := w_of_p p0 `ᵥ⊛ᵣ r_cat_r `ᵥ⊛ e).
+ change (w_sub_ren r_cat_l e _ _) with H1.
+ change (w_sub_ren r_cat_r e _ _) with H2.
+ change (_ `ᵥ⊛ᵣ r_cat_l `ᵥ⊛ e) with x1 in H1 |- *.
+ change (_ `ᵥ⊛ᵣ r_cat_r `ᵥ⊛ e) with x2 in H2 |- *.
+ rewrite H1, H2; clear H1 H2 x1 x2; cbn - [ r_cat_l r_cat_r ].
+ pose (H1 := p_of_w_eq_aux_p A3 p (r_cat_l ᵣ⊛ e));
+ change (p_of_w_eq_aux_p A3 _ _) with H1 in IHA1 |- *.
+ pose (H2 := p_of_w_eq_aux_p B p0 (r_cat_r ᵣ⊛ e));
+ change (p_of_w_eq_aux_p B _ _) with H2 in IHA2 |- *.
+ transitivity ([ rew [funp : pat `+A3 => p_dom p =[ val ]> Γ] H1
+ in p_dom_of_w_0p A3 (w_of_p p `ᵥ⊛ r_cat_l ᵣ⊛ e) ,
+ rew [funp : pat `+B => p_dom p =[ val ]> Γ] H2
+ in p_dom_of_w_0p B (w_of_p p0 `ᵥ⊛ r_cat_r ᵣ⊛ e) ])%asgn.
+ nowrewrite <- H1, <- H2, eq_refl_map2_distr.
+ refine (a_cat_uniq _ _ _ _ _); [ apply IHA1 | apply IHA2 ] .
+ - pose (H1 := w_sub_ren r_cat_l e `-A4 (w_of_p p1)) .
+ pose (H2 := w_sub_ren r_cat_r e `-B0 (w_of_p p2)) .
+ pose (x1 := w_of_p p1 `ᵥ⊛ᵣ r_cat_l `ᵥ⊛ e).
+ pose (x2 := w_of_p p2 `ᵥ⊛ᵣ r_cat_r `ᵥ⊛ e).
+ change (w_sub_ren r_cat_l e _ _) with H1.
+ change (w_sub_ren r_cat_r e _ _) with H2.
+ change (_ `ᵥ⊛ᵣ r_cat_l `ᵥ⊛ e) with x1 in H1 |- *.
+ change (_ `ᵥ⊛ᵣ r_cat_r `ᵥ⊛ e) with x2 in H2 |- *.
+ rewrite H1, H2; clear H1 H2 x1 x2; cbn - [ r_cat_l r_cat_r ].
+ pose (H1 := p_of_w_eq_aux_n A4 p1 (r_cat_l ᵣ⊛ e));
+ change (p_of_w_eq_aux_n A4 _ _) with H1 in IHA1 |- *.
+ pose (H2 := p_of_w_eq_aux_n B0 p2 (r_cat_r ᵣ⊛ e));
+ change (p_of_w_eq_aux_n B0 _ _) with H2 in IHA2 |- *.
+ transitivity ([ rew [funp : pat `-A4 => p_dom p =[ val ]> Γ] H1
+ in p_dom_of_w_0n A4 (w_of_p p1 `ᵥ⊛ r_cat_l ᵣ⊛ e) ,
+ rew [funp : pat `-B0 => p_dom p =[ val ]> Γ] H2
+ in p_dom_of_w_0n B0 (w_of_p p2 `ᵥ⊛ r_cat_r ᵣ⊛ e) ])%asgn.
+ nowrewrite <- H1, <- H2, eq_refl_map2_distr.
+ refine (a_cat_uniq _ _ _ _ _); [ apply IHA1 | apply IHA2 ] .
+ - match goal with | |- ?s ≡ₐ e => pose (xx := s); change (_ ≡ₐ e) with (xx ≡ₐ e) end .
+ remember xx as xx'; unfold xx in Heqxx'; clear xx.
+ rewrite (eq_trans Heqxx' (eq_sym (rew_map _ POr1 _ _))).
+ apply IHA1.
+ - match goal with | |- ?s ≡ₐ e => pose (xx := s); change (_ ≡ₐ e) with (xx ≡ₐ e) end .
+ remember xx as xx'; unfold xx in Heqxx'; clear xx.
+ rewrite (eq_trans Heqxx' (eq_sym (rew_map _ POr2 _ _))).
+ apply IHA2.
+ - match goal with | |- ?s ≡ₐ e => pose (xx := s); change (_ ≡ₐ e) with (xx ≡ₐ e) end .
+ remember xx as xx'; unfold xx in Heqxx'; clear xx.
+ rewrite (eq_trans Heqxx' (eq_sym (rew_map _ PAnd1 _ _))).
+ apply IHA1.
+ - match goal with | |- ?s ≡ₐ e => pose (xx := s); change (_ ≡ₐ e) with (xx ≡ₐ e) end .
+ remember xx as xx'; unfold xx in Heqxx'; clear xx.
+ rewrite (eq_trans Heqxx' (eq_sym (rew_map _ PAnd2 _ _))).
+ apply IHA2.
+ - intros ? v; repeat (dependent elimination v; auto).
+ - intros ? v; repeat (dependent elimination v; auto).
+ - match goal with | |- ?s ≡ₐ e => pose (xx := s); change (_ ≡ₐ e) with (xx ≡ₐ e) end .
+ remember xx as xx'; unfold xx in Heqxx'; clear xx.
+ rewrite (eq_trans Heqxx' (eq_sym (rew_map _ PNegP _ _))).
+ apply IHA.
+ - match goal with | |- ?s ≡ₐ e => pose (xx := s); change (_ ≡ₐ e) with (xx ≡ₐ e) end .
+ remember xx as xx'; unfold xx in Heqxx'; clear xx.
+ rewrite (eq_trans Heqxx' (eq_sym (rew_map _ PNegN _ _))).
+ apply IHA.
+Qed.
+
+(* coq unification drives me crazy!! *)
+Definitionupg_vp {Γ} {A : pre_ty pos} : whn Γ `+A -> val Γ `+A := funv => v.
+Definitionupg_kp {Γ} {A : pre_ty pos} : term Γ `-A -> val Γ `-A := funv => v.
+Definitionupg_vn {Γ} {A : pre_ty neg} : term Γ `+A -> val Γ `+A := funv => v.
+Definitionupg_kn {Γ} {A : pre_ty neg} : whn Γ `-A -> val Γ `-A := funv => v.
+Definitiondwn_vp {Γ} {A : pre_ty pos} : val Γ `+A -> whn Γ `+A := funv => v.
+Definitiondwn_kp {Γ} {A : pre_ty pos} : val Γ `-A -> term Γ `-A := funv => v.
+Definitiondwn_vn {Γ} {A : pre_ty neg} : val Γ `+A -> term Γ `+A := funv => v.
+Definitiondwn_kn {Γ} {A : pre_ty neg} : val Γ `-A -> whn Γ `-A := funv => v.
+
+Lemmanf_eq_split_p {Γ : neg_ctx} {A : pre_ty pos} (i : Γ ∋ `-A) (p : pat `+A) γ
+ : nf_eq (i ⋅ v_split_p (dwn_vp (p `ᵥ⊛ γ)))
+ (i ⋅ (p : o_op obs_op `-A) ⦇ γ ⦈).
+ unfold v_split_p, dwn_vp; cbn.
+ pose proof (p_dom_of_v_eq A p γ).
+ pose (H' := p_of_w_eq_aux_p A p γ); fold H' in H.
+ pose (a := p_dom_of_w_0p A (w_of_p p `ᵥ⊛ γ)); fold a in H |- *.
+ remember a as a'; clear a Heqa'.
+ revert a' H; rewrite H'; intros; noweconstructor.
+Qed.
+Lemmanf_eq_split_n {Γ : neg_ctx} {A : pre_ty neg} (i : Γ ∋ `+A) (p : pat `-A) γ
+ : nf_eq (i ⋅ v_split_n (dwn_kn (p `ᵥ⊛ γ)))
+ (i ⋅ (p : o_op obs_op `+A) ⦇ γ ⦈).
+ unfold v_split_n, dwn_kn; cbn.
+ pose proof (p_dom_of_v_eq A p γ).
+ pose (H' := p_of_w_eq_aux_n A p γ); fold H' in H.
+ pose (a := p_dom_of_w_0n A (w_of_p p `ᵥ⊛ γ)); fold a in H |- *.
+ remember a as a'; clear a Heqa'.
+ revert a' H; rewrite H'; intros; noweconstructor.
+Qed.
Finally we can return to saner pursuits.
+
+
+
OGS Instanciation
+
The notion of values and configurations we will pass to the generic OGS construction
+are our mu-mu-tilde values and states, but restricted to negative typing contexts. As
+such, while we have proven all the metatheory for arbitrary typing contexts, we still
+need a tiny bit of work to provide the laws in this special case. Once again, thanks to
+our tricky notion of subset context (Ctx/Subset.v), there is no need to cross a
+complex isomorphism (between contexts of negative types and contexts of types where
+all types are negative). We start by instanciating the substitution structures and
+their laws.
It now suffices to prove the remaining assumptions on the language machine: that
+evaluation respects substitution and that the 'bad-instanciation' relation has finite
+chains. For this we pull in some tooling for coinductive reasoning.
+
From Coinduction Require Import coinduction lattice rel tactics.
+
+Ltacrefold_eval :=
+ change (Structure.iter _ _ ?a) with (eval a);
+ change (Structure.subst (funpat : T1 => let'T1_0 := pat in?f) T1_0 ?u)
+ with (bind_delay' u f).
First we have easy lemmata on observation application.
+
esplit.
+-intros; apply p_app_eq.
+-intros ?? [[]|[]] ????; cbn.
+1,4: change (?a `ₜ⊛ ?b) with (a ᵥ⊛ b); nowrewrite (w_sub_sub _ _ _ _).
+ all: change (?a `ᵥ⊛ ?b) with (a ᵥ⊛ b) at1; nowrewrite (w_sub_sub _ _ _ _).
Next we prove the core argument of OGS soundness: evaluating a substitution is like
+evaluating the configuration, substituting the result and evaluating again.
+While tedious the proof is basically direct, going by induction on the structure of
+one-step evaluation.
+
-cbn; intros Γ Δ; unfold comp_eq, it_eq; coinduction R CIH; intros c a.
+cbn; funelim (eval_aux c); trynowdestruct (s_prf i).
++cbn; econstructor; refold_eval.
+change (?t `ₜ⊛ ?a) with (upg_kp t ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?t `ᵥ⊛ ?a) with (upg_vp t ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (RecL (?t `ₜ⊛ a_shift1 ?a)) with (upg_kp (RecL t) ᵥ⊛ a); rewrite t_sub1_sub.
+ apply CIH.
+ + change (it_eqF _ ?RX?RY _ _ _) with
+ (it_eq_map ∅ₑ RX RY T1_0
+ (eval (Cut _ (Whn (v `ᵥ⊛ a)) (a `- A i)))
+ (eval (Cut _
+ (Whn ((w_of_p (p_of_w_0p A v) `ᵥ⊛ nf_args (s_var_upg i ⋅ v_split_p v)) `ᵥ⊛ a))
+ (var (s_var_upg i) `ₜ⊛ a)))).
+ unfold nf_args, v_split_p, cut_r, fill_args.
+ nowrewrite (refold_id_aux A v).
+ + cbn; econstructor; refold_eval; apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_vp v ᵥ⊛ a); rewrite s_sub2_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_vp v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_vp v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ₜ⊛ ?a) with (upg_vn v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_kn v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ₜ⊛ ?a) with (upg_vn v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_kn v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (RecR (?t `ₜ⊛ a_shift1 ?a)) with (upg_vn (RecR t) ᵥ⊛ a); rewrite t_sub1_sub.
+ apply CIH.
+ + change (it_eqF _ ?RX?RY _ _ _) with
+ (it_eq_map ∅ₑ RX RY T1_0
+ (eval (Cut _ (a `+ A i) (Whn (k `ᵥ⊛ a))))
+ (eval (Cut _
+ (var (s_var_upg i) `ₜ⊛ a)
+ (Whn ((w_of_p (p_of_w_0n A k) `ᵥ⊛ nf_args (s_var_upg i ⋅ v_split_n k))
+ `ᵥ⊛ a))))).
+ unfold nf_args, v_split_n, cut_r, fill_args.
+ nowrewrite (refold_id_aux A k).
+ + cbn; econstructor; refold_eval; apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_kn v ᵥ⊛ a); rewrite s_sub2_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_kn v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_kn v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ₜ⊛ ?a) with (upg_kp v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
+ + cbn; econstructor; refold_eval.
+ change (?v `ᵥ⊛ ?a) with (upg_vp v ᵥ⊛ a); rewrite s_sub1_sub.
+ apply CIH.
Next we prove that evaluating a normal form is just returning this normal form. This
+goes by approximately induction on observations.
+
-cbn; intros ? [ A i [ o γ ]]; cbn; unfold p_app, nf_args, cut_r, fill_args.
+cbnin o; funelim (w_of_p o); simpl_depind; inversion eqargs.
+all: match goal with
+ | H : _ = ?A† |- _ => destruct A; dependent destruction H
+ end.
+1-2: unfold c_var in i; destruct (s_prf i).
+all: dependent destruction eqargs; cbn.
+all: apply it_eq_unstep; cbn -[w_of_p]; econstructor.
+1-2: noweconstructor.
+all: match goal with
+ | γ : dom ?p =[val_n]> _ |- _ => first [ exact (nf_eq_split_p _ p γ)
+ | exact (nf_eq_split_n _ p γ) ]
+ end.
Finally we prove the finite chain property. The proof here is quite tedious again, with
+numerous cases, but it is still by brutally walking through the structure of one-step
+evaluation and finding the needed redex.
+
-intros A; econstructor; intros [ B m ] H; dependent elimination H;
+ cbn [projT1 projT2] in i, i0.
+destruct y as [ [] | [] ].
+all: dependent elimination v; trynowdestruct (t0 (Vvar _)).
+all: clear t0.
+all: dependent elimination o; cbnin i0.
+all: match goal with
+ | u : dom _ =[val_n]> _ |- _ =>
+ cbnin i0;
+ pose (vv := u _ Ctx.top); change (u _ Ctx.top) with vv in i0;
+ remember vv as v'; clear u vv Heqv'; cbnin v'
+ | _ => idtac
+ end.
+ 7-18,25-36: apply it_eq_step in i0; nowinversion i0.
+ 1-9,19-24: match goal with
+ | t : term _ _ |- _ =>
+ dependent elimination t;
+ [ apply it_eq_step in i0; nowinversion i0
+ | apply it_eq_step in i0; nowinversion i0
+ | ]
+ | _ => idtac
+ end.
+ all:
+ match goal with
+ | t : evalₒ (Cut _ (Whn ?w) _) ≊ _ |- _ =>
+ dependent elimination w;
+ [ | apply it_eq_step in i0; nowinversion i0 ]
+ | t : evalₒ (Cut _ _ (Whn ?w)) ≊ _ |- _ =>
+ dependent elimination w;
+ [ | apply it_eq_step in i0; nowinversion i0 ]
+ | _ => idtac
+ end.
+ all:
+ apply it_eq_step in i0; cbnin i0; dependent elimination i0; cbnin r_rel;
+ apply noConfusion_inv in r_rel; unfold v_split_n,v_split_p in r_rel;
+ cbnin r_rel; unfold NoConfusionHom_f_cut,s_var_upg in r_rel; cbnin r_rel;
+ pose proof (H := f_equal pr1 r_rel); cbnin H; dependent destruction H;
+ apply DepElim.pr2_uip in r_rel;
+ pose proof (H := f_equal pr1 r_rel); cbnin H; dependent destruction H;
+ apply DepElim.pr2_uip in r_rel; dependent destruction r_rel.
+ all:
+ econstructor; intros [ t o ] H; cbnin t,o; dependent elimination H;
+ dependent elimination v;
+ [ apply it_eq_step in i0; cbnin i0; nowinversion i0
+ | apply it_eq_step in i0; cbnin i0; nowinversion i0
+ | ].
+ all:
+ match goal with
+ | H : is_var (Whn ?w) -> _ |- _ =>
+ dependent elimination w;
+ trynowdestruct (H (Vvar _))
+ end.
+ all: apply it_eq_step in i0; cbnin i0; dependent elimination i0.
+Qed.
+
+
Final theorem
+
We have finished instanciating our generic OGS construction on the System D calculus. For
+good measure we give here the concrete instanciation of the soundness theorem, stating
+that bisimilarity of the OGS strategies substitution equivalence.
+
Definitionsubst_eq (Δ : neg_ctx) {Γ} : relation (state Γ) :=
+ funuv => forall (σ : Γ =[val]> Δ), evalₒ (u ₜ⊛ σ : state_n Δ) ≈ evalₒ (v ₜ⊛ σ : state_n Δ) .
+Notation"x ≈⟦sub Δ ⟧≈ y" := (subst_eq Δ x y) (at level50).
+
+Theoremsubst_correct (Δ : neg_ctx) {Γ : neg_ctx} (xy : state Γ)
+ : x ≈⟦ogs Δ ⟧≈ y -> x ≈⟦sub Δ ⟧≈ y.
+ exact (ogs_correction _ x y).
+Qed.
Finally it does not cost much to recover the more standard CIU equivalence, which we
+derive here for the case of terms (not co-terms).
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
Built with Alectryon, running Coq+SerAPI v8.17.0+0.17.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
+
Un(i)typed Call-by-value Lambda Calculus
+
+
This file follows very closely its sibling CBVTyped.v. It is recommended to read that
+one first: the comments on this one will be quite more terse and focus on the differences.
+
+
Syntax
+
As our framework is "well-typed well-scoped", when we mean untyped, we really mean
+unityped, i.e., we will take the single element set as set of types. Although, the
+typing rules will not be as uninteresting as it would seem. Indeed, as we adopt a
+single-sided sequent-calculus style of presentation, we do give types to continuations
+(formally negated types), and we thus have not one type, but two: the type of terms
+⊕ and the type of continuations ⊖.
+
Typing contexts, terms, values, evaluation contexts and configurations work
+straightforwardly.
The reference Ctx.ctx was not found in the current
+environment.
+Bind Scope ctx_scope with t_ctx .
Note that is is the proper pure untyped lambda calculus: in constrast with our STLC example,
+the lambda is not recursive and there is no unit value nor base type, only functions.
+
Inductiveterm (Γ : t_ctx) : Type :=
+| Val : val Γ -> term Γ
+| App : term Γ -> term Γ -> term Γ
+with val (Γ : t_ctx) : Type :=
+| Var : Γ ∋ ⊕ -> val Γ
+| Lam : term (Γ ▶ₓ ⊕) -> val Γ
+.
Definitionassign1 {Γa} v : (Γ ▶ₓ a) =[val_m]> Γ := a_id ▶ₐ v .
+Definitiont_subst1 {Γa} (t : term _) v := t `ₜ⊛ @assign1 Γ a v.
+Notation"t /[ v ]" := (t_subst1 t v) (at level50, left associativity).
+
+
+
An Evaluator
+
+
Patterns and Observations
+
As before we define observations (copatterns), as there are only functions and
+continuations there is exactly one pattern at each type. Knowing this, we could
+have made this type (obs) disappear, but it is kept here for the sake of being
+more explicit.
+
Variantobs : ty -> Type :=
+| ORet : obs (⊖)
+| OApp : obs (⊕)
+.
Observation still behave as binders, returning bind a term (what we are returning) and
+applying binds a term (the argument) and a continuation.
We can now obtain a correctness theorem with respect to standard
+CIU-equivalence by embedding terms into states. Proving that CIU-equivalence
+entails our substitution equivalence is left to the reader!
+
Theoremulc_ciu_correctΔ {Γ} (xy : term Γ)
+ : c_init x ≈⟦ogs Δ⟧≈ c_init y -> x ≈⟦ciu Δ⟧≈ y .
+ intros H σ k; rewrite2 sub_init.
+ nowapply ulc_subst_correct.
+Qed.
+
+
+
Example terms
+
Following a trick by Conor McBride we provide a cool notation for writing terms
+without any DeBruijn indices, instead leveraging Coq's binders. For this we need
+a bit of infrastructure, for manipulating terms that only have term variables.
+
Equationsn_ctx : nat -> t_ctx :=
+ n_ctx O := ∅ ;
+ n_ctx (S n) := n_ctx n ▶ₓ ⊕ .
+
+Notationvar' n := (n_ctx n ∋ ⊕).
+Notationterm' n := (term (n_ctx n)).
+
+Equationsv_emb (n : nat) : var' (S n) :=
+ v_emb O := Ctx.top ;
+ v_emb (S n) := pop (v_emb n) .
+
+Equationsv_lift {n} : var' n -> var' (S n) :=
+ @v_lift (S _) Ctx.top := Ctx.top ;
+ @v_lift (S _) (pop i) := pop (v_lift i) .
Here starts the magic.
+
Equationsmk_var (m : nat) : foralln, var' (m + S n) :=
+ mk_var O := v_emb ;
+ mk_var (S m) := v_lift ∘ mk_var m .
+
+Definitionmk_lam {m : nat} (f : (foralln, term' (m + S n)) -> term' (S m)) : term' m
+ := Val (Lam (f (Val ∘ Var ∘ mk_var m))).
+Arguments mk_lam {_} & _ .
Now a bit of syntactic sugar.
+
Declare Custom Entry lambda.
+Notation"✨ e ✨" := (e : term' 0) (e custom lambda at level2).
+Notation"x" := (x _) (in custom lambda at level0, x ident).
+Notation"( x )" := x (in custom lambda at level0, x at level2).
+Notation"x y" := (App x y)
+ (in custom lambda at level1, left associativity).
+Notation"'λ' x .. y ⇒ U" :=
+ (mk_lam (funx => .. (mk_lam (funy => U)) ..))
+ (in custom lambda at level1, x binder, y binder, U at level2).
Aren't these beautiful?
+
DefinitionDelta := ✨ (λx ⇒ x x) ✨ .
+DefinitionOmega := ✨ (λx ⇒ x x) (λx ⇒ x x) ✨ .
+DefinitionUpsilon := ✨ λf ⇒ (λx ⇒ f (x x)) (λx ⇒ f (x x)) ✨.
+DefinitionTheta := ✨ (λxf ⇒ f (x x f)) (λxf ⇒ f (x x f)) ✨.
You can check the actual lambda-term they unwrap to
+by running eg: