jwiegley / coq-cds4ltl

A formalization of finite, constructive log analysis using linear temporal logic

Geek Repo:Geek Repo

Github PK Tool:Github PK Tool

Linear Temporal Logic in Coq

This project is an implementation in Coq of the axiomatic system for linear temporal logic described in: A Calculational Deductive System for Linear Temporal Logic, ACM Computing Surveys, Vol. 53, No. 3, June 2020.

With many thanks to Professor Stan Warford and Dr. Scott Staley for answering my questions and offering their support and encouragement.

Overview of Contents

The following is a brief overview of what is provided by this repository. Note that for axiomatic foundations, both a Module Type is given that offers the axioms, and a Module based on that Module Type which develops all of the theorems that follow from those axiom.

MinBool

MinBool.v defines Boolean logic, based on ¬ (not), ∨ (or) and three axioms, using the foundation discovered by Edward Huntington. Note that this is a classical development, and as such requires assuming the law of the excluded middle for some of the basic proofs to go through. Future work will include reworking this foundation on a constructive basis.

Bool

Bool.v extends MinBool.v to include the ∧ (and) operation and related theorems.

MinLTL

MinLTL.v defines the axioms and theorems of Linear Temporal Logic in terms of the ◯ (next) and U (until) modalities.

LTL

LTL.v extends MinLTL.v to include the ◇ (eventually), □ (always) and W (wait) derived modalities and their theorems. There is also another module type defining R (release) and M (strong release), although this is relatively undeveloped at the moment.

Model

The system of axioms given for LTL above is proved to be sound using a semantic model of infinite streams, where the propositions of the logic are taken to be propositional predicates on streams – that is, mathematical sets of streams – where a propoosition is “true” when its set is inhabited, meaning the predicate holds for all inhabitants of that set at offset zero. This is expressed using the a satisfaction operator:

[σ, j] ⊨ P

which states that P holds at position j in stream σ. Thus, modalities such as ◯ may be easily defined semantically as:

Definition next := λ p σ, [σ, 1] ⊨ p.

Syntax and Step

The semantic Model above denotes propositions directly into a propositional logic of sets, which although sound is neither decidable nor computable. To make practical use of LTL, we require a more reified development. This is given by the two modules Syntax.v and Step.v. Syntax defines a deep syntactic embedding of LTL formula using Positive Normal Form (PNF), although with a denotation of such constructions into propositional logic similar to Model.v, and a proof that the axioms are sound over this denotation. Step.v allows for a “single step” of reduction given a formula and and a single element of the input stream, such that it results in either a definitive answer or the next formula to use on future inputs.

Work list

Complete the stream semantic model of LinearTemporalLogic

This may be found in Model.v

Typeset new proof (57): □ ¬p ∧ p U q ⇒ q

Write a short paper on the Coq formulation of CDS4LTL

Read the ACL2 paper that Scott provided

  • Note taken on [2021-02-11 Thu 08:57]
    Scott to provide a copy of his ACL2 paper to show John a model for what he might write up about coq the we could include in the book.

Starting a new working document

… to cover new work on CDS4LTL/Coq proof discovery; new, reduced axiomatic foundations for CDS4LTL; and new interesting proofs in LTL. Prof. Warford has agreed to set up a document template and host it on his GitHub site.

“Freeze” the current Vega-paper

… and Staley “additional theorems” paper as is; And begin putting all new work and results into the above new working document.

Build a list of publications and venues our work could be sent to

Look into SPOT as time permits

Build up theorems for Release (R) and Strong Release (M)

Explore bi-directional temporal flows

Our semantic model currently is based on infinite streams, with a satisfiability proposition with j being of type nat:

(σ, j) ⊨ p

If we use a pair or tuple of infinite streams, where one represents time looking back, and the other represents time looking forward, then we could use the same satisfiability condition but with j being of type integer, such that:

type BiStream = (Stream, Stream)

((σ₁, σ₂), n) ⊨ p <-> if n < 0 then (σ₁, -n) ⊨ p else (σ₂, n) ⊨ p

Review the Mathematical Philosophy courses at Pepperdine

Write a program to produce an ideal ordering of the LTL theorems

This would first build a dependency graph, and then for items at the same level, order them based on term complexity.

Prove the new axiom relating individual terms to until

p ∧ q U ¬p ⇒ (q ∧ p) U (q ∧ p ∧ ¬◯ p)

Typeset the proof of (55) in LaTeX

