Documentation

Cslib.Logics.LinearLogic.CLL.PhaseSemantics.Basic

Phase semantics for Classical Linear Logic #

This file develops the phase semantics for Classical Linear Logic (CLL), providing an algebraic interpretation of linear logic propositions in terms of phase spaces.

Phase semantics is a denotational semantics for linear logic where propositions are interpreted as subsets of a commutative monoid, and logical operations correspond to specific set-theoretic operations.

Main definitions #

Main results #

Several lemmas about facts and orthogonality useful in the proof of soundness are proven here.

TODO #

References #

class Cslib.CLL.PhaseSpace (M : Type u) extends CommMonoid M :

A phase space is a commutative monoid M equipped with a distinguished subset ⊥. This forms the algebraic foundation for interpreting linear logic propositions.

Instances

    Basic operations #

    def Cslib.CLL.PhaseSpace.imp {M : Type u_2} [PhaseSpace M] (X Y : Set M) :
    Set M

    Implication between two setsin a phase space: the set of elements m such that for all x ∈ X, we have m * x ∈ Y.

    Equations
    Instances For

      The orthogonal X⫠ of a set X: the set of elements that map X into ⊥ under multiplication.

      Equations
      Instances For

        The orthogonal X⫠ of a set X: the set of elements that map X into ⊥ under multiplication.

        Equations
        Instances For

          Properties of orthogonality #

          @[simp]
          theorem Cslib.CLL.PhaseSpace.orthogonal_def {P : Type u_1} [PhaseSpace P] (X : Set P) :
          orthogonal X = {m : P | xX, m * x bot}
          theorem Cslib.CLL.PhaseSpace.orth_antitone {P : Type u_1} [PhaseSpace P] {X Y : Set P} (hXY : X Y) :

          The orthogonal operation is antitone: if X ⊆ Y then Y⫠ ⊆ X⫠.

          The biorthogonal operation is extensive: X ⊆ X⫠⫠ for any set X.

          The triple orthogonal equals the orthogonal: X⫠⫠⫠ = X⫠.

          The biorthogonal closure operator on sets in a phase space.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            Basic theory of phase spaces #

            theorem Cslib.CLL.PhaseSpace.orth_iUnion {P : Type u_1} [PhaseSpace P] {ι : Sort u_2} (G : ιSet P) :
            orthogonal (⋃ (i : ι), G i) = ⋂ (i : ι), orthogonal (G i)

            Given a phase space (P, ⊥) and a set of subsets (Gᵢ)_{i ∈ I} of P, we have that (⋃ᵢ Gᵢ)⫠ = ⋂ᵢ Gᵢ⫠.

            theorem Cslib.CLL.PhaseSpace.iInter_biorth_eq_orth_iUnion_orth {P : Type u_1} [PhaseSpace P] {ι : Sort u_2} (G : ιSet P) :
            ⋂ (i : ι), orthogonal (orthogonal (G i)) = orthogonal (⋃ (i : ι), orthogonal (G i))

            Given a phase space (P, ⊥) and a set of subsets (Gᵢ)_{i ∈ I} of P, we have that ∩ᵢ Gᵢ⫠⫠ = (∪ᵢ Gᵢ⫠)⫠.

            Facts #

            A fact is a subset of a phase space that equals its biorthogonal closure.

            Equations
            Instances For
              structure Cslib.CLL.PhaseSpace.Fact (P : Type u_2) [PhaseSpace P] :
              Type u_2

              The type of facts in a phase space.

              Instances For
                theorem Cslib.CLL.PhaseSpace.Fact.eq {P : Type u_1} [PhaseSpace P] (G : Fact P) :
                G = orthogonal (orthogonal G)
                theorem Cslib.CLL.PhaseSpace.mem_dual {P : Type u_1} [PhaseSpace P] {p : P} {G : Fact P} :
                p orthogonal G qG, p * q bot
                theorem Cslib.CLL.PhaseSpace.of_Fact {P : Type u_1} [PhaseSpace P] {G : Fact P} {p : P} (hp : ∀ (q : P), (∀ rG, q * r bot)p * q bot) :
                p G
                @[simp]
                theorem Cslib.CLL.PhaseSpace.mem_carrier {P : Type u_1} [PhaseSpace P] (G : Fact P) :
                G.carrier = G

                Construct a fact from a set G and a proof that its biorthogonal closure is contained in G.

                Equations
                Instances For
                  @[simp]
                  theorem Cslib.CLL.PhaseSpace.Fact.mk_subset_coe {P : Type u_1} [PhaseSpace P] (G : Set P) (h : orthogonal (orthogonal G) G) :
                  (mk_subset G h) = G
                  def Cslib.CLL.PhaseSpace.Fact.mk_dual {P : Type u_1} [PhaseSpace P] (G H : Set P) (h : G = orthogonal H) :

                  Construct a fact from a set G and a proof that G equals the orthogonal of some set H.

                  Equations
                  Instances For
                    @[simp]
                    theorem Cslib.CLL.PhaseSpace.Fact.mk_dual_coe {P : Type u_1} [PhaseSpace P] (G H : Set P) (h : G = orthogonal H) :
                    (mk_dual G H h) = G
                    theorem Cslib.CLL.PhaseSpace.coe_mk {P : Type u_1} [PhaseSpace P] {X : Set P} {h : isFact X} :
                    { carrier := X, property := h } = X
                    @[simp]
                    theorem Cslib.CLL.PhaseSpace.closed {P : Type u_1} [PhaseSpace P] (F : Fact P) :
                    isFact F

                    In any phase space, {1}⫠ = ⊥.

                    The fact given by the dual of G.

                    Equations
                    Instances For
                      @[simp]
                      theorem Cslib.CLL.PhaseSpace.dualFact_coe {P : Type u_1} [PhaseSpace P] (G : Set P) :
                      (dualFact G) = {m : P | xG, m * x bot}
                      @[simp]
                      theorem Cslib.CLL.PhaseSpace.mem_one {P : Type u_1} [PhaseSpace P] {p : P} :
                      p 1 qbot, p * q bot
                      theorem Cslib.CLL.PhaseSpace.mul_mem_one {P : Type u_1} [PhaseSpace P] {p q : P} (hp : p 1) (hq : q 1) :
                      p * q 1
                      @[simp]
                      theorem Cslib.CLL.PhaseSpace.mem_top {P : Type u_1} [PhaseSpace P] {x : P} :
                      theorem Cslib.CLL.PhaseSpace.mem_zero {P : Type u_1} [PhaseSpace P] {p : P} :
                      p 0 ∀ (q : P), p * q bot
                      theorem Cslib.CLL.PhaseSpace.biorth_least_fact {P : Type u_1} [PhaseSpace P] (G : Set P) {F : Set P} :
                      isFact FG Forthogonal (orthogonal G) F

                      In a phase space, G⫠⫠ is the smallest fact containing G.

                      0 is the least fact (w.r.t. inclusion).

                      theorem Cslib.CLL.PhaseSpace.sInf_isFact {P : Type u_1} [PhaseSpace P] {S : Set (Fact P)} :
                      isFact (sInf ((fun (F : Fact P) => F) '' S))

                      Arbitrary intersections of facts are facts.

                      Intersection of the carriers of a set of facts.

                      Equations
                      Instances For
                        theorem Cslib.CLL.PhaseSpace.inter_isFact_of_isFact {P : Type u_1} [PhaseSpace P] {A B : Set P} (hA : isFact A) (hB : isFact B) :
                        isFact (A B)

                        Binary intersections of facts are facts.

                        theorem Cslib.CLL.PhaseSpace.iInter_eq_sInf_image {P : Type u_1} {α : Type u_2} (S : Set α) (f : αSet P) :
                        xS, f x = sInf (f '' S)
                        @[simp]
                        Equations
                        • One or more equations did not get rendered due to their size.
                        @[simp]
                        theorem Cslib.CLL.PhaseSpace.coe_min {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                        (GH) = G H

                        The idempotent elements within a given set X.

                        Equations
                        Instances For

                          The set I of idempotents that "belong to 1" in the phase semantics.

                          Equations
                          Instances For

                            Interpretation of the connectives #

                            Linear negation of a fact, given by taking its orthogonal.

                            Equations
                            Instances For

                              Linear negation of a fact, given by taking its orthogonal.

                              Equations
                              Instances For
                                @[simp]
                                theorem Cslib.CLL.PhaseSpace.Fact.coe_neg {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                G = orthogonal G
                                @[simp]
                                theorem Cslib.CLL.PhaseSpace.Fact.neg_neg {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                @[simp]
                                theorem Cslib.CLL.PhaseSpace.Fact.neg_inj {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                G = H G = H

                                The tensor product X ⊗ Y of two facts, defined as the dual of the orthogonal of the pointwise product.

                                Equations
                                Instances For

                                  The tensor product X ⊗ Y of two facts, defined as the dual of the orthogonal of the pointwise product.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    def Cslib.CLL.PhaseSpace.Fact.parr {P : Type u_1} [PhaseSpace P] (X Y : Fact P) :

                                    The par (multiplicative disjunction) X ⅋ Y of two facts, defined as the dual of the pointwise product of the orthogonals.

                                    Equations
                                    Instances For

                                      The par (multiplicative disjunction) X ⅋ Y of two facts, defined as the dual of the pointwise product of the orthogonals.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        @[simp]
                                        theorem Cslib.CLL.PhaseSpace.Fact.one_tensor {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                        1 G = G
                                        theorem Cslib.CLL.PhaseSpace.Fact.tensor_comm {P : Type u_1} [PhaseSpace P] {X Y : Fact P} :
                                        X Y = Y X
                                        @[simp]
                                        theorem Cslib.CLL.PhaseSpace.Fact.tensor_one {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                        G 1 = G
                                        theorem Cslib.CLL.PhaseSpace.Fact.coe_tensor_assoc {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                        ↑((G H) K) = orthogonal (orthogonal (G * (H * K)))
                                        @[simp]
                                        theorem Cslib.CLL.PhaseSpace.Fact.tensor_assoc {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                        (G H) K = G H K
                                        theorem Cslib.CLL.PhaseSpace.Fact.tensor_rotate {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                        G H K = H K G
                                        theorem Cslib.CLL.PhaseSpace.Fact.tensor_le_tensor {P : Type u_1} [PhaseSpace P] {G K H L : Fact P} (hGK : G K) (hHL : H L) :
                                        G H K L
                                        theorem Cslib.CLL.PhaseSpace.Fact.par_le_par {P : Type u_1} [PhaseSpace P] {G H K L : Fact P} (hGK : G K) (hHL : H L) :
                                        G H K L
                                        @[simp]
                                        theorem Cslib.CLL.PhaseSpace.Fact.bot_par {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                        G = G
                                        @[simp]
                                        theorem Cslib.CLL.PhaseSpace.Fact.par_bot {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                        G = G
                                        @[simp]
                                        theorem Cslib.CLL.PhaseSpace.Fact.par_assoc {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                        (G H) K = G H K
                                        theorem Cslib.CLL.PhaseSpace.Fact.par_comm {P : Type u_1} [PhaseSpace P] (G H : Fact P) :
                                        G H = H G

                                        Linear implication between facts, defined as the dual of the orthogonal of the pointwise product.

                                        Equations
                                        Instances For

                                          Linear implication between facts, defined as the dual of the orthogonal of the pointwise product.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            theorem Cslib.CLL.PhaseSpace.Fact.neg_par {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                            (G H) = G H
                                            theorem Cslib.CLL.PhaseSpace.Fact.mul_subset_tensor {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                            G * H ↑(G H)
                                            @[simp]
                                            theorem Cslib.CLL.PhaseSpace.Fact.one_linImpl {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                            1 G = G
                                            @[simp]
                                            @[simp]
                                            theorem Cslib.CLL.PhaseSpace.Fact.tensor_linImpl {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                            G H K = G (H K)
                                            theorem Cslib.CLL.PhaseSpace.Fact.linImpl_par {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                            G H K = (G H) K
                                            theorem Cslib.CLL.PhaseSpace.Fact.linImpl_par' {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                            G H K = H (G K)
                                            @[simp]
                                            theorem Cslib.CLL.PhaseSpace.Fact.linImpl_iff_implies {P : Type u_1} [PhaseSpace P] {p : P} {G H : Fact P} :
                                            p G H imp (↑G) (↑H) p

                                            The with (additive conjunction) X & Y of two facts, defined as the intersection of the two facts.

                                            Equations
                                            • X & Y = XY
                                            Instances For

                                              The with (additive conjunction) X & Y of two facts, defined as the intersection of the two facts.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For

                                                The oplus (additive disjunction) X ⊕ Y of two facts, defined as the dual of the orthogonal of the union.

                                                Equations
                                                Instances For

                                                  The oplus (additive disjunction) X ⊕ Y of two facts, defined as the dual of the orthogonal of the union.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For

                                                    The exponential !X (of course) of a fact, defined as the dual of the orthogonal of the intersection with the idempotents.

                                                    Equations
                                                    Instances For

                                                      The exponential !X (of course) of a fact, defined as the dual of the orthogonal of the intersection with the idempotents.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For

                                                        The exponential ?X (why not) of a fact, defined as the dual of the intersection of the orthogonal with the idempotents.

                                                        Equations
                                                        Instances For

                                                          The exponential ?X (why not) of a fact, defined as the dual of the intersection of the orthogonal with the idempotents.

                                                          Equations
                                                          • One or more equations did not get rendered due to their size.
                                                          Instances For

                                                            Properties of Additives #

                                                            theorem Cslib.CLL.PhaseSpace.Fact.neg_plus {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                                            (G H) = G & H
                                                            theorem Cslib.CLL.PhaseSpace.Fact.neg_with {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                                            (G & H) = G H
                                                            theorem Cslib.CLL.PhaseSpace.Fact.with_comm {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                                            G & H = H & G
                                                            @[simp]
                                                            theorem Cslib.CLL.PhaseSpace.Fact.with_assoc {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                                            (G & H) & K = G & (H & K)
                                                            @[simp]
                                                            theorem Cslib.CLL.PhaseSpace.Fact.top_with {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                                            & G = G
                                                            @[simp]
                                                            theorem Cslib.CLL.PhaseSpace.Fact.with_top {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                                            G & = G
                                                            theorem Cslib.CLL.PhaseSpace.Fact.plus_comm {P : Type u_1} [PhaseSpace P] {G H : Fact P} :
                                                            G H = H G
                                                            @[simp]
                                                            theorem Cslib.CLL.PhaseSpace.Fact.plus_assoc {P : Type u_1} [PhaseSpace P] {G H K : Fact P} :
                                                            (G H) K = G (H K)
                                                            @[simp]
                                                            theorem Cslib.CLL.PhaseSpace.Fact.zero_plus {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                                            0 G = G
                                                            @[simp]
                                                            theorem Cslib.CLL.PhaseSpace.Fact.plus_zero {P : Type u_1} [PhaseSpace P] {G : Fact P} :
                                                            G 0 = G
                                                            @[reducible, inline]

                                                            A fact G is valid if the unit 1 belongs to G.

                                                            Equations
                                                            Instances For

                                                              Interpretation of propositions #

                                                              The interpretation of a CLL proposition in a phase space, given a valuation of atoms to facts.

                                                              Equations
                                                              • One or more equations did not get rendered due to their size.
                                                              Instances For