GithubHelp home page GithubHelp logo

leanprover-community / con-nf Goto Github PK

View Code? Open in Web Editor NEW
57.0 17.0 3.0 6.74 MB

A formal consistency proof of Quine's set theory New Foundations

Home Page: https://leanprover-community.github.io/con-nf/

Lean 87.43% Batchfile 0.01% Shell 0.01% TeX 12.17% CSS 0.09% Python 0.28%
lean4 set-theory new-foundations

con-nf's Introduction

New Foundations is consistent

Con(NF)

.github/workflows/push_main.yml

In 1937, Quine proposed a set theory called "New Foundations", and since 2010, Randall Holmes has claimed to have a proof of its consistency. In this repository, we use the interactive theorem prover Lean to verify the difficult part of his proof, thus proving that New Foundations is indeed consistent. The proof is now complete, and the theorem statements can be found in ConNF/Model/Result.lean (source, docs).

See our website for more information, the documentation of our Lean code, and the deformalisation paper that translates the Lean definitions into English.

To run our code locally, install elan, clone the repository, and run the following command in a terminal in the repository's root directory.

lake exe cache get

The code can then be viewed in an editor such as Visual Studio Code, or compiled directly from the command-line using lake build.

Objective

It is known that New Foundations is consistent if and only if a theory called Tangled Type Theory (TTT) is consistent (see theorem 1 here). We have formally constructed a model of TTT in Lean, thus proving (on paper) that New Foundations is consistent, or in short, Con(NF). We are working from various versions of the paper proof by Holmes:

but many alterations and additions have been made to make the proof compatible with Lean's type theory.

This project depends on mathlib, the community mathematical library written in Lean. This allows us to use familiar results about things like cardinals and groups without having to prove them ourselves.

Every definition and theorem in mathlib and this project have been checked by Lean's trusted kernel, which computationally verifies that the proofs we have constructed are indeed correct. However, Lean cannot check that the statements of the definitions and theorems match their intended English equivalents, so when drawing conclusions from the code in this project, translation to and from English must be done with care.

Tangled type theory

TTT is a many-sorted set theory with equality "=" and the membership relation "∈". The sorts are indexed by a limit ordinal λ, and elements of λ are called type indices. A formula "x = y" is well-formed if x and y have the same type, and a formula "x ∈ y" is well-formed if x has any type less than y.

One of the axioms of tangled type theory is extensionality, which stipulates that a set of type α is uniquely determined by its elements of any type β < α. This is strange: for example, if two sets of type α differ, they have different type β elements for every β < α. This property makes it difficult to construct models of TTT.

Strategy

Our construction of the model uses the following rough strategy.

Construction of the base type

Let λ be a limit ordinal, κ > λ be a regular ordinal, and μ > κ be a strong limit cardinal with cofinality at least κ. Sets of size less than κ are called small.

We first construct an auxiliary type at level -1, called the base type, below all types that will eventually become part of the model. Elements of this type are called atoms (although they are not atoms in the ZFU or NFU sense, for instance). There are μ atoms, partitioned into litters of size κ.

Constructing each type

At each type level α, we will produce a collection of model elements for our intended model of TTT, which we will sometimes call t-sets. We also produce a group of permutations, called allowable permutations, which act on the t-sets. The membership relation is preserved under the action of allowable permutations. Each t-set is stipulated to have a support under the action of allowable permutations; this is a small set of objects called addresses, such that whenever an allowable permutation fixes all elements of a support, it also fixes the t-set.

Each t-set at level α will be given a preferred extension of some type β < α, and we can recover from a t-set's elements which extension it prefers. The extensions of such a t-set in other lower types can be deduced from its β-extension. This allows us to satisfy TTT's extensionality axiom.

Controlling the size of each type

Each type α can only be constructed under the assumption that all types β < α are of size exactly μ (among other hypotheses). It is easy to prove that the collection of t-sets at level α has cardinality at least μ, so we need to show that there are at most μ elements of this set. We do this by showing that there are not that many fundamentally different descriptions of tangles under the action of allowable permutations. This requires the freedom of action theorem, which is a technical lemma that allows us to construct allowable permutations. The main result of this section is here.

Finishing the induction

We can then run the above process recursively to produce the types of tangles at all type levels α. This is an easy step to perform in set theory, but requires a lot of work in type theory because of the interconnectedness of the various inductive hypotheses we need. We then check that our construction indeed produces a model of TTT by checking that it satisfies a finite axiomatisation of the theory. We have chosen to convert Hailperin's finite axiomatisation of NF's comprehension scheme into a finite axiomatisation of TTT, which we present in our results file. Note, however, that this choice is arbitrary, and any other finite axiomatisation can be easily proven with the infrastructure already in place.

Dependency graph

dependency graph

con-nf's People

Contributors

amm299 avatar darby79797 avatar leanprover-community-bot avatar m3hgu5t4 avatar nomeata avatar peterlefanulumsdaine avatar randall-holmes avatar trixelian avatar vihdzp avatar yaeldillies avatar zeramorphic 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

Watchers

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

con-nf's Issues

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Allowability condition for flexible litters

