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 #
PhaseSpace: A commutative monoid equipped with a distinguished subset ⊥PhaseSpace.imp: Linear implicationX ⊸ Ybetween sets in a phase spacePhaseSpace.orthogonal: The orthogonalX⫠of a set XPhaseSpace.isFact: A fact is a set that equals its biorthogonal closureFact: The type of facts in a phase spacePhaseSpace.FactExpr: Inductive type for representing operations on factsPhaseSpace.interpret: Interpretation of the connectives on factsPhaseSpace.interpProp: Interpretation of CLL propositions as facts in a phase space
Main results #
PhaseSpace.biorthogonalClosure: The biorthogonal operation forms a closure operatorPhaseSpace.orth_iUnion: Orthogonal of union equals intersection of orthogonalsPhaseSpace.sInf_isFactandPhaseSpace.inter_isFact_of_isFact: Facts are closed under intersectionsPhaseSpace.biorth_least_fact: The biorthogonal closure gives the smallest fact containing a set
Several lemmas about facts and orthogonality useful in the proof of soundness are proven here.
TODO #
- Soundness theorem
- Completeness theorem
References #
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 #
Implication between two setsin a phase space: the set of elements m such that for all x ∈ X, we have m * x ∈ Y.
Instances For
The orthogonal X⫠ of a set X: the set of elements that map X into ⊥ under multiplication.
Instances For
The orthogonal X⫠ of a set X: the set of elements that map X into ⊥ under multiplication.
Equations
- Cslib.CLL.PhaseSpace.«term_⫠» = Lean.ParserDescr.trailingNode `Cslib.CLL.PhaseSpace.«term_⫠» 1024 1024 (Lean.ParserDescr.symbol "⫠")
Instances For
Properties of orthogonality #
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 #
Given a phase space (P, ⊥) and a set of subsets (Gᵢ)_{i ∈ I} of P, we have that (⋃ᵢ Gᵢ)⫠ = ⋂ᵢ Gᵢ⫠.
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
Equations
- Cslib.CLL.PhaseSpace.instSetLikeFact = { coe := Cslib.CLL.PhaseSpace.Fact.carrier, coe_injective' := ⋯ }
Equations
- Cslib.CLL.PhaseSpace.instHasSubsetFact = { Subset := fun (A B : Cslib.CLL.PhaseSpace.Fact P) => ↑A ⊆ ↑B }
Equations
Construct a fact from a set G and a proof that its biorthogonal closure is contained in G.
Equations
- Cslib.CLL.PhaseSpace.Fact.mk_subset G h = { carrier := G, property := ⋯ }
Instances For
Construct a fact from a set G and a proof that G equals the orthogonal of some set H.
Equations
Instances For
In any phase space, {1}⫠ = ⊥.
The fact given by the dual of G.
Equations
Instances For
Equations
Equations
Equations
In a phase space, G⫠⫠ is the smallest fact containing G.
0 is the least fact (w.r.t. inclusion).
Arbitrary intersections of facts are facts.
Intersection of the carriers of a set of facts.
Equations
- Cslib.CLL.PhaseSpace.carriersInf S = sInf ((fun (F : Cslib.CLL.PhaseSpace.Fact P) => ↑F) '' S)
Instances For
Binary intersections of facts are facts.
Equations
- Cslib.CLL.PhaseSpace.instInfSetFact = { sInf := fun (S : Set (Cslib.CLL.PhaseSpace.Fact P)) => { carrier := Cslib.CLL.PhaseSpace.carriersInf S, property := ⋯ } }
Equations
- One or more equations did not get rendered due to their size.
The idempotent elements within a given set X.
Equations
- Cslib.CLL.PhaseSpace.idempotentsIn X = {m : M | IsIdempotentElem m ∧ m ∈ X}
Instances For
The set I of idempotents that "belong to 1" in the phase semantics.
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
- Cslib.CLL.PhaseSpace.Fact.«term_ᗮ» = Lean.ParserDescr.trailingNode `Cslib.CLL.PhaseSpace.Fact.«term_ᗮ» 1024 1024 (Lean.ParserDescr.symbol "ᗮ")
Instances For
The tensor product X ⊗ Y of two facts,
defined as the dual of the orthogonal of the pointwise product.
Equations
- X ⊗ Y = Cslib.CLL.PhaseSpace.dualFact (Cslib.CLL.PhaseSpace.orthogonal (↑X * ↑Y))
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
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
Linear implication between facts, defined as the dual of the orthogonal of the pointwise product.
Equations
- X ⊸ Y = Cslib.CLL.PhaseSpace.dualFact (↑X * Cslib.CLL.PhaseSpace.orthogonal ↑Y)
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
The with (additive conjunction) X & Y of two facts,
defined as the intersection of the two facts.
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
- X ⊕ Y = Cslib.CLL.PhaseSpace.dualFact (Cslib.CLL.PhaseSpace.orthogonal (↑X ∪ ↑Y))
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 #
A fact G is valid if the unit 1 belongs to G.
Instances For
Interpretation of propositions #
The interpretation of a CLL proposition in a phase space, given a valuation of atoms to facts.
Equations
- Cslib.CLL.PhaseSpace.interpProp v (Cslib.CLL.Proposition.atom a) = v a
- Cslib.CLL.PhaseSpace.interpProp v (Cslib.CLL.Proposition.atomDual a) = (v a)ᗮ
- Cslib.CLL.PhaseSpace.interpProp v Cslib.CLL.Proposition.one = 1
- Cslib.CLL.PhaseSpace.interpProp v Cslib.CLL.Proposition.zero = 0
- Cslib.CLL.PhaseSpace.interpProp v Cslib.CLL.Proposition.top = ⊤
- Cslib.CLL.PhaseSpace.interpProp v Cslib.CLL.Proposition.bot = ⊥
- Cslib.CLL.PhaseSpace.interpProp v (A.tensor B) = Cslib.CLL.PhaseSpace.interpProp v A ⊗ Cslib.CLL.PhaseSpace.interpProp v B
- Cslib.CLL.PhaseSpace.interpProp v (A.parr B) = Cslib.CLL.PhaseSpace.interpProp v A ⅋ Cslib.CLL.PhaseSpace.interpProp v B
- Cslib.CLL.PhaseSpace.interpProp v (A.with B) = Cslib.CLL.PhaseSpace.interpProp v A & Cslib.CLL.PhaseSpace.interpProp v B
- Cslib.CLL.PhaseSpace.interpProp v (A.oplus B) = Cslib.CLL.PhaseSpace.interpProp v A ⊕ Cslib.CLL.PhaseSpace.interpProp v B
- Cslib.CLL.PhaseSpace.interpProp v A.bang = ! Cslib.CLL.PhaseSpace.interpProp v A
- Cslib.CLL.PhaseSpace.interpProp v A.quest = ʔ Cslib.CLL.PhaseSpace.interpProp v A
Instances For
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.