Documentation

Mathlib.Data.WSeq.Relation

Relations between and equivalence of weak sequences #

This file defines a relation between weak sequences as a relation between their some elements, ignoring computation time (none elements). Equivalence is then defined in the obvious way.

Main definitions #

def Stream'.WSeq.LiftRelO {α : Type u} {β : Type v} (R : αβProp) (C : WSeq αWSeq βProp) :
Option (α × WSeq α)Option (β × WSeq β)Prop

lift a relation to a relation over weak sequences

Equations
    Instances For
      theorem Stream'.WSeq.LiftRelO.imp {α : Type u} {β : Type v} {R S : αβProp} {C D : WSeq αWSeq βProp} (H1 : ∀ (a : α) (b : β), R a bS a b) (H2 : ∀ (s : WSeq α) (t : WSeq β), C s tD s t) {o : Option (α × WSeq α)} {p : Option (β × WSeq β)} :
      LiftRelO R C o pLiftRelO S D o p
      theorem Stream'.WSeq.LiftRelO.imp_right {α : Type u} {β : Type v} (R : αβProp) {C D : WSeq αWSeq βProp} (H : ∀ (s : WSeq α) (t : WSeq β), C s tD s t) {o : Option (α × WSeq α)} {p : Option (β × WSeq β)} :
      LiftRelO R C o pLiftRelO R D o p
      theorem Stream'.WSeq.LiftRelO.swap {α : Type u} {β : Type v} (R : αβProp) (C : WSeq αWSeq βProp) :
      def Stream'.WSeq.BisimO {α : Type u} (R : WSeq αWSeq αProp) :
      Option (α × WSeq α)Option (α × WSeq α)Prop

      Definition of bisimilarity for weak sequences

      Equations
        Instances For
          theorem Stream'.WSeq.BisimO.imp {α : Type u} {R S : WSeq αWSeq αProp} (H : ∀ (s t : WSeq α), R s tS s t) {o p : Option (α × WSeq α)} :
          BisimO R o pBisimO S o p
          def Stream'.WSeq.LiftRel {α : Type u} {β : Type v} (R : αβProp) (s : WSeq α) (t : WSeq β) :

          Two weak sequences are LiftRel R related if they are either both empty, or they are both nonempty and the heads are R related and the tails are LiftRel R related. (This is a coinductive definition.)

          Equations
            Instances For
              theorem Stream'.WSeq.liftRel_destruct {α : Type u} {β : Type v} {R : αβProp} {s : WSeq α} {t : WSeq β} :
              theorem Stream'.WSeq.liftRel_destruct_iff {α : Type u} {β : Type v} {R : αβProp} {s : WSeq α} {t : WSeq β} :
              theorem Stream'.WSeq.LiftRel.swap_lem {α : Type u} {β : Type v} {R : αβProp} {s1 : WSeq α} {s2 : WSeq β} (h : LiftRel R s1 s2) :
              theorem Stream'.WSeq.LiftRel.swap {α : Type u} {β : Type v} (R : αβProp) :
              theorem Stream'.WSeq.LiftRel.refl {α : Type u} (R : ααProp) (H : Reflexive R) :
              theorem Stream'.WSeq.LiftRel.symm {α : Type u} (R : ααProp) (H : Symmetric R) :
              theorem Stream'.WSeq.LiftRel.trans {α : Type u} (R : ααProp) (H : Transitive R) :
              theorem Stream'.WSeq.LiftRel.equiv {α : Type u} (R : ααProp) :
              def Stream'.WSeq.Equiv {α : Type u} :
              WSeq αWSeq αProp

              If two sequences are equivalent, then they have the same values and the same computational behavior (i.e. if one loops forever then so does the other), although they may differ in the number of thinks needed to arrive at the answer.

              Equations
                Instances For

                  If two sequences are equivalent, then they have the same values and the same computational behavior (i.e. if one loops forever then so does the other), although they may differ in the number of thinks needed to arrive at the answer.

                  Equations
                    Instances For
                      theorem Stream'.WSeq.Equiv.refl {α : Type u} (s : WSeq α) :
                      s s
                      theorem Stream'.WSeq.Equiv.symm {α : Type u} {s t : WSeq α} :
                      s tt s
                      theorem Stream'.WSeq.Equiv.trans {α : Type u} {s t u : WSeq α} :
                      s tt us u
                      theorem Stream'.WSeq.destruct_congr {α : Type u} {s t : WSeq α} :
                      s tComputation.LiftRel (BisimO fun (x1 x2 : WSeq α) => x1 x2) s.destruct t.destruct
                      theorem Stream'.WSeq.destruct_congr_iff {α : Type u} {s t : WSeq α} :
                      s t Computation.LiftRel (BisimO fun (x1 x2 : WSeq α) => x1 x2) s.destruct t.destruct
                      theorem Stream'.WSeq.liftRel_dropn_destruct {α : Type u} {β : Type v} {R : αβProp} {s : WSeq α} {t : WSeq β} (H : LiftRel R s t) (n : ) :
                      theorem Stream'.WSeq.exists_of_liftRel_left {α : Type u} {β : Type v} {R : αβProp} {s : WSeq α} {t : WSeq β} (H : LiftRel R s t) {a : α} (h : a s) :
                      (b : β), b t R a b
                      theorem Stream'.WSeq.exists_of_liftRel_right {α : Type u} {β : Type v} {R : αβProp} {s : WSeq α} {t : WSeq β} (H : LiftRel R s t) {b : β} (h : b t) :
                      (a : α), a s R a b
                      @[simp]
                      theorem Stream'.WSeq.liftRel_nil {α : Type u} {β : Type v} (R : αβProp) :
                      @[simp]
                      theorem Stream'.WSeq.liftRel_cons {α : Type u} {β : Type v} (R : αβProp) (a : α) (b : β) (s : WSeq α) (t : WSeq β) :
                      LiftRel R (cons a s) (cons b t) R a b LiftRel R s t
                      @[simp]
                      theorem Stream'.WSeq.liftRel_think_left {α : Type u} {β : Type v} (R : αβProp) (s : WSeq α) (t : WSeq β) :
                      LiftRel R s.think t LiftRel R s t
                      @[simp]
                      theorem Stream'.WSeq.liftRel_think_right {α : Type u} {β : Type v} (R : αβProp) (s : WSeq α) (t : WSeq β) :
                      LiftRel R s t.think LiftRel R s t
                      theorem Stream'.WSeq.cons_congr {α : Type u} {s t : WSeq α} (a : α) (h : s t) :
                      cons a s cons a t
                      theorem Stream'.WSeq.think_equiv {α : Type u} (s : WSeq α) :
                      theorem Stream'.WSeq.think_congr {α : Type u} {s t : WSeq α} (h : s t) :
                      theorem Stream'.WSeq.head_congr {α : Type u} {s t : WSeq α} :
                      s ts.head.Equiv t.head
                      theorem Stream'.WSeq.flatten_equiv {α : Type u} {c : Computation (WSeq α)} {s : WSeq α} (h : s c) :
                      theorem Stream'.WSeq.liftRel_flatten {α : Type u} {β : Type v} {R : αβProp} {c1 : Computation (WSeq α)} {c2 : Computation (WSeq β)} (h : Computation.LiftRel (LiftRel R) c1 c2) :
                      LiftRel R (flatten c1) (flatten c2)
                      theorem Stream'.WSeq.tail_congr {α : Type u} {s t : WSeq α} (h : s t) :
                      theorem Stream'.WSeq.dropn_congr {α : Type u} {s t : WSeq α} (h : s t) (n : ) :
                      s.drop n t.drop n
                      theorem Stream'.WSeq.get?_congr {α : Type u} {s t : WSeq α} (h : s t) (n : ) :
                      (s.get? n).Equiv (t.get? n)
                      theorem Stream'.WSeq.mem_congr {α : Type u} {s t : WSeq α} (h : s t) (a : α) :
                      a s a t
                      theorem Stream'.WSeq.Equiv.ext {α : Type u} {s t : WSeq α} (h : ∀ (n : ), (s.get? n).Equiv (t.get? n)) :
                      s t
                      theorem Stream'.WSeq.liftRel_map {α : Type u} {β : Type v} {γ : Type w} {δ : Type u_1} (R : αβProp) (S : γδProp) {s1 : WSeq α} {s2 : WSeq β} {f1 : αγ} {f2 : βδ} (h1 : LiftRel R s1 s2) (h2 : ∀ {a : α} {b : β}, R a bS (f1 a) (f2 b)) :
                      LiftRel S (map f1 s1) (map f2 s2)
                      theorem Stream'.WSeq.map_congr {α : Type u} {β : Type v} (f : αβ) {s t : WSeq α} (h : s t) :
                      map f s map f t
                      theorem Stream'.WSeq.liftRel_append {α : Type u} {β : Type v} (R : αβProp) {s1 s2 : WSeq α} {t1 t2 : WSeq β} (h1 : LiftRel R s1 t1) (h2 : LiftRel R s2 t2) :
                      LiftRel R (s1.append s2) (t1.append t2)
                      theorem Stream'.WSeq.liftRel_join.lem {α : Type u} {β : Type v} (R : αβProp) {S : WSeq (WSeq α)} {T : WSeq (WSeq β)} {U : WSeq αWSeq βProp} (ST : LiftRel (LiftRel R) S T) (HU : ∀ (s1 : WSeq α) (s2 : WSeq β), ( (s : WSeq α), (t : WSeq β), (S : WSeq (WSeq α)), (T : WSeq (WSeq β)), s1 = s.append S.join s2 = t.append T.join LiftRel R s t LiftRel (LiftRel R) S T) → U s1 s2) {a : Option (α × WSeq α)} (ma : a S.join.destruct) :
                      (b : Option (β × WSeq β)), b T.join.destruct LiftRelO R U a b
                      theorem Stream'.WSeq.liftRel_join {α : Type u} {β : Type v} (R : αβProp) {S : WSeq (WSeq α)} {T : WSeq (WSeq β)} (h : LiftRel (LiftRel R) S T) :
                      theorem Stream'.WSeq.join_congr {α : Type u} {S T : WSeq (WSeq α)} (h : LiftRel Equiv S T) :
                      theorem Stream'.WSeq.liftRel_bind {α : Type u} {β : Type v} {γ : Type w} {δ : Type u_1} (R : αβProp) (S : γδProp) {s1 : WSeq α} {s2 : WSeq β} {f1 : αWSeq γ} {f2 : βWSeq δ} (h1 : LiftRel R s1 s2) (h2 : ∀ {a : α} {b : β}, R a bLiftRel S (f1 a) (f2 b)) :
                      LiftRel S (s1.bind f1) (s2.bind f2)
                      theorem Stream'.WSeq.bind_congr {α : Type u} {β : Type v} {s1 s2 : WSeq α} {f1 f2 : αWSeq β} (h1 : s1 s2) (h2 : ∀ (a : α), f1 a f2 a) :
                      s1.bind f1 s2.bind f2
                      @[simp]
                      theorem Stream'.WSeq.join_ret {α : Type u} (s : WSeq α) :
                      (ret s).join s
                      @[simp]
                      theorem Stream'.WSeq.join_map_ret {α : Type u} (s : WSeq α) :
                      (map ret s).join s
                      @[simp]
                      theorem Stream'.WSeq.join_append {α : Type u} (S T : WSeq (WSeq α)) :
                      @[simp]
                      theorem Stream'.WSeq.bind_ret {α : Type u} {β : Type v} (f : αβ) (s : WSeq α) :
                      s.bind (ret f) map f s
                      @[simp]
                      theorem Stream'.WSeq.ret_bind {α : Type u} {β : Type v} (a : α) (f : αWSeq β) :
                      (ret a).bind f f a
                      @[simp]
                      theorem Stream'.WSeq.join_join {α : Type u} (SS : WSeq (WSeq (WSeq α))) :
                      @[simp]
                      theorem Stream'.WSeq.bind_assoc_comp {α : Type u} {β : Type v} {γ : Type w} (s : WSeq α) (f : αWSeq β) (g : βWSeq γ) :
                      (s.bind f).bind g s.bind ((fun (y : WSeq β) => y.bind g) f)
                      @[simp]
                      theorem Stream'.WSeq.bind_assoc {α : Type u} {β : Type v} {γ : Type w} (s : WSeq α) (f : αWSeq β) (g : βWSeq γ) :
                      (s.bind f).bind g s.bind fun (x : α) => (f x).bind g