GithubHelp home page GithubHelp logo

paper-proof / paperproof Goto Github PK

View Code? Open in Web Editor NEW
332.0 332.0 9.0 75.39 MB

Lean theorem proving interface which feels like pen-and-paper proofs.

License: MIT License

JavaScript 2.64% TypeScript 59.32% Lean 30.71% CSS 7.33%
lean4 mathematics proof-tree

paperproof's Introduction

Paperproof

A new proof interface for Lean 4.

Paperproof vscode

Paperproof will inspect how the hypotheses and goals were changing throughout the Lean 4 proof, and display this history - making it equivalent to how we think of a mathematical proof on paper.

Open in GitHub Codespaces

You will need to wait ~3-4 minutes for the codespace to be ready.

We created a few videos about Paperproof:

Here you can read about Paperproof in context of other proof trees: lakesare.brick.do/lean-coq-isabel-and-their-proof-trees.

And here you can read how Paperproof analyzes Lean's InfoTree to build the trees you see in the user interface: link.

In the following tables, you can see what tactics such as apply, rw, or cases look like in Paperproof; and how Paperproof renders real proofs from well-known repos.

Common tactics
Lean Paperproof

apply

theorem apply (a b: ℝ) : a = b := by
  apply le_antisymm
image

have

theorem have (a b: ℝ)
(h1: a ≤ b) (h2: b ≤ a) : True := by
  have hi := le_antisymm h1 h2
image

intro

theorem intro
: ∀ (N: ℕ), ∃ M, N + N = M := by
  intro n
image

rw

theorem rw (a b: ℕ)
(h1: a = b) : (10 * a = 666) := by
  rw [h1]
image

by_contra

theorem by_contra (m: ℕ)
: 2 ≤ m := by
  by_contra h
image

use

theorem use
: ∃ x: ℕ, x = 5 := by
  use 42
image

induction

theorem induction (n: ℕ)
: Nat.mul 0 n = 0 := by
  induction' n with k ih
image

cases

theorem casesN (n: ℕ)
: Nat.mul 0 n = 0 := by
  cases' n with m
image
theorem casesAnd (A B C: Prop)
(h: A ∧ B) : C := by
  cases' h with a b
image
theorem casesOr (A B C: Prop)
(h: A ∨ B) : C := by
  cases' h with a b
image
inductive Random where
  | hi: ℕ → String → Random
  | hello: (2 + 2 = 4) → Random
  | wow: Random
theorem casesRandom (C: Prop)
(h: Random) : C := by
  cases' h with a b c
image
Full-fledged proofs

Mathematics in Lean (Jeremy Avigad, Patrick Massot)
(mathematics_in_lean/MIL/C05_Elementary_Number_Theory/solutions/Solutions_S03_Infinitely_Many_Primes.lean:155)

Mathematics in Lean - Paperproof

Mathlib
(mathlib4/Mathlib/Algebra/Field/Power.lean:30)

Mathlib - Paperproof

Hitchhiker's Guide to Logical Verification
(Anne Baanen, Alexander Bentkamp, Jasmin Blanchette, Johannes Hölzl, Jannis Limperg)
(logical_verification_2023/blob/main/lean/LoVe/LoVe05_FunctionalProgramming_Demo.lean:316)

Hitchhiker's Guide to Logical Verification - Paperproof

Installation

  1. Install the "Paperproof" vscode extension (link).

  2. [TEMPORARY STEP] Downgrade "lean4" vscode extension to version 0.0.144

    image

    Explanation: Paperproof depends on lean4 extension in order to avoid loading your computer with excessive Lean server instances; however lean4 api regularly updates in a way that introduces breaking changes, resulting in a blank screen in Paperproof. Hopefully their api stabilizes soon and we can remove this step, but at the moment - please downgrade lean4, and turn off automatic extension updates for lean4.

  3. In your lakefile.lean, write:

    require Paperproof from git "https://github.com/Paper-Proof/paperproof.git"@"main"/"lean"
  4. Then, in your terminal, run:

    lake update Paperproof

    Note: if you're getting "error: unexpected arguments: Paperproof", it means you're on the older version of Lean, and it doesn't support per-package updates. In that case, just run lake build.

  5. In a Lean file with your theorems, write:

    import Paperproof
  6. You're done!

    Now, click on the paperproof icon (after you installed the Paperproof extension, it should appear in all .lean files), this will open a Paperproof panel within vscode.

    You can click on any theorem now (well, only tactic-based proofs, those starting with by, are supported now) - you should see your proof tree rendered.