(55) □ (p ⇒ (◯ p ∧ q) ∨ r) ⟹ p ⇒ □ q ∨ q U r

  Proof:

      p ⇒ □ q ∨ q U r
    =   ⟨¬¬p ≡ p⟩
      ¬¬(p ⇒ □ q ∨ q U r)
    =   ⟨(54) Definition of □⟩
      ¬¬(p ⇒ ¬◇ ¬q ∨ q U r)
    =   ⟨(38) Definition of ◇⟩
      ¬¬(p ⇒ ¬(⊤ U ¬q) ∨ q U r)
    =   ⟨p ⇒ q ≡ ¬p ∨ q, ¬(p ∨ q) ≡ ¬p ∧ ¬q⟩
      ¬(¬¬p ∧ ¬¬(⊤ U ¬q) ∧ ¬(q U r))
    =   ⟨¬¬p ≡ p⟩
      ¬(p ∧ ⊤ U ¬q ∧ ¬(q U r))
    =   ⟨(170) Axiom, Distributivity of ¬over W⟩
      ¬(p ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨p ≡ ⊤ ∧ p⟩
      ¬((⊤ ∧ p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨(90) □ excluded middle⟩
      ¬(((□ p ∨ ◇ ¬p) ∧ p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨(p ∨ q) ∧ r ≡ (p ∧ r) ∨ (p ∧ q)⟩
      ¬((□ p ∧ p ∨ ◇ ¬p ∧ p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨(p ∨ q) ∧ r ≡ (p ∧ r) ∨ (p ∧ q)⟩
      ¬((□ p ∧ p) ∧ ¬r U (¬q ∧ ¬r) ∨ (◇ ¬p ∧ p) ∧ ¬r U (¬q ∧ ¬r))
    ⇒   ⟨p ∧ q ⇒ p⟩
      ¬(□ p ∧ ¬r U (¬q ∧ ¬r) ∨ (◇ ¬p ∧ p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨p ∧ q ≡ q ∧ p⟩
      ¬(□ p ∧ ¬r U (¬q ∧ ¬r) ∨ (p ∧ ◇ ¬p) ∧ ¬r U (¬q ∧ ¬r))
    ⇒   ⟨(83) Distributivity of ∧ over U⟩
      ¬((p ∧ ¬r) U (p ∧ ¬q ∧ ¬r) ∨ (p ∧ ◇ ¬p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨(42) Eventuality⟩
      ¬(◇ (p ∧ ¬q ∧ ¬r) ∨ (p ∧ ◇ ¬p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨p ≡ p ∨ q, p ⇒ q ≡ ¬p ∨ q⟩
      ¬(◇ (p ∧ (q ⇒ ¬◯ p) ∧ ¬r) ∨ (p ∧ ◇ ¬p) ∧ ¬r U (¬q ∧ ¬r))
    =   ⟨p ∨ q ≡ q ∨ p⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ (p ∧ ◇ ¬p) ∧ ¬r U (¬q ∧ ¬r))
    ⇒   ⟨(75) p ∧ ◇ ¬p ⇒ ◇ (p ∧ ◯ ¬p)⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ p U (p ∧ ¬◯ p) ∧ ¬r U (¬q ∧ ¬r))
    ⇒   ⟨(NEW) Axiom, Distributivity of ∧ over U⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ (p ∧ ¬r) U ((p ∧ ¬◯ p) ∧ ¬r ∨ p ∧ ¬q ∧ ¬r ∨ (p ∧ ¬◯ p) ∧ ¬q ∧ ¬r))
    ⇒   ⟨(42) Eventuality⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ ◇ ((p ∧ ¬◯ p) ∧ ¬r ∨ p ∧ ¬q ∧ ¬r ∨ (p ∧ ¬◯ p) ∧ ¬q ∧ ¬r))
    ⇒   ⟨p ∧ q ⇒ p⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ ◇ ((p ∧ ¬◯ p) ∧ ¬r ∨ p ∧ ¬q ∧ ¬r ∨ p ∧ ¬q ∧ ¬r))
    =   ⟨p ∨ p ≡ p⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ ◇ ((p ∧ ¬◯ p) ∧ ¬r ∨ p ∧ ¬q ∧ ¬r))
    =   ⟨p ∧ (q ∧ r) ≡ (p ∧ q) ∧ r⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ ◇ (p ∧ ¬◯ p ∧ ¬r ∨ p ∧ ¬q ∧ ¬r))
    =   ⟨(p ∧ q) ∨ (p ∧ r) ≡ p ∧ (q ∨ r)⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ ◇ (p ∧ (¬◯ p ∧ ¬r ∨ ¬q ∧ ¬r)))
    =   ⟨(p ∧ r) ∨ (q ∧ r) ≡ (p ∨ q) ∧ r⟩
      ¬(◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r) ∨ ◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r))
    =   ⟨p ∨ p ≡ p⟩
      ¬◇ (p ∧ (◯ p ⇒ ¬q) ∧ ¬r)
    =   ⟨p ∧ q ≡ ¬(¬p ∨ ¬q)⟩
      ¬◇ ¬(p ⇒ ¬¬((◯ p ⇒ ¬q) ⇒ ¬¬r))
    =   ⟨¬¬p ≡ p⟩
      ¬◇ ¬(p ⇒ (◯ p ⇒ ¬q) ⇒ r)
    =   ⟨¬(p ∨ q) ≡ ¬p ∧ ¬q⟩
      ¬◇ ¬(p ⇒ ¬¬◯ p ∧ ¬¬q ∨ r)
    =   ⟨¬¬p ≡ p⟩
      ¬◇ ¬(p ⇒ ◯ p ∧ q ∨ r)
    =   ⟨(54) Definition of □⟩
      □ (p ⇒ ◯ p ∧ q ∨ r)    ∎

Typeset the proof of (56) in LaTeX

  • Note taken on [2021-02-11 Thu 08:57]
    Scott to provide a template LaTeX file for John to typeset his new (56) Theorem proof.
(56) □ (p ⇒ ◯ (p ∨ q)) ⇒ p ⇒ □ p ∨ p U q

 Proof:

     true
   =   ⟨(55) Axiom, U Induction, with p := p, q := ◯ p, r := ◯ q⟩
     □ (p ⇒ (◯ p ∧ ◯ p) ∨ ◯ q) ⇒ (p ⇒ □ ◯ p ∨ ◯ p U ◯ q)
   =   ⟨idempotency of ∧⟩
     □ (p ⇒ ◯ p ∨ ◯ q) ⇒ (p ⇒ □ ◯ p ∨ ◯ p U ◯ q)
   =   ⟨(4) Distributivity of ◯ over ∨⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ □ ◯ p ∨ ◯ p U ◯ q)
   =   ⟨⊤ ∧ p ≡ p⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ ⊤ ∧ (p ⇒ □ ◯ p ∨ ◯ p U ◯ q)
   =   ⟨p ⇒ p ≡ ⊤⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ p) ∧ (p ⇒ □ ◯ p ∨ (◯ p U ◯ q))
   =   ⟨p ⇒ q ≡ ¬p ∨ q, ∨ distributes over ∧⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ p ∧ (□ ◯ p ∨ ◯ p U ◯ q))
   =   ⟨∧ distributes over ∨⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ (p ∧ □ ◯ p) ∨ (p ∧ ◯ p U ◯ q))
   =>  ⟨p ⇒ p ∨ q, with p := p ∧ b, q := q⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ ((p ∧ □ ◯ p) ∨ q) ∨ (p ∧ ◯ p U ◯ q))
   =   ⟨commutativity of ∨⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ (q ∨ (p ∧ □ ◯ p)) ∨ (p ∧ ◯ p U ◯ q))
   =   ⟨(9) Axiom, Distributivity of ◯ over U⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ (q ∨ (p ∧ □ ◯ p)) ∨ (p ∧ ◯ (p U q)))
   =   ⟨(10) Axiom, Expansion of U⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ (p ∧ □ ◯ p) ∨ (p U q))
   =   ⟨(73) Exchange of ◯ and □⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ (p ∧ ◯ □ p) ∨ (p U q))
   =   ⟨(66) Expansion of □⟩
     □ (p ⇒ ◯ (p ∨ q)) ⇒ (p ⇒ (□ p ∨ (p U q)))    ∎