The allowability condition for flexible litters is currently incorrect, since the restriction lemma cannot be proven with this definition. As it stands, the condition states:

For all extended indices $A$, either all $A$-flexible litters are in both the domain and range of a specification $\sigma$, or there are $\mu$-many not in the domain and $\mu$-many not in the range.

In the 'all' case in the restriction lemma, with this definition, we cannot prove that under a restriction, either all flexible litters remain in the specification, or $\mu$-many are removed.

Images of flexible and non-flexible things

There are a number of places where we seem to need the fact that the image of a flexible litter (or a near-litter to a flexible litter) is mapped to (something near) a flexible litter under an allowable partial permutation. The condition on non-flexible litters does not give everything that we need, however. It reads (currently):

Whenever $σ$ contains some condition $⟨⟨f_{γ,δ}^A(g), N⟩, [-1,δ,A]⟩$, then every $A$-allowable permutation $ρ$ extending $σ$ has $N \sim f_{γ,δ}^A(ρ • g)$.

We are often in a position where we either do not know if such a ρ even exists, or we have a permutation that is of too low a level, and therefore we can't use this condition. (Just noticed that in the code, this talks about $\gamma$-allowable permutations not $A$-allowable ones, this may be one of the problems?)

Non-flexible case

Flexible allowability condition

In this case, we add the condition $⟨f_{γ,δ}^A(t), \pi(f_{γ,δ}^A(t))⟩$ to a specification, where $\pi$ is $\delta$-allowable and $t \in \tau_\gamma$. We need to show that $\pi(f_{γ,δ}^A(t))$ is a non-flexible litter. However, because $\pi$ is $\delta$-allowable, we can't use the unpacked coherence condition on it.

Range is support-closed

We need to show that $\sigma$ along with the new constraint is support-closed under the action of $\gamma$-allowable permutations. These permutations are too low to use the unpacked coherence condition.

Carefully extends: all flexible litters lie in the range

Here we also need that the newly added near-litter is non-flexible.

There may be other cases where we need this, I've looked quite thoroughly at the code but I don't understand all of the partially-completed proofs.

Well-foundedness of constrains

The 'constrains' relation $\prec$ on support conditions is currently not well-founded. There a number of potential fixes for this, including de-identifying tangles and their $f$-images, and restructuring the $f$-maps themselves.

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Non-flexibility of near-litter image

Suppose that we are adding a near-litter $N$ to $\sigma$, where $N^\circ$ and $N\ \triangle\ N^\circ$ lie in $\sigma$ already. Then, the image $M$ of $N$ is already determined. How can we tell that $M$ is not (precisely) a flexible litter? This is required for a few conditions in this case. We may additionally be allowed to assume $N^\circ$ is non-flexible, but I haven't looked at this too deeply.

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Initial goals

We need to do a number of things in order to convert tangled_web.lean into a form ready for use in our project. There are also other small initial things we need to do before really starting the project proper. These include, but are not limited to:

  • κ should be made uncountable (is this already here and implicit given κ > λ?)
  • μ cofinality can be >= κ
  • (-1)-allowable should be defined in terms of near litters mapping to near litters, and then show from that that local cardinals map to local cardinals; this definition is way more composable and demonstrates the group structure easily
  • local_cards is incorrectly defined, it should have something to do with set (set μ)
  • redefine atoms := μ and then possibly use set (set atoms) for local cardinals, essentially we could just use atoms instead of μ when we're talking about atoms and just use μ when we're talking about cardinalities of types and so on
  • get rid of clan
  • prove that small symmetric difference is transitive and hence an equivalence relation, construct setoid out of near litters?
  • define codes in lean, maybe making a type distinction between codes in model and not in model
  • look at schröder bernstein implementation in mathlib to look at inspiration
  • define typed atoms
  • make a setoid for representative codes and maybe a quotient so that we can rw easier?
  • make an extensionality lemma for clan_perm on π
  • fix badge on README.md
  • stop GitHub actions from trying to apply a Jekyll theme to the gh-pages branch - it's already fully rendered

We also need to make two blueprints, one for the model and one for proofs about it - this will not be tracked as part of this issue.

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

Only deploy blueprint on pushes to main

Currently the CI workflow pushes blueprint data and static web data to gh-pages when a PR is submitted. We should instead have two workflows - one which runs on a push to main, and one which runs on a push to a PR.

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

The error could also be caused by upgrading Lean from leanprover-community/lean:3.42.0 to leanprover-community/lean:3.42.1.

You can see the errors by running:

leanproject up
leanproject build

Model blueprint

We need to construct a blueprint for the construction of the model of TTT.

Add documentation to lean files

Documentation includes references to untangled.pdf. The aim is that people inexperienced with lean should be able to understand the files, given sufficient understanding of relevant source material.

  • src/litter.lean
  • src/params.lean

Automatic upgrade has failed

Oh no! We have failed to automatically upgrade your project to the latest versions of Lean and its dependencies.

If your project currently builds, this is probably because of changes made in its dependencies:

You can see the errors by running:

leanproject up
leanproject build

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.