Propositions.
- atom {Atom : Type u} (x : Atom) : Proposition Atom
- atomDual {Atom : Type u} (x : Atom) : Proposition Atom
- one {Atom : Type u} : Proposition Atom
- zero {Atom : Type u} : Proposition Atom
- top {Atom : Type u} : Proposition Atom
- bot {Atom : Type u} : Proposition Atom
- tensor
{Atom : Type u}
(a b : Proposition Atom)
: Proposition Atom
The multiplicative conjunction connective.
- parr
{Atom : Type u}
(a b : Proposition Atom)
: Proposition Atom
The multiplicative disjunction connective.
- oplus
{Atom : Type u}
(a b : Proposition Atom)
: Proposition Atom
The additive disjunction connective.
- with
{Atom : Type u}
(a b : Proposition Atom)
: Proposition Atom
The additive conjunction connective.
- bang
{Atom : Type u}
(a : Proposition Atom)
: Proposition Atom
The "of course" exponential.
- quest
{Atom : Type u}
(a : Proposition Atom)
: Proposition Atom
The "why not" exponential. This is written as ʔ, or _?, to distinguish itself from the lean syntatical hole ? syntax
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Equations
Equations
Equations
Equations
The multiplicative conjunction connective.
Equations
- Cslib.CLL.«term_⊗_» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_⊗_» 35 36 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⊗ ") (Lean.ParserDescr.cat `term 36))
Instances For
The additive disjunction connective.
Equations
- Cslib.CLL.«term_⊕_» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_⊕_» 35 36 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⊕ ") (Lean.ParserDescr.cat `term 36))
Instances For
The multiplicative disjunction connective.
Equations
- Cslib.CLL.term_⅋_ = Lean.ParserDescr.trailingNode `Cslib.CLL.term_⅋_ 30 31 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⅋ ") (Lean.ParserDescr.cat `term 31))
Instances For
The additive conjunction connective.
Equations
- Cslib.CLL.«term_&_» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_&_» 30 31 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " & ") (Lean.ParserDescr.cat `term 31))
Instances For
The "of course" exponential.
Equations
- Cslib.CLL.term!_ = Lean.ParserDescr.node `Cslib.CLL.term!_ 95 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "!") (Lean.ParserDescr.cat `term 95))
Instances For
The "why not" exponential. This is written as ʔ, or _?, to distinguish itself from the lean syntatical hole ? syntax
Equations
- Cslib.CLL.«termʔ_» = Lean.ParserDescr.node `Cslib.CLL.«termʔ_» 95 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "ʔ") (Lean.ParserDescr.cat `term 95))
Instances For
Positive propositions.
Equations
Instances For
Negative propositions.
Equations
Instances For
Whether a Proposition is positive is decidable.
Equations
Whether a Proposition is negative is decidable.
Equations
Propositional duality.
Equations
- (Cslib.CLL.Proposition.atom x_1).dual = Cslib.CLL.Proposition.atomDual x_1
- (Cslib.CLL.Proposition.atomDual x_1).dual = Cslib.CLL.Proposition.atom x_1
- Cslib.CLL.Proposition.one.dual = Cslib.CLL.Proposition.bot
- Cslib.CLL.Proposition.bot.dual = Cslib.CLL.Proposition.one
- Cslib.CLL.Proposition.zero.dual = Cslib.CLL.Proposition.top
- Cslib.CLL.Proposition.top.dual = Cslib.CLL.Proposition.zero
- (a.tensor b).dual = a.dual.parr b.dual
- (a.parr b).dual = a.dual.tensor b.dual
- (a.oplus b).dual = a.dual.with b.dual
- (a.with b).dual = a.dual.oplus b.dual
- a.bang.dual = a.dual.quest
- a.quest.dual = a.dual.bang
Instances For
Propositional duality.
Equations
- Cslib.CLL.«term_⫠» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_⫠» 1024 1024 (Lean.ParserDescr.symbol "⫠")
Instances For
No proposition is equal to its dual.
Two propositions are equal iff their respective duals are equal.
Duality is an involution.
Linear implication.
Instances For
Linear implication.
Equations
- Cslib.CLL.«term_⊸_» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_⊸_» 25 26 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⊸ ") (Lean.ParserDescr.cat `term 26))
Instances For
A sequent in CLL is a multiset of propositions.
Equations
- Cslib.CLL.Sequent Atom = Multiset (Cslib.CLL.Proposition Atom)
Instances For
Checks that all propositions in a sequent Γ are question marks.
Equations
- Γ.allQuest = Multiset.fold and true (Multiset.map (fun (x : Cslib.CLL.Proposition Atom) => match x with | a.quest => true | x => false) Γ)
Instances For
A proof in the sequent calculus for classical linear logic.
- ax {Atom : Type u} {a : Proposition Atom} : Proof {a, a.dual}
- cut {Atom : Type u} {a : Proposition Atom} {Γ Δ : Multiset (Proposition Atom)} : Proof (a ::ₘ Γ) → Proof (a.dual ::ₘ Δ) → Proof (Γ + Δ)
- one {Atom : Type u} : Proof {1}
- bot {Atom : Type u} {Γ : Sequent Atom} : Proof Γ → Proof (⊥ ::ₘ Γ)
- parr {Atom : Type u} {a b : Proposition Atom} {Γ : Multiset (Proposition Atom)} : Proof (a ::ₘ b ::ₘ Γ) → Proof (a.parr b ::ₘ Γ)
- tensor {Atom : Type u} {a : Proposition Atom} {Γ : Multiset (Proposition Atom)} {b : Proposition Atom} {Δ : Multiset (Proposition Atom)} : Proof (a ::ₘ Γ) → Proof (b ::ₘ Δ) → Proof (a.tensor b ::ₘ (Γ + Δ))
- oplus₁ {Atom : Type u} {a : Proposition Atom} {Γ : Multiset (Proposition Atom)} {b : Proposition Atom} : Proof (a ::ₘ Γ) → Proof (a.oplus b ::ₘ Γ)
- oplus₂ {Atom : Type u} {b : Proposition Atom} {Γ : Multiset (Proposition Atom)} {a : Proposition Atom} : Proof (b ::ₘ Γ) → Proof (a.oplus b ::ₘ Γ)
- with {Atom : Type u} {a : Proposition Atom} {Γ : Multiset (Proposition Atom)} {b : Proposition Atom} : Proof (a ::ₘ Γ) → Proof (b ::ₘ Γ) → Proof (a.with b ::ₘ Γ)
- top {Atom : Type u} {Γ : Multiset (Proposition Atom)} : Proof (⊤ ::ₘ Γ)
- quest {Atom : Type u} {a : Proposition Atom} {Γ : Multiset (Proposition Atom)} : Proof (a ::ₘ Γ) → Proof (a.quest ::ₘ Γ)
- weaken {Atom : Type u} {Γ : Sequent Atom} {a : Proposition Atom} : Proof Γ → Proof (a.quest ::ₘ Γ)
- contract {Atom : Type u} {a : Proposition Atom} {Γ : Multiset (Proposition Atom)} : Proof (a.quest ::ₘ a.quest ::ₘ Γ) → Proof (a.quest ::ₘ Γ)
- bang {Atom : Type u} {Γ : Sequent Atom} {a : Proposition Atom} : Γ.allQuest = true → Proof (a ::ₘ Γ) → Proof (a.bang ::ₘ Γ)
Instances For
A proof in the sequent calculus for classical linear logic.
Equations
- Cslib.CLL.«term⇓_» = Lean.ParserDescr.node `Cslib.CLL.«term⇓_» 1022 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "⇓") (Lean.ParserDescr.cat `term 90))
Instances For
Rewrites the conclusion of a proof into an equal one.
Equations
- Cslib.CLL.Proof.sequent_rw h p = h ▸ p
Instances For
A sequent is provable if there exists a proof that concludes it.
Equations
- Γ.Provable = Nonempty (Cslib.CLL.Proof Γ)
Instances For
Having a proof of Γ shows that it is provable.
Equations
- p.toProof = Classical.choice p
Instances For
The axiom, but where the order of propositions is reversed.
Equations
Instances For
Inversion of the ⅋ rule.
Equations
- h.parr_inversion = ⋯ ▸ (⋯ ▸ Cslib.CLL.Proof.ax'.tensor Cslib.CLL.Proof.ax').cut' h
Instances For
Inversion of the & rule, first component.
Equations
Instances For
Inversion of the & rule, second component.
Equations
Instances For
Logical equivalences #
Two propositions are equivalent if one implies the other and vice versa. Proof-relevant version.
Instances For
Conversion from proof-relevant to proof-irrelevant versions of propositional equivalence.
Two propositions are equivalent if one implies the other and vice versa. Proof-relevant version.
Equations
- Cslib.CLL.«term_≡⇓_» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_≡⇓_» 29 30 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ≡⇓ ") (Lean.ParserDescr.cat `term 30))
Instances For
Propositional equivalence, proof-irrelevant version (Prop).
Equations
- Cslib.CLL.«term_≡_» = Lean.ParserDescr.trailingNode `Cslib.CLL.«term_≡_» 29 30 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ≡ ") (Lean.ParserDescr.cat `term 30))
Instances For
Proof-relevant equivalence is coerciable into proof-irrelevant equivalence.
Equations
- Cslib.CLL.instCoeEquivEquiv = { coe := ⋯ }
Proof-relevant equivalence is reflexive.
Instances For
Proof-relevant equivalence is symmetric.
Equations
- Cslib.CLL.Proposition.equiv.symm a h = (h.2, h.1)
Instances For
Proof-irrelevant equivalence is reflexive.
Proof-irrelevant equivalence is symmetric.
Proof-irrelevant equivalence is transitive.
Transforms a proof-irrelevant equivalence into a proof-relevant one (this is not computable).
Instances For
The canonical equivalence relation for propositions.
Equations
- Cslib.CLL.Proposition.propositionSetoid = { r := Cslib.CLL.Proposition.Equiv, iseqv := ⋯ }
Instances For
!⊤ ≡⇓ 1
Equations
Instances For
ʔ0 ≡⇓ ⊥
Equations
Instances For
a ⊗ 0 ≡⇓ 0
Equations
Instances For
a ⅋ ⊤ ≡⇓ ⊤
Equations
Instances For
⊗ distributed over ⊕.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The proposition at the head of a proof can be substituted by an equivalent proposition.
Equations
- Cslib.CLL.Proposition.subst_eqv_head heqv p = ⋯ ▸ p.cut heqv.1
Instances For
Any proposition in a proof (regardless of its position) can be substituted by an equivalent proposition.
Equations
- Cslib.CLL.Proposition.subst_eqv heqv p = ⋯ ▸ Cslib.CLL.Proposition.subst_eqv_head heqv (⋯ ▸ p)
Instances For
Tensor is commutative.
Equations
Instances For
⊗ is associative.
Equations
- One or more equations did not get rendered due to their size.
Instances For
⊕ is idempotent.
Equations
Instances For
& is idempotent.