Typeset the proof of (58) in LaTeX

(58) □ (◯ p ⇒ p) ⇒ (◇ p ⇒ p)

 Proof:

     true
   =   ⟨(56) Axiom, U Induction, with p := ¬p, q := ⊥⟩
     □ (¬p ⇒ ◯ (¬p ∨ ⊥)) ⇒ (¬p ⇒ □ ¬p ∨ (¬p U ⊥))
   =   ⟨(11) Axiom, Right zero of U⟩
     □ (¬p ⇒ ◯ (¬p ∨ ⊥)) ⇒ (¬p ⇒ □ ¬p ∨ ⊥)
   =   ⟨p ∨ ⊥ ≡ p⟩
     □ (¬p ⇒ ◯ ¬p) ⇒ (¬p ⇒ □ ¬p)
   =   ⟨(3) Linearity⟩
     □ (¬p ⇒ ¬(◯ ¬¬p)) ⇒ (¬p ⇒ □ ¬p)
   =   ⟨¬¬p ≡ p, ¬p ⇒ q ≡ p ∨ q⟩
     □ (p ∨ ¬ ◯ p) ⇒ p ∨ □ ¬p
   =   ⟨commutativity of ∨, p ⇒ q ≡ ¬p ∨ q⟩
     □ (◯ p ⇒ p) ⇒ □ ¬p ∨ p
   =   ⟨(54) Definition of □⟩
     □ (◯ p ⇒ p) ⇒ (¬◇ ¬¬p ∨ p)
   =>  ⟨¬¬p ≡ p, p ⇒ q ≡ ¬p ∨ q⟩
     □ (◯ p ⇒ p) ⇒ (◇ p ⇒ p)    ∎

Prove monotonicity properties of the modal operators

Find an appropriate statement of (82)

Is Proper (impl ==> impl) eventually a valid statement?

Find a more minimal set of until axioms

Take a fresh-eyes look at the Dummett formula

Find a better axiomatic basis for proving (82)

Rebuild the syntactic, Mealy-machine based evaluation model

Typeset new proof (238)

About

A formalization of finite, constructive log analysis using linear temporal logic


Languages

Language:Coq 98.9%Language:Nix 0.6%Language:Perl 0.3%Language:Makefile 0.2%