perga/README.org

13 KiB
Raw Blame History

Perga

perga is a basic proof assistant based on a dependently typed lambda calculus (calculus of constructions). This implementation is based on the exposition in Nederpelt and Geuvers' Type Theory and Formal Proof. Right now it is a perfectly capable higher order logic proof checker, though there is lots of room for improved ergonomics and usability, which I intend to work on. At the moment, perga is comparable to Automath in terms of power and ease of use, being slightly more powerful than Automath, and a touch less ergonomic.

Syntax

The syntax is fairly flexible and should work as you expect. Identifiers can be Unicode as long as megaparsec calls them alphanumeric. λ and Π abstractions can be written in many obvious ways that should be clear from the examples below. Additionally, arrows can be used as an abbreviation for a Π type where the parameter doesn't appear in the body as usual.

All of the following example terms correctly parse, and should look familiar if you are used to standard lambda calculus notation or Coq syntax.

  λ (α : *) . λ (β : *) . λ (x : α) . λ (y : β) . x
  fun (A B C : *) (g : → C) (f : A → B) (x : A) ⇒ g (f x)
  fun (S : *) (P Q : S -> *) (H : Π (x : S) . P x -> Q x) (HP : forall (x : S), P x) => fun (x : S) => H x (HP x)

I mostly stick to Coq syntax throughout this file and the examples, as that is what I'm most used to and is easiest to type. I will probably make the syntax more strict in the future, as this level of flexibility is really not necessary.

Definitions work similarly, having abstract syntax as shown below.

  <ident> (<ident> : <type>)* : <type>? := <term> | axiom;

(The distinction between <type> and <term> is for emphasis; they are the exact same syntactic category.) Here's a couple definitions of the const function from above showing the options with the syntax.

  const := λ (α : *) . λ (β : *) . λ (x : α) . λ (y : β) . x;
  const : forall (α β : *), α → β → α := fun (α β : *) (x : α) (y : β) ⇒ x;
  const (α β : *) (x : α) (y : β) : α := x;

Type ascriptions are optional. If included, perga will check to make sure your definition matches the ascription, and, if so, will remember the way your wrote the type when printing inferred types, which is particularly handy when using abbreviations for complex types. perga has no problem inferring the types of top-level definitions, as they are completely determined by the term, but I recommend including ascriptions most of the time, as they serve as a nice piece of documentation, help guide the implementation process, and make sure you are implementing the type you think you are.

If the RHS of a definition is axiom, then perga will assume that the identifier is an inhabitant of the type ascribed to it (as such when using axioms, a type ascription is required). This allows you to use axioms.

Line comments are -- like in Haskell, and block comments are [* *] somewhat like ML (and nest properly). There is no significant whitespace, so you are free to format code as you wish.

Usage

Running perga without any arguments drops you into a basic repl. From here, you can type in definitions which perga will typecheck. Previous definitions are accessible in future definitions. The usual readline keybindings are available, including navigating history, which is saved between sessions (in ~/.cache/perga/history). In the repl, you can enter ":q", press C-c, or press C-d to quit. Entering ":e" shows everything that has been defined along with their types. Entering ":t <ident>" prints the type of a particular identifier. Entering ":n <expr>" will fully normalize (including unfolding definitions) an expression.

Here's an example session defining Church numerals and doing some computations with them to show the capabilities of the repl.

  > nat : * := forall (A : *), (A → A) → A → A;
  > zero : nat := fun (A : *) (f : A -> A) (x : A) ⇒ x;
  > :t zero
  nat
  > suc : nat → nat := fun (n : nat) (A : *) (f : A → A) (x : A) ⇒ f (n A f x);
  > one : nat := suc zero;
  > add (n m : nat) : nat := fun (A : *) (f : A → A) (x : A) ⇒ n A f (m A f x);
  > :e
  add : nat -> nat -> nat
  nat : *
  one : nat
  suc : nat -> nat
  zero : nat
  > two : nat := suc one;
  > :n add one one
  λ (A : *) (f : A -> A) (x : A) . f (f x)
  > :n two
  λ (A : *) (f : A -> A) (x : A) . f (f x)
  > :q

You can also give perga a filename as an argument, in which case it will typecheck every definition in the file. Upon finishing, which should be nearly instantaneous, it will print out all the definitions it parsed along with their types (like you had typed ":e" in the repl) so you can verify that it worked.

Simple Example

There are many very well commented examples in the /wball/perga/src/commit/f915663f94bc5b1b40831eb14a110a4b23c15e49/examples folder. These include

I intend to extend these examples further.

Here is an example file defining Leibniz equality and proving that it is reflexive, symmetric, and transitive.

  -- file: equality.pg
  
  -- Defining Leibniz equality
  -- Note that we can leave the ascription off
  eq (A : *) (x y : A) := forall (P : A -> *), P x -> P y;

  -- Equality is reflexive, which is easy to prove
  -- Here we give an ascription so that when `perga` reports the type,
  -- it references `eq` rather than inferring the type.
  eq_refl (A : *) (x : A) : eq A x x := fun (P : A -> *) (Hx : P x) => Hx;

  -- Equality is symmetric. This one's a little harder to prove.
  eq_sym (A : *) (x y : A) (Hxy : eq A x y) : eq A y x :=
      fun (P : A -> *) (Hy : P y) =>
          Hxy (fun (z : A) => P z -> P x) (fun (Hx : P x) => Hx) Hy;

  -- Equality is transitive.
  eq_trans (A : *) (x y z : A) (Hxy : eq A x y) (Hyz : eq A y z) : eq A x z :=
      fun (P : A -> *) (Hx : P x) => Hyz P (Hxy P Hx);