Tutorial

If you worked with formal proofs before, you might find Paperproof most similar to proof trees/Gentzen trees. The resemblance is not spurious, we can easily mimic Semantic Tableaux and Natural Deduction trees with Paperproof. All of these interfaces show "the history of a proof" - the way hypotheses and nodes were changing throughout the proof.

Unlike Gentzen, we can make use of css and javascript - so there are many visual syntax sugars on top of what would be a formal proof tree:

  • hypotheses aren't repeated when used multiple times,
  • goals and hypotheses are visually differentiated,
  • variable scopes are shown as darkening backgrounds,
  • available hypotheses are indicated via node transparencies,
  • and so on.

Below, you will see a table with the main features of Paperproof.

Paperproof walkthrough
Lean Paperproof
Hypotheses are displayed as green nodes, goals are displayed as red nodes, tactics are displayed as transparent nodes with dashed borders.
image image
A proof should be read "towards the middle" - so, hypotheses should be read from top to bottom; and goals should be read bottom up.
image image
If you dragged tactic/goal/hypothesis nodes around, you would see arrows; however we stack these nodes on top of each other and collapse these arrows into invisible "0-length" arrows for cleaner UI.
image
Opaque nodes represent a focused goal, and currently available hypotheses.
In general - slightly darker backgrounds denote variable scopes - you can only use hypotheses that are in or outside of your scope box, you can never dive into another box. Don't overthink this however, we'll always highlight the available hypotheses as you're writing the proof, consider backgrounds a visual hint that will eventually become second nature.
image
To zoom in on a particular dark box, you can click on it.
image

Updating

To update Paperproof, you only need to rerun lake update Paperproof. This will fetch the newest version of the Paperproof Lean library from this github repo, and build it.

Vscode extensions are automatically updated, however you can check for new updates with
cmd+shift+p => "Extensions: Show Extension Updates".

Paperproof is a package that's usually only used during development, so you might want to remove it from your lakefile.lean when you're pushing to production. In order to do that, just remove the Paperproof require from lakefile.lean, and run lake update Paperproof. This will clean up lake-manifest.json and lake-packages for you.

Development

You're welcome to contribute to Paperproof, see the instructions in CONTRIBUTING.md.

paperproof's People

Contributors

antonkov avatar ddrone avatar lakesare avatar pitmonticone avatar seasawher avatar

Stargazers

 avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar

Watchers

 avatar  avatar  avatar  avatar  avatar  avatar

paperproof's Issues

app - how should focusing work

Maybe we should transparentize all tactic nodes, but leave the tactics that led to the current state opacity: 1.

Currently:

image

Suggesting:

image

lean - how to best let people update/install/remove Paperproof

Decided to go for lake update Paperproof.


Related

Ultimate instructions for fixing your Lean dependencies:

  1. cp lake-packages/mathlib/lean-toolchain lean-toolchain (lake build will keep building mathlib if mathlib depends on a lean version different from yours!)
  2. rm -rf lake-packages; rm -rf build
  3. lake exe cache clean!
  4. lake exe cache get
  5. lake build

Support other editors

Hello,

Paperproof looks really interesting :)
Do you plan to support other editors besides VS Code like Neovim & Emacs?

Thanks

Can't see the paperproof render on Linux

I am using Linux. I can see the Lean Infoview however when I click the Paperproof window, there is nothing.

I've followed the tutorial: 1. Installed extension 2. Required paperproof 3. Imported paperproof 4. Used "lake update" and made sure my lean version is up to date

pic-selected-240530-1611-35

Some notes