Running perga equality.pg yields the following output.

  eq : ∏ (A : *) . A -> A -> *
  eq_refl : ∏ (A : *) (x : A) . eq A x x
  eq_sym : ∏ (A : *) (x y : A) . eq A x y -> eq A y x
  eq_trans : ∏ (A : *) (x y z : A) . eq A x y -> eq A y z -> eq A x z

This means our proofs were accepted. Furthermore, as a sanity check, we can see that the types correspond exactly to what we wanted to prove.

Future Goals

Substantive

TODO Let-expressions

I vaguely remember from reading Coq'art that Coq does something special with let expressions, so I'll maybe want to check that out. I tried implementing let as syntax sugar for an immediately called function, but that proved to be a massive mess with how I'm handling things. let expressions would definitely be handy for factoring out complex sub expressions.

TODO Inference

Obviously not fully decidable, but I might be able to implement some basic unification algorithm. This isn't super necessary though, I find leaving off the types of arguments to generally be a bad idea, but in some cases it can be handy, especially not at the top level.

TODO Implicits

Much, much more useful than inference, implicit arguments would be amazing. It also seems a lot more complicated, but any system for dealing with implicit arguments is far better than none. Getting rid of stuff like lines 213-215 of the example file would be amazing.

TODO Module System

A proper module system would be wonderful. To me, ML style modules with structures, signatures, and functors seems like the right way to handle algebraic structures for a relatively simple language, rather than records (or a bunch of and's like I currently have) or type classes (probably much harder, but could be nicer), but any way of managing scope, importing files, etc. is a necessity.

TODO Universes?

Not really all that necessary, especially without inductive definitions, but could be fun.

TODO Inductive Definitions

This is definitely a stretch goal. It would be cool though, and would turn this proof checker into a much more competent programming language. It's not necessary for the math, but inductive definitions let you leverage computation in proofs, which is amazing. They also make certain definitions way easier, by avoiding needing to manually stipulate elimination rules, including induction principles, and let you keep more math constructive and understandable to the computer.

Cosmetic/usage/technical

TODO Prettier pretty printing

Right now, everything defaults to one line, which can be a problem with how large the proof terms get. Probably want to use prettyprinter to be able to nicely handle indentation and line breaks.

TODO Smarter normalization/beta-equivalence checking

I had what I thought was a smarter way to check for β-equivalence than just fully normalizing both terms and checking if they are the same, but it turned out to be a little wrong, which isn't too surprising since I just made it up. It's probably salvageable, but I'd also like to look into other forms of normalization and checking for β-equivalence.

TODO Better repl

The repl is decent, but implementing something like this would be awesome. I'd also at least like to add a new command ":l <filename>" to load the definitions from a file.

TODO Improve error messages

Error messages are decent, but a little buggy. Syntax error messages are pretty ok, but could have better labeling. The type check error messages are decent, but could do with better location information. Right now, the location defaults to the end of the current definition, which is often good enough, but more detail can't hurt.

TODO Better testing

I would like to avoid regressions as I keep working on this, and a suite of unit tests would make me feel much more comfortable. I made unit tests, then added a ton of stuff. Most of the unit tests are kind of pointless now. For now, I think running the code on the example file is pretty sufficient.

TODO Alternate syntax

I've had a bunch of ideas for a more mathematician-friendly syntax bouncing around my head for a while. Implementing one of them would be awesome, but probably quite tricky.

Something like

  Theorem basic (S : *) (P : S → *) :
      (∀ (x : S), P x → Q x) → (∀ (x : S), P x) → ∀ (x : S), Q x.
  Proof
          1. Suppose ∀ (x : S), P x → Q x
          2. Suppose ∀ (x : S), P x
          3. Let x : S
          4. P x by [2 x]
          5. Q x by [1 x 4]
  Qed

I think could be reliably translated into

  basic (S : *) (P : S → *) : (Π (x : S), P x → Q x) → (Π (x : S), P x) → Π (x : S), Q x :=
        fun (a1 : Π (x : S), P x → Q x) ⇒
            fun (a2 : Π (x : S), P x) ⇒
                fun (x : S) ⇒
                    a1 x (a2 x);

and is more intuitively understandable to a mathematician not familiar with type theory, while the latter would be utter nonsense.

I'm imagining the parser could be chosen based on the file extension or something. Some way to mix the syntaxes could be nice too.

TODO treesitter parser and/or emacs mode

Really not necessary, especially while the syntax is in a bit of flux, but would eventually be nice. The syntax is simple enough that a treesitter grammar shouldn't be too hard to write. An emacs mode would especially be nice if I ever get end up implementing an alternate syntax, to better handle indentation, automatically adjust line numbers, etc.

TODO TUI

This is definitely a stretch goal, and I'm not sure how good of an idea it would be, but I'm imagining a TUI split into two panels. On the left you can see the term you are building with holes in it. On the right you have the focused hole's type as well as the types of everything in scope (like Coq and Lean show while you're in the middle of a proof). Then you can interact with the system by entering commands (e.g. intros, apply, etc.) which changes the proof term on the left. You'd also just be able to type in the left window as well, and edit the proof term directly. This way you'd get the benefits of working with tactics, making it way faster to construct proof terms, and the benefits of working with proof terms directly, namely transparency and simplicity. I'll probably want to look into brick if I want to make this happen.