[Some notes that we could forget if we don't write them down]


UPDATE: done, top vs bottom rendering is added to instructions

  • It is not super obvious to the new user that the arrows from goals to goals have ↑ direction, and the arrows from hypotheses to hypotheses have ↓ direction. That you are supposed to read the proof "from the outside in" (rather than top-down, or bottom-up).
    We should probably insert a gif/[pic with arrows] in the readme at some point about this (I'm writing this down because I already think it's obvious, and of course it isn't). A gif that shows us typing in some proof with the tree appearing on the right should work as such demonstration, bc will show how hypotheses appear from the top, and goals appear from the bottom.
image

lake update failed

The output of command lake update

error: > git rev-parse --verify HEAD    # in directory ./lake-packages/Paperproof
error: stderr:
fatal: Needed a single revision
error: external command `git` exited with code 128
  1. Lean (version 4.1.0, commit a832f398b80a, Release)
  2. git version 2.42.0

"Paperproof => image" exports

There should be an easy way to export paperproof proofs into svgs to enable paperproof-rendered proof insertion into lean textbooks/blog posts.
It's possible to copypaste paperproof proofs as a screenshot, however for the uses mentioned it's usually important that text in the image is selectable.

Build error in CheckIfUserIsStillTyping.lean

Hiya,

I'm following the instructions provided by both Mathematics in Lean and Paperproof and it seems to be completely broken at the moment. After adding require Paperproof from git "https://github.com/Paper-Proof/paperproof.git"@"main"/"lean" to the lakefile.lean file, and import Paperproof to one of the .lean files, and then restarting the file as indicated by VS Code, I get the following error:

[1980/1980] Building Paperproof
Some build steps logged failures:
- Building CheckIfUserIsStillTyping
error: build failed

I have tried re-installing everything and have even wiped my entire set of elan toolchains but nothing seems to help. Can anyone here shed some light on this?

How can I suppress "parseTree XXXms" output?

I have always used the VS Code extension. I like it very much. But the frequent appearance of notifications like parseTree XXXms is a bit depressing. Is there any way to turn it off?

Re:window resizing issue, arrow deletion issue

Summary: both of these are caused by tldraw doing if (shape.type === 'frame') (certainly) and if (shape.type === 'arrow') (likely) in its source code.


Our resizing issue (that makes all shapes scale) happens because tldraw source code references if (shape.type === 'frame') in the 2.0.0-alpha.12 tldraw's Resizing.ts. And our shape's type is 'window'.
Tldraw 2.0.0-alpha.14 changes that to this.editor.isShapeOfType(shape, FrameShapeUtil) (link), but I don't think that would help (needs to be checked though).

We get the same issue with arrows most likely, that's why we get the error when we delete the node without deleting the arrows first.

I think we shouldn't focus on either of these issues, because:

  1. Tldraw changes actively. There are a lot of imports/methods to rename to check if the upgrade fixes anything. Better to do this when tldraw is more stable.
  2. More importantly - it wouldn't help much, resizing the window is not very important, and, arguable, should be disabled altogether. What we want instead is the level-based moving-nodes-around UI we discussed. The window should resize when we move the nodes within it.

Parser.lean - parse tactic combinators

Deal with backtracking tactic combinators issue

There is also a similar issue for repeat, can be checked for this proof. Again - can be fixed in the Converter, but better fixed in the Parser if possible.

import Mathlib.Tactic
import Mathlib.Data.Real.Basic

variable (a b c d : ℝ)

example : min a b = min b a := by
  apply le_antisymm
  repeat
    apply le_min
    apply min_le_right
    apply min_le_left
image

app - make zooming better

  1. This is what happens on zoomToFit(), notice the padding. We'd like the content to go fully to the borders of the screen.

    image
  2. make cmd+ have smaller steps, the jumps are huge now

  3. rezoom when the proof changes

Also: how about the proof is always zoomed to content. That might be better. Right now I have to zoom and drag and rezoom all the time as I'm writing the proof.

When the proof is small, it's in the middle of the screen. When it becomes bigger, it just always fits within the contents of the screen without any actions from us.

Actually: click on a to make it fill up the entire screen. This could be a better UI. Zooming is disabled altogether.

app - add bgs for deeper windows

image
theorem theorem_7 (p q r : Prop) : p ∧ (q ∨ r) ↔ (p ∧ q) ∨ (p ∧ r) := by
  apply Iff.intro
  intro h
  cases h.right
  exact Or.inl ⟨h.left, ‹q›⟩

  exact Or.inr ⟨h.left, ‹r›⟩
  intro h
  cases h

  rename_i hpq
  exact ⟨hpq.left, Or.inl hpq.right⟩

  rename_i hpr
  -- exact ⟨hpr.left, Or.inr hpr.right⟩
  apply And.intro
  exact hpr.left

Mathlib `have` tactic support

The Mathlib have tactic does not seem to be supported yet, is that right?

The following two code snippets are mathematically equivalent, as far as I'm concerned, and it would be nice if they produced the same display:

import Mathlib.Tactic.Positivity
import Mathlib.Tactic.Linarith
import Paperproof

example {x y : ℚ} : 2 * x * y ≤ x ^ 2 + y ^ 2 := by
  have H : 0 ≤ (x - y) ^ 2 := by
    positivity
  linarith

example {x y : ℚ} : 2 * x * y ≤ x ^ 2 + y ^ 2 := by
  have H : 0 ≤ (x - y) ^ 2 
  · positivity
  linarith

But they generate different Paperproof output -- I think the one above (for core have) is much better than the one below (for Mathlib have).
Screenshot 2024-03-06 at 11 05 11 PM
Screenshot 2024-03-06 at 11 04 49 PM
I'm not sure if there is a principled reason for the difference or if it's just that Mathlib have is not yet explicitly supported.

lean - theorem from mathlib that doesn't fully work

theorem pow_dvd_pow_of_dvd {a b : α} (h : a ∣ b) : ∀ n : ℕ, a ^ n ∣ b ^ n
  | 0 => by rw [pow_zero, pow_zero]
  | n + 1 => by
    rw [pow_succ, pow_succ]
    exact mul_dvd_mul h (pow_dvd_pow_of_dvd h n)
image

Looks like it's because of our standard issue that only by + tactics get parsed

Syntax assigned to tactic is incorrect

Screenshot 2024-01-04 at 01 42 01
For example for cases the syntax to the cases tactic is everything with all insides, but we split by \n and therefore we display only
cases h with rw[hep].
Screenshot 2024-01-04 at 01 42 47
Instead we should assign "the remaining" syntax, i.e.g only cases h with since rw[hep] is consumed by tactic.

Or maybe it should be generated by completely different algorithm. Syntax assignment heuristic works ok-ish but definitely not always correct

`calc` tactic display

The calc tactic has a slightly distracting Paperproof display which privileges the first line of the calc (printing this explicitly) over the later lines. Term-mode calc does not seem to have this issue.

Example:

import Mathlib.Tactic.Ring
import Paperproof

example {a b : ℕ} (h : b = 2 * a) : ∃ x : ℕ, b ^ 3 + 1 = 2 * x + 1 := by
  have H :=
  -- term-mode `calc`
  calc b ^ 3 + 1 = (2 * a) ^ 3 + 1 := by rw [h]
    _ = 2 * (4 * a ^ 3) + 1 := by ring
  exact ⟨_, H⟩

example {a b : ℕ} (h : b = 2 * a) : ∃ x : ℕ, b ^ 3 + 1 = 2 * x + 1 := by
  use 4 * a ^ 3
  -- tactic-mode `calc`
  calc b ^ 3 + 1 = (2 * a) ^ 3 + 1 := by rw [h]
    _ = 2 * (4 * a ^ 3) + 1 := by ring

Here's the Paperproof output, in order:
Screenshot 2024-03-06 at 11 47 11 PM
Screenshot 2024-03-06 at 11 46 13 PM
In the second one (tactic-mode calc) this line is printed:

  calc b ^ 3 + 1 = (2 * a) ^ 3 + 1 := by rw [h]

This line is only part of the full calc, and the later line(s) of the calc don't appear; the asymmetry is a bit confusing. Is there a way to suppress this line from the Paperproof output?

Recommend Projects

  • React photo React

    A declarative, efficient, and flexible JavaScript library for building user interfaces.

  • Vue.js photo Vue.js

    🖖 Vue.js is a progressive, incrementally-adoptable JavaScript framework for building UI on the web.

  • Typescript photo Typescript

    TypeScript is a superset of JavaScript that compiles to clean JavaScript output.

  • TensorFlow photo TensorFlow

    An Open Source Machine Learning Framework for Everyone

  • Django photo Django

    The Web framework for perfectionists with deadlines.

  • D3 photo D3

    Bring data to life with SVG, Canvas and HTML. 📊📈🎉

Recommend Topics

  • javascript

    JavaScript (JS) is a lightweight interpreted programming language with first-class functions.

  • web

    Some thing interesting about web. New door for the world.

  • server

    A server is a program made to process requests and deliver data to clients.

  • Machine learning

    Machine learning is a way of modeling and interpreting data that allows a piece of software to respond intelligently.

  • Game

    Some thing interesting about game, make everyone happy.

Recommend Org

  • Facebook photo Facebook

    We are working to build community through open source technology. NB: members must have two-factor auth.

  • Microsoft photo Microsoft

    Open source projects and samples from Microsoft.

  • Google photo Google

    Google ❤️ Open Source for everyone.

  • D3 photo D3

    Data-Driven Documents codes.