├── .docker └── gitpod │ └── Dockerfile ├── .github ├── pull_request_template.md └── workflows │ ├── update_versions.yml │ └── upgrade.yml ├── .gitignore ├── .gitpod.yml ├── .vscode └── settings.json ├── LICENSE ├── README.md ├── Tutorials.lean ├── Tutorials ├── Exercises │ ├── 00FirstProofs.lean │ ├── 01EqualityRewriting.lean │ ├── 02IffIfAnd.lean │ ├── 03ForallOr.lean │ ├── 04Exists.lean │ ├── 05SequenceLimits.lean │ ├── 06SubSequences.lean │ ├── 07FirstNegations.lean │ ├── 07bisAbstractNegations.lean │ ├── 08LimitsNegation.lean │ └── 09LimitsFinal.lean ├── Solutions │ ├── 00FirstProofs.lean │ ├── 01EqualityRewriting.lean │ ├── 02IffIfAnd.lean │ ├── 03ForallOr.lean │ ├── 04Exists.lean │ ├── 05SequenceLimits.lean │ ├── 06SubSequences.lean │ ├── 07FirstNegations.lean │ ├── 07bisAbstractNegations.lean │ ├── 08LimitsNegation.lean │ └── 09LimitsFinal.lean └── TutoLib.lean ├── lake-manifest.json ├── lakefile.lean ├── lean-toolchain └── mk_exercises.py /.docker/gitpod/Dockerfile: -------------------------------------------------------------------------------- 1 | # This is the Dockerfile for `leanprovercommunity/mathlib:gitpod`. 2 | 3 | # gitpod doesn't support multiple FROM statements, (or rather, you can't copy from one to another) 4 | # so we just install everything in one go 5 | FROM ubuntu:jammy 6 | 7 | USER root 8 | 9 | RUN apt-get update && apt-get install sudo git curl git bash-completion -y && apt-get clean 10 | 11 | RUN useradd -l -u 33333 -G sudo -md /home/gitpod -s /bin/bash -p gitpod gitpod \ 12 | # passwordless sudo for users in the 'sudo' group 13 | && sed -i.bkp -e 's/%sudo\s\+ALL=(ALL\(:ALL\)\?)\s\+ALL/%sudo ALL=NOPASSWD:ALL/g' /etc/sudoers 14 | USER gitpod 15 | WORKDIR /home/gitpod 16 | 17 | SHELL ["/bin/bash", "-c"] 18 | 19 | # gitpod bash prompt 20 | RUN { echo && echo "PS1='\[\033[01;32m\]\u\[\033[00m\] \[\033[01;34m\]\w\[\033[00m\]\$(__git_ps1 \" (%s)\") $ '" ; } >> .bashrc 21 | 22 | # install elan 23 | RUN curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none 24 | 25 | RUN . ~/.profile && elan toolchain install leanprover/lean4:v4.1.0-rc1 26 | 27 | ENV PATH="/home/gitpod/.local/bin:/home/gitpod/.elan/bin:${PATH}" 28 | 29 | # fix the infoview when the container is used on gitpod: 30 | ENV VSCODE_API_VERSION="1.50.0" 31 | 32 | # ssh to github once to bypass the unknown fingerprint warning 33 | RUN ssh -o StrictHostKeyChecking=no github.com || true 34 | 35 | # run sudo once to suppress usage info 36 | RUN sudo echo finished 37 | -------------------------------------------------------------------------------- /.github/pull_request_template.md: -------------------------------------------------------------------------------- 1 | Please make sure you modified the files in `src/solutions` and used the `mk_exercises.py` script to recreate the files in `src/exercises`. Otherwise your modifications will soon be automatically overwritten. 2 | -------------------------------------------------------------------------------- /.github/workflows/update_versions.yml: -------------------------------------------------------------------------------- 1 | on: 2 | push: 3 | branches: 4 | - master 5 | 6 | jobs: 7 | update_lean_xyz_branch: 8 | runs-on: ubuntu-latest 9 | name: Update lean-x.y.z branch 10 | steps: 11 | 12 | - name: checkout project 13 | uses: actions/checkout@v2 14 | with: 15 | fetch-depth: 0 16 | 17 | - name: update branch 18 | uses: leanprover-contrib/update-versions-action@master -------------------------------------------------------------------------------- /.github/workflows/upgrade.yml: -------------------------------------------------------------------------------- 1 | on: 2 | schedule: 3 | - cron: '0 0 * * 1' 4 | 5 | jobs: 6 | upgrade_lean: 7 | runs-on: ubuntu-latest 8 | name: Bump Lean and dependency versions 9 | steps: 10 | - name: checkout project 11 | uses: actions/checkout@v2 12 | - name: upgrade Lean and dependencies 13 | uses: leanprover-contrib/lean-upgrade-action@master 14 | with: 15 | repo: ${{ github.repository }} 16 | access-token: ${{ secrets.GITHUB_TOKEN }} -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | lakefile.olean 2 | /build 3 | /lake-packages/* 4 | /.cache 5 | .DS_Store 6 | 7 | -------------------------------------------------------------------------------- /.gitpod.yml: -------------------------------------------------------------------------------- 1 | image: 2 | file: .docker/gitpod/Dockerfile 3 | 4 | vscode: 5 | extensions: 6 | - leanprover.lean4 7 | 8 | tasks: 9 | - init: lake exe cache get 10 | - command: sudo apt-get --assume-yes install gcc && lake exe cache get 11 | -------------------------------------------------------------------------------- /.vscode/settings.json: -------------------------------------------------------------------------------- 1 | { 2 | "editor.minimap.enabled": false, 3 | "editor.acceptSuggestionOnEnter": "off", 4 | "editor.insertSpaces": true, 5 | "editor.tabSize": 2, 6 | "files.encoding": "utf8", 7 | "files.eol": "\n", 8 | "files.insertFinalNewline": true, 9 | "files.trimFinalNewlines": true, 10 | "files.trimTrailingWhitespace": true, 11 | "editor.unicodeHighlight.ambiguousCharacters": false, 12 | "editor.unicodeHighlight.nonBasicASCII": false, 13 | "files.exclude": { 14 | "**/.git": true, 15 | "**/.svn": true, 16 | "**/.hg": true, 17 | "**/CVS": true, 18 | "**/.DS_Store": true, 19 | "**/Thumbs.db": true, 20 | "**/*.olean": true 21 | "**/.gitpod.yml": true, 22 | "**/.vscode": true, 23 | ".docker": true, 24 | "build": true, 25 | "lake-packages": true, 26 | ".gitignore": true, 27 | ".github": true, 28 | "mk_exercises.py": true, 29 | "lake-manifest.json": true, 30 | "lakefile.lean": true, 31 | "lean-toolchain": true, 32 | "Tutorials.lean": true, 33 | "README.md":true, 34 | }, 35 | "lean4.infoview.showTooltipOnHover": false, 36 | "lean4.infoview.emphasizeFirstGoal": true, 37 | "lean4.infoview.showExpectedType": false 38 | } 39 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # lean 4 tutorials 2 | 3 | [![Gitpod Ready-to-Code](https://img.shields.io/badge/Gitpod-ready--to--code-blue?logo=gitpod)](https://gitpod.io/#https://github.com/leanprover-community/tutorials4) 4 | 5 | The goal of this project is to quickly teach you how to use Lean 4 for 6 | mathematics using a very hands-on approach. It can be used alongside 7 | [Theorem proving in Lean 4](https://leanprover.github.io/theorem_proving_in_lean4/) 8 | or independently. 9 | 10 | Currently, these tutorials do not cover creating your own theories, only 11 | proving things in elementary real analysis. All exercises are adapted 12 | from a first year undergraduate course by Patrick Massot in Orsay, 13 | adding only explanations about compressing proofs using slightly advanced 14 | tactics like `rintros` and `rcases`. 15 | 16 | What you do need first is to install Lean 4 and get this project. [todo: add instructions] 17 | 18 | Then, in the `Tutorials4/Tutorials` folder, create a copy of the exercises folder for you work. 19 | This way it won't be overwritten if you update the project to get new exercises. 20 | You can then open the tutorials folder in VS code. 21 | For instance you can type: 22 | ``` 23 | cp -r tutorials4/Tutorials/Exercises tutorials4/Tutorials/MyExercises 24 | code tutorials4 25 | ``` 26 | VSCode has a file explorer that you can open by clicking the top icon in 27 | the icon column on the left. In this explorer, you can navigate to 28 | `Tutorials/MyExercises` and click on `00FirstProofs.lean`. 29 | This file does not contain any exercise, it is meant as an 30 | overview of the basics. You can skip it if you are really eager to start 31 | coding, but this is not recommended. You don't need to understand 32 | everything while reading this file, only try to get a feel for what it's 33 | looking like, and maybe start picking up some key words. 34 | 35 | There are solutions for all the exercises in `Tutorials/Solutions`. If you 36 | need help about any specific exercise, you can come on 37 | [Zulip](https://leanprover.zulipchat.com) in the "new members" stream 38 | and look for a thread called "tutorials4 NNNN" where NNNN is the exercise 39 | number. If no such thread exists, you can create one! 40 | -------------------------------------------------------------------------------- /Tutorials.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.Solutions.«00FirstProofs» 2 | import Tutorials.Solutions.«01EqualityRewriting» 3 | import Tutorials.Solutions.«02IffIfAnd» 4 | import Tutorials.Solutions.«03ForallOr» 5 | import Tutorials.Solutions.«04Exists» 6 | import Tutorials.Solutions.«05SequenceLimits» 7 | import Tutorials.Solutions.«06SubSequences» 8 | import Tutorials.Solutions.«07bisAbstractNegations» 9 | import Tutorials.Solutions.«07FirstNegations» 10 | import Tutorials.Solutions.«08LimitsNegation» 11 | import Tutorials.Solutions.«09LimitsFinal» 12 | -------------------------------------------------------------------------------- /Tutorials/Exercises/00FirstProofs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | This file is intended for Lean beginners. The goal is to demonstrate what it feels like to prove 3 | things using Lean and mathlib. Complicated definitions and theory building are not covered. 4 | Everything is covered again more slowly and with exercises in the next files. 5 | -/ 6 | -- We want real numbers and their basic properties 7 | import Mathlib.Data.Real.Basic 8 | -- We want to be able to use Lean's built-in "help" functionality 9 | import Mathlib.Tactic.LibrarySearch 10 | 11 | 12 | -- We want to be able to define functions using the law of excluded middle 13 | noncomputable section 14 | 15 | 16 | /- 17 | Our first goal is to define the set of upper bounds of a set of real numbers. 18 | This is already defined in mathlib (in a more general context), but we repeat 19 | it for the sake of exposition. Right-click "upperBounds" below to get offered 20 | to jump to mathlib's version 21 | -/ 22 | #check upperBounds 23 | 24 | /-- The set of upper bounds of a set of real numbers ℝ -/ 25 | def upBounds (A : Set ℝ) := 26 | { x : ℝ | ∀ a ∈ A, a ≤ x } 27 | 28 | /-- Predicate `is_maximum a A` means `a` is a maximum of `A` -/ 29 | def IsMaximum (a : ℝ) (A : Set ℝ) := 30 | a ∈ A ∧ a ∈ upBounds A 31 | 32 | /- 33 | In the above definition, the symbol `∧` means "and". We also see the most 34 | visible difference between set theoretic foundations and type theoretic ones 35 | (used by almost all proof assistants). In set theory, everything is a set, and the 36 | only relation you get from foundations are `=` and `∈`. In type theory, there is 37 | a meta-theoretic relation of "typing": `a : ℝ` reads "`a` is a real number" or, 38 | more precisely, "the type of `a` is `ℝ`". Here "meta-theoretic" means this is not a 39 | statement you can prove or disprove inside the theory, it's a fact that is true or 40 | not. Here we impose this fact, in other circumstances, it would be checked by the 41 | Lean kernel. 42 | By contrast, `a ∈ A` is a statement inside the theory. Here it's part of the 43 | definition, in other circumstances it could be something proven inside Lean. 44 | -/ 45 | /- For illustrative purposes, we now define an infix version of the above predicate. 46 | It will allow us to write `a is_a_max_of A`, which is closer to a sentence. 47 | -/ 48 | infixl:55 " is_a_max_of " => IsMaximum 49 | 50 | /- 51 | Let's prove something now! A set of real numbers has at most one maximum. Here 52 | everything left of the final `:` is introducing the objects and assumption. The equality 53 | `x = y` right of the colon is the conclusion. 54 | -/ 55 | theorem unique_max (A : Set ℝ) (x y : ℝ) (hx : x is_a_max_of A) (hy : y is_a_max_of A) : x = y := by 56 | -- We first break our assumptions in their two constituent pieces. 57 | -- We are free to choose the name following `with` 58 | rcases hx with ⟨x_in, x_up⟩ 59 | rcases hy with ⟨y_in, y_up⟩ 60 | -- Assumption `x_up` means x isn't less than elements of A, let's apply this to y 61 | specialize x_up y 62 | -- Assumption `x_up` now needs the information that `y` is indeed in `A`. 63 | specialize x_up y_in 64 | -- Let's do this quicker with roles swapped 65 | specialize y_up x x_in 66 | -- We explained to Lean the idea of this proof. 67 | -- Now we know `x ≤ y` and `y ≤ x`, and Lean shouldn't need more help. 68 | -- `linarith` proves equalities and inequalities that follow linearly from 69 | -- the assumption we have. 70 | linarith 71 | 72 | /- 73 | The above proof is too long, even if you remove comments. We don't really need the 74 | unpacking steps at the beginning; we can access both parts of the assumption 75 | `hx : x is_a_max_of A` using shortcuts `hx.1` and `hx.2`. We can also improve 76 | readability without assistance from the tactic state display, clearly announcing 77 | intermediate goals using `have`. This way we get to the following version of the 78 | same proof. 79 | -/ 80 | example (A : Set ℝ) (x y : ℝ) (hx : x is_a_max_of A) (hy : y is_a_max_of A) : x = y := by 81 | have : x ≤ y := hy.2 x hx.1 82 | have : y ≤ x := hx.2 y hy.1 83 | linarith 84 | 85 | /- 86 | Notice how mathematics based on type theory treats the assumption 87 | `∀ a ∈ A, a ≤ y` as a function turning an element `a` of `A` into the statement 88 | `a ≤ y`. More precisely, this assumption is the abbreviation of 89 | `∀ a : ℝ, a ∈ A → a ≤ y`. The expression `hy.2 x` appearing in the above proof 90 | is then the statement `x ∈ A → x ≤ y`, which itself is a function turning a 91 | statement `x ∈ A` into `x ≤ y` so that the full expression `hy.2 x hx.1` is 92 | indeed a proof of `x ≤ y`. 93 | 94 | One could argue a three-line-long proof of this lemma is still two lines too long. 95 | This is debatable, but mathlib's style is to write very short proofs for trivial 96 | lemmas. Those proofs are not easy to read but they are meant to indicate that the 97 | proof is probably not worth reading. 98 | 99 | In order to reach this stage, we need to know what `linarith` did for us. It invoked 100 | the lemma `le_antisymm` which says: `x ≤ y → y ≤ x → x = y`. This arrow, which 101 | is used both for function and implication, is right associative. So the statement is 102 | `x ≤ y → (y ≤ x → x = y)` which reads: I will send a proof `p` of `x ≤ y` to a function 103 | sending a proof `q'` of `y ≤ x` to a proof of `x = y`. Hence `le_antisymm p q'` is a 104 | proof of `x = y`. 105 | 106 | Using this we can get our one-line proof: 107 | -/ 108 | example (A : Set ℝ) (x y : ℝ) (hx : x is_a_max_of A) (hy : y is_a_max_of A) : x = y := 109 | le_antisymm (hy.2 x hx.1) (hx.2 y hy.1) 110 | 111 | /- 112 | Such a proof is called a proof term (or a "term mode" proof). Notice it has no `by`. 113 | It is directly the kind of low level proof that the Lean kernel is 114 | consuming. Commands like `rcases`, `specialize` or `linarith` are called tactics, they 115 | help users constructing proof terms that could be very tedious to write directly. 116 | The most efficient proof style combines tactics with proof terms like our previous 117 | `have : x ≤ y := hy.2 x hx.1` where `hy.2 x hx.1` is a proof term embeded inside 118 | a tactic mode proof. 119 | 120 | In the remaining of this file, we'll be characterizing infima of sets of real numbers 121 | in term of sequences. 122 | -/ 123 | /-- The set of lower bounds of a set of real numbers ℝ -/ 124 | def lowBounds (A : Set ℝ) := 125 | { x : ℝ | ∀ a ∈ A, x ≤ a } 126 | 127 | /- 128 | We now define `a` is an infimum of `A`. Again there is already a more general version 129 | in mathlib. 130 | -/ 131 | def IsInf (x : ℝ) (A : Set ℝ) := 132 | x is_a_max_of lowBounds A 133 | 134 | -- Let's define it also as an infix operator 135 | infixl:55 " is_an_inf_of " => IsInf 136 | 137 | /- 138 | We need to prove that any number which is greater than the infimum of A is greater 139 | than some element of A. 140 | -/ 141 | theorem inf_lt {A : Set ℝ} {x : ℝ} (hx : x is_an_inf_of A) : ∀ y, x < y → ∃ a ∈ A, a < y := by 142 | -- Let `y` be any real number. 143 | intro y 144 | -- Let's prove the contrapositive 145 | contrapose 146 | -- The symbol `¬` means negation. Let's ask Lean to rewrite the goal without negation, 147 | -- pushing negation through quantifiers and inequalities 148 | push_neg 149 | -- Let's assume the premise, calling the assumption `h` 150 | intro h 151 | -- `h` is exactly saying `y` is a lower bound of `A` so the second part of 152 | -- the infimum assumption `hx` applied to `y` and `h` is exactly what we want. 153 | exact hx.2 y h 154 | 155 | /- 156 | In the above proof, the sequence `contrapose, push_neg` is so common that it can be 157 | abbreviated to `contrapose!`. With these commands, we enter the gray zone between 158 | proof checking and proof finding. Practical computer proof checking crucially needs 159 | the computer to handle tedious proof steps. In the next proof, we'll start using 160 | `linarith` a bit more seriously, going one step further into automation. 161 | 162 | Our next real goal is to prove inequalities for limits of sequences. We extract the 163 | following lemma: if `y ≤ x + ε` for all positive `ε` then `y ≤ x`. 164 | -/ 165 | theorem le_of_le_add_eps {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 166 | -- Let's prove the contrapositive, asking Lean to push negations right away. 167 | contrapose! 168 | -- Assume `h : x < y`. 169 | intro h 170 | -- We need to find `ε` such that `ε` is positive and `x + ε < y`. 171 | -- Let's use `(y-x)/2` 172 | use (y - x) / 2 173 | -- we now have two properties to prove. Let's do both in turn, using `linarith` 174 | constructor 175 | linarith 176 | linarith 177 | 178 | /- 179 | Note how `linarith` was used for both sub-goals at the end of the above proof. 180 | We could have shortened that using the semi-colon combinator instead of comma, 181 | writing `constructor <;> linarith`. 182 | 183 | Next we will study a compressed version of that proof: 184 | -/ 185 | example {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 186 | contrapose! 187 | exact fun h => ⟨(y - x) / 2, by linarith, by linarith⟩ 188 | 189 | /- 190 | The angle brackets `⟨` and `⟩` introduce compound data or proofs. A proof 191 | of a `∃ z, P z` statemement is composed of a witness `z₀` and a proof `h` of 192 | `P z₀`. The compound is denoted by `⟨z₀, h⟩`. In the example above, the predicate is 193 | itself compound, it is a conjunction `P z ∧ Q z`. So the proof term should read 194 | `⟨z₀, ⟨h₁, h₂⟩⟩` where `h₁` (resp. `h₂`) is a proof of `P z₀` (resp. `Q z₀`). 195 | But these so-called "anonymous constructor" brackets are right-associative, so we can 196 | get rid of the nested brackets. 197 | 198 | Note also how we can use `by` to enter tactics anywhere a term is expected. 199 | 200 | Going all the way to a proof term would make the proof much longer, because we 201 | crucially use automation with `contrapose!` and `linarith`. 202 | -/ 203 | 204 | /- 205 | One could argue that the above proof is a bit too terse, and we are relying too much 206 | on linarith. Let's have more `linarith` calls for smaller steps. For the sake 207 | of (tiny) variation, we will also assume the premise and argue by contradiction 208 | instead of contraposing. 209 | -/ 210 | example {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 211 | intro h 212 | -- Assume the conclusion is false, and call this assumption H. 213 | by_contra H 214 | push_neg at H 215 | -- Now let's compute. 216 | have key := calc 217 | -- Each line must end with `:=` followed by a proof term 218 | -- We want to specialize our assumption `h` to `ε = (y-x)/2` but this is long to 219 | -- type, so let's put a hole `_` that Lean will fill in by comparing the 220 | -- statement we want to prove and our proof term with a hole. As usual, 221 | -- positivity of `(y-x)/2` is proved by `linarith` 222 | y ≤ x + (y - x) / 2 := h _ (by linarith) 223 | _ = x / 2 + y / 2 := by ring 224 | _ < y := by linarith 225 | 226 | -- our key now says `y < y` (notice how the sequence `≤`, `=`, `<` was correctly 227 | -- merged into a `<`). Let `linarith` find the desired contradiction now. 228 | linarith 229 | -- alternatively, we could have provided the proof term 230 | -- `exact lt_irrefl y key` 231 | 232 | /- 233 | Now we are ready for some analysis. Let's define convergence of sequences of real numbers 234 | (of course there is a much more general definition in mathlib). 235 | -/ 236 | 237 | /-- The sequence `u` tends to `l` -/ 238 | def Limit (u : ℕ → ℝ) (l : ℝ) := 239 | ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε 240 | 241 | /- 242 | In the above definition, `u n` denotes the n-th term of the sequence. We can 243 | add parentheses to get `u (n)` but we try to avoid parentheses because they pile up 244 | very quickly (and note the space between `u` and `(` is required). 245 | -/ 246 | 247 | -- If y ≤ u n for all n and u n goes to x then y ≤ x 248 | theorem le_lim {x y : ℝ} {u : ℕ → ℝ} (hu : Limit u x) (ineq : ∀ n, y ≤ u n) : y ≤ x := by 249 | -- Let's apply our previous lemma 250 | apply le_of_le_add_eps 251 | -- We need to prove y ≤ x + ε for all positive ε. 252 | -- Let ε be any positive real 253 | intro ε ε_pos 254 | -- we now specialize our limit assumption to this `ε`, and immediately 255 | -- fix a `N` as promised by the definition. 256 | rcases hu ε ε_pos with ⟨N, HN⟩ 257 | -- Now we only need to compute until reaching the conclusion 258 | calc 259 | y ≤ u N := ineq N 260 | _ = x + (u N - x) := by linarith 261 | -- In the next step we use the `gcongr` tactic which uses "generalized congruence" lemmas 262 | -- to zoom on the relevant part of the inequality goal, in this case `u N - x ≤ |u N - x|`. 263 | -- We then need a lemma saying `z ≤ |z|`. Because we don't know the name of this lemma, 264 | -- let's use `exact?`. Because searching through the library can be slow, 265 | -- Lean will write what it found in the Lean message window when cursor is on 266 | -- that line, so that we can replace it by the lemma. We see `le_abs_self`, which 267 | -- says `a ≤ |a|`, exactly what we're looking for. 268 | _ ≤ x + |u N - x| := by gcongr ; exact? 269 | _ ≤ x + ε := by gcongr ; apply HN; linarith 270 | 271 | /- 272 | The next lemma has been extracted from the main proof in order to discuss numbers. 273 | In ordinary maths, we know that ℕ is *not* contained in `ℝ`, whatever the 274 | construction of real numbers that we use. For instance a natural number is not 275 | an equivalence class of Cauchy sequences. But it's very easy to 276 | pretend otherwise. Formal maths requires slightly more care. In the statement below, 277 | the "type ascription" `(n + 1 : ℝ)` forces Lean to convert the natural number 278 | `n+1` into a real number. The "inclusion" map will be displayed in tactic state 279 | as `↑`. There are various lemmas asserting this map is compatible with addition and 280 | monotone, but we don't want to bother writing their names. The `norm_cast` 281 | tactic is designed to wisely apply those lemmas for us. 282 | -/ 283 | theorem inv_succ_pos : ∀ n : ℕ, 1 / (n + 1 : ℝ) > 0 := by 284 | -- Let `n` be any integer 285 | intro n 286 | -- Since we don't know the name of the relevant lemma, asserting that the inverse of 287 | -- a positive number is positive, let's state that is suffices 288 | -- to prove that `n+1`, seen as a real number, is positive, and ask `exact?` 289 | suffices (n + 1 : ℝ) > 0 by exact? 290 | -- Now we want to reduce to a statement about natural numbers, not real numbers 291 | -- coming from natural numbers. 292 | norm_cast 293 | -- and then get the usual help from `linarith` 294 | linarith 295 | 296 | /- 297 | That was a pretty long proof for an obvious fact. And stating it as a lemma feels 298 | stupid, so let's find a way to write it on one line in case we want to include it 299 | in some other proof without stating a lemma. First the `exact?` call 300 | above displays the name of the relevant lemma: `one_div_pos`. We can also 301 | replace the `linarith` call on the last line by `exact?` to learn the name 302 | of the lemma `Nat.succ_pos` asserting that the successor of a natural number is 303 | positive. There is also a variant on `norm_cast` that combines it with `exact`. 304 | The term mode analogue of `intro` is `fun`. We get down to: 305 | -/ 306 | example : ∀ n : ℕ, 1 / (n + 1 : ℝ) > 0 := 307 | fun n ↦ one_div_pos.mpr (by exact_mod_cast Nat.succ_pos n) 308 | 309 | /- 310 | The next proof uses mostly known things, so we will commment only new aspects. 311 | -/ 312 | theorem limit_inv_succ : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) ≤ ε := by 313 | intro ε ε_pos 314 | suffices ∃ N : ℕ, 1 / ε ≤ N by 315 | -- Because we didn't provide a name for the above statement, Lean called it `this`. 316 | -- Let's fix an `N` that works. 317 | rcases this with ⟨N, HN⟩ 318 | use N 319 | intro n Hn 320 | -- Now we want to rewrite the goal using lemmas 321 | -- `div_le_iff' : 0 < b → (a / b ≤ c ↔ a ≤ b * c)` 322 | -- `div_le_iff : 0 < b → (a / b ≤ c ↔ a ≤ c * b)` 323 | -- the second one will be rewritten from right to left, as indicated by `←`. 324 | -- Lean will create a side goal for the required positivity assumption that 325 | -- we don't provide for `div_le_iff'`. 326 | rw [div_le_iff', ← div_le_iff ε_pos] 327 | -- We want to replace assumption `Hn` by its real counter-part so that 328 | -- linarith can find what it needs. 329 | replace Hn : (N : ℝ) ≤ n 330 | exact_mod_cast Hn 331 | linarith 332 | -- we are still left with the positivity assumption. We already discussed 333 | -- how to prove it in the preceding lemma, but we could alternatively use 334 | -- the `positivity` tactic whose job is to prove obvious positivity statements. 335 | positivity 336 | -- Now we need to prove that sufficient statement. 337 | -- We want to use that `ℝ` is archimedean. So we start typing 338 | -- `exact archimedean_` and hit Ctrl-space to see what completion Lean proposes 339 | -- the lemma `archimedean_iff_nat_le` sounds promising. We select the left to 340 | -- right implication using `.1`. This a generic lemma for fields equiped with 341 | -- a linear (ie total) order. We need to provide a proof that `ℝ` is indeed 342 | -- archimedean. This is done using the `infer_instance` tactic that will be 343 | -- covered elsewhere. 344 | exact archimedean_iff_nat_le.1 (by infer_instance) (1 / ε) 345 | 346 | /- 347 | We can now put all pieces together, with almost no new things to explain. 348 | -/ 349 | theorem inf_seq (A : Set ℝ) (x : ℝ) : 350 | x is_an_inf_of A ↔ x ∈ lowBounds A ∧ ∃ u : ℕ → ℝ, Limit u x ∧ ∀ n, u n ∈ A := by 351 | constructor 352 | · intro h 353 | constructor 354 | · exact h.1 355 | -- On the next line, we don't need to tell Lean to treat `n+1` as a real number because 356 | -- we add `x` to it, so Lean knows there is only one way to make sense of this expression. 357 | have key : ∀ n : ℕ, ∃ a ∈ A, a < x + 1 / (n + 1) := by 358 | intro n 359 | -- we can use the lemma we proved above 360 | apply inf_lt h 361 | -- and another one we proved! 362 | have : 0 < 1 / (n + 1 : ℝ) := inv_succ_pos n 363 | linarith 364 | -- Now we need to use axiom of (countable) choice 365 | choose u hu using key 366 | use u 367 | constructor 368 | · intro ε ε_pos 369 | -- again we use a lemma we proved, specializing it to our fixed `ε`, and fixing a `N` 370 | rcases limit_inv_succ ε ε_pos with ⟨N, H⟩ 371 | use N 372 | intro n hn 373 | have : x ≤ u n := h.1 _ (hu n).1 374 | have := 375 | calc 376 | u n < x + 1 / (n + 1) := (hu n).2 377 | _ ≤ x + ε := add_le_add (le_refl x) (H n hn) 378 | 379 | rw [abs_of_nonneg] <;> linarith 380 | · intro n 381 | exact (hu n).1 382 | · intro h 383 | -- Assumption `h` is made of nested compound statements. We can use 384 | -- `rcases` to unpack it in one go. 385 | rcases h with ⟨x_min, u, lim, huA⟩ 386 | constructor 387 | exact x_min 388 | intro y y_mino 389 | apply le_lim lim 390 | intro n 391 | exact y_mino (u n) (huA n) 392 | 393 | -------------------------------------------------------------------------------- /Tutorials/Exercises/01EqualityRewriting.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | /- 4 | One of the earliest kind of proofs one encounters while learning mathematics is proving by 5 | a calculation. It may not sound like a proof, but this is actually using lemmas expressing 6 | properties of operations on numbers. It also uses the fundamental property of equality: if two 7 | mathematical objects A and B are equal then, in any statement involving A, one can replace A 8 | by B. This operation is called rewriting, and the Lean "tactic" for this is `rw`. 9 | 10 | In the following exercises, we will use the following two lemmas: 11 | `mul_assoc a b c : a * b * c = a * (b * c)` 12 | `mul_comm a b : a*b = b*a` 13 | 14 | Hence the command 15 | `rw [mul_assoc a b c]`, 16 | will replace `a*b*c` by `a*(b*c)` in the current goal. 17 | 18 | In order to replace backward, we use 19 | `rw [← mul_assoc a b c]`, 20 | replacing `a*(b*c)` by `a*b*c` in the current goal. 21 | 22 | Of course we don't want to constantly invoke those lemmas, and we will eventually introduce 23 | more powerful solutions. 24 | -/ 25 | 26 | example (a b c : ℝ) : a * b * c = b * (a * c) := by 27 | rw [mul_comm a b] 28 | rw [mul_assoc b a c] 29 | 30 | -- 0001 31 | example (a b c : ℝ) : c * b * a = b * (a * c) := by 32 | sorry 33 | 34 | -- 0002 35 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 36 | sorry 37 | 38 | /- 39 | Now let's return to the preceding example to experiment with what happens 40 | if we don't give arguments to mul_assoc or mul_comm. 41 | For instance, you can start the next proof with 42 | `rw [← mul_assoc]`, 43 | Try to figure out what happens. 44 | -/ 45 | -- 0003 46 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 47 | sorry 48 | 49 | /- 50 | We can also perform rewriting in an assumption of the local context, using for instance 51 | `rw [mul_comm a b] at hyp`, 52 | in order to replace a*b by b*a in assumption hyp. 53 | 54 | The next example will use a third lemma: 55 | `two_mul a : 2*a = a + a` 56 | 57 | Also we use the `exact` tactic, which allows to provide a direct proof term. 58 | -/ 59 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 60 | rw [hyp'] at hyp 61 | rw [mul_comm d a] at hyp 62 | rw [← two_mul (a * d)] at hyp 63 | rw [← mul_assoc 2 a d] at hyp 64 | exact hyp 65 | 66 | -- Our assumption hyp is now exactly what we have to prove 67 | /- 68 | And the next one can use: 69 | `sub_self x : x - x = 0` 70 | -/ 71 | -- 0004 72 | example (a b c d : ℝ) (hyp : c = b * a - d) (hyp' : d = a * b) : c = 0 := by 73 | sorry 74 | 75 | /- 76 | What is written in the two preceding example is very far away from what we would write on 77 | paper. Let's now see how to get a more natural layout. 78 | Inside each pair of curly braces below, the goal is to prove equality with the preceding line. 79 | -/ 80 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 81 | calc 82 | c = d * a + b := by rw [hyp] 83 | _ = d * a + a * d := by rw [hyp'] 84 | _ = a * d + a * d := by rw [mul_comm d a] 85 | _ = 2 * (a * d) := by rw [two_mul] 86 | _ = 2 * a * d := by rw [mul_assoc] 87 | 88 | 89 | /- 90 | From a practical point of view, when writing such a proof, it is convenient to: 91 | * pause the tactic state view update in VScode by clicking the Pause icon button 92 | in the top right corner of the Lean Goal buffer 93 | * write the full calculation, ending each line with ":= by {}" 94 | * resume tactic state update by clicking the Play icon button and fill in proofs between 95 | curly braces. 96 | 97 | Let's return to the other example using this method. 98 | -/ 99 | -- 0005 100 | example (a b c d : ℝ) (hyp : c = b * a - d) (hyp' : d = a * b) : c = 0 := by 101 | sorry 102 | 103 | /- 104 | The preceding proofs have exhausted our supply of "mul_comm" patience. Now it's time 105 | to get the computer to work harder. The `ring` tactic will prove any goal that follows by 106 | applying only the axioms of commutative (semi-)rings, in particular commutativity and 107 | associativity of addition and multiplication, as well as distributivity. 108 | -/ 109 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 110 | calc 111 | c = d * a + b := by rw [hyp] 112 | _ = d * a + a * d := by rw [hyp'] 113 | _ = 2 * a * d := by ring 114 | 115 | 116 | /- 117 | Of course we can use `ring` outside of `calc`. Let's do the next one in one line. 118 | -/ 119 | -- 0006 120 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 121 | sorry 122 | 123 | /- 124 | This is too much fun. Let's do it again. 125 | -/ 126 | -- 0007 127 | example (a b : ℝ) : a + b + a = 2 * a + b := by 128 | sorry 129 | 130 | /- 131 | Maybe this is cheating. Let's try to do the next computation without ring. 132 | We could use: 133 | `pow_two x : x^2 = x*x` 134 | `mul_sub a b c : a*(b-c) = a*b - a*c` 135 | `add_mul a b c : (a+b)*c = a*c + b*c` 136 | `add_sub a b c : a + (b - c) = (a + b) - c` 137 | `sub_sub a b c : a - b - c = a - (b + c)` 138 | `add_zero a : a + 0 = a` 139 | -/ 140 | -- 0008 141 | example (a b : ℝ) : (a + b) * (a - b) = a ^ 2 - b ^ 2 := by 142 | sorry 143 | 144 | -- Let's stick to ring in the end. 145 | 146 | -------------------------------------------------------------------------------- /Tutorials/Exercises/02IffIfAnd.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | /- 4 | In the previous file, we saw how to rewrite using equalities. 5 | The analogue operation with mathematical statements is rewriting using 6 | equivalences. This is also done using the `rw` tactic. 7 | Lean uses `↔` to denote equivalence instead of `⇔` (increase font size if you don't see a difference). 8 | 9 | In the following exercises we will use the lemma: 10 | 11 | `sub_nonneg {x y : ℝ} : 0 ≤ y - x ↔ x ≤ y` 12 | 13 | The curly braces around x and y instead of parentheses mean Lean will always try to figure out what 14 | x and y are from context, unless we really insist on telling it (we'll see how to insist much later). 15 | Let's not worry about that for now. 16 | 17 | In order to announce an intermediate statement we use: 18 | 19 | `have my_name : my statement := by` 20 | 21 | This triggers the apparition of a new goal: proving the statement. After this is done, 22 | the statement becomes available under the name `my_name`. 23 | -/ 24 | 25 | example {a b c : ℝ} (hab : a ≤ b) : c + a ≤ c + b := by 26 | rw [← sub_nonneg] 27 | have key : c + b - (c + a) = b - a := by-- Here we introduce an intermediate statement named key 28 | -- and prove it in an idented code block (or on the same line if the proof is very short) 29 | ring 30 | -- we can now use the key statement 31 | rw [key] 32 | rw [sub_nonneg] 33 | exact hab 34 | 35 | /- 36 | Of course the previous lemma is already in the core library, named `add_le_add_left`, so we can use it below. 37 | 38 | Let's prove a variation (without invoking commutativity of addition to reduce to the previous statement 39 | since this would spoil our fun). 40 | -/ 41 | -- 0009 42 | example {a b : ℝ} (hab : a ≤ b) (c : ℝ) : a + c ≤ b + c := by 43 | sorry 44 | 45 | /- 46 | Let's see how we could use this lemma. It is already in the core library, under the name `add_le_add_right`: 47 | 48 | `add_le_add_right {a b : ℝ} (hab : a ≤ b) (c : ℝ) : a + c ≤ b + c` 49 | 50 | This can be read as: "add_le_add_right is a function that will take as input real numbers a and b, an 51 | assumption `hab` claiming a ≤ b and a real number c, and will output a proof of a + c ≤ b + c". 52 | 53 | In addition, recall that curly braces around `a b` mean Lean will figure out those arguments unless we 54 | insist to help. This is because they can be deduced from the next argument `hab`. 55 | So it will be sufficient to feed `hab` and `c` to this function. 56 | -/ 57 | example {a b : ℝ} (ha : 0 ≤ a) : b ≤ a + b := by 58 | calc 59 | b = 0 + b := by ring 60 | _ ≤ a + b := by exact add_le_add_right ha b 61 | 62 | 63 | /- 64 | In the second line of the above proof, we need to prove `0 + b ≤ a + b`. 65 | The proof after the colon says: this is exactly lemma `add_le_add_right` applied to ha and b. 66 | Actually the `calc` block expects proof terms, and the `by` keyword is used to tell Lean we will use tactics 67 | to build such a proof term. But since the only tactic used in this block is `exact`, we can skip 68 | tactics entirely, and write: 69 | -/ 70 | example (a b : ℝ) (ha : 0 ≤ a) : b ≤ a + b := by 71 | calc 72 | b = 0 + b := by ring 73 | _ ≤ a + b := add_le_add_right ha b 74 | 75 | 76 | -- Let's do a variant. 77 | -- 0010 78 | example (a b : ℝ) (hb : 0 ≤ b) : a ≤ a + b := by 79 | sorry 80 | 81 | /- 82 | The two preceding examples are in the core library : 83 | 84 | `le_add_of_nonneg_left {a b : ℝ} (ha : 0 ≤ a) : b ≤ a + b` 85 | `le_add_of_nonneg_right {a b : ℝ} (hb : 0 ≤ b) : a ≤ a + b` 86 | 87 | Again, there won't be any need to memorize those names, we will 88 | soon see how to get rid of such goals automatically. 89 | But we can already try to understand how their names are built: 90 | "le_add" describe the conclusion "less or equal than some addition" 91 | It comes first because we are focussed on proving stuff, and 92 | auto-completion works by looking at the beginning of words. 93 | "of" introduces assumptions. "nonneg" is Lean's abbreviation for non-negative. 94 | "left" or "right" disambiguates between the two variations. 95 | 96 | Let's use those lemmas by hand for now. 97 | 98 | Note that you can have several inequalities steps in a `calc` block, 99 | transitivity of inequalities will be used automatically to assemble 100 | the pieces. 101 | -/ 102 | -- 0011 103 | example (a b : ℝ) (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := by 104 | sorry 105 | 106 | -- And let's combine with our earlier lemmas. 107 | -- 0012 108 | example (a b c d : ℝ) (hab : a ≤ b) (hcd : c ≤ d) : a + c ≤ b + d := by 109 | sorry 110 | 111 | /- 112 | In the above examples, we prepared proofs of assumptions of our lemmas beforehand, so 113 | that we could feed them to the lemmas. This is called forward reasoning. 114 | The `calc` proofs also belong to this category. 115 | 116 | We can also announce the use of a lemma, and provide proofs after the fact, using 117 | the `apply` tactic. This is called backward reasoning because we get the conclusion 118 | first, and provide proofs later. Using `rw` on the goal (rather than on an assumption 119 | from the local context) is also backward reasoning. 120 | 121 | Let's do that using the lemma 122 | 123 | `mul_nonneg {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : 0 ≤ x*y` 124 | -/ 125 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 126 | rw [← sub_nonneg] 127 | have key : b * c - a * c = (b - a) * c := by ring 128 | rw [key] 129 | apply mul_nonneg 130 | -- Here we don't provide proofs for the lemma's assumptions 131 | -- Now we need to provide the proofs. 132 | -- There are now two things to prove. We use the center dot (typed using `\.`) to 133 | -- focus on the current first goal. 134 | · rw [sub_nonneg] 135 | exact hab 136 | · exact hc 137 | 138 | /- 139 | Let's prove the same statement using only forward reasoning: announcing stuff, 140 | proving it by working with known facts, moving forward. 141 | -/ 142 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 143 | have hab' : 0 ≤ b - a := by 144 | rw [← sub_nonneg] at hab 145 | exact hab 146 | have h₁ : 0 ≤ (b - a) * c := mul_nonneg hab' hc 147 | have h₂ : (b - a) * c = b * c - a * c := by ring 148 | have h₃ : 0 ≤ b * c - a * c := by 149 | rw [h₂] at h₁ 150 | exact h₁ 151 | rw [sub_nonneg] at h₃ 152 | exact h₃ 153 | 154 | /- 155 | One reason why the backward reasoning proof is shorter is because Lean can 156 | infer of lot of things by comparing the goal and the lemma statement. Indeed 157 | in the `apply mul_nonneg` line, we didn't need to tell Lean that `x = b - a` 158 | and `y = c` in the lemma. It was infered by "unification" between the lemma 159 | statement and the goal. 160 | 161 | To be fair to the forward reasoning version, we should introduce a convenient 162 | variation on `rw`. The `rwa` tactic performs rewrite and then looks for an 163 | assumption matching the goal. We can use it to rewrite our latest proof as: 164 | -/ 165 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 166 | have hab' : 0 ≤ b - a := by rwa [← sub_nonneg] at hab 167 | have h₁ : 0 ≤ (b - a) * c := mul_nonneg hab' hc 168 | have h₂ : (b - a) * c = b * c - a * c := by ring 169 | have h₃ : 0 ≤ b * c - a * c := by rwa [h₂] at h₁ 170 | rwa [sub_nonneg] at h₃ 171 | 172 | /- 173 | Let's now combine forward and backward reasoning, to get our most 174 | efficient proof of this statement. Note in particular how unification is used 175 | to know what to prove inside the parentheses in the `mul_nonneg` arguments. 176 | -/ 177 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 178 | rw [← sub_nonneg] 179 | calc 180 | 0 ≤ (b - a) * c := mul_nonneg (by rwa [sub_nonneg]) hc 181 | _ = b * c - a * c := by ring 182 | 183 | 184 | /- 185 | Let's now practice all three styles using: 186 | 187 | `mul_nonneg_of_nonpos_of_nonpos {a b : α} (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b` 188 | 189 | `sub_nonpos {a b : α} : a - b ≤ 0 ↔ a ≤ b` 190 | -/ 191 | -- First using mostly backward reasoning 192 | -- 0013 193 | example (a b c : ℝ) (hc : c ≤ 0) (hab : a ≤ b) : b * c ≤ a * c := by 194 | sorry 195 | 196 | -- Using forward reasoning 197 | -- 0014 198 | example (a b c : ℝ) (hc : c ≤ 0) (hab : a ≤ b) : b * c ≤ a * c := by 199 | sorry 200 | 201 | -- 0015 202 | /-- Using a combination of both, with a `calc` block -/ 203 | example (a b c : ℝ) (hc : c ≤ 0) (hab : a ≤ b) : b * c ≤ a * c := by 204 | sorry 205 | 206 | /- 207 | Let's now move to proving implications. Lean denotes implications using 208 | a simple arrow `→`, the same it uses for functions (say denoting the type of functions 209 | from `ℕ` to `ℕ` by `ℕ → ℕ`). This is because it sees a proof of `P ⇒ Q` as a function turning 210 | a proof of `P` into a proof `Q`. 211 | 212 | Many of the examples that we already met are implications under the hood. For instance we proved 213 | 214 | `le_add_of_nonneg_left (a b : ℝ) (ha : 0 ≤ a) : b ≤ a + b` 215 | 216 | But this can be rephrased as 217 | 218 | `le_add_of_nonneg_left (a b : ℝ) : 0 ≤ a → b ≤ a + b` 219 | 220 | In order to prove `P → Q`, we use the tactic `intro`, followed by an assumption name. 221 | This creates an assumption with that name asserting that `P` holds, and turns the goal into `Q`. 222 | 223 | Let's check we can go from our old version of `le_add_of_nonneg_left` to the new one. 224 | 225 | -/ 226 | example (a b : ℝ) : 0 ≤ a → b ≤ a + b := by 227 | intro ha 228 | exact le_add_of_nonneg_left ha 229 | 230 | /- 231 | Actually Lean doesn't make any difference between those two versions. It is also happy with 232 | -/ 233 | example (a b : ℝ) : 0 ≤ a → b ≤ a + b := 234 | le_add_of_nonneg_left 235 | 236 | /- No tactic state is shown in the above line because we don't even need to enter 237 | tactic mode using `by`. 238 | 239 | Let's practise using `intro`. -/ 240 | -- 0016 241 | example (a b : ℝ) : 0 ≤ b → a ≤ a + b := by 242 | sorry 243 | 244 | /- 245 | What about lemmas having more than one assumption? For instance: 246 | 247 | `add_nonneg {a b : ℝ} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b` 248 | 249 | A natural idea is to use the conjunction operator (logical AND), which Lean denotes 250 | by ∧. Assumptions built using this operator can be decomposed using the `rcases` tactic, 251 | which is a very general assumption-decomposing tactic. 252 | -/ 253 | example {a b : ℝ} : 0 ≤ a ∧ 0 ≤ b → 0 ≤ a + b := by 254 | intro hyp 255 | rcases hyp with ⟨ha, hb⟩ 256 | exact add_nonneg ha hb 257 | 258 | /- 259 | Needing that intermediate line invoking `rcases` shows this formulation is not what is used by 260 | Lean. It rather sees `add_nonneg` as two nested implications: 261 | if a is non-negative then if b is non-negative then a+b is non-negative. 262 | It reads funny, but it is much more convenient to use in practice. 263 | -/ 264 | example {a b : ℝ} : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := 265 | add_nonneg 266 | 267 | /- 268 | The above pattern is so common that implications are defined as right-associative operators, 269 | hence parentheses are not needed above. 270 | 271 | Let's prove that the naive conjunction version implies the funny Lean version. For this we need 272 | to know how to prove a conjunction. The `constructor` tactic creates two goals from a conjunction goal. 273 | It can also be used to create two implication goals from an equivalence goal. 274 | -/ 275 | example {a b : ℝ} (H : 0 ≤ a ∧ 0 ≤ b → 0 ≤ a + b) : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := by 276 | intro ha 277 | intro hb 278 | apply H 279 | constructor 280 | exact ha 281 | exact hb 282 | 283 | /- 284 | Let's practice `rcases` and `constructor`. In the next exercise, `P`, `Q` and `R` denote 285 | unspecified mathematical statements. 286 | -/ 287 | -- 0017 288 | example (P Q : Prop) : P ∧ Q → Q ∧ P := by 289 | sorry 290 | 291 | /- 292 | Of course using `constructor` only to be able to use `exact` twice in a row feels silly. One can 293 | also use the anonymous constructor syntax: ⟨ ⟩ 294 | Beware those are not parentheses but angle brackets. This is a generic way of providing 295 | compound objects to Lean when Lean already has a very clear idea of what it is waiting for. 296 | 297 | So we could have replaced the last three lines by: 298 | exact ⟨hQ, hP⟩ 299 | 300 | We can also combine the `intros` steps. We can now compress our earlier proof to: 301 | -/ 302 | example {a b : ℝ} (H : 0 ≤ a ∧ 0 ≤ b → 0 ≤ a + b) : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := by 303 | intro ha hb 304 | exact H ⟨ha, hb⟩ 305 | 306 | /- 307 | The anonymous contructor trick actually also works in `intro`. So we can replace 308 | intro h, 309 | rcases h with ⟨h₁, h₂⟩ 310 | by 311 | intro ⟨h₁, h₂⟩, 312 | Now redo the previous exercise using all those compressing techniques, in exactly two lines. -/ 313 | -- 0018 314 | example (P Q : Prop) : P ∧ Q → Q ∧ P := by 315 | sorry 316 | 317 | /- 318 | We are ready to come back to the equivalence between the different formulations of 319 | lemmas having two assumptions. Remember the `constructor` tactic can be used to split 320 | an equivalence into two implications. 321 | -/ 322 | -- 0019 323 | example (P Q R : Prop) : P ∧ Q → R ↔ P → Q → R := by 324 | sorry 325 | 326 | /- 327 | If you used more than five lines in the above exercise then try to compress things 328 | (without simply removing line ends). 329 | 330 | One last compression technique: given a proof `h` of a conjunction `P ∧ Q`, one can get 331 | a proof of `P` using `h.left` and a proof of `Q` using `h.right`, without using `rcases`. 332 | One can also use the more generic (but less legible) names `h.1` and `h.2`. 333 | 334 | Similarly, given a proof `h` of `P ↔ Q`, one can get a proof of `P → Q` using `h.mp` 335 | and a proof of `Q → P` using `h.mpr` (or the generic `h.1` and `h.2` that are even less legible 336 | in this case). 337 | 338 | Before the final exercise in this file, let's make sure we'll be able to leave 339 | without learning 10 lemma names. The `linarith` tactic will prove any equality or 340 | inequality or contradiction that follows by linear combinations of assumptions from the 341 | context (with constant coefficients). 342 | -/ 343 | example (a b : ℝ) (hb : 0 ≤ b) : a ≤ a + b := by linarith 344 | 345 | /- 346 | Now let's enjoy this for a while. 347 | -/ 348 | -- 0020 349 | example (a b : ℝ) (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := by 350 | sorry 351 | 352 | -- And let's combine with our earlier lemmas. 353 | -- 0021 354 | example (a b c d : ℝ) (hab : a ≤ b) (hcd : c ≤ d) : a + c ≤ b + d := by 355 | sorry 356 | 357 | /- 358 | Final exercise 359 | 360 | In the last exercise of this file, we will use the divisibility relation on ℕ, 361 | denoted by `∣` (beware this is a unicode divisibility bar, not the ASCII pipe character), 362 | and the `gcd` function. 363 | 364 | The definitions are the usual ones, but our goal is to avoid using these definitions and 365 | only use the following three lemmas: 366 | 367 | `dvd_refl (a : ℕ) : a ∣ a` 368 | 369 | `dvd_antisymm {a b : ℕ} : a ∣ b → b ∣ a → a = b :=` 370 | 371 | `dvd_gcd_iff {a b c : ℕ} : c ∣ gcd a b ↔ c ∣ a ∧ c ∣ b` 372 | -/ 373 | -- All functions and lemmas below are about natural numbers. 374 | open Nat 375 | 376 | -- 0022 377 | example (a b : ℕ) : a ∣ b ↔ gcd a b = a := by 378 | sorry 379 | 380 | -------------------------------------------------------------------------------- /Tutorials/Exercises/03ForallOr.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | In this file, we'll learn about the `∀` quantifier, and the disjunction 5 | operator `∨` (logical OR). 6 | 7 | # The universal quantifier 8 | 9 | Let `P` be a predicate on a type `X`. This means for every mathematical 10 | object `x` with type `X`, we get a mathematical statement `P x`. 11 | In Lean, `P x` has type `Prop`. 12 | 13 | Lean sees a proof `h` of `∀ x, P x` as a function sending any `x : X` to 14 | a proof `h x` of `P x`. 15 | This already explains the main way to use an assumption or lemma which 16 | starts with a `∀`. 17 | 18 | In order to prove `∀ x, P x`, we use `intro x` to fix an arbitrary object 19 | with type `X`, and call it `x`. 20 | 21 | Note also we don't need to give the type of `x` in the expression `∀ x, P x` 22 | as long as the type of `P` is clear to Lean, which can then infer the type of `x`. 23 | 24 | Let's consider two predicates to play with `∀`. 25 | 26 | `EvenFun (f : ℝ → ℝ) : ∀ x, f (-x) = f x` 27 | 28 | `OddFun (f : ℝ → ℝ) : ∀ x, f (-x) = -f x`. 29 | 30 | -/ 31 | 32 | /- 33 | In the next proof, we also take the opportunity to introduce the 34 | `unfold` tactic, which simply unfolds definitions. Here this is purely 35 | for didactic reason, Lean doesn't need those `unfold` invocations. 36 | We will also use `rfl` which is a term proving equalities that are true 37 | by definition (in a very strong sense to be discussed later). 38 | -/ 39 | example (f g : ℝ → ℝ) : EvenFun f → EvenFun g → EvenFun (f + g) := by 40 | -- Assume f is even 41 | intro hf 42 | -- which means ∀ x, f (-x) = f x 43 | unfold EvenFun at hf 44 | -- and the same for g 45 | intro hg 46 | unfold EvenFun at hg 47 | -- We need to prove ∀ x, (f+g)(-x) = (f+g)(x) 48 | unfold EvenFun 49 | -- Let x be any real number 50 | intro x 51 | -- and let's compute 52 | calc 53 | (f + g) (-x) = f (-x) + g (-x) := rfl 54 | _ = f x + g (-x) := by rw [hf x] 55 | _ = f x + g x := by rw [hg x] 56 | _ = (f + g) x := rfl 57 | 58 | 59 | /- 60 | In the preceding proof, all `unfold` lines are purely for 61 | psychological comfort. 62 | 63 | Sometimes unfolding is necessary because we want to apply a tactic 64 | that operates purely on the syntactical level. 65 | The main such tactic is `rw`. 66 | 67 | The same property of `rw` explain why the first computation line 68 | is necessary, although its proof is simply `rfl`. 69 | Before that line, `rw [hf x]` won't find anything like `f (-x)` hence 70 | will give up. 71 | The last line is not necessary however, since it only proves 72 | something that is true by definition, and is not followed by 73 | a `rw`. 74 | 75 | Also, Lean doesn't need to be told that hf should be specialized to 76 | x before rewriting, exactly as in the first file 01_equality_rewriting. 77 | We can also gather several rewrites using a list of expressions. 78 | 79 | One last trick is that `rw` can take a list of expressions to use for 80 | rewriting. For instance `rw [h₁, h₂, h₃]` is equivalent to three 81 | lines `rw [h₁]`, `rw [h₂]` and `rw [h₃]`. Note that you can inspect the tactic 82 | state between those rewrites when reading a proof using this syntax. You 83 | simply need to move the cursor inside the list. 84 | 85 | Hence we can compress the above proof to: 86 | -/ 87 | example (f g : ℝ → ℝ) : EvenFun f → EvenFun g → EvenFun (f + g) := by 88 | intro hf hg x 89 | calc 90 | (f + g) (-x) = f (-x) + g (-x) := rfl 91 | _ = f x + g x := by rw [hf, hg] 92 | 93 | /- 94 | Now let's practice. If you need to learn how to type a unicode symbol, 95 | you can put your mouse cursor above the symbol and wait for one second. 96 | -/ 97 | -- 0023 98 | example (f g : ℝ → ℝ) : EvenFun f → EvenFun (g ∘ f) := by 99 | sorry 100 | 101 | -- 0024 102 | example (f g : ℝ → ℝ) : OddFun f → OddFun g → OddFun (g ∘ f) := by 103 | sorry 104 | 105 | /- 106 | Let's have more quantifiers, and play with forward and backward reasoning. 107 | 108 | In the next definitions, note how `∀ x₁, ∀ x₂` is abreviated to `∀ x₁ x₂`. 109 | 110 | `NonDecreasing (f : ℝ → ℝ) := ∀ x₁ x₂, x₁ ≤ x₂ → f x₁ ≤ f x₂` 111 | 112 | `NonIncreasing (f : ℝ → ℝ) := ∀ x₁ x₂, x₁ ≤ x₂ → f x₁ ≥ f x₂` 113 | 114 | -/ 115 | 116 | 117 | -- Let's be very explicit and use forward reasoning first. 118 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 119 | -- Let x₁ and x₂ be real numbers such that x₁ ≤ x₂ 120 | intro x₁ x₂ h 121 | -- Since f is non-decreasing, f x₁ ≤ f x₂. 122 | have step₁ : f x₁ ≤ f x₂ := hf x₁ x₂ h 123 | -- Since g is non-decreasing, we then get g (f x₁) ≤ g (f x₂). 124 | exact hg (f x₁) (f x₂) step₁ 125 | 126 | /- 127 | In the above proof, note how inconvenient it is to specify x₁ and x₂ in `hf x₁ x₂ h` since 128 | they could be inferred from the type of h. 129 | We could have written `hf _ _ h` and Lean would have filled the holes denoted by _. 130 | 131 | Even better we could have written the definition 132 | of `NonDecreasing` as: ∀ {x₁ x₂}, x₁ ≤ x₂ → f x₁ ≤ f x₂, with curly braces to denote 133 | implicit arguments. 134 | 135 | But let's leave that aside for now. One possible variation on the above proof is to 136 | use the `specialize` tactic to replace hf by its specialization to the relevant value. 137 | -/ 138 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 139 | intro x₁ x₂ h 140 | specialize hf x₁ x₂ h 141 | exact hg (f x₁) (f x₂) hf 142 | 143 | /- 144 | This `specialize` tactic is mostly useful for exploration, or in preparation for rewriting 145 | in the assumption. One can very often replace its use by using more complicated expressions 146 | directly involving the original assumption, as in the next variation: 147 | -/ 148 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 149 | intro x₁ x₂ h 150 | exact hg (f x₁) (f x₂) (hf x₁ x₂ h) 151 | 152 | /- 153 | Since the above proof uses only `intros` and `exact`, we could very easily replace it by the 154 | raw proof term: 155 | -/ 156 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := 157 | fun x₁ x₂ h ↦ hg (f x₁) (f x₂) (hf x₁ x₂ h) 158 | 159 | /- 160 | Of course the above proof is difficult to decipher. The principle in mathlib is to use 161 | such a proof when the result is obvious and you don't want to read the proof anyway. 162 | 163 | Instead of pursuing this style, let's see how backward reasoning would look like here. 164 | As usual with this style, we use `apply` and enjoy Lean specializing assumptions for us 165 | using unification. 166 | -/ 167 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 168 | -- Let x₁ and x₂ be real numbers such that x₁ ≤ x₂ 169 | intro x₁ x₂ h 170 | -- We need to prove (g ∘ f) x₁ ≤ (g ∘ f) x₂. 171 | -- Since g is non-decreasing, it suffices to prove f x₁ ≤ f x₂ 172 | apply hg 173 | -- which follows from our assumption on f 174 | apply hf 175 | -- and on x₁ and x₂ 176 | exact h 177 | 178 | -- 0025 179 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonIncreasing g) : NonIncreasing (g ∘ f) := by 180 | sorry 181 | 182 | /- # Disjunctions 183 | 184 | Let's switch to disjunctions now. Lean denotes by `∨` the 185 | logical OR operator. 186 | 187 | In order to make use of an assumption 188 | `hyp : P ∨ Q` 189 | we use the cases tactic: 190 | `rcases hyp with hP | hQ` 191 | which creates two proof branches: one branch assuming `hP : P`, 192 | and one branch assuming `hQ : Q`. 193 | 194 | In order to directly prove a goal `P ∨ Q`, 195 | we use either the `left` tactic and prove `P` or the `right` 196 | tactic and prove `Q`. 197 | 198 | In the next proof we use `ring` and `linarith` to get rid of 199 | easy computations or inequalities, as well as one lemma: 200 | 201 | `mul_eq_zero : a*b = 0 ↔ a = 0 ∨ b = 0` 202 | -/ 203 | example (a b : ℝ) : a = a * b → a = 0 ∨ b = 1 := by 204 | intro hyp 205 | have H : a * (1 - b) = 0 := by 206 | calc 207 | a * (1 - b) = a - a * b := by ring 208 | _ = 0 := by linarith 209 | 210 | rw [mul_eq_zero] at H 211 | rcases H with Ha | Hb 212 | · left 213 | exact Ha 214 | · right 215 | linarith 216 | 217 | -- 0026 218 | example (x y : ℝ) : x ^ 2 = y ^ 2 → x = y ∨ x = -y := by 219 | sorry 220 | 221 | /- 222 | In the next exercise, we can use: 223 | `eq_or_lt_of_le : x ≤ y → x = y ∨ x < y` 224 | -/ 225 | -- 0027 226 | example (f : ℝ → ℝ) : NonDecreasing f ↔ ∀ x y, x < y → f x ≤ f y := by 227 | sorry 228 | 229 | /- 230 | In the next exercise, we can use: 231 | `le_total x y : x ≤ y ∨ y ≤ x` 232 | -/ 233 | -- 0028 234 | example (f : ℝ → ℝ) (h : NonDecreasing f) (h' : ∀ x, f (f x) = x) : ∀ x, f x = x := by 235 | sorry 236 | 237 | -------------------------------------------------------------------------------- /Tutorials/Exercises/04Exists.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Mathlib.Data.Int.Parity 3 | 4 | /- # Existential quantifiers 5 | 6 | In this file, we learn how to handle the ∃ quantifier. 7 | 8 | In order to prove `∃ x, P x`, we give some `x₀` using tactic `use x₀` and 9 | then prove `P x₀`. This `x₀` can be an object from the local context 10 | or a more complicated expression. 11 | 12 | In the example below, `P x₀` is `8 = 2*4` which is true by definition so 13 | Lean does not ask for a proof. 14 | -/ 15 | 16 | example : ∃ n : ℕ, 8 = 2 * n := by 17 | use 4 18 | 19 | /- 20 | In order to use `h : ∃ x, P x`, we use the `rcases` tactic to fix 21 | one `x₀` that works. 22 | 23 | Again `h` can come straight from the local context or can be a more 24 | complicated expression. 25 | -/ 26 | example (n : ℕ) (h : ∃ k : ℕ, n = k + 1) : n > 0 := by 27 | -- Let's fix k₀ such that n = k₀ + 1. 28 | rcases h with ⟨k₀, hk₀⟩ 29 | -- It now suffices to prove k₀ + 1 > 0. 30 | rw [hk₀] 31 | -- and we have a lemma about this 32 | exact Nat.succ_pos k₀ 33 | 34 | /- 35 | The next exercises use divisibility in ℤ (beware the ∣ symbol which is 36 | not ASCII). 37 | 38 | By definition, a ∣ b ↔ ∃ k, b = a*k, so you can prove a ∣ b using the 39 | `use` tactic. 40 | -/ 41 | -- Until the end of this file, a, b and c will denote integers, unless 42 | -- explicitly stated otherwise 43 | variable (a b c : ℤ) 44 | 45 | -- 0029 46 | example (h₁ : a ∣ b) (h₂ : b ∣ c) : a ∣ c := by 47 | sorry 48 | 49 | /- 50 | A very common pattern is to have an assumption or lemma asserting 51 | `h : ∃ x, y = ...` 52 | and this is used through the combo: 53 | rcases h with ⟨x, hx⟩ 54 | rw hx at ... 55 | The tactic `rcases` allows us to simplify the above combo when the 56 | name `hx` is replaced by the special name `rfl`, as in the following example. 57 | -/ 58 | example (h1 : a ∣ b) (h2 : a ∣ c) : a ∣ b + c := by 59 | rcases h1 with ⟨k, rfl⟩ 60 | rcases h2 with ⟨l, rfl⟩ 61 | use k + l 62 | ring 63 | 64 | /- 65 | You can use the same `rfl` trick with the `rintro` tactic which 66 | is a power powerful variant of `intro`. 67 | -/ 68 | example : a ∣ b → a ∣ c → a ∣ b + c := by 69 | rintro ⟨k, rfl⟩ ⟨l, rfl⟩ 70 | use k + l 71 | ring 72 | 73 | -- 0030 74 | example : 0 ∣ a ↔ a = 0 := by 75 | sorry 76 | 77 | /- 78 | We can now start combining quantifiers, using the definition 79 | 80 | `Surjective (f : X → Y) := ∀ y, ∃ x, f x = y` 81 | 82 | -/ 83 | open Function 84 | 85 | -- In the remaining of this file, f and g will denote functions from 86 | -- ℝ to ℝ. 87 | variable (f g : ℝ → ℝ) 88 | 89 | -- 0031 90 | example (h : Surjective (g ∘ f)) : Surjective g := by 91 | sorry 92 | 93 | /- 94 | The above exercise can be done in three lines. Try to do the 95 | next exercise in four lines. 96 | -/ 97 | -- 0032 98 | example (hf : Surjective f) (hg : Surjective g) : Surjective (g ∘ f) := by 99 | sorry 100 | 101 | -------------------------------------------------------------------------------- /Tutorials/Exercises/05SequenceLimits.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Mathlib.Algebra.Group.Pi 3 | import Tutorials.TutoLib 4 | /- 5 | In this file we manipulate the elementary definition of limits of 6 | sequences of real numbers. 7 | mathlib has a much more general definition of limits, but here 8 | we want to practice using the logical operators and relations 9 | covered in the previous files. 10 | 11 | A sequence u is a function from ℕ to ℝ, hence Lean says 12 | `u : ℕ → ℝ` 13 | The definition we'll be using is: 14 | 15 | -- Definition of « u tends to l » 16 | `def seq_limit (u : ℕ → ℝ) (l : ℝ) : Prop := ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε` 17 | 18 | Note the use of `∀ ε > 0, ...` which is an abbreviation of 19 | `∀ ε, ε > 0 → ... ` 20 | 21 | In particular, a statement like `h : ∀ ε > 0, ...` 22 | can be specialized to a given ε₀ by 23 | `specialize h ε₀ hε₀` 24 | where `hε₀` is a proof of `ε₀ > 0`. 25 | 26 | Also recall that, wherever Lean expects some proof term, we can 27 | start a tactic mode proof using the keyword `by` (followed by curly braces 28 | if you need more than one tactic invocation). 29 | For instance, if the local context contains: 30 | 31 | δ : ℝ 32 | δ_pos : δ > 0 33 | h : ∀ ε > 0, ... 34 | 35 | then we can specialize h to the real number δ/2 using: 36 | `specialize h (δ/2) (by linarith)` 37 | where `by linarith` will provide the proof of `δ/2 > 0` expected by Lean. 38 | 39 | We'll take this opportunity to use two new tactics: 40 | 41 | `norm_num` will perform numerical normalization on the goal and `norm_num at h` 42 | will do the same in assumption `h`. This will get rid of trivial calculations on numbers, 43 | like replacing `|l - l|` by zero in the next exercise. 44 | Its name stands for "normalize numbers". 45 | 46 | `congr` will try to prove equalities between applications of functions by recursively 47 | proving the arguments are the same. Its name stands for "congruence". 48 | For instance, if the goal is `f x + g y = f z + g t` then congr will replace it by 49 | two goals: `x = z` and `y = t`. 50 | You can limit the recursion depth by specifying a natural number after `congr`. 51 | For instance, in the above example, `congr 1` will give new goals 52 | `f x = f z` and `g y = g t`, which only inspect arguments of the addition and not deeper. 53 | -/ 54 | variable (u v w : ℕ → ℝ) (l l' : ℝ) 55 | 56 | -- If u is constant with value l then u tends to l 57 | -- 0033 58 | example : (∀ n, u n = l) → SeqLimit u l := by 59 | sorry 60 | 61 | /- When dealing with absolute values, we'll use lemmas: 62 | 63 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 64 | 65 | `abs_add (x y : ℝ) : |x + y| ≤ |x| + |y|` 66 | 67 | `abs_sub_comm (x y : ℝ) : |x - y| = |y - x|` 68 | 69 | You should probably write them down on a sheet of paper that you keep at 70 | hand since they are used in many exercises. 71 | -/ 72 | -- Assume l > 0. Then u tends to l implies u n ≥ l/2 for large enough n 73 | -- 0034 74 | example (hl : l > 0) : SeqLimit u l → ∃ N, ∀ n ≥ N, u n ≥ l / 2 := by 75 | sorry 76 | 77 | /- 78 | When dealing with max, you can use 79 | 80 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 81 | 82 | `le_max_left p q : p ≤ max p q` 83 | 84 | `le_max_right p q : q ≤ max p q` 85 | 86 | You should probably add them to the sheet of paper where you wrote 87 | the `abs` lemmas since they are used in many exercises. 88 | 89 | Let's see an example. 90 | -/ 91 | -- If u tends to l and v tends l' then u+v tends to l+l' 92 | example (hu : SeqLimit u l) (hv : SeqLimit v l') : SeqLimit (u + v) (l + l') := by 93 | intro ε ε_pos 94 | rcases hu (ε / 2) (by linarith) with ⟨N₁, hN₁⟩ 95 | rcases hv (ε / 2) (by linarith) with ⟨N₂, hN₂⟩ 96 | use max N₁ N₂ 97 | intro n hn 98 | rcases ge_max_iff.mp hn with ⟨hn₁, hn₂⟩ 99 | have fact₁ : |u n - l| ≤ ε / 2 := hN₁ n (by linarith) 100 | have fact₂ : |v n - l'| ≤ ε / 2 := hN₂ n (by linarith) 101 | calc 102 | |(u + v) n - (l + l')| = |u n + v n - (l + l')| := rfl 103 | _ = |u n - l + (v n - l')| := by congr ; ring 104 | _ ≤ |u n - l| + |v n - l'| := by apply abs_add 105 | _ ≤ ε := by linarith 106 | 107 | 108 | /- 109 | In the above proof, we used `have` to prepare facts for `linarith` consumption in the last line. 110 | Since we have direct proof terms for them, we can feed them directly to `linarith` as in the next proof 111 | of the same statement. 112 | Another variation we introduce is rewriting using `ge_max_iff` and letting `linarith` handle the 113 | conjunction, instead of creating two new assumptions. 114 | -/ 115 | example (hu : SeqLimit u l) (hv : SeqLimit v l') : SeqLimit (u + v) (l + l') := by 116 | intro ε ε_pos 117 | rcases hu (ε / 2) (by linarith) with ⟨N₁, hN₁⟩ 118 | rcases hv (ε / 2) (by linarith) with ⟨N₂, hN₂⟩ 119 | use max N₁ N₂ 120 | intro n hn 121 | rw [ge_max_iff] at hn 122 | calc 123 | |(u + v) n - (l + l')| = |u n + v n - (l + l')| := rfl 124 | _ = |u n - l + (v n - l')| := by ring_nf 125 | _ ≤ |u n - l| + |v n - l'| := by apply abs_add 126 | _ ≤ ε := by linarith [hN₁ n (by linarith), hN₂ n (by linarith)] 127 | 128 | 129 | -- Let's do something similar: the squeezing theorem. 130 | -- 0035 131 | example (hu : SeqLimit u l) (hw : SeqLimit w l) (h : ∀ n, u n ≤ v n) (h' : ∀ n, v n ≤ w n) : 132 | SeqLimit v l := by 133 | sorry 134 | 135 | -- What about < ε? 136 | -- 0036 137 | example (u l) : SeqLimit u l ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| < ε := by 138 | sorry 139 | 140 | /- In the next exercise, we'll use 141 | 142 | `eq_of_abs_sub_le_all (x y : ℝ) : (∀ ε > 0, |x - y| ≤ ε) → x = y` 143 | -/ 144 | -- A sequence admits at most one limit 145 | -- 0037 146 | example : SeqLimit u l → SeqLimit u l' → l = l' := by 147 | sorry 148 | 149 | /- 150 | Let's now practice deciphering definitions before proving. 151 | 152 | -/ 153 | def NonDecreasingSeq (u : ℕ → ℝ) := 154 | ∀ n m, n ≤ m → u n ≤ u m 155 | 156 | def IsSeqSup (M : ℝ) (u : ℕ → ℝ) := 157 | (∀ n, u n ≤ M) ∧ ∀ ε > 0, ∃ n₀, u n₀ ≥ M - ε 158 | 159 | -- 0038 160 | example (M : ℝ) (h : IsSeqSup M u) (h' : NonDecreasingSeq u) : SeqLimit u M := by 161 | sorry 162 | 163 | -------------------------------------------------------------------------------- /Tutorials/Exercises/06SubSequences.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | This file continues the elementary study of limits of sequences. 5 | It can be skipped if the previous file was too easy, it won't introduce 6 | any new tactic or trick. 7 | 8 | Remember useful lemmas: 9 | 10 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 11 | 12 | `abs_add (x y : ℝ) : |x + y| ≤ |x| + |y|` 13 | 14 | `abs_sub_comm (x y : ℝ) : |x - y| = |y - x|` 15 | 16 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 17 | 18 | `le_max_left p q : p ≤ max p q` 19 | 20 | `le_max_right p q : q ≤ max p q` 21 | 22 | and the definition: 23 | 24 | `def SeqLimit (u : ℕ → ℝ) (l : ℝ) : Prop := ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε` 25 | 26 | You can also use a property proved in the previous file: 27 | 28 | `unique_limit : SeqLimit u l → SeqLimit u l' → l = l'` 29 | 30 | `def extraction (φ : ℕ → ℕ) := ∀ n m, n < m → φ n < φ m` 31 | -/ 32 | variable {φ : ℕ → ℕ} 33 | 34 | /- 35 | The next lemma is proved by an easy induction, but we haven't seen induction 36 | in this tutorial. If you did the natural number game then you can delete 37 | the proof below and try to reconstruct it. 38 | -/ 39 | /-- An extraction is greater than id -/ 40 | theorem id_le_extraction' : Extraction φ → ∀ n, n ≤ φ n := by 41 | intro hyp n 42 | induction' n with n hn 43 | · exact Nat.zero_le _ 44 | · exact Nat.succ_le_of_lt (by linarith [hyp n (n + 1) (by linarith)]) 45 | 46 | -- 0039 47 | /-- Extractions take arbitrarily large values for arbitrarily large 48 | inputs. -/ 49 | theorem extraction_ge : Extraction φ → ∀ N N', ∃ n ≥ N', φ n ≥ N := by 50 | sorry 51 | 52 | variable {u : ℕ → ℝ} {a l : ℝ} 53 | 54 | /- 55 | In the exercise, we use `∃ n ≥ N, ...` which is the abbreviation of 56 | `∃ n, n ≥ N ∧ ...`. 57 | Lean can read this abbreviation, but does not it when displaying the goal. 58 | -/ 59 | -- 0040 60 | /-- If `a` is a cluster point of `u` then there are values of 61 | `u` arbitrarily close to `a` for arbitrarily large input. -/ 62 | theorem near_cluster : ClusterPoint u a → ∀ ε > 0, ∀ N, ∃ n ≥ N, |u n - a| ≤ ε := by 63 | sorry 64 | 65 | /- 66 | The above exercice can be done in five lines. 67 | Hint: you can use the anonymous constructor syntax when proving 68 | existential statements. 69 | -/ 70 | -- 0041 71 | /-- If `u` tends to `l` then its subsequences tend to `l`. -/ 72 | theorem subseq_tendsto_of_tendsto' (h : SeqLimit u l) (hφ : Extraction φ) : SeqLimit (u ∘ φ) l := by 73 | sorry 74 | 75 | -- 0042 76 | /-- If `u` tends to `l` all its cluster points are equal to `l`. -/ 77 | theorem cluster_limit (hl : SeqLimit u l) (ha : ClusterPoint u a) : a = l := by 78 | sorry 79 | 80 | /-- Cauchy_sequence sequence -/ 81 | def CauchySequence (u : ℕ → ℝ) := 82 | ∀ ε > 0, ∃ N, ∀ p q, p ≥ N → q ≥ N → |u p - u q| ≤ ε 83 | 84 | -- 0043 85 | example : (∃ l, SeqLimit u l) → CauchySequence u := by 86 | sorry 87 | 88 | /- 89 | In the next exercise, you can reuse 90 | `near_cluster : cluster_point u a → ∀ ε > 0, ∀ N, ∃ n ≥ N, |u n - a| ≤ ε` 91 | -/ 92 | -- 0044 93 | example (hu : CauchySequence u) (hl : ClusterPoint u l) : SeqLimit u l := by 94 | sorry 95 | 96 | -------------------------------------------------------------------------------- /Tutorials/Exercises/07FirstNegations.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | Negations, proof by contradiction and contraposition. 5 | 6 | This file introduces the logical rules and tactics related to negation: 7 | exfalso, by_contra, contrapose, by_cases and push_neg. 8 | 9 | There is a special statement denoted by `False` which, by definition, 10 | has no proof. 11 | 12 | So `False` implies everything. Indeed `False → P` means any proof of 13 | `False` could be turned into a proof of P. 14 | This fact is known by its latin name 15 | "ex falso quod libet" (from False follows whatever you want). 16 | Hence Lean's tactic to invoke this is called `exfalso`. 17 | -/ 18 | example : False → 0 = 1 := by 19 | intro h 20 | exfalso 21 | exact h 22 | 23 | /- 24 | The preceding example suggests that this definition of `False` isn't very useful. 25 | But actually it allows us to define the negation of a statement P as 26 | "P implies False" that we can read as "if P were true, we would get 27 | a contradiction". Lean denotes this by `¬ P`. 28 | 29 | One can prove that (¬ P) ↔ (P ↔ False). But in practice we directly 30 | use the definition of `¬ P`. 31 | -/ 32 | example {x : ℝ} : ¬x < x := by 33 | intro hyp 34 | rw [lt_iff_le_and_ne] at hyp 35 | rcases hyp with ⟨hyp_inf, hyp_non⟩ 36 | clear hyp_inf -- we won't use that one, so let's discard it 37 | change x = x → False at hyp_non -- Lean doesn't need this psychological line 38 | apply hyp_non 39 | rfl 40 | 41 | open Int 42 | 43 | -- 0045 44 | example (n : ℤ) (h_even : Even n) (h_not_even : ¬Even n) : 0 = 1 := by 45 | sorry 46 | 47 | -- 0046 48 | example (P Q : Prop) (h₁ : P ∨ Q) (h₂ : ¬(P ∧ Q)) : ¬P ↔ Q := by 49 | sorry 50 | 51 | /- 52 | The definition of negation easily implies that, for every statement P, 53 | `P → ¬ ¬ P`. 54 | 55 | The excluded middle axiom, which asserts `P ∨ ¬ P` allows us to 56 | prove the converse implication. 57 | 58 | Together those two implications form the principle of double negation elimination. 59 | `not_not {P : Prop} : (¬ ¬ P) ↔ P` 60 | 61 | The implication `¬ ¬ P → P` is the basis for proofs by contradiction: 62 | in order to prove `P`, it suffices to prove `¬ ¬ P`, ie `¬ P → False`. 63 | 64 | Of course there is no need to keep explaining all this. The tactic 65 | `by_contra Hyp` will transform any goal `P` into `False` and 66 | add `Hyp : ¬ P` to the local context. 67 | 68 | Let's return to a proof from the 5th file: uniqueness of limits for a sequence. 69 | This cannot be proved without using some version of the excluded middle 70 | axiom. We used it secretely in 71 | 72 | `eq_of_abs_sub_le_all (x y : ℝ) : (∀ ε > 0, |x - y| ≤ ε) → x = y` 73 | 74 | (we'll prove a variation on this lemma below). 75 | 76 | In the proof below, we also take the opportunity to introduce the `let` tactic 77 | which creates a local definition. If needed, it can be unfolded using `unfold_let`. 78 | For instance after using `let N₀ := max N N'`, you could write 79 | `unfold_let N₀ at h` to replace `N₀` by its definition in some local assumption `h`. 80 | -/ 81 | example (u : ℕ → ℝ) (l l' : ℝ) : SeqLimit u l → SeqLimit u l' → l = l' := by 82 | intro hl hl' 83 | by_contra H 84 | change l ≠ l' at H -- Lean does not need this line, but it makes H easier to read 85 | have ineg : |l - l'| > 0 := abs_pos.mpr (sub_ne_zero_of_ne H) -- details about that line are not important 86 | rcases hl (|l - l'| / 4) (by linarith) with ⟨N, hN⟩ 87 | rcases hl' (|l - l'| / 4) (by linarith) with ⟨N', hN'⟩ 88 | let N₀ := max N N' 89 | specialize hN N₀ (le_max_left _ _) 90 | specialize hN' N₀ (le_max_right _ _) 91 | have clef : |l - l'| < |l - l'| 92 | calc 93 | |l - l'| = |l - u N₀ + (u N₀ - l')| := by ring_nf 94 | _ ≤ |l - u N₀| + |u N₀ - l'| := by apply abs_add 95 | _ = |u N₀ - l| + |u N₀ - l'| := by rw [abs_sub_comm] 96 | _ < |l - l'| := by linarith 97 | 98 | linarith -- linarith can also find simple numerical contradictions 99 | 100 | /- 101 | Another incarnation of the excluded middle axiom is the principle of 102 | contraposition: in order to prove P ⇒ Q, it suffices to prove 103 | not Q ⇒ not P. 104 | -/ 105 | 106 | -- Using a proof by contradiction, let's prove the contraposition principle 107 | -- 0047 108 | example (P Q : Prop) (h : ¬Q → ¬P) : P → Q := by 109 | sorry 110 | 111 | /- 112 | Again Lean doesn't need this principle explained to it. We can use the 113 | `contrapose` tactic. 114 | -/ 115 | example (P Q : Prop) (h : ¬Q → ¬P) : P → Q := by 116 | contrapose 117 | exact h 118 | 119 | /- 120 | In the next exercise, we'll use 121 | Odd n : ∃ k, n = 2*k + 1 122 | Int.odd_iff_not_even {n : ℤ} : odd n ↔ ¬ even n 123 | -/ 124 | -- 0048 125 | example (n : ℤ) : Even (n ^ 2) ↔ Even n := by 126 | sorry 127 | 128 | /- 129 | As a last step on our law of the excluded middle tour, let's notice that, especially 130 | in pure logic exercises, it can sometimes be useful to use the 131 | excluded middle axiom in its original form: 132 | `Classical.em : ∀ P, P ∨ ¬ P` 133 | 134 | Instead of applying this lemma and then using the `rcases` tactic, we 135 | have the shortcut 136 | `by_cases h : P` 137 | 138 | combining both steps to create two proof branches: one assuming 139 | `h : P`, and the other assuming `h : ¬ P`. 140 | 141 | For instance, let's prove a reformulation of this implication relation, 142 | which is sometimes used as a definition in other logical foundations, 143 | especially those based on truth tables (hence very strongly using 144 | excluded middle from the very beginning). 145 | -/ 146 | variable (P Q : Prop) 147 | 148 | example : P → Q ↔ ¬P ∨ Q := by 149 | constructor 150 | · intro h 151 | by_cases hP : P 152 | · right 153 | exact h hP 154 | · left 155 | exact hP 156 | · intro h hP 157 | rcases h with hnP | hQ 158 | · exfalso 159 | exact hnP hP 160 | · exact hQ 161 | 162 | -- 0049 163 | example : ¬(P ∧ Q) ↔ ¬P ∨ ¬Q := by 164 | sorry 165 | 166 | /- 167 | It is crucial to understand negation of quantifiers. 168 | Let's do it by hand for a little while. 169 | In the first exercise, only the definition of negation is needed. 170 | -/ 171 | -- 0050 172 | example (n : ℤ) : (¬∃ k, n = 2 * k) ↔ ∀ k, n ≠ 2 * k := by 173 | sorry 174 | 175 | /- 176 | Contrary to negation of the existential quantifier, negation of the 177 | universal quantifier requires excluded middle for the first implication. 178 | In order to prove this, we can use either 179 | * a double proof by contradiction 180 | * a contraposition, `not_not : (¬ ¬ P) ↔ P` and a proof by contradiction. 181 | 182 | Recall we have 183 | `EvenFun (f : ℝ → ℝ) := ∀ x, f (-x) = f x` 184 | 185 | -/ 186 | 187 | -- 0051 188 | example (f : ℝ → ℝ) : ¬EvenFun f ↔ ∃ x, f (-x) ≠ f x := by 189 | sorry 190 | 191 | /- 192 | Of course we can't keep repeating the above proofs, especially the second one. 193 | So we use the `push_neg` tactic. 194 | -/ 195 | example : ¬EvenFun fun x => 2 * x := by 196 | unfold EvenFun 197 | -- Here unfolding is important because push_neg won't do it. 198 | push_neg 199 | use 42 200 | linarith 201 | 202 | -- 0052 203 | example (f : ℝ → ℝ) : ¬EvenFun f ↔ ∃ x, f (-x) ≠ f x := by 204 | sorry 205 | 206 | def BoundedAbove (f : ℝ → ℝ) := 207 | ∃ M, ∀ x, f x ≤ M 208 | 209 | example : ¬BoundedAbove fun x => x := by 210 | unfold BoundedAbove 211 | push_neg 212 | intro M 213 | use M + 1 214 | linarith 215 | 216 | -- Let's contrapose 217 | -- 0053 218 | example (x : ℝ) : (∀ ε > 0, x ≤ ε) → x ≤ 0 := by 219 | sorry 220 | 221 | /- 222 | The "contrapose, push_neg" combo is so common that we can abreviate it to 223 | `contrapose!` 224 | 225 | Let's use this trick, together with: 226 | `eq_or_lt_of_le : a ≤ b → a = b ∨ a < b` 227 | -/ 228 | -- 0054 229 | example (f : ℝ → ℝ) : (∀ x y, x < y → f x < f y) ↔ ∀ x y, x ≤ y ↔ f x ≤ f y := by 230 | sorry 231 | 232 | -------------------------------------------------------------------------------- /Tutorials/Exercises/07bisAbstractNegations.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | open Classical 4 | 5 | /- 6 | Theoretical negations. 7 | 8 | This file is for people interested in logic who want to fully understand 9 | negations. It comes after the file `07FirstNegations`. 10 | 11 | Here we don't use `contrapose` or `push_neg`. The goal is to prove lemmas 12 | that are used by those tactics. Of course we can use 13 | `exfalso`, `by_contra` and `by_cases`. 14 | 15 | If this doesn't sound like fun then skip ahead to the next file. 16 | -/ 17 | section NegationProp 18 | 19 | variable (P Q : Prop) 20 | 21 | -- 0055 22 | example : P → Q ↔ ¬Q → ¬P := by 23 | sorry 24 | 25 | -- 0056 26 | theorem non_imp (P Q : Prop) : ¬(P → Q) ↔ P ∧ ¬Q := by 27 | sorry 28 | 29 | -- In the next one, let's use the axiom 30 | -- `propext {P Q : Prop} : (P ↔ Q) → P = Q` 31 | -- 0057 32 | example (P : Prop) : ¬P ↔ P = False := by 33 | sorry 34 | 35 | end NegationProp 36 | 37 | section NegationQuantifiers 38 | 39 | variable (X : Type) (P : X → Prop) 40 | 41 | -- 0058 42 | example : (¬∀ x, P x) ↔ ∃ x, ¬P x := by 43 | sorry 44 | 45 | -- 0059 46 | example : (¬∃ x, P x) ↔ ∀ x, ¬P x := by 47 | sorry 48 | 49 | -- 0060 50 | example (P : ℝ → Prop) : (¬∃ ε > 0, P ε) ↔ ∀ ε > 0, ¬P ε := by 51 | sorry 52 | 53 | -- 0061 54 | example (P : ℝ → Prop) : (¬∀ x > 0, P x) ↔ ∃ x > 0, ¬P x := by 55 | sorry 56 | 57 | end NegationQuantifiers 58 | 59 | -------------------------------------------------------------------------------- /Tutorials/Exercises/08LimitsNegation.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | 4 | section 5 | 6 | /- 7 | The first part of this file makes sure you can negate quantified statements 8 | in your head without the help of `push_neg`. 9 | 10 | You need to complete the statement and then use the `check_me` tactic 11 | to check your answer. This tactic exists only for those exercises, 12 | it mostly calls `push_neg` and then cleans up a bit. 13 | 14 | `def SeqLimit (u : ℕ → ℝ) (l : ℝ) : Prop := ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε` 15 | -/ 16 | -- In this section, u denotes a sequence of real numbers 17 | -- f is a function from ℝ to ℝ 18 | -- x₀ and l are real numbers 19 | variable (u : ℕ → ℝ) (f : ℝ → ℝ) (x₀ l : ℝ) 20 | 21 | 22 | -- Negation of "u tends to l" 23 | -- 0062 24 | example : 25 | sorry 26 | check_me 27 | (¬∀ ε > 0, ∃ δ > 0, ∀ x, |x - x₀| ≤ δ → |f x - f x₀| ≤ ε) ↔sorry 28 | ∃ ε > 0, ∀ δ > 0, ∃ x, |x - x₀| ≤ δ ∧ |f x - f x₀| > ε := by 29 | sorry 30 | 31 | /- 32 | In the next exercise, we need to keep in mind that 33 | `∀ x x', ...` is the abbreviation of 34 | `∀ x, ∀ x', ... `. 35 | 36 | Also, `∃ x x', ...` is the abbreviation of `∃ x, ∃ x', ...`. 37 | -/ 38 | -- Negation of "f is uniformly continuous on ℝ" 39 | -- 0064 40 | example : 41 | sorry 42 | check_me 43 | ∀ ε > 0, ∃ N, ∀ n ≥ N, |(f ∘ u) n - f x₀| ≤ ε) ↔sorry 44 | ∃ u : ℕ → ℝ, 45 | (∀ δ > 0, ∃ N, ∀ n ≥ N, |u n - x₀| ≤ δ) ∧ ∃ ε > 0, ∀ N, ∃ n ≥ N, |f (u n) - f x₀| > ε := by 46 | sorry 47 | 48 | end 49 | 50 | /- 51 | We now turn to elementary applications of negations to limits of sequences. 52 | Remember that `linarith` can find easy numerical contradictions. 53 | 54 | Also recall the following lemmas: 55 | 56 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 57 | 58 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 59 | 60 | `le_max_left p q : p ≤ max p q` 61 | 62 | `le_max_right p q : q ≤ max p q` 63 | 64 | /-- The sequence `u` tends to `+∞`. -/ 65 | `def TendstoInfinity (u : ℕ → ℝ) := ∀ A, ∃ N, ∀ n ≥ N, u n ≥ A` 66 | -/ 67 | 68 | -- 0066 69 | example {u : ℕ → ℝ} : TendstoInfinity u → ∀ l, ¬SeqLimit u l := by 70 | sorry 71 | 72 | def NondecreasingSeq (u : ℕ → ℝ) := 73 | ∀ n m, n ≤ m → u n ≤ u m 74 | 75 | -- 0067 76 | example (u : ℕ → ℝ) (l : ℝ) (h : SeqLimit u l) (h' : NondecreasingSeq u) : ∀ n, u n ≤ l := by 77 | sorry 78 | 79 | /- 80 | In the following exercises, `A : set ℝ` means that A is a set of real numbers. 81 | We can use the usual notation `x ∈ A`. 82 | 83 | The notation `∀ x ∈ A, ...` is the abbreviation of `∀ x, x ∈ A → ... ` 84 | 85 | The notation `∃ x ∈ A, ...` is the abbreviation of `∃ x, x ∈ A ∧ ... `. 86 | 87 | We'll work with upper bounds and supremums. 88 | Again we'll introduce specialized definitions for the sake of exercises, but mathlib 89 | has more general versions. 90 | 91 | `def UpperBound (A : set ℝ) (x : ℝ) := ∀ a ∈ A, a ≤ x` 92 | 93 | `def IsSup (A : set ℝ) (x : ℝ) := UpperBound A x ∧ ∀ y, UpperBound A y → x ≤ y` 94 | 95 | 96 | Remark: one can easily show that a set of real numbers has at most one sup, 97 | but we won't need this. 98 | -/ 99 | -- 0068 100 | example {A : Set ℝ} {x : ℝ} (hx : IsSup A x) : ∀ y, y < x → ∃ a ∈ A, y < a := by 101 | sorry 102 | 103 | /- 104 | Let's do a variation on an example from file 07 that will be useful in the last 105 | exercise below. 106 | -/ 107 | -- 0069 108 | theorem le_of_le_add_all' {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 109 | sorry 110 | 111 | -- 0070 112 | example {x y : ℝ} {u : ℕ → ℝ} (hu : SeqLimit u x) (ineg : ∀ n, u n ≤ y) : x ≤ y := by 113 | sorry 114 | 115 | -------------------------------------------------------------------------------- /Tutorials/Exercises/09LimitsFinal.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | This is the final file in the series. Here we use everything covered 5 | in previous files to prove a couple of famous theorems from 6 | elementary real analysis. Of course they all have more general versions 7 | in mathlib. 8 | 9 | As usual, keep in mind the following: 10 | 11 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 12 | 13 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 14 | 15 | `le_max_left p q : p ≤ max p q` 16 | 17 | `le_max_right p q : q ≤ max p q` 18 | 19 | as well as a lemma from the previous file: 20 | 21 | `le_of_le_add_all : (∀ ε > 0, y ≤ x + ε) → y ≤ x` 22 | 23 | Let's start with a variation on a known exercise. 24 | -/ 25 | -- 0071 26 | theorem le_lim' {x y : ℝ} {u : ℕ → ℝ} (hu : SeqLimit u x) (ineg : ∃ N, ∀ n ≥ N, y ≤ u n) : y ≤ x := by 27 | sorry 28 | 29 | /- 30 | Let's now return to the result proved in the `00_` file of this series, 31 | and prove again the sequential characterization of upper bounds (with a slighly 32 | different proof). 33 | 34 | For this, and other exercises below, we'll need many things that we proved in previous files, 35 | and a couple of extras. 36 | 37 | From the 5th file: 38 | 39 | `limit_const (x : ℝ) : SeqLimit (λ n, x) x` 40 | 41 | 42 | `squeeze (lim_u : SeqLimit u l) (lim_w : SeqLimit w l) (hu : ∀ n, u n ≤ v n) (hw : ∀ n, v n ≤ w n) : SeqLimit v l` 43 | 44 | From the 8th: 45 | 46 | `def UpperBound (A : set ℝ) (x : ℝ) := ∀ a ∈ A, a ≤ x` 47 | 48 | `def IsSup (A : set ℝ) (x : ℝ) := UpperBound A x ∧ ∀ y, UpperBound A y → x ≤ y` 49 | 50 | `lt_sup (hx : is_sup A x) : ∀ y, y < x → ∃ a ∈ A, y < a ` 51 | 52 | You can also use: 53 | 54 | `Nat.one_div_pos_of_nat {n : ℕ} : 0 < 1 / (n + 1 : ℝ)` 55 | 56 | `inv_succ_le_all : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1/(n + 1 : ℝ) ≤ ε` 57 | 58 | and their easy consequences: 59 | 60 | `limit_of_sub_le_inv_succ (h : ∀ n, |u n - x| ≤ 1/(n+1)) : SeqLimit u x` 61 | 62 | `limit_const_add_inv_succ (x : ℝ) : SeqLimit (λ n, x + 1/(n+1)) x` 63 | 64 | `limit_const_sub_inv_succ (x : ℝ) : SeqLimit (λ n, x - 1/(n+1)) x` 65 | 66 | as well as: 67 | 68 | `lim_le (hu : SeqLimit u x) (ineg : ∀ n, u n ≤ y) : x ≤ y` 69 | 70 | The structure of the proof is offered. It features a new tactic: 71 | `choose` which invokes the axiom of choice (observing the tactic state before and 72 | after using it should be enough to understand everything). 73 | -/ 74 | -- 0072 75 | theorem isSup_iff (A : Set ℝ) (x : ℝ) : 76 | IsSup A x ↔ UpperBound A x ∧ ∃ u : ℕ → ℝ, SeqLimit u x ∧ ∀ n, u n ∈ A := by 77 | constructor 78 | · intro h 79 | constructor 80 | sorry 81 | · have : ∀ n : ℕ, ∃ a ∈ A, x - 1 / (n + 1) < a := by 82 | intro n 83 | have : 1 / (n + 1 : ℝ) > 0 := Nat.one_div_pos_of_nat 84 | sorry 85 | choose u hu using this 86 | sorry 87 | · rintro ⟨maj, u, limu, u_in⟩ 88 | sorry 89 | 90 | /-- Continuity of a function at a point -/ 91 | def ContinuousAtPt (f : ℝ → ℝ) (x₀ : ℝ) : Prop := 92 | ∀ ε > 0, ∃ δ > 0, ∀ x, |x - x₀| ≤ δ → |f x - f x₀| ≤ ε 93 | 94 | variable {f : ℝ → ℝ} {x₀ : ℝ} {u : ℕ → ℝ} 95 | 96 | -- 0073 97 | theorem seq_continuous_of_continuous (hf : ContinuousAtPt f x₀) (hu : SeqLimit u x₀) : 98 | SeqLimit (f ∘ u) (f x₀) := by 99 | sorry 100 | 101 | -- 0074 102 | example : (∀ u : ℕ → ℝ, SeqLimit u x₀ → SeqLimit (f ∘ u) (f x₀)) → ContinuousAtPt f x₀ := by 103 | sorry 104 | 105 | /- 106 | Recall from the 6th file: 107 | 108 | 109 | `def extraction (φ : ℕ → ℕ) := ∀ n m, n < m → φ n < φ m` 110 | 111 | `def ClusterPoint (u : ℕ → ℝ) (a : ℝ) := ∃ φ, extraction φ ∧ SeqLimit (u ∘ φ) a` 112 | 113 | 114 | `id_le_extraction : extraction φ → ∀ n, n ≤ φ n` 115 | 116 | and from the 8th file: 117 | 118 | `def TendstoInfinity (u : ℕ → ℝ) := ∀ A, ∃ N, ∀ n ≥ N, u n ≥ A` 119 | 120 | `not_seqLimit_of_tendstoInfinity : TendstoInfinity u → ∀ l, ¬ SeqLimit u l` 121 | -/ 122 | variable {φ : ℕ → ℕ} 123 | 124 | -- 0075 125 | theorem subseq_tendstoInfinity (h : TendstoInfinity u) (hφ : Extraction φ) : 126 | TendstoInfinity (u ∘ φ) := by 127 | sorry 128 | 129 | -- 0076 130 | theorem squeeze_infinity {u v : ℕ → ℝ} (hu : TendstoInfinity u) (huv : ∀ n, u n ≤ v n) : 131 | TendstoInfinity v := by 132 | sorry 133 | 134 | /- 135 | We will use segments: `Icc a b := { x | a ≤ x ∧ x ≤ b }` 136 | The notation stands for Interval-closed-closed. Variations exist with 137 | o or i instead of c, where o stands for open and i for infinity. 138 | 139 | We will use the following version of Bolzano-Weierstrass 140 | 141 | `bolzano_weierstrass (h : ∀ n, u n ∈ Icc a b) : ∃ c ∈ Icc a b, ClusterPoint u c` 142 | 143 | as well as the obvious 144 | 145 | `seqLimit_id : TendstoInfinity (λ n, n)` 146 | -/ 147 | open Set 148 | 149 | -- 0077 150 | theorem bdd_above_segment {f : ℝ → ℝ} {a b : ℝ} (hf : ∀ x ∈ Icc a b, ContinuousAtPt f x) : 151 | ∃ M, ∀ x ∈ Icc a b, f x ≤ M := by 152 | sorry 153 | 154 | /- 155 | In the next exercise, we can use: 156 | 157 | `abs_neg x : |-x| = |x|` 158 | -/ 159 | -- 0078 160 | theorem continuous_opposite {f : ℝ → ℝ} {x₀ : ℝ} (h : ContinuousAtPt f x₀) : 161 | ContinuousAtPt (fun x => -f x) x₀ := by 162 | sorry 163 | 164 | /- 165 | Now let's combine the two exercises above 166 | -/ 167 | -- 0079 168 | theorem bdd_below_segment {f : ℝ → ℝ} {a b : ℝ} (hf : ∀ x ∈ Icc a b, ContinuousAtPt f x) : 169 | ∃ m, ∀ x ∈ Icc a b, m ≤ f x := by 170 | sorry 171 | 172 | /- 173 | Remember from the 5th file: 174 | 175 | `unique_limit : SeqLimit u l → SeqLimit u l' → l = l'` 176 | 177 | and from the 6th one: 178 | 179 | `subseq_tendsto_of_tendsto (h : SeqLimit u l) (hφ : extraction φ) : SeqLimit (u ∘ φ) l` 180 | 181 | We now admit the following version of the least upper bound theorem 182 | (that cannot be proved without discussing the construction of real numbers 183 | or admitting another strong theorem). 184 | 185 | `sup_segment {a b : ℝ} {A : set ℝ} (hnonvide : ∃ x, x ∈ A) (h : A ⊆ Icc a b) :` 186 | `∃ x ∈ Icc a b, is_sup A x` 187 | 188 | In the next exercise, it can be useful to prove inclusions of sets of real number. 189 | By definition, `A ⊆ B` means : `∀ x, x ∈ A → x ∈ B`. 190 | Hence one can start a proof of `A ⊆ B` by `intro x x_in`, 191 | which brings `x : ℝ` and `x_in : x ∈ A` in the local context, 192 | and then prove `x ∈ B`. 193 | 194 | Note also the use of 195 | `{x | P x}` 196 | which denotes the set of `x` satisfying predicate `P`. 197 | 198 | Hence `x' ∈ { x | P x} ↔ P x'`, by definition. 199 | -/ 200 | -- 0080 201 | example {a b : ℝ} (hab : a ≤ b) (hf : ∀ x ∈ Icc a b, ContinuousAtPt f x) : 202 | ∃ x₀ ∈ Icc a b, ∀ x ∈ Icc a b, f x ≤ f x₀ := by 203 | sorry 204 | 205 | -- The following stupid lemma can be used below. 206 | lemma stupid {a b x : ℝ} (h : x ∈ Icc a b) (h' : x ≠ b) : x < b := 207 | lt_of_le_of_ne h.right h' 208 | 209 | /- 210 | And now the final boss... 211 | -/ 212 | def I := (Icc 0 1 : Set ℝ) -- the type ascription makes sure 0 and 1 are real numbers here 213 | 214 | -- 0081 215 | example (f : ℝ → ℝ) (hf : ∀ x, ContinuousAtPt f x) (h₀ : f 0 < 0) (h₁ : f 1 > 0) : 216 | ∃ x₀ ∈ I, f x₀ = 0 := by 217 | let A := { x | x ∈ I ∧ f x < 0 } 218 | have ex_x₀ : ∃ x₀ ∈ I, IsSup A x₀ := by 219 | sorry 220 | rcases ex_x₀ with ⟨x₀, x₀_in, x₀_sup⟩ 221 | use x₀, x₀_in 222 | have : f x₀ ≤ 0 := by 223 | sorry 224 | have x₀_1 : x₀ < 1 := by 225 | sorry 226 | have : f x₀ ≥ 0 := by 227 | have in_I : ∃ N : ℕ, ∀ n ≥ N, x₀ + 1 / (n + 1) ∈ I := by 228 | have : ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) ≤ 1 - x₀ := by 229 | sorry 230 | sorry 231 | have not_in : ∀ n : ℕ, x₀ + 1 / (n + 1) ∉ A := by 232 | -- By definition, x ∉ A means ¬ (x ∈ A). 233 | sorry 234 | dsimp at not_in 235 | sorry 236 | linarith 237 | 238 | -------------------------------------------------------------------------------- /Tutorials/Solutions/00FirstProofs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | This file is intended for Lean beginners. The goal is to demonstrate what it feels like to prove 3 | things using Lean and mathlib. Complicated definitions and theory building are not covered. 4 | Everything is covered again more slowly and with exercises in the next files. 5 | -/ 6 | -- We want real numbers and their basic properties 7 | import Mathlib.Data.Real.Basic 8 | -- We want to be able to use Lean's built-in "help" functionality 9 | import Mathlib.Tactic.LibrarySearch 10 | 11 | 12 | -- We want to be able to define functions using the law of excluded middle 13 | noncomputable section 14 | 15 | 16 | /- 17 | Our first goal is to define the set of upper bounds of a set of real numbers. 18 | This is already defined in mathlib (in a more general context), but we repeat 19 | it for the sake of exposition. Right-click "upperBounds" below to get offered 20 | to jump to mathlib's version 21 | -/ 22 | #check upperBounds 23 | 24 | /-- The set of upper bounds of a set of real numbers ℝ -/ 25 | def upBounds (A : Set ℝ) := 26 | { x : ℝ | ∀ a ∈ A, a ≤ x } 27 | 28 | /-- Predicate `is_maximum a A` means `a` is a maximum of `A` -/ 29 | def IsMaximum (a : ℝ) (A : Set ℝ) := 30 | a ∈ A ∧ a ∈ upBounds A 31 | 32 | /- 33 | In the above definition, the symbol `∧` means "and". We also see the most 34 | visible difference between set theoretic foundations and type theoretic ones 35 | (used by almost all proof assistants). In set theory, everything is a set, and the 36 | only relation you get from foundations are `=` and `∈`. In type theory, there is 37 | a meta-theoretic relation of "typing": `a : ℝ` reads "`a` is a real number" or, 38 | more precisely, "the type of `a` is `ℝ`". Here "meta-theoretic" means this is not a 39 | statement you can prove or disprove inside the theory, it's a fact that is true or 40 | not. Here we impose this fact, in other circumstances, it would be checked by the 41 | Lean kernel. 42 | By contrast, `a ∈ A` is a statement inside the theory. Here it's part of the 43 | definition, in other circumstances it could be something proven inside Lean. 44 | -/ 45 | /- For illustrative purposes, we now define an infix version of the above predicate. 46 | It will allow us to write `a is_a_max_of A`, which is closer to a sentence. 47 | -/ 48 | infixl:55 " is_a_max_of " => IsMaximum 49 | 50 | /- 51 | Let's prove something now! A set of real numbers has at most one maximum. Here 52 | everything left of the final `:` is introducing the objects and assumption. The equality 53 | `x = y` right of the colon is the conclusion. 54 | -/ 55 | theorem unique_max (A : Set ℝ) (x y : ℝ) (hx : x is_a_max_of A) (hy : y is_a_max_of A) : x = y := by 56 | -- We first break our assumptions in their two constituent pieces. 57 | -- We are free to choose the name following `with` 58 | rcases hx with ⟨x_in, x_up⟩ 59 | rcases hy with ⟨y_in, y_up⟩ 60 | -- Assumption `x_up` means x isn't less than elements of A, let's apply this to y 61 | specialize x_up y 62 | -- Assumption `x_up` now needs the information that `y` is indeed in `A`. 63 | specialize x_up y_in 64 | -- Let's do this quicker with roles swapped 65 | specialize y_up x x_in 66 | -- We explained to Lean the idea of this proof. 67 | -- Now we know `x ≤ y` and `y ≤ x`, and Lean shouldn't need more help. 68 | -- `linarith` proves equalities and inequalities that follow linearly from 69 | -- the assumption we have. 70 | linarith 71 | 72 | /- 73 | The above proof is too long, even if you remove comments. We don't really need the 74 | unpacking steps at the beginning; we can access both parts of the assumption 75 | `hx : x is_a_max_of A` using shortcuts `hx.1` and `hx.2`. We can also improve 76 | readability without assistance from the tactic state display, clearly announcing 77 | intermediate goals using `have`. This way we get to the following version of the 78 | same proof. 79 | -/ 80 | example (A : Set ℝ) (x y : ℝ) (hx : x is_a_max_of A) (hy : y is_a_max_of A) : x = y := by 81 | have : x ≤ y := hy.2 x hx.1 82 | have : y ≤ x := hx.2 y hy.1 83 | linarith 84 | 85 | /- 86 | Notice how mathematics based on type theory treats the assumption 87 | `∀ a ∈ A, a ≤ y` as a function turning an element `a` of `A` into the statement 88 | `a ≤ y`. More precisely, this assumption is the abbreviation of 89 | `∀ a : ℝ, a ∈ A → a ≤ y`. The expression `hy.2 x` appearing in the above proof 90 | is then the statement `x ∈ A → x ≤ y`, which itself is a function turning a 91 | statement `x ∈ A` into `x ≤ y` so that the full expression `hy.2 x hx.1` is 92 | indeed a proof of `x ≤ y`. 93 | 94 | One could argue a three-line-long proof of this lemma is still two lines too long. 95 | This is debatable, but mathlib's style is to write very short proofs for trivial 96 | lemmas. Those proofs are not easy to read but they are meant to indicate that the 97 | proof is probably not worth reading. 98 | 99 | In order to reach this stage, we need to know what `linarith` did for us. It invoked 100 | the lemma `le_antisymm` which says: `x ≤ y → y ≤ x → x = y`. This arrow, which 101 | is used both for function and implication, is right associative. So the statement is 102 | `x ≤ y → (y ≤ x → x = y)` which reads: I will send a proof `p` of `x ≤ y` to a function 103 | sending a proof `q'` of `y ≤ x` to a proof of `x = y`. Hence `le_antisymm p q'` is a 104 | proof of `x = y`. 105 | 106 | Using this we can get our one-line proof: 107 | -/ 108 | example (A : Set ℝ) (x y : ℝ) (hx : x is_a_max_of A) (hy : y is_a_max_of A) : x = y := 109 | le_antisymm (hy.2 x hx.1) (hx.2 y hy.1) 110 | 111 | /- 112 | Such a proof is called a proof term (or a "term mode" proof). Notice it has no `by`. 113 | It is directly the kind of low level proof that the Lean kernel is 114 | consuming. Commands like `rcases`, `specialize` or `linarith` are called tactics, they 115 | help users constructing proof terms that could be very tedious to write directly. 116 | The most efficient proof style combines tactics with proof terms like our previous 117 | `have : x ≤ y := hy.2 x hx.1` where `hy.2 x hx.1` is a proof term embeded inside 118 | a tactic mode proof. 119 | 120 | In the remaining of this file, we'll be characterizing infima of sets of real numbers 121 | in term of sequences. 122 | -/ 123 | /-- The set of lower bounds of a set of real numbers ℝ -/ 124 | def lowBounds (A : Set ℝ) := 125 | { x : ℝ | ∀ a ∈ A, x ≤ a } 126 | 127 | /- 128 | We now define `a` is an infimum of `A`. Again there is already a more general version 129 | in mathlib. 130 | -/ 131 | def IsInf (x : ℝ) (A : Set ℝ) := 132 | x is_a_max_of lowBounds A 133 | 134 | -- Let's define it also as an infix operator 135 | infixl:55 " is_an_inf_of " => IsInf 136 | 137 | /- 138 | We need to prove that any number which is greater than the infimum of A is greater 139 | than some element of A. 140 | -/ 141 | theorem inf_lt {A : Set ℝ} {x : ℝ} (hx : x is_an_inf_of A) : ∀ y, x < y → ∃ a ∈ A, a < y := by 142 | -- Let `y` be any real number. 143 | intro y 144 | -- Let's prove the contrapositive 145 | contrapose 146 | -- The symbol `¬` means negation. Let's ask Lean to rewrite the goal without negation, 147 | -- pushing negation through quantifiers and inequalities 148 | push_neg 149 | -- Let's assume the premise, calling the assumption `h` 150 | intro h 151 | -- `h` is exactly saying `y` is a lower bound of `A` so the second part of 152 | -- the infimum assumption `hx` applied to `y` and `h` is exactly what we want. 153 | exact hx.2 y h 154 | 155 | /- 156 | In the above proof, the sequence `contrapose, push_neg` is so common that it can be 157 | abbreviated to `contrapose!`. With these commands, we enter the gray zone between 158 | proof checking and proof finding. Practical computer proof checking crucially needs 159 | the computer to handle tedious proof steps. In the next proof, we'll start using 160 | `linarith` a bit more seriously, going one step further into automation. 161 | 162 | Our next real goal is to prove inequalities for limits of sequences. We extract the 163 | following lemma: if `y ≤ x + ε` for all positive `ε` then `y ≤ x`. 164 | -/ 165 | theorem le_of_le_add_eps {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 166 | -- Let's prove the contrapositive, asking Lean to push negations right away. 167 | contrapose! 168 | -- Assume `h : x < y`. 169 | intro h 170 | -- We need to find `ε` such that `ε` is positive and `x + ε < y`. 171 | -- Let's use `(y-x)/2` 172 | use (y - x) / 2 173 | -- we now have two properties to prove. Let's do both in turn, using `linarith` 174 | constructor 175 | linarith 176 | linarith 177 | 178 | /- 179 | Note how `linarith` was used for both sub-goals at the end of the above proof. 180 | We could have shortened that using the semi-colon combinator instead of comma, 181 | writing `constructor <;> linarith`. 182 | 183 | Next we will study a compressed version of that proof: 184 | -/ 185 | example {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 186 | contrapose! 187 | exact fun h => ⟨(y - x) / 2, by linarith, by linarith⟩ 188 | 189 | /- 190 | The angle brackets `⟨` and `⟩` introduce compound data or proofs. A proof 191 | of a `∃ z, P z` statemement is composed of a witness `z₀` and a proof `h` of 192 | `P z₀`. The compound is denoted by `⟨z₀, h⟩`. In the example above, the predicate is 193 | itself compound, it is a conjunction `P z ∧ Q z`. So the proof term should read 194 | `⟨z₀, ⟨h₁, h₂⟩⟩` where `h₁` (resp. `h₂`) is a proof of `P z₀` (resp. `Q z₀`). 195 | But these so-called "anonymous constructor" brackets are right-associative, so we can 196 | get rid of the nested brackets. 197 | 198 | Note also how we can use `by` to enter tactics anywhere a term is expected. 199 | 200 | Going all the way to a proof term would make the proof much longer, because we 201 | crucially use automation with `contrapose!` and `linarith`. 202 | -/ 203 | 204 | /- 205 | One could argue that the above proof is a bit too terse, and we are relying too much 206 | on linarith. Let's have more `linarith` calls for smaller steps. For the sake 207 | of (tiny) variation, we will also assume the premise and argue by contradiction 208 | instead of contraposing. 209 | -/ 210 | example {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 211 | intro h 212 | -- Assume the conclusion is false, and call this assumption H. 213 | by_contra H 214 | push_neg at H 215 | -- Now let's compute. 216 | have key := calc 217 | -- Each line must end with `:=` followed by a proof term 218 | -- We want to specialize our assumption `h` to `ε = (y-x)/2` but this is long to 219 | -- type, so let's put a hole `_` that Lean will fill in by comparing the 220 | -- statement we want to prove and our proof term with a hole. As usual, 221 | -- positivity of `(y-x)/2` is proved by `linarith` 222 | y ≤ x + (y - x) / 2 := h _ (by linarith) 223 | _ = x / 2 + y / 2 := by ring 224 | _ < y := by linarith 225 | 226 | -- our key now says `y < y` (notice how the sequence `≤`, `=`, `<` was correctly 227 | -- merged into a `<`). Let `linarith` find the desired contradiction now. 228 | linarith 229 | -- alternatively, we could have provided the proof term 230 | -- `exact lt_irrefl y key` 231 | 232 | /- 233 | Now we are ready for some analysis. Let's define convergence of sequences of real numbers 234 | (of course there is a much more general definition in mathlib). 235 | -/ 236 | 237 | /-- The sequence `u` tends to `l` -/ 238 | def Limit (u : ℕ → ℝ) (l : ℝ) := 239 | ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε 240 | 241 | /- 242 | In the above definition, `u n` denotes the n-th term of the sequence. We can 243 | add parentheses to get `u (n)` but we try to avoid parentheses because they pile up 244 | very quickly (and note the space between `u` and `(` is required). 245 | -/ 246 | 247 | -- If y ≤ u n for all n and u n goes to x then y ≤ x 248 | theorem le_lim {x y : ℝ} {u : ℕ → ℝ} (hu : Limit u x) (ineq : ∀ n, y ≤ u n) : y ≤ x := by 249 | -- Let's apply our previous lemma 250 | apply le_of_le_add_eps 251 | -- We need to prove y ≤ x + ε for all positive ε. 252 | -- Let ε be any positive real 253 | intro ε ε_pos 254 | -- we now specialize our limit assumption to this `ε`, and immediately 255 | -- fix a `N` as promised by the definition. 256 | rcases hu ε ε_pos with ⟨N, HN⟩ 257 | -- Now we only need to compute until reaching the conclusion 258 | calc 259 | y ≤ u N := ineq N 260 | _ = x + (u N - x) := by linarith 261 | -- In the next step we use the `gcongr` tactic which uses "generalized congruence" lemmas 262 | -- to zoom on the relevant part of the inequality goal, in this case `u N - x ≤ |u N - x|`. 263 | -- We then need a lemma saying `z ≤ |z|`. Because we don't know the name of this lemma, 264 | -- let's use `exact?`. Because searching through the library can be slow, 265 | -- Lean will write what it found in the Lean message window when cursor is on 266 | -- that line, so that we can replace it by the lemma. We see `le_abs_self`, which 267 | -- says `a ≤ |a|`, exactly what we're looking for. 268 | _ ≤ x + |u N - x| := by gcongr ; exact? 269 | _ ≤ x + ε := by gcongr ; apply HN; linarith 270 | 271 | /- 272 | The next lemma has been extracted from the main proof in order to discuss numbers. 273 | In ordinary maths, we know that ℕ is *not* contained in `ℝ`, whatever the 274 | construction of real numbers that we use. For instance a natural number is not 275 | an equivalence class of Cauchy sequences. But it's very easy to 276 | pretend otherwise. Formal maths requires slightly more care. In the statement below, 277 | the "type ascription" `(n + 1 : ℝ)` forces Lean to convert the natural number 278 | `n+1` into a real number. The "inclusion" map will be displayed in tactic state 279 | as `↑`. There are various lemmas asserting this map is compatible with addition and 280 | monotone, but we don't want to bother writing their names. The `norm_cast` 281 | tactic is designed to wisely apply those lemmas for us. 282 | -/ 283 | theorem inv_succ_pos : ∀ n : ℕ, 1 / (n + 1 : ℝ) > 0 := by 284 | -- Let `n` be any integer 285 | intro n 286 | -- Since we don't know the name of the relevant lemma, asserting that the inverse of 287 | -- a positive number is positive, let's state that is suffices 288 | -- to prove that `n+1`, seen as a real number, is positive, and ask `exact?` 289 | suffices (n + 1 : ℝ) > 0 by exact? 290 | -- Now we want to reduce to a statement about natural numbers, not real numbers 291 | -- coming from natural numbers. 292 | norm_cast 293 | -- and then get the usual help from `linarith` 294 | linarith 295 | 296 | /- 297 | That was a pretty long proof for an obvious fact. And stating it as a lemma feels 298 | stupid, so let's find a way to write it on one line in case we want to include it 299 | in some other proof without stating a lemma. First the `exact?` call 300 | above displays the name of the relevant lemma: `one_div_pos`. We can also 301 | replace the `linarith` call on the last line by `exact?` to learn the name 302 | of the lemma `Nat.succ_pos` asserting that the successor of a natural number is 303 | positive. There is also a variant on `norm_cast` that combines it with `exact`. 304 | The term mode analogue of `intro` is `fun`. We get down to: 305 | -/ 306 | example : ∀ n : ℕ, 1 / (n + 1 : ℝ) > 0 := 307 | fun n ↦ one_div_pos.mpr (by exact_mod_cast Nat.succ_pos n) 308 | 309 | /- 310 | The next proof uses mostly known things, so we will commment only new aspects. 311 | -/ 312 | theorem limit_inv_succ : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) ≤ ε := by 313 | intro ε ε_pos 314 | suffices ∃ N : ℕ, 1 / ε ≤ N by 315 | -- Because we didn't provide a name for the above statement, Lean called it `this`. 316 | -- Let's fix an `N` that works. 317 | rcases this with ⟨N, HN⟩ 318 | use N 319 | intro n Hn 320 | -- Now we want to rewrite the goal using lemmas 321 | -- `div_le_iff' : 0 < b → (a / b ≤ c ↔ a ≤ b * c)` 322 | -- `div_le_iff : 0 < b → (a / b ≤ c ↔ a ≤ c * b)` 323 | -- the second one will be rewritten from right to left, as indicated by `←`. 324 | -- Lean will create a side goal for the required positivity assumption that 325 | -- we don't provide for `div_le_iff'`. 326 | rw [div_le_iff', ← div_le_iff ε_pos] 327 | -- We want to replace assumption `Hn` by its real counter-part so that 328 | -- linarith can find what it needs. 329 | replace Hn : (N : ℝ) ≤ n 330 | exact_mod_cast Hn 331 | linarith 332 | -- we are still left with the positivity assumption. We already discussed 333 | -- how to prove it in the preceding lemma, but we could alternatively use 334 | -- the `positivity` tactic whose job is to prove obvious positivity statements. 335 | positivity 336 | -- Now we need to prove that sufficient statement. 337 | -- We want to use that `ℝ` is archimedean. So we start typing 338 | -- `exact archimedean_` and hit Ctrl-space to see what completion Lean proposes 339 | -- the lemma `archimedean_iff_nat_le` sounds promising. We select the left to 340 | -- right implication using `.1`. This a generic lemma for fields equiped with 341 | -- a linear (ie total) order. We need to provide a proof that `ℝ` is indeed 342 | -- archimedean. This is done using the `infer_instance` tactic that will be 343 | -- covered elsewhere. 344 | exact archimedean_iff_nat_le.1 (by infer_instance) (1 / ε) 345 | 346 | /- 347 | We can now put all pieces together, with almost no new things to explain. 348 | -/ 349 | theorem inf_seq (A : Set ℝ) (x : ℝ) : 350 | x is_an_inf_of A ↔ x ∈ lowBounds A ∧ ∃ u : ℕ → ℝ, Limit u x ∧ ∀ n, u n ∈ A := by 351 | constructor 352 | · intro h 353 | constructor 354 | · exact h.1 355 | -- On the next line, we don't need to tell Lean to treat `n+1` as a real number because 356 | -- we add `x` to it, so Lean knows there is only one way to make sense of this expression. 357 | have key : ∀ n : ℕ, ∃ a ∈ A, a < x + 1 / (n + 1) := by 358 | intro n 359 | -- we can use the lemma we proved above 360 | apply inf_lt h 361 | -- and another one we proved! 362 | have : 0 < 1 / (n + 1 : ℝ) := inv_succ_pos n 363 | linarith 364 | -- Now we need to use axiom of (countable) choice 365 | choose u hu using key 366 | use u 367 | constructor 368 | · intro ε ε_pos 369 | -- again we use a lemma we proved, specializing it to our fixed `ε`, and fixing a `N` 370 | rcases limit_inv_succ ε ε_pos with ⟨N, H⟩ 371 | use N 372 | intro n hn 373 | have : x ≤ u n := h.1 _ (hu n).1 374 | have := 375 | calc 376 | u n < x + 1 / (n + 1) := (hu n).2 377 | _ ≤ x + ε := add_le_add (le_refl x) (H n hn) 378 | 379 | rw [abs_of_nonneg] <;> linarith 380 | · intro n 381 | exact (hu n).1 382 | · intro h 383 | -- Assumption `h` is made of nested compound statements. We can use 384 | -- `rcases` to unpack it in one go. 385 | rcases h with ⟨x_min, u, lim, huA⟩ 386 | constructor 387 | exact x_min 388 | intro y y_mino 389 | apply le_lim lim 390 | intro n 391 | exact y_mino (u n) (huA n) 392 | -------------------------------------------------------------------------------- /Tutorials/Solutions/01EqualityRewriting.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | /- 4 | One of the earliest kind of proofs one encounters while learning mathematics is proving by 5 | a calculation. It may not sound like a proof, but this is actually using lemmas expressing 6 | properties of operations on numbers. It also uses the fundamental property of equality: if two 7 | mathematical objects A and B are equal then, in any statement involving A, one can replace A 8 | by B. This operation is called rewriting, and the Lean "tactic" for this is `rw`. 9 | 10 | In the following exercises, we will use the following two lemmas: 11 | `mul_assoc a b c : a * b * c = a * (b * c)` 12 | `mul_comm a b : a*b = b*a` 13 | 14 | Hence the command 15 | `rw [mul_assoc a b c]`, 16 | will replace `a*b*c` by `a*(b*c)` in the current goal. 17 | 18 | In order to replace backward, we use 19 | `rw [← mul_assoc a b c]`, 20 | replacing `a*(b*c)` by `a*b*c` in the current goal. 21 | 22 | Of course we don't want to constantly invoke those lemmas, and we will eventually introduce 23 | more powerful solutions. 24 | -/ 25 | 26 | example (a b c : ℝ) : a * b * c = b * (a * c) := by 27 | rw [mul_comm a b] 28 | rw [mul_assoc b a c] 29 | 30 | -- 0001 31 | example (a b c : ℝ) : c * b * a = b * (a * c) := by 32 | -- sorry 33 | rw [mul_comm c b] 34 | rw [mul_assoc b c a] 35 | rw [mul_comm c a] 36 | -- sorry 37 | 38 | -- 0002 39 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 40 | -- sorry 41 | rw [← mul_assoc a b c] 42 | rw [mul_comm a b] 43 | rw [mul_assoc b a c] 44 | -- sorry 45 | 46 | /- 47 | Now let's return to the preceding example to experiment with what happens 48 | if we don't give arguments to mul_assoc or mul_comm. 49 | For instance, you can start the next proof with 50 | `rw [← mul_assoc]`, 51 | Try to figure out what happens. 52 | -/ 53 | -- 0003 54 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 55 | -- sorry 56 | rw [← mul_assoc] 57 | -- "rw mul_comm," doesn't do what we want. 58 | rw [mul_comm a b] 59 | rw [mul_assoc] 60 | -- sorry 61 | 62 | /- 63 | We can also perform rewriting in an assumption of the local context, using for instance 64 | `rw [mul_comm a b] at hyp`, 65 | in order to replace a*b by b*a in assumption hyp. 66 | 67 | The next example will use a third lemma: 68 | `two_mul a : 2*a = a + a` 69 | 70 | Also we use the `exact` tactic, which allows to provide a direct proof term. 71 | -/ 72 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 73 | rw [hyp'] at hyp 74 | rw [mul_comm d a] at hyp 75 | rw [← two_mul (a * d)] at hyp 76 | rw [← mul_assoc 2 a d] at hyp 77 | exact hyp 78 | 79 | -- Our assumption hyp is now exactly what we have to prove 80 | /- 81 | And the next one can use: 82 | `sub_self x : x - x = 0` 83 | -/ 84 | -- 0004 85 | example (a b c d : ℝ) (hyp : c = b * a - d) (hyp' : d = a * b) : c = 0 := by 86 | -- sorry 87 | rw [hyp'] at hyp 88 | rw [mul_comm b a] at hyp 89 | rw [sub_self (a * b)] at hyp 90 | exact hyp 91 | -- sorry 92 | 93 | /- 94 | What is written in the two preceding example is very far away from what we would write on 95 | paper. Let's now see how to get a more natural layout. 96 | Inside each pair of curly braces below, the goal is to prove equality with the preceding line. 97 | -/ 98 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 99 | calc 100 | c = d * a + b := by rw [hyp] 101 | _ = d * a + a * d := by rw [hyp'] 102 | _ = a * d + a * d := by rw [mul_comm d a] 103 | _ = 2 * (a * d) := by rw [two_mul] 104 | _ = 2 * a * d := by rw [mul_assoc] 105 | 106 | 107 | /- 108 | From a practical point of view, when writing such a proof, it is convenient to: 109 | * pause the tactic state view update in VScode by clicking the Pause icon button 110 | in the top right corner of the Lean Goal buffer 111 | * write the full calculation, ending each line with ":= by {}" 112 | * resume tactic state update by clicking the Play icon button and fill in proofs between 113 | curly braces. 114 | 115 | Let's return to the other example using this method. 116 | -/ 117 | -- 0005 118 | example (a b c d : ℝ) (hyp : c = b * a - d) (hyp' : d = a * b) : c = 0 := by 119 | -- sorry 120 | calc 121 | c = b * a - d := by rw [hyp] 122 | _ = b * a - a * b := by rw [hyp'] 123 | _ = a * b - a * b := by rw [mul_comm a b] 124 | _ = 0 := by rw [sub_self (a * b)] 125 | -- sorry 126 | 127 | /- 128 | The preceding proofs have exhausted our supply of "mul_comm" patience. Now it's time 129 | to get the computer to work harder. The `ring` tactic will prove any goal that follows by 130 | applying only the axioms of commutative (semi-)rings, in particular commutativity and 131 | associativity of addition and multiplication, as well as distributivity. 132 | -/ 133 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 134 | calc 135 | c = d * a + b := by rw [hyp] 136 | _ = d * a + a * d := by rw [hyp'] 137 | _ = 2 * a * d := by ring 138 | 139 | 140 | /- 141 | Of course we can use `ring` outside of `calc`. Let's do the next one in one line. 142 | -/ 143 | -- 0006 144 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 145 | -- sorry 146 | ring 147 | -- sorry 148 | 149 | /- 150 | This is too much fun. Let's do it again. 151 | -/ 152 | -- 0007 153 | example (a b : ℝ) : a + b + a = 2 * a + b := by 154 | -- sorry 155 | ring 156 | -- sorry 157 | 158 | /- 159 | Maybe this is cheating. Let's try to do the next computation without ring. 160 | We could use: 161 | `pow_two x : x^2 = x*x` 162 | `mul_sub a b c : a*(b-c) = a*b - a*c` 163 | `add_mul a b c : (a+b)*c = a*c + b*c` 164 | `add_sub a b c : a + (b - c) = (a + b) - c` 165 | `sub_sub a b c : a - b - c = a - (b + c)` 166 | `add_zero a : a + 0 = a` 167 | -/ 168 | -- 0008 169 | example (a b : ℝ) : (a + b) * (a - b) = a ^ 2 - b ^ 2 := by 170 | -- sorry 171 | rw [pow_two a] 172 | rw [pow_two b] 173 | rw [mul_sub (a + b) a b] 174 | rw [add_mul a b a] 175 | rw [add_mul a b b] 176 | rw [mul_comm b a] 177 | rw [← sub_sub] 178 | rw [← add_sub] 179 | rw [sub_self] 180 | rw [add_zero] 181 | -- sorry 182 | 183 | -- Let's stick to ring in the end. 184 | -------------------------------------------------------------------------------- /Tutorials/Solutions/02IffIfAnd.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | /- 4 | In the previous file, we saw how to rewrite using equalities. 5 | The analogue operation with mathematical statements is rewriting using 6 | equivalences. This is also done using the `rw` tactic. 7 | Lean uses `↔` to denote equivalence instead of `⇔` (increase font size if you don't see a difference). 8 | 9 | In the following exercises we will use the lemma: 10 | 11 | `sub_nonneg {x y : ℝ} : 0 ≤ y - x ↔ x ≤ y` 12 | 13 | The curly braces around x and y instead of parentheses mean Lean will always try to figure out what 14 | x and y are from context, unless we really insist on telling it (we'll see how to insist much later). 15 | Let's not worry about that for now. 16 | 17 | In order to announce an intermediate statement we use: 18 | 19 | `have my_name : my statement := by` 20 | 21 | This triggers the apparition of a new goal: proving the statement. After this is done, 22 | the statement becomes available under the name `my_name`. 23 | -/ 24 | 25 | example {a b c : ℝ} (hab : a ≤ b) : c + a ≤ c + b := by 26 | rw [← sub_nonneg] 27 | have key : c + b - (c + a) = b - a := by-- Here we introduce an intermediate statement named key 28 | -- and prove it in an idented code block (or on the same line if the proof is very short) 29 | ring 30 | -- we can now use the key statement 31 | rw [key] 32 | rw [sub_nonneg] 33 | exact hab 34 | 35 | /- 36 | Of course the previous lemma is already in the core library, named `add_le_add_left`, so we can use it below. 37 | 38 | Let's prove a variation (without invoking commutativity of addition to reduce to the previous statement 39 | since this would spoil our fun). 40 | -/ 41 | -- 0009 42 | example {a b : ℝ} (hab : a ≤ b) (c : ℝ) : a + c ≤ b + c := by 43 | -- sorry 44 | have key : b + c - (a + c) = b - a := by ring 45 | rw [← sub_nonneg] 46 | rw [key] 47 | rw [sub_nonneg] 48 | exact hab 49 | -- sorry 50 | 51 | /- 52 | Let's see how we could use this lemma. It is already in the core library, under the name `add_le_add_right`: 53 | 54 | `add_le_add_right {a b : ℝ} (hab : a ≤ b) (c : ℝ) : a + c ≤ b + c` 55 | 56 | This can be read as: "add_le_add_right is a function that will take as input real numbers a and b, an 57 | assumption `hab` claiming a ≤ b and a real number c, and will output a proof of a + c ≤ b + c". 58 | 59 | In addition, recall that curly braces around `a b` mean Lean will figure out those arguments unless we 60 | insist to help. This is because they can be deduced from the next argument `hab`. 61 | So it will be sufficient to feed `hab` and `c` to this function. 62 | -/ 63 | example {a b : ℝ} (ha : 0 ≤ a) : b ≤ a + b := by 64 | calc 65 | b = 0 + b := by ring 66 | _ ≤ a + b := by exact add_le_add_right ha b 67 | 68 | 69 | /- 70 | In the second line of the above proof, we need to prove `0 + b ≤ a + b`. 71 | The proof after the colon says: this is exactly lemma `add_le_add_right` applied to ha and b. 72 | Actually the `calc` block expects proof terms, and the `by` keyword is used to tell Lean we will use tactics 73 | to build such a proof term. But since the only tactic used in this block is `exact`, we can skip 74 | tactics entirely, and write: 75 | -/ 76 | example (a b : ℝ) (ha : 0 ≤ a) : b ≤ a + b := by 77 | calc 78 | b = 0 + b := by ring 79 | _ ≤ a + b := add_le_add_right ha b 80 | 81 | 82 | -- Let's do a variant. 83 | -- 0010 84 | example (a b : ℝ) (hb : 0 ≤ b) : a ≤ a + b := by 85 | -- sorry 86 | calc 87 | a = a + 0 := by ring 88 | _ ≤ a + b := add_le_add_left hb a 89 | -- sorry 90 | 91 | /- 92 | The two preceding examples are in the core library : 93 | 94 | `le_add_of_nonneg_left {a b : ℝ} (ha : 0 ≤ a) : b ≤ a + b` 95 | `le_add_of_nonneg_right {a b : ℝ} (hb : 0 ≤ b) : a ≤ a + b` 96 | 97 | Again, there won't be any need to memorize those names, we will 98 | soon see how to get rid of such goals automatically. 99 | But we can already try to understand how their names are built: 100 | "le_add" describe the conclusion "less or equal than some addition" 101 | It comes first because we are focussed on proving stuff, and 102 | auto-completion works by looking at the beginning of words. 103 | "of" introduces assumptions. "nonneg" is Lean's abbreviation for non-negative. 104 | "left" or "right" disambiguates between the two variations. 105 | 106 | Let's use those lemmas by hand for now. 107 | 108 | Note that you can have several inequalities steps in a `calc` block, 109 | transitivity of inequalities will be used automatically to assemble 110 | the pieces. 111 | -/ 112 | -- 0011 113 | example (a b : ℝ) (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := by 114 | -- sorry 115 | calc 116 | 0 ≤ a := ha 117 | _ ≤ a + b := le_add_of_nonneg_right hb 118 | -- sorry 119 | 120 | -- And let's combine with our earlier lemmas. 121 | -- 0012 122 | example (a b c d : ℝ) (hab : a ≤ b) (hcd : c ≤ d) : a + c ≤ b + d := by 123 | -- sorry 124 | calc 125 | a + c ≤ b + c := add_le_add_right hab c 126 | _ ≤ b + d := add_le_add_left hcd b 127 | -- sorry 128 | 129 | /- 130 | In the above examples, we prepared proofs of assumptions of our lemmas beforehand, so 131 | that we could feed them to the lemmas. This is called forward reasoning. 132 | The `calc` proofs also belong to this category. 133 | 134 | We can also announce the use of a lemma, and provide proofs after the fact, using 135 | the `apply` tactic. This is called backward reasoning because we get the conclusion 136 | first, and provide proofs later. Using `rw` on the goal (rather than on an assumption 137 | from the local context) is also backward reasoning. 138 | 139 | Let's do that using the lemma 140 | 141 | `mul_nonneg {x y : ℝ} (hx : 0 ≤ x) (hy : 0 ≤ y) : 0 ≤ x*y` 142 | -/ 143 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 144 | rw [← sub_nonneg] 145 | have key : b * c - a * c = (b - a) * c := by ring 146 | rw [key] 147 | apply mul_nonneg 148 | -- Here we don't provide proofs for the lemma's assumptions 149 | -- Now we need to provide the proofs. 150 | -- There are now two things to prove. We use the center dot (typed using `\.`) to 151 | -- focus on the current first goal. 152 | · rw [sub_nonneg] 153 | exact hab 154 | · exact hc 155 | 156 | /- 157 | Let's prove the same statement using only forward reasoning: announcing stuff, 158 | proving it by working with known facts, moving forward. 159 | -/ 160 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 161 | have hab' : 0 ≤ b - a := by 162 | rw [← sub_nonneg] at hab 163 | exact hab 164 | have h₁ : 0 ≤ (b - a) * c := mul_nonneg hab' hc 165 | have h₂ : (b - a) * c = b * c - a * c := by ring 166 | have h₃ : 0 ≤ b * c - a * c := by 167 | rw [h₂] at h₁ 168 | exact h₁ 169 | rw [sub_nonneg] at h₃ 170 | exact h₃ 171 | 172 | /- 173 | One reason why the backward reasoning proof is shorter is because Lean can 174 | infer of lot of things by comparing the goal and the lemma statement. Indeed 175 | in the `apply mul_nonneg` line, we didn't need to tell Lean that `x = b - a` 176 | and `y = c` in the lemma. It was infered by "unification" between the lemma 177 | statement and the goal. 178 | 179 | To be fair to the forward reasoning version, we should introduce a convenient 180 | variation on `rw`. The `rwa` tactic performs rewrite and then looks for an 181 | assumption matching the goal. We can use it to rewrite our latest proof as: 182 | -/ 183 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 184 | have hab' : 0 ≤ b - a := by rwa [← sub_nonneg] at hab 185 | have h₁ : 0 ≤ (b - a) * c := mul_nonneg hab' hc 186 | have h₂ : (b - a) * c = b * c - a * c := by ring 187 | have h₃ : 0 ≤ b * c - a * c := by rwa [h₂] at h₁ 188 | rwa [sub_nonneg] at h₃ 189 | 190 | /- 191 | Let's now combine forward and backward reasoning, to get our most 192 | efficient proof of this statement. Note in particular how unification is used 193 | to know what to prove inside the parentheses in the `mul_nonneg` arguments. 194 | -/ 195 | example (a b c : ℝ) (hc : 0 ≤ c) (hab : a ≤ b) : a * c ≤ b * c := by 196 | rw [← sub_nonneg] 197 | calc 198 | 0 ≤ (b - a) * c := mul_nonneg (by rwa [sub_nonneg]) hc 199 | _ = b * c - a * c := by ring 200 | 201 | 202 | /- 203 | Let's now practice all three styles using: 204 | 205 | `mul_nonneg_of_nonpos_of_nonpos {a b : α} (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b` 206 | 207 | `sub_nonpos {a b : α} : a - b ≤ 0 ↔ a ≤ b` 208 | -/ 209 | -- First using mostly backward reasoning 210 | -- 0013 211 | example (a b c : ℝ) (hc : c ≤ 0) (hab : a ≤ b) : b * c ≤ a * c := by 212 | -- sorry 213 | rw [← sub_nonneg] 214 | have fact : a * c - b * c = (a - b) * c 215 | ring 216 | rw [fact] 217 | apply mul_nonneg_of_nonpos_of_nonpos 218 | · rwa [sub_nonpos] 219 | · exact hc 220 | -- sorry 221 | 222 | -- Using forward reasoning 223 | -- 0014 224 | example (a b c : ℝ) (hc : c ≤ 0) (hab : a ≤ b) : b * c ≤ a * c := by 225 | -- sorry 226 | have hab' : a - b ≤ 0 := by rwa [← sub_nonpos] at hab 227 | have h₁ : 0 ≤ (a - b) * c := mul_nonneg_of_nonpos_of_nonpos hab' hc 228 | have h₂ : (a - b) * c = a * c - b * c := by ring 229 | have h₃ : 0 ≤ a * c - b * c := by rwa [h₂] at h₁ 230 | rwa [sub_nonneg] at h₃ 231 | -- sorry 232 | 233 | -- 0015 234 | /-- Using a combination of both, with a `calc` block -/ 235 | example (a b c : ℝ) (hc : c ≤ 0) (hab : a ≤ b) : b * c ≤ a * c := by 236 | -- sorry 237 | have hab' : a - b ≤ 0 := by rwa [sub_nonpos] 238 | rw [← sub_nonneg] 239 | calc 240 | 0 ≤ (a - b) * c := mul_nonneg_of_nonpos_of_nonpos hab' hc 241 | _ = a * c - b * c := by ring 242 | -- sorry 243 | 244 | /- 245 | Let's now move to proving implications. Lean denotes implications using 246 | a simple arrow `→`, the same it uses for functions (say denoting the type of functions 247 | from `ℕ` to `ℕ` by `ℕ → ℕ`). This is because it sees a proof of `P ⇒ Q` as a function turning 248 | a proof of `P` into a proof `Q`. 249 | 250 | Many of the examples that we already met are implications under the hood. For instance we proved 251 | 252 | `le_add_of_nonneg_left (a b : ℝ) (ha : 0 ≤ a) : b ≤ a + b` 253 | 254 | But this can be rephrased as 255 | 256 | `le_add_of_nonneg_left (a b : ℝ) : 0 ≤ a → b ≤ a + b` 257 | 258 | In order to prove `P → Q`, we use the tactic `intro`, followed by an assumption name. 259 | This creates an assumption with that name asserting that `P` holds, and turns the goal into `Q`. 260 | 261 | Let's check we can go from our old version of `le_add_of_nonneg_left` to the new one. 262 | 263 | -/ 264 | example (a b : ℝ) : 0 ≤ a → b ≤ a + b := by 265 | intro ha 266 | exact le_add_of_nonneg_left ha 267 | 268 | /- 269 | Actually Lean doesn't make any difference between those two versions. It is also happy with 270 | -/ 271 | example (a b : ℝ) : 0 ≤ a → b ≤ a + b := 272 | le_add_of_nonneg_left 273 | 274 | /- No tactic state is shown in the above line because we don't even need to enter 275 | tactic mode using `by`. 276 | 277 | Let's practise using `intro`. -/ 278 | -- 0016 279 | example (a b : ℝ) : 0 ≤ b → a ≤ a + b := by 280 | -- sorry 281 | intro hb 282 | calc 283 | a = a + 0 := by ring 284 | _ ≤ a + b := add_le_add_left hb a 285 | -- sorry 286 | 287 | /- 288 | What about lemmas having more than one assumption? For instance: 289 | 290 | `add_nonneg {a b : ℝ} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b` 291 | 292 | A natural idea is to use the conjunction operator (logical AND), which Lean denotes 293 | by ∧. Assumptions built using this operator can be decomposed using the `rcases` tactic, 294 | which is a very general assumption-decomposing tactic. 295 | -/ 296 | example {a b : ℝ} : 0 ≤ a ∧ 0 ≤ b → 0 ≤ a + b := by 297 | intro hyp 298 | rcases hyp with ⟨ha, hb⟩ 299 | exact add_nonneg ha hb 300 | 301 | /- 302 | Needing that intermediate line invoking `rcases` shows this formulation is not what is used by 303 | Lean. It rather sees `add_nonneg` as two nested implications: 304 | if a is non-negative then if b is non-negative then a+b is non-negative. 305 | It reads funny, but it is much more convenient to use in practice. 306 | -/ 307 | example {a b : ℝ} : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := 308 | add_nonneg 309 | 310 | /- 311 | The above pattern is so common that implications are defined as right-associative operators, 312 | hence parentheses are not needed above. 313 | 314 | Let's prove that the naive conjunction version implies the funny Lean version. For this we need 315 | to know how to prove a conjunction. The `constructor` tactic creates two goals from a conjunction goal. 316 | It can also be used to create two implication goals from an equivalence goal. 317 | -/ 318 | example {a b : ℝ} (H : 0 ≤ a ∧ 0 ≤ b → 0 ≤ a + b) : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := by 319 | intro ha 320 | intro hb 321 | apply H 322 | constructor 323 | exact ha 324 | exact hb 325 | 326 | /- 327 | Let's practice `rcases` and `constructor`. In the next exercise, `P`, `Q` and `R` denote 328 | unspecified mathematical statements. 329 | -/ 330 | -- 0017 331 | example (P Q : Prop) : P ∧ Q → Q ∧ P := by 332 | -- sorry 333 | intro hyp 334 | rcases hyp with ⟨hP, hQ⟩ 335 | constructor 336 | · exact hQ 337 | · exact hP 338 | -- sorry 339 | 340 | /- 341 | Of course using `constructor` only to be able to use `exact` twice in a row feels silly. One can 342 | also use the anonymous constructor syntax: ⟨ ⟩ 343 | Beware those are not parentheses but angle brackets. This is a generic way of providing 344 | compound objects to Lean when Lean already has a very clear idea of what it is waiting for. 345 | 346 | So we could have replaced the last three lines by: 347 | exact ⟨hQ, hP⟩ 348 | 349 | We can also combine the `intros` steps. We can now compress our earlier proof to: 350 | -/ 351 | example {a b : ℝ} (H : 0 ≤ a ∧ 0 ≤ b → 0 ≤ a + b) : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := by 352 | intro ha hb 353 | exact H ⟨ha, hb⟩ 354 | 355 | /- 356 | The anonymous contructor trick actually also works in `intro`. So we can replace 357 | intro h, 358 | rcases h with ⟨h₁, h₂⟩ 359 | by 360 | intro ⟨h₁, h₂⟩, 361 | Now redo the previous exercise using all those compressing techniques, in exactly two lines. -/ 362 | -- 0018 363 | example (P Q : Prop) : P ∧ Q → Q ∧ P := by 364 | -- sorry 365 | intro ⟨hP, hQ⟩ 366 | exact ⟨hQ, hP⟩ 367 | -- sorry 368 | 369 | /- 370 | We are ready to come back to the equivalence between the different formulations of 371 | lemmas having two assumptions. Remember the `constructor` tactic can be used to split 372 | an equivalence into two implications. 373 | -/ 374 | -- 0019 375 | example (P Q R : Prop) : P ∧ Q → R ↔ P → Q → R := by 376 | -- sorry 377 | constructor 378 | · intro hyp hP hQ 379 | exact hyp ⟨hP, hQ⟩ 380 | · rintro hyp ⟨hP, hQ⟩ 381 | exact hyp hP hQ 382 | -- sorry 383 | 384 | /- 385 | If you used more than five lines in the above exercise then try to compress things 386 | (without simply removing line ends). 387 | 388 | One last compression technique: given a proof `h` of a conjunction `P ∧ Q`, one can get 389 | a proof of `P` using `h.left` and a proof of `Q` using `h.right`, without using `rcases`. 390 | One can also use the more generic (but less legible) names `h.1` and `h.2`. 391 | 392 | Similarly, given a proof `h` of `P ↔ Q`, one can get a proof of `P → Q` using `h.mp` 393 | and a proof of `Q → P` using `h.mpr` (or the generic `h.1` and `h.2` that are even less legible 394 | in this case). 395 | 396 | Before the final exercise in this file, let's make sure we'll be able to leave 397 | without learning 10 lemma names. The `linarith` tactic will prove any equality or 398 | inequality or contradiction that follows by linear combinations of assumptions from the 399 | context (with constant coefficients). 400 | -/ 401 | example (a b : ℝ) (hb : 0 ≤ b) : a ≤ a + b := by linarith 402 | 403 | /- 404 | Now let's enjoy this for a while. 405 | -/ 406 | -- 0020 407 | example (a b : ℝ) (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a + b := by 408 | -- sorry 409 | linarith 410 | -- sorry 411 | 412 | -- And let's combine with our earlier lemmas. 413 | -- 0021 414 | example (a b c d : ℝ) (hab : a ≤ b) (hcd : c ≤ d) : a + c ≤ b + d := by 415 | -- sorry 416 | linarith 417 | -- sorry 418 | 419 | /- 420 | Final exercise 421 | 422 | In the last exercise of this file, we will use the divisibility relation on ℕ, 423 | denoted by `∣` (beware this is a unicode divisibility bar, not the ASCII pipe character), 424 | and the `gcd` function. 425 | 426 | The definitions are the usual ones, but our goal is to avoid using these definitions and 427 | only use the following three lemmas: 428 | 429 | `dvd_refl (a : ℕ) : a ∣ a` 430 | 431 | `dvd_antisymm {a b : ℕ} : a ∣ b → b ∣ a → a = b :=` 432 | 433 | `dvd_gcd_iff {a b c : ℕ} : c ∣ gcd a b ↔ c ∣ a ∧ c ∣ b` 434 | -/ 435 | -- All functions and lemmas below are about natural numbers. 436 | open Nat 437 | 438 | -- 0022 439 | example (a b : ℕ) : a ∣ b ↔ gcd a b = a := by 440 | -- sorry 441 | have fact : gcd a b ∣ a ∧ gcd a b ∣ b := by rw [← dvd_gcd_iff] 442 | constructor 443 | · intro h 444 | apply dvd_antisymm fact.left 445 | rw [dvd_gcd_iff] 446 | exact ⟨dvd_refl a, h⟩ 447 | · intro h 448 | rw [← h] 449 | exact fact.right 450 | -- sorry 451 | -------------------------------------------------------------------------------- /Tutorials/Solutions/03ForallOr.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | In this file, we'll learn about the `∀` quantifier, and the disjunction 5 | operator `∨` (logical OR). 6 | 7 | # The universal quantifier 8 | 9 | Let `P` be a predicate on a type `X`. This means for every mathematical 10 | object `x` with type `X`, we get a mathematical statement `P x`. 11 | In Lean, `P x` has type `Prop`. 12 | 13 | Lean sees a proof `h` of `∀ x, P x` as a function sending any `x : X` to 14 | a proof `h x` of `P x`. 15 | This already explains the main way to use an assumption or lemma which 16 | starts with a `∀`. 17 | 18 | In order to prove `∀ x, P x`, we use `intro x` to fix an arbitrary object 19 | with type `X`, and call it `x`. 20 | 21 | Note also we don't need to give the type of `x` in the expression `∀ x, P x` 22 | as long as the type of `P` is clear to Lean, which can then infer the type of `x`. 23 | 24 | Let's consider two predicates to play with `∀`. 25 | 26 | `EvenFun (f : ℝ → ℝ) : ∀ x, f (-x) = f x` 27 | 28 | `OddFun (f : ℝ → ℝ) : ∀ x, f (-x) = -f x`. 29 | 30 | -/ 31 | 32 | /- 33 | In the next proof, we also take the opportunity to introduce the 34 | `unfold` tactic, which simply unfolds definitions. Here this is purely 35 | for didactic reason, Lean doesn't need those `unfold` invocations. 36 | We will also use `rfl` which is a term proving equalities that are true 37 | by definition (in a very strong sense to be discussed later). 38 | -/ 39 | example (f g : ℝ → ℝ) : EvenFun f → EvenFun g → EvenFun (f + g) := by 40 | -- Assume f is even 41 | intro hf 42 | -- which means ∀ x, f (-x) = f x 43 | unfold EvenFun at hf 44 | -- and the same for g 45 | intro hg 46 | unfold EvenFun at hg 47 | -- We need to prove ∀ x, (f+g)(-x) = (f+g)(x) 48 | unfold EvenFun 49 | -- Let x be any real number 50 | intro x 51 | -- and let's compute 52 | calc 53 | (f + g) (-x) = f (-x) + g (-x) := rfl 54 | _ = f x + g (-x) := by rw [hf x] 55 | _ = f x + g x := by rw [hg x] 56 | _ = (f + g) x := rfl 57 | 58 | 59 | /- 60 | In the preceding proof, all `unfold` lines are purely for 61 | psychological comfort. 62 | 63 | Sometimes unfolding is necessary because we want to apply a tactic 64 | that operates purely on the syntactical level. 65 | The main such tactic is `rw`. 66 | 67 | The same property of `rw` explain why the first computation line 68 | is necessary, although its proof is simply `rfl`. 69 | Before that line, `rw [hf x]` won't find anything like `f (-x)` hence 70 | will give up. 71 | The last line is not necessary however, since it only proves 72 | something that is true by definition, and is not followed by 73 | a `rw`. 74 | 75 | Also, Lean doesn't need to be told that hf should be specialized to 76 | x before rewriting, exactly as in the first file 01_equality_rewriting. 77 | We can also gather several rewrites using a list of expressions. 78 | 79 | One last trick is that `rw` can take a list of expressions to use for 80 | rewriting. For instance `rw [h₁, h₂, h₃]` is equivalent to three 81 | lines `rw [h₁]`, `rw [h₂]` and `rw [h₃]`. Note that you can inspect the tactic 82 | state between those rewrites when reading a proof using this syntax. You 83 | simply need to move the cursor inside the list. 84 | 85 | Hence we can compress the above proof to: 86 | -/ 87 | example (f g : ℝ → ℝ) : EvenFun f → EvenFun g → EvenFun (f + g) := by 88 | intro hf hg x 89 | calc 90 | (f + g) (-x) = f (-x) + g (-x) := rfl 91 | _ = f x + g x := by rw [hf, hg] 92 | 93 | /- 94 | Now let's practice. If you need to learn how to type a unicode symbol, 95 | you can put your mouse cursor above the symbol and wait for one second. 96 | -/ 97 | -- 0023 98 | example (f g : ℝ → ℝ) : EvenFun f → EvenFun (g ∘ f) := by 99 | -- sorry 100 | intro hf x 101 | calc 102 | (g ∘ f) (-x) = g (f (-x)) := rfl 103 | _ = g (f x) := by rw [hf] 104 | -- sorry 105 | 106 | -- 0024 107 | example (f g : ℝ → ℝ) : OddFun f → OddFun g → OddFun (g ∘ f) := by 108 | -- sorry 109 | intro hf hg x 110 | calc 111 | (g ∘ f) (-x) = g (f (-x)) := rfl 112 | _ = -g (f x) := by rw [hf, hg] 113 | _ = -(g ∘ f) x := rfl 114 | -- sorry 115 | 116 | /- 117 | Let's have more quantifiers, and play with forward and backward reasoning. 118 | 119 | In the next definitions, note how `∀ x₁, ∀ x₂` is abreviated to `∀ x₁ x₂`. 120 | 121 | `NonDecreasing (f : ℝ → ℝ) := ∀ x₁ x₂, x₁ ≤ x₂ → f x₁ ≤ f x₂` 122 | 123 | `NonIncreasing (f : ℝ → ℝ) := ∀ x₁ x₂, x₁ ≤ x₂ → f x₁ ≥ f x₂` 124 | 125 | -/ 126 | 127 | 128 | -- Let's be very explicit and use forward reasoning first. 129 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 130 | -- Let x₁ and x₂ be real numbers such that x₁ ≤ x₂ 131 | intro x₁ x₂ h 132 | -- Since f is non-decreasing, f x₁ ≤ f x₂. 133 | have step₁ : f x₁ ≤ f x₂ := hf x₁ x₂ h 134 | -- Since g is non-decreasing, we then get g (f x₁) ≤ g (f x₂). 135 | exact hg (f x₁) (f x₂) step₁ 136 | 137 | /- 138 | In the above proof, note how inconvenient it is to specify x₁ and x₂ in `hf x₁ x₂ h` since 139 | they could be inferred from the type of h. 140 | We could have written `hf _ _ h` and Lean would have filled the holes denoted by _. 141 | 142 | Even better we could have written the definition 143 | of `NonDecreasing` as: ∀ {x₁ x₂}, x₁ ≤ x₂ → f x₁ ≤ f x₂, with curly braces to denote 144 | implicit arguments. 145 | 146 | But let's leave that aside for now. One possible variation on the above proof is to 147 | use the `specialize` tactic to replace hf by its specialization to the relevant value. 148 | -/ 149 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 150 | intro x₁ x₂ h 151 | specialize hf x₁ x₂ h 152 | exact hg (f x₁) (f x₂) hf 153 | 154 | /- 155 | This `specialize` tactic is mostly useful for exploration, or in preparation for rewriting 156 | in the assumption. One can very often replace its use by using more complicated expressions 157 | directly involving the original assumption, as in the next variation: 158 | -/ 159 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 160 | intro x₁ x₂ h 161 | exact hg (f x₁) (f x₂) (hf x₁ x₂ h) 162 | 163 | /- 164 | Since the above proof uses only `intros` and `exact`, we could very easily replace it by the 165 | raw proof term: 166 | -/ 167 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := 168 | fun x₁ x₂ h ↦ hg (f x₁) (f x₂) (hf x₁ x₂ h) 169 | 170 | /- 171 | Of course the above proof is difficult to decipher. The principle in mathlib is to use 172 | such a proof when the result is obvious and you don't want to read the proof anyway. 173 | 174 | Instead of pursuing this style, let's see how backward reasoning would look like here. 175 | As usual with this style, we use `apply` and enjoy Lean specializing assumptions for us 176 | using unification. 177 | -/ 178 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonDecreasing g) : NonDecreasing (g ∘ f) := by 179 | -- Let x₁ and x₂ be real numbers such that x₁ ≤ x₂ 180 | intro x₁ x₂ h 181 | -- We need to prove (g ∘ f) x₁ ≤ (g ∘ f) x₂. 182 | -- Since g is non-decreasing, it suffices to prove f x₁ ≤ f x₂ 183 | apply hg 184 | -- which follows from our assumption on f 185 | apply hf 186 | -- and on x₁ and x₂ 187 | exact h 188 | 189 | -- 0025 190 | example (f g : ℝ → ℝ) (hf : NonDecreasing f) (hg : NonIncreasing g) : NonIncreasing (g ∘ f) := by 191 | -- sorry 192 | intro x₁ x₂ h 193 | apply hg 194 | exact hf x₁ x₂ h 195 | -- sorry 196 | 197 | /- # Disjunctions 198 | 199 | Let's switch to disjunctions now. Lean denotes by `∨` the 200 | logical OR operator. 201 | 202 | In order to make use of an assumption 203 | `hyp : P ∨ Q` 204 | we use the cases tactic: 205 | `rcases hyp with hP | hQ` 206 | which creates two proof branches: one branch assuming `hP : P`, 207 | and one branch assuming `hQ : Q`. 208 | 209 | In order to directly prove a goal `P ∨ Q`, 210 | we use either the `left` tactic and prove `P` or the `right` 211 | tactic and prove `Q`. 212 | 213 | In the next proof we use `ring` and `linarith` to get rid of 214 | easy computations or inequalities, as well as one lemma: 215 | 216 | `mul_eq_zero : a*b = 0 ↔ a = 0 ∨ b = 0` 217 | -/ 218 | example (a b : ℝ) : a = a * b → a = 0 ∨ b = 1 := by 219 | intro hyp 220 | have H : a * (1 - b) = 0 := by 221 | calc 222 | a * (1 - b) = a - a * b := by ring 223 | _ = 0 := by linarith 224 | 225 | rw [mul_eq_zero] at H 226 | rcases H with Ha | Hb 227 | · left 228 | exact Ha 229 | · right 230 | linarith 231 | 232 | -- 0026 233 | example (x y : ℝ) : x ^ 2 = y ^ 2 → x = y ∨ x = -y := by 234 | -- sorry 235 | intro hyp 236 | have H : (x - y) * (x + y) = 0 237 | calc 238 | (x - y) * (x + y) = x ^ 2 - y ^ 2 := by ring 239 | _ = y ^ 2 - y ^ 2 := by rw [hyp] 240 | _ = 0 := by ring 241 | 242 | rw [mul_eq_zero] at H 243 | rcases H with h1 | h2 244 | · left 245 | linarith 246 | · right 247 | linarith 248 | -- sorry 249 | 250 | /- 251 | In the next exercise, we can use: 252 | `eq_or_lt_of_le : x ≤ y → x = y ∨ x < y` 253 | -/ 254 | -- 0027 255 | example (f : ℝ → ℝ) : NonDecreasing f ↔ ∀ x y, x < y → f x ≤ f y := by 256 | -- sorry 257 | constructor 258 | · intro hf x y hxy 259 | apply hf 260 | linarith 261 | · intro hf x y hxy 262 | have clef : x = y ∨ x < y := eq_or_lt_of_le hxy 263 | rcases clef with hxy | hxy 264 | · rw [hxy] 265 | · exact hf x y hxy 266 | -- sorry 267 | 268 | /- 269 | In the next exercise, we can use: 270 | `le_total x y : x ≤ y ∨ y ≤ x` 271 | -/ 272 | -- 0028 273 | example (f : ℝ → ℝ) (h : NonDecreasing f) (h' : ∀ x, f (f x) = x) : ∀ x, f x = x := by 274 | -- sorry 275 | intro x 276 | have : f (f x) = x := by rw [h'] 277 | have : f x ≤ x ∨ x ≤ f x := le_total (f x) x 278 | rcases this with hx | hx 279 | · have f1 : f (f x) ≤ f x := h (f x) x hx 280 | rw [h'] at f1 281 | linarith 282 | · have f1 : f x ≤ f (f x) := h x (f x) hx 283 | rw [h' x] at f1 284 | linarith 285 | -- sorry 286 | -------------------------------------------------------------------------------- /Tutorials/Solutions/04Exists.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Mathlib.Data.Int.Parity 3 | 4 | /- # Existential quantifiers 5 | 6 | In this file, we learn how to handle the ∃ quantifier. 7 | 8 | In order to prove `∃ x, P x`, we give some `x₀` using tactic `use x₀` and 9 | then prove `P x₀`. This `x₀` can be an object from the local context 10 | or a more complicated expression. 11 | 12 | In the example below, `P x₀` is `8 = 2*4` which is true by definition so 13 | Lean does not ask for a proof. 14 | -/ 15 | 16 | example : ∃ n : ℕ, 8 = 2 * n := by 17 | use 4 18 | 19 | /- 20 | In order to use `h : ∃ x, P x`, we use the `rcases` tactic to fix 21 | one `x₀` that works. 22 | 23 | Again `h` can come straight from the local context or can be a more 24 | complicated expression. 25 | -/ 26 | example (n : ℕ) (h : ∃ k : ℕ, n = k + 1) : n > 0 := by 27 | -- Let's fix k₀ such that n = k₀ + 1. 28 | rcases h with ⟨k₀, hk₀⟩ 29 | -- It now suffices to prove k₀ + 1 > 0. 30 | rw [hk₀] 31 | -- and we have a lemma about this 32 | exact Nat.succ_pos k₀ 33 | 34 | /- 35 | The next exercises use divisibility in ℤ (beware the ∣ symbol which is 36 | not ASCII). 37 | 38 | By definition, a ∣ b ↔ ∃ k, b = a*k, so you can prove a ∣ b using the 39 | `use` tactic. 40 | -/ 41 | -- Until the end of this file, a, b and c will denote integers, unless 42 | -- explicitly stated otherwise 43 | variable (a b c : ℤ) 44 | 45 | -- 0029 46 | example (h₁ : a ∣ b) (h₂ : b ∣ c) : a ∣ c := by 47 | -- sorry 48 | rcases h₁ with ⟨k, hk⟩ 49 | rcases h₂ with ⟨l, hl⟩ 50 | use k * l 51 | calc 52 | c = b * l := hl 53 | _ = a * k * l := by rw [hk] 54 | _ = a * (k * l) := by ring 55 | -- sorry 56 | 57 | /- 58 | A very common pattern is to have an assumption or lemma asserting 59 | `h : ∃ x, y = ...` 60 | and this is used through the combo: 61 | rcases h with ⟨x, hx⟩ 62 | rw hx at ... 63 | The tactic `rcases` allows us to simplify the above combo when the 64 | name `hx` is replaced by the special name `rfl`, as in the following example. 65 | -/ 66 | example (h1 : a ∣ b) (h2 : a ∣ c) : a ∣ b + c := by 67 | rcases h1 with ⟨k, rfl⟩ 68 | rcases h2 with ⟨l, rfl⟩ 69 | use k + l 70 | ring 71 | 72 | /- 73 | You can use the same `rfl` trick with the `rintro` tactic which 74 | is a power powerful variant of `intro`. 75 | -/ 76 | example : a ∣ b → a ∣ c → a ∣ b + c := by 77 | rintro ⟨k, rfl⟩ ⟨l, rfl⟩ 78 | use k + l 79 | ring 80 | 81 | -- 0030 82 | example : 0 ∣ a ↔ a = 0 := by 83 | -- sorry 84 | constructor 85 | · rintro ⟨k, rfl⟩ 86 | ring 87 | · rintro rfl 88 | use 0 89 | rfl 90 | -- sorry 91 | 92 | /- 93 | We can now start combining quantifiers, using the definition 94 | 95 | `Surjective (f : X → Y) := ∀ y, ∃ x, f x = y` 96 | 97 | -/ 98 | open Function 99 | 100 | -- In the remaining of this file, f and g will denote functions from 101 | -- ℝ to ℝ. 102 | variable (f g : ℝ → ℝ) 103 | 104 | -- 0031 105 | example (h : Surjective (g ∘ f)) : Surjective g := by 106 | -- sorry 107 | intro y 108 | rcases h y with ⟨w, rfl⟩ 109 | use f w 110 | rfl 111 | -- sorry 112 | 113 | /- 114 | The above exercise can be done in three lines. Try to do the 115 | next exercise in four lines. 116 | -/ 117 | -- 0032 118 | example (hf : Surjective f) (hg : Surjective g) : Surjective (g ∘ f) := by 119 | -- sorry 120 | intro z 121 | rcases hg z with ⟨y, rfl⟩ 122 | rcases hf y with ⟨x, rfl⟩ 123 | use x 124 | rfl 125 | -- sorry 126 | -------------------------------------------------------------------------------- /Tutorials/Solutions/05SequenceLimits.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Mathlib.Algebra.Group.Pi 3 | import Tutorials.TutoLib 4 | /- 5 | In this file we manipulate the elementary definition of limits of 6 | sequences of real numbers. 7 | mathlib has a much more general definition of limits, but here 8 | we want to practice using the logical operators and relations 9 | covered in the previous files. 10 | 11 | A sequence u is a function from ℕ to ℝ, hence Lean says 12 | `u : ℕ → ℝ` 13 | The definition we'll be using is: 14 | 15 | -- Definition of « u tends to l » 16 | `def seq_limit (u : ℕ → ℝ) (l : ℝ) : Prop := ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε` 17 | 18 | Note the use of `∀ ε > 0, ...` which is an abbreviation of 19 | `∀ ε, ε > 0 → ... ` 20 | 21 | In particular, a statement like `h : ∀ ε > 0, ...` 22 | can be specialized to a given ε₀ by 23 | `specialize h ε₀ hε₀` 24 | where `hε₀` is a proof of `ε₀ > 0`. 25 | 26 | Also recall that, wherever Lean expects some proof term, we can 27 | start a tactic mode proof using the keyword `by` (followed by curly braces 28 | if you need more than one tactic invocation). 29 | For instance, if the local context contains: 30 | 31 | δ : ℝ 32 | δ_pos : δ > 0 33 | h : ∀ ε > 0, ... 34 | 35 | then we can specialize h to the real number δ/2 using: 36 | `specialize h (δ/2) (by linarith)` 37 | where `by linarith` will provide the proof of `δ/2 > 0` expected by Lean. 38 | 39 | We'll take this opportunity to use two new tactics: 40 | 41 | `norm_num` will perform numerical normalization on the goal and `norm_num at h` 42 | will do the same in assumption `h`. This will get rid of trivial calculations on numbers, 43 | like replacing `|l - l|` by zero in the next exercise. 44 | Its name stands for "normalize numbers". 45 | 46 | `congr` will try to prove equalities between applications of functions by recursively 47 | proving the arguments are the same. Its name stands for "congruence". 48 | For instance, if the goal is `f x + g y = f z + g t` then congr will replace it by 49 | two goals: `x = z` and `y = t`. 50 | You can limit the recursion depth by specifying a natural number after `congr`. 51 | For instance, in the above example, `congr 1` will give new goals 52 | `f x = f z` and `g y = g t`, which only inspect arguments of the addition and not deeper. 53 | -/ 54 | variable (u v w : ℕ → ℝ) (l l' : ℝ) 55 | 56 | -- If u is constant with value l then u tends to l 57 | -- 0033 58 | example : (∀ n, u n = l) → SeqLimit u l := by 59 | -- sorry 60 | intro h ε ε_pos 61 | use 0 62 | intro n _hn 63 | rw [h] 64 | norm_num 65 | linarith 66 | -- sorry 67 | 68 | /- When dealing with absolute values, we'll use lemmas: 69 | 70 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 71 | 72 | `abs_add (x y : ℝ) : |x + y| ≤ |x| + |y|` 73 | 74 | `abs_sub_comm (x y : ℝ) : |x - y| = |y - x|` 75 | 76 | You should probably write them down on a sheet of paper that you keep at 77 | hand since they are used in many exercises. 78 | -/ 79 | -- Assume l > 0. Then u tends to l implies u n ≥ l/2 for large enough n 80 | -- 0034 81 | example (hl : l > 0) : SeqLimit u l → ∃ N, ∀ n ≥ N, u n ≥ l / 2 := by 82 | -- sorry 83 | intro h 84 | rcases h (l / 2) (by linarith) with ⟨N, hN⟩ 85 | use N 86 | intro n hn 87 | specialize hN n hn 88 | rw [abs_le] at hN 89 | linarith 90 | -- sorry 91 | 92 | /- 93 | When dealing with max, you can use 94 | 95 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 96 | 97 | `le_max_left p q : p ≤ max p q` 98 | 99 | `le_max_right p q : q ≤ max p q` 100 | 101 | You should probably add them to the sheet of paper where you wrote 102 | the `abs` lemmas since they are used in many exercises. 103 | 104 | Let's see an example. 105 | -/ 106 | -- If u tends to l and v tends l' then u+v tends to l+l' 107 | example (hu : SeqLimit u l) (hv : SeqLimit v l') : SeqLimit (u + v) (l + l') := by 108 | intro ε ε_pos 109 | rcases hu (ε / 2) (by linarith) with ⟨N₁, hN₁⟩ 110 | rcases hv (ε / 2) (by linarith) with ⟨N₂, hN₂⟩ 111 | use max N₁ N₂ 112 | intro n hn 113 | rcases ge_max_iff.mp hn with ⟨hn₁, hn₂⟩ 114 | have fact₁ : |u n - l| ≤ ε / 2 := hN₁ n (by linarith) 115 | have fact₂ : |v n - l'| ≤ ε / 2 := hN₂ n (by linarith) 116 | calc 117 | |(u + v) n - (l + l')| = |u n + v n - (l + l')| := rfl 118 | _ = |u n - l + (v n - l')| := by congr ; ring 119 | _ ≤ |u n - l| + |v n - l'| := by apply abs_add 120 | _ ≤ ε := by linarith 121 | 122 | 123 | /- 124 | In the above proof, we used `have` to prepare facts for `linarith` consumption in the last line. 125 | Since we have direct proof terms for them, we can feed them directly to `linarith` as in the next proof 126 | of the same statement. 127 | Another variation we introduce is rewriting using `ge_max_iff` and letting `linarith` handle the 128 | conjunction, instead of creating two new assumptions. 129 | -/ 130 | example (hu : SeqLimit u l) (hv : SeqLimit v l') : SeqLimit (u + v) (l + l') := by 131 | intro ε ε_pos 132 | rcases hu (ε / 2) (by linarith) with ⟨N₁, hN₁⟩ 133 | rcases hv (ε / 2) (by linarith) with ⟨N₂, hN₂⟩ 134 | use max N₁ N₂ 135 | intro n hn 136 | rw [ge_max_iff] at hn 137 | calc 138 | |(u + v) n - (l + l')| = |u n + v n - (l + l')| := rfl 139 | _ = |u n - l + (v n - l')| := by ring_nf 140 | _ ≤ |u n - l| + |v n - l'| := by apply abs_add 141 | _ ≤ ε := by linarith [hN₁ n (by linarith), hN₂ n (by linarith)] 142 | 143 | 144 | -- Let's do something similar: the squeezing theorem. 145 | -- 0035 146 | example (hu : SeqLimit u l) (hw : SeqLimit w l) (h : ∀ n, u n ≤ v n) (h' : ∀ n, v n ≤ w n) : 147 | SeqLimit v l := by 148 | -- sorry 149 | intro ε ε_pos 150 | rcases hu ε ε_pos with ⟨N, hN⟩ 151 | rcases hw ε ε_pos with ⟨N', hN'⟩ 152 | use max N N' 153 | intro n hn 154 | rw [ge_max_iff] at hn 155 | specialize hN n (by linarith) 156 | specialize hN' n (by linarith) 157 | specialize h n 158 | specialize h' n 159 | rw [abs_le] at * 160 | constructor 161 | -- Here `linarith` can finish, but on paper we would write 162 | calc 163 | -ε ≤ u n - l := by linarith 164 | _ ≤ v n - l := by linarith 165 | 166 | calc 167 | v n - l ≤ w n - l := by linarith 168 | _ ≤ ε := by linarith 169 | -- sorry 170 | 171 | -- What about < ε? 172 | -- 0036 173 | example (u l) : SeqLimit u l ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| < ε := by 174 | -- sorry 175 | constructor 176 | · intro hyp ε ε_pos 177 | rcases hyp (ε / 2) (by linarith) with ⟨N, hN⟩ 178 | use N 179 | intro n hn 180 | calc 181 | |u n - l| ≤ ε / 2 := hN n hn 182 | _ < ε := by linarith 183 | 184 | · intro hyp ε ε_pos 185 | rcases hyp ε ε_pos with ⟨N, hN⟩ 186 | use N 187 | intro n hn 188 | specialize hN n hn 189 | linarith 190 | -- sorry 191 | 192 | /- In the next exercise, we'll use 193 | 194 | `eq_of_abs_sub_le_all (x y : ℝ) : (∀ ε > 0, |x - y| ≤ ε) → x = y` 195 | -/ 196 | -- A sequence admits at most one limit 197 | -- 0037 198 | example : SeqLimit u l → SeqLimit u l' → l = l' := by 199 | -- sorry 200 | intro hl hl' 201 | apply eq_of_abs_sub_le_all 202 | intro ε ε_pos 203 | rcases hl (ε / 2) (by linarith) with ⟨N, hN⟩ 204 | rcases hl' (ε / 2) (by linarith) with ⟨N', hN'⟩ 205 | calc 206 | |l - l'| = |l - u (max N N') + (u (max N N') - l')| := by ring_nf 207 | _ ≤ |l - u (max N N')| + |u (max N N') - l'| := by apply abs_add 208 | _ = |u (max N N') - l| + |u (max N N') - l'| := by rw [abs_sub_comm] 209 | _ ≤ ε := by linarith [hN (max N N') (le_max_left _ _), hN' (max N N') (le_max_right _ _)] 210 | -- sorry 211 | 212 | /- 213 | Let's now practice deciphering definitions before proving. 214 | 215 | -/ 216 | def NonDecreasingSeq (u : ℕ → ℝ) := 217 | ∀ n m, n ≤ m → u n ≤ u m 218 | 219 | def IsSeqSup (M : ℝ) (u : ℕ → ℝ) := 220 | (∀ n, u n ≤ M) ∧ ∀ ε > 0, ∃ n₀, u n₀ ≥ M - ε 221 | 222 | -- 0038 223 | example (M : ℝ) (h : IsSeqSup M u) (h' : NonDecreasingSeq u) : SeqLimit u M := by 224 | -- sorry 225 | intro ε ε_pos 226 | rcases h with ⟨inf_M, sup_M_ep⟩ 227 | rcases sup_M_ep ε ε_pos with ⟨n₀, hn₀⟩ 228 | use n₀ 229 | intro n hn 230 | rw [abs_le] 231 | constructor <;> linarith [inf_M n, h' n₀ n hn] 232 | -- sorry 233 | -------------------------------------------------------------------------------- /Tutorials/Solutions/06SubSequences.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | This file continues the elementary study of limits of sequences. 5 | It can be skipped if the previous file was too easy, it won't introduce 6 | any new tactic or trick. 7 | 8 | Remember useful lemmas: 9 | 10 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 11 | 12 | `abs_add (x y : ℝ) : |x + y| ≤ |x| + |y|` 13 | 14 | `abs_sub_comm (x y : ℝ) : |x - y| = |y - x|` 15 | 16 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 17 | 18 | `le_max_left p q : p ≤ max p q` 19 | 20 | `le_max_right p q : q ≤ max p q` 21 | 22 | and the definition: 23 | 24 | `def SeqLimit (u : ℕ → ℝ) (l : ℝ) : Prop := ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε` 25 | 26 | You can also use a property proved in the previous file: 27 | 28 | `unique_limit : SeqLimit u l → SeqLimit u l' → l = l'` 29 | 30 | `def extraction (φ : ℕ → ℕ) := ∀ n m, n < m → φ n < φ m` 31 | -/ 32 | variable {φ : ℕ → ℕ} 33 | 34 | /- 35 | The next lemma is proved by an easy induction, but we haven't seen induction 36 | in this tutorial. If you did the natural number game then you can delete 37 | the proof below and try to reconstruct it. 38 | -/ 39 | /-- An extraction is greater than id -/ 40 | theorem id_le_extraction' : Extraction φ → ∀ n, n ≤ φ n := by 41 | intro hyp n 42 | induction' n with n hn 43 | · exact Nat.zero_le _ 44 | · exact Nat.succ_le_of_lt (by linarith [hyp n (n + 1) (by linarith)]) 45 | 46 | -- 0039 47 | /-- Extractions take arbitrarily large values for arbitrarily large 48 | inputs. -/ 49 | theorem extraction_ge : Extraction φ → ∀ N N', ∃ n ≥ N', φ n ≥ N := by 50 | -- sorry 51 | intro h N N' 52 | use max N N' 53 | constructor 54 | apply le_max_right 55 | calc 56 | N ≤ max N N' := by apply le_max_left 57 | _ ≤ φ (max N N') := by apply id_le_extraction' h 58 | -- sorry 59 | 60 | variable {u : ℕ → ℝ} {a l : ℝ} 61 | 62 | /- 63 | In the exercise, we use `∃ n ≥ N, ...` which is the abbreviation of 64 | `∃ n, n ≥ N ∧ ...`. 65 | Lean can read this abbreviation, but does not it when displaying the goal. 66 | -/ 67 | -- 0040 68 | /-- If `a` is a cluster point of `u` then there are values of 69 | `u` arbitrarily close to `a` for arbitrarily large input. -/ 70 | theorem near_cluster : ClusterPoint u a → ∀ ε > 0, ∀ N, ∃ n ≥ N, |u n - a| ≤ ε := by 71 | -- sorry 72 | intro hyp ε ε_pos N 73 | rcases hyp with ⟨φ, φ_extr, hφ⟩ 74 | rcases hφ ε ε_pos with ⟨N', hN'⟩ 75 | rcases extraction_ge φ_extr N N' with ⟨q, hq, hq'⟩ 76 | exact ⟨φ q, hq', hN' _ hq⟩ 77 | -- sorry 78 | 79 | /- 80 | The above exercice can be done in five lines. 81 | Hint: you can use the anonymous constructor syntax when proving 82 | existential statements. 83 | -/ 84 | -- 0041 85 | /-- If `u` tends to `l` then its subsequences tend to `l`. -/ 86 | theorem subseq_tendsto_of_tendsto' (h : SeqLimit u l) (hφ : Extraction φ) : SeqLimit (u ∘ φ) l := by 87 | -- sorry 88 | intro ε ε_pos 89 | rcases h ε ε_pos with ⟨N, hN⟩ 90 | use N 91 | intro n hn 92 | apply hN 93 | calc 94 | N ≤ n := hn 95 | _ ≤ φ n := id_le_extraction' hφ n 96 | -- sorry 97 | 98 | -- 0042 99 | /-- If `u` tends to `l` all its cluster points are equal to `l`. -/ 100 | theorem cluster_limit (hl : SeqLimit u l) (ha : ClusterPoint u a) : a = l := by 101 | -- sorry 102 | rcases ha with ⟨φ, φ_extr, lim_u_φ⟩ 103 | have lim_u_φ' : SeqLimit (u ∘ φ) l := subseq_tendsto_of_tendsto' hl φ_extr 104 | exact unique_limit lim_u_φ lim_u_φ' 105 | -- sorry 106 | 107 | /-- Cauchy_sequence sequence -/ 108 | def CauchySequence (u : ℕ → ℝ) := 109 | ∀ ε > 0, ∃ N, ∀ p q, p ≥ N → q ≥ N → |u p - u q| ≤ ε 110 | 111 | -- 0043 112 | example : (∃ l, SeqLimit u l) → CauchySequence u := by 113 | -- sorry 114 | intro hyp 115 | rcases hyp with ⟨l, hl⟩ 116 | intro ε ε_pos 117 | rcases hl (ε / 2) (by linarith) with ⟨N, hN⟩ 118 | use N 119 | intro p q hp hq 120 | calc 121 | |u p - u q| = |u p - l + (l - u q)| := by congr; ring 122 | _ ≤ |u p - l| + |l - u q| := by apply abs_add 123 | _ = |u p - l| + |u q - l| := by rw [abs_sub_comm (u q) l] 124 | _ ≤ ε := by linarith [hN p hp, hN q hq] 125 | -- sorry 126 | 127 | /- 128 | In the next exercise, you can reuse 129 | `near_cluster : cluster_point u a → ∀ ε > 0, ∀ N, ∃ n ≥ N, |u n - a| ≤ ε` 130 | -/ 131 | -- 0044 132 | example (hu : CauchySequence u) (hl : ClusterPoint u l) : SeqLimit u l := by 133 | -- sorry 134 | intro ε ε_pos 135 | rcases hu (ε / 2) (by linarith) with ⟨N, hN⟩ 136 | use N 137 | have clef : ∃ N' ≥ N, |u N' - l| ≤ ε / 2 138 | apply near_cluster hl (ε / 2) (by linarith) 139 | rcases clef with ⟨N', h⟩ 140 | rcases h with ⟨hNN', hN'⟩ 141 | intro n hn 142 | calc 143 | |u n - l| = |u n - u N' + (u N' - l)| := by congr; ring 144 | _ ≤ |u n - u N'| + |u N' - l| := by apply abs_add 145 | _ ≤ ε := by linarith [hN n N' (by linarith) hNN'] 146 | -- sorry 147 | -------------------------------------------------------------------------------- /Tutorials/Solutions/07FirstNegations.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | Negations, proof by contradiction and contraposition. 5 | 6 | This file introduces the logical rules and tactics related to negation: 7 | exfalso, by_contra, contrapose, by_cases and push_neg. 8 | 9 | There is a special statement denoted by `False` which, by definition, 10 | has no proof. 11 | 12 | So `False` implies everything. Indeed `False → P` means any proof of 13 | `False` could be turned into a proof of P. 14 | This fact is known by its latin name 15 | "ex falso quod libet" (from False follows whatever you want). 16 | Hence Lean's tactic to invoke this is called `exfalso`. 17 | -/ 18 | example : False → 0 = 1 := by 19 | intro h 20 | exfalso 21 | exact h 22 | 23 | /- 24 | The preceding example suggests that this definition of `False` isn't very useful. 25 | But actually it allows us to define the negation of a statement P as 26 | "P implies False" that we can read as "if P were true, we would get 27 | a contradiction". Lean denotes this by `¬ P`. 28 | 29 | One can prove that (¬ P) ↔ (P ↔ False). But in practice we directly 30 | use the definition of `¬ P`. 31 | -/ 32 | example {x : ℝ} : ¬x < x := by 33 | intro hyp 34 | rw [lt_iff_le_and_ne] at hyp 35 | rcases hyp with ⟨hyp_inf, hyp_non⟩ 36 | clear hyp_inf -- we won't use that one, so let's discard it 37 | change x = x → False at hyp_non -- Lean doesn't need this psychological line 38 | apply hyp_non 39 | rfl 40 | 41 | open Int 42 | 43 | -- 0045 44 | example (n : ℤ) (h_even : Even n) (h_not_even : ¬Even n) : 0 = 1 := by 45 | -- sorry 46 | exfalso 47 | exact h_not_even h_even 48 | -- sorry 49 | 50 | -- 0046 51 | example (P Q : Prop) (h₁ : P ∨ Q) (h₂ : ¬(P ∧ Q)) : ¬P ↔ Q := by 52 | -- sorry 53 | constructor 54 | · intro hnP 55 | rcases h₁ with hP | hQ 56 | · exfalso 57 | exact hnP hP 58 | · exact hQ 59 | · intro hQ hP 60 | exact h₂ ⟨hP, hQ⟩ 61 | -- sorry 62 | 63 | /- 64 | The definition of negation easily implies that, for every statement P, 65 | `P → ¬ ¬ P`. 66 | 67 | The excluded middle axiom, which asserts `P ∨ ¬ P` allows us to 68 | prove the converse implication. 69 | 70 | Together those two implications form the principle of double negation elimination. 71 | `not_not {P : Prop} : (¬ ¬ P) ↔ P` 72 | 73 | The implication `¬ ¬ P → P` is the basis for proofs by contradiction: 74 | in order to prove `P`, it suffices to prove `¬ ¬ P`, ie `¬ P → False`. 75 | 76 | Of course there is no need to keep explaining all this. The tactic 77 | `by_contra Hyp` will transform any goal `P` into `False` and 78 | add `Hyp : ¬ P` to the local context. 79 | 80 | Let's return to a proof from the 5th file: uniqueness of limits for a sequence. 81 | This cannot be proved without using some version of the excluded middle 82 | axiom. We used it secretely in 83 | 84 | `eq_of_abs_sub_le_all (x y : ℝ) : (∀ ε > 0, |x - y| ≤ ε) → x = y` 85 | 86 | (we'll prove a variation on this lemma below). 87 | 88 | In the proof below, we also take the opportunity to introduce the `let` tactic 89 | which creates a local definition. If needed, it can be unfolded using `unfold_let`. 90 | For instance after using `let N₀ := max N N'`, you could write 91 | `unfold_let N₀ at h` to replace `N₀` by its definition in some local assumption `h`. 92 | -/ 93 | example (u : ℕ → ℝ) (l l' : ℝ) : SeqLimit u l → SeqLimit u l' → l = l' := by 94 | intro hl hl' 95 | by_contra H 96 | change l ≠ l' at H -- Lean does not need this line, but it makes H easier to read 97 | have ineg : |l - l'| > 0 := abs_pos.mpr (sub_ne_zero_of_ne H) -- details about that line are not important 98 | rcases hl (|l - l'| / 4) (by linarith) with ⟨N, hN⟩ 99 | rcases hl' (|l - l'| / 4) (by linarith) with ⟨N', hN'⟩ 100 | let N₀ := max N N' 101 | specialize hN N₀ (le_max_left _ _) 102 | specialize hN' N₀ (le_max_right _ _) 103 | have clef : |l - l'| < |l - l'| 104 | calc 105 | |l - l'| = |l - u N₀ + (u N₀ - l')| := by ring_nf 106 | _ ≤ |l - u N₀| + |u N₀ - l'| := by apply abs_add 107 | _ = |u N₀ - l| + |u N₀ - l'| := by rw [abs_sub_comm] 108 | _ < |l - l'| := by linarith 109 | 110 | linarith -- linarith can also find simple numerical contradictions 111 | 112 | /- 113 | Another incarnation of the excluded middle axiom is the principle of 114 | contraposition: in order to prove P ⇒ Q, it suffices to prove 115 | not Q ⇒ not P. 116 | -/ 117 | 118 | -- Using a proof by contradiction, let's prove the contraposition principle 119 | -- 0047 120 | example (P Q : Prop) (h : ¬Q → ¬P) : P → Q := by 121 | -- sorry 122 | intro hP 123 | by_contra hnQ 124 | exact h hnQ hP 125 | -- sorry 126 | 127 | /- 128 | Again Lean doesn't need this principle explained to it. We can use the 129 | `contrapose` tactic. 130 | -/ 131 | example (P Q : Prop) (h : ¬Q → ¬P) : P → Q := by 132 | contrapose 133 | exact h 134 | 135 | /- 136 | In the next exercise, we'll use 137 | Odd n : ∃ k, n = 2*k + 1 138 | Int.odd_iff_not_even {n : ℤ} : odd n ↔ ¬ even n 139 | -/ 140 | -- 0048 141 | example (n : ℤ) : Even (n ^ 2) ↔ Even n := by 142 | -- sorry 143 | constructor 144 | · contrapose 145 | rw [← Int.odd_iff_not_even, ← Int.odd_iff_not_even] 146 | rintro ⟨k, rfl⟩ 147 | use 2 * k * (k + 1) 148 | ring 149 | · rintro ⟨k, rfl⟩ 150 | use 2 * k ^ 2 151 | ring 152 | -- sorry 153 | 154 | /- 155 | As a last step on our law of the excluded middle tour, let's notice that, especially 156 | in pure logic exercises, it can sometimes be useful to use the 157 | excluded middle axiom in its original form: 158 | `Classical.em : ∀ P, P ∨ ¬ P` 159 | 160 | Instead of applying this lemma and then using the `rcases` tactic, we 161 | have the shortcut 162 | `by_cases h : P` 163 | 164 | combining both steps to create two proof branches: one assuming 165 | `h : P`, and the other assuming `h : ¬ P`. 166 | 167 | For instance, let's prove a reformulation of this implication relation, 168 | which is sometimes used as a definition in other logical foundations, 169 | especially those based on truth tables (hence very strongly using 170 | excluded middle from the very beginning). 171 | -/ 172 | variable (P Q : Prop) 173 | 174 | example : P → Q ↔ ¬P ∨ Q := by 175 | constructor 176 | · intro h 177 | by_cases hP : P 178 | · right 179 | exact h hP 180 | · left 181 | exact hP 182 | · intro h hP 183 | rcases h with hnP | hQ 184 | · exfalso 185 | exact hnP hP 186 | · exact hQ 187 | 188 | -- 0049 189 | example : ¬(P ∧ Q) ↔ ¬P ∨ ¬Q := by 190 | -- sorry 191 | constructor 192 | · intro h 193 | by_cases hP : P 194 | · right 195 | intro hQ 196 | exact h ⟨hP, hQ⟩ 197 | · left 198 | exact hP 199 | · rintro h ⟨hP, hQ⟩ 200 | rcases h with hnP | hnQ 201 | · exact hnP hP 202 | · exact hnQ hQ 203 | -- sorry 204 | 205 | /- 206 | It is crucial to understand negation of quantifiers. 207 | Let's do it by hand for a little while. 208 | In the first exercise, only the definition of negation is needed. 209 | -/ 210 | -- 0050 211 | example (n : ℤ) : (¬∃ k, n = 2 * k) ↔ ∀ k, n ≠ 2 * k := by 212 | -- sorry 213 | constructor 214 | · intro hyp k hk 215 | exact hyp ⟨k, hk⟩ 216 | · rintro hyp ⟨k, rfl⟩ 217 | exact hyp k rfl 218 | -- sorry 219 | 220 | /- 221 | Contrary to negation of the existential quantifier, negation of the 222 | universal quantifier requires excluded middle for the first implication. 223 | In order to prove this, we can use either 224 | * a double proof by contradiction 225 | * a contraposition, `not_not : (¬ ¬ P) ↔ P` and a proof by contradiction. 226 | 227 | Recall we have 228 | `EvenFun (f : ℝ → ℝ) := ∀ x, f (-x) = f x` 229 | 230 | -/ 231 | 232 | -- 0051 233 | example (f : ℝ → ℝ) : ¬EvenFun f ↔ ∃ x, f (-x) ≠ f x := by 234 | -- sorry 235 | constructor 236 | · contrapose 237 | intro h 238 | rw [not_not] 239 | intro x 240 | by_contra H 241 | apply h 242 | use x 243 | /- Alternative version 244 | intro h 245 | by_contra H 246 | apply h 247 | intro x 248 | by_contra H' 249 | apply H 250 | use x -/ 251 | · intro ⟨x, hx⟩ h' 252 | exact hx (h' x) 253 | -- sorry 254 | 255 | /- 256 | Of course we can't keep repeating the above proofs, especially the second one. 257 | So we use the `push_neg` tactic. 258 | -/ 259 | example : ¬EvenFun fun x => 2 * x := by 260 | unfold EvenFun 261 | -- Here unfolding is important because push_neg won't do it. 262 | push_neg 263 | use 42 264 | linarith 265 | 266 | -- 0052 267 | example (f : ℝ → ℝ) : ¬EvenFun f ↔ ∃ x, f (-x) ≠ f x := by 268 | -- sorry 269 | unfold EvenFun 270 | push_neg 271 | rfl 272 | -- sorry 273 | 274 | def BoundedAbove (f : ℝ → ℝ) := 275 | ∃ M, ∀ x, f x ≤ M 276 | 277 | example : ¬BoundedAbove fun x => x := by 278 | unfold BoundedAbove 279 | push_neg 280 | intro M 281 | use M + 1 282 | linarith 283 | 284 | -- Let's contrapose 285 | -- 0053 286 | example (x : ℝ) : (∀ ε > 0, x ≤ ε) → x ≤ 0 := by 287 | -- sorry 288 | contrapose 289 | push_neg 290 | intro h 291 | use x / 2 292 | constructor <;> linarith 293 | -- sorry 294 | 295 | /- 296 | The "contrapose, push_neg" combo is so common that we can abreviate it to 297 | `contrapose!` 298 | 299 | Let's use this trick, together with: 300 | `eq_or_lt_of_le : a ≤ b → a = b ∨ a < b` 301 | -/ 302 | -- 0054 303 | example (f : ℝ → ℝ) : (∀ x y, x < y → f x < f y) ↔ ∀ x y, x ≤ y ↔ f x ≤ f y := by 304 | -- sorry 305 | constructor 306 | · intro hf x y 307 | constructor 308 | · intro hxy 309 | rcases eq_or_lt_of_le hxy with hxy | hxy 310 | · rw [hxy] 311 | · linarith [hf x y hxy] 312 | · contrapose! 313 | apply hf 314 | · intro hf x y 315 | contrapose! 316 | intro h 317 | rwa [hf] 318 | -- sorry 319 | -------------------------------------------------------------------------------- /Tutorials/Solutions/07bisAbstractNegations.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | open Classical 4 | 5 | /- 6 | Theoretical negations. 7 | 8 | This file is for people interested in logic who want to fully understand 9 | negations. It comes after the file `07FirstNegations`. 10 | 11 | Here we don't use `contrapose` or `push_neg`. The goal is to prove lemmas 12 | that are used by those tactics. Of course we can use 13 | `exfalso`, `by_contra` and `by_cases`. 14 | 15 | If this doesn't sound like fun then skip ahead to the next file. 16 | -/ 17 | section NegationProp 18 | 19 | variable (P Q : Prop) 20 | 21 | -- 0055 22 | example : P → Q ↔ ¬Q → ¬P := by 23 | -- sorry 24 | constructor 25 | · intro h hnQ hP 26 | exact hnQ (h hP) 27 | · intro h hP 28 | by_contra hnQ 29 | exact h hnQ hP 30 | -- sorry 31 | 32 | -- 0056 33 | theorem non_imp (P Q : Prop) : ¬(P → Q) ↔ P ∧ ¬Q := by 34 | -- sorry 35 | constructor 36 | · intro h 37 | by_contra H 38 | apply h 39 | intro hP 40 | by_contra H' 41 | apply H 42 | exact ⟨hP, H'⟩ 43 | · intro h h' 44 | rcases h with ⟨hP, hnQ⟩ 45 | exact hnQ (h' hP) 46 | -- sorry 47 | 48 | -- In the next one, let's use the axiom 49 | -- `propext {P Q : Prop} : (P ↔ Q) → P = Q` 50 | -- 0057 51 | example (P : Prop) : ¬P ↔ P = False := by 52 | -- sorry 53 | constructor 54 | · intro h 55 | apply propext 56 | constructor 57 | · intro h' 58 | exact h h' 59 | · intro h 60 | exfalso 61 | exact h 62 | · intro h 63 | rw [h] 64 | exact id 65 | -- sorry 66 | 67 | end NegationProp 68 | 69 | section NegationQuantifiers 70 | 71 | variable (X : Type) (P : X → Prop) 72 | 73 | -- 0058 74 | example : (¬∀ x, P x) ↔ ∃ x, ¬P x := by 75 | -- sorry 76 | constructor 77 | · intro h 78 | by_contra H 79 | apply h 80 | intro x 81 | by_contra H' 82 | apply H 83 | use x 84 | · rintro ⟨x, hx⟩ h' 85 | exact hx (h' x) 86 | -- sorry 87 | 88 | -- 0059 89 | example : (¬∃ x, P x) ↔ ∀ x, ¬P x := by 90 | -- sorry 91 | constructor 92 | · intro h x h' 93 | apply h 94 | use x 95 | · rintro h ⟨x, hx⟩ 96 | exact h x hx 97 | -- sorry 98 | 99 | -- 0060 100 | example (P : ℝ → Prop) : (¬∃ ε > 0, P ε) ↔ ∀ ε > 0, ¬P ε := by 101 | -- sorry 102 | constructor 103 | · intro h ε ε_pos hP 104 | apply h 105 | use ε 106 | · rintro h ⟨ε, ε_pos, hP⟩ 107 | exact h ε ε_pos hP 108 | -- sorry 109 | 110 | -- 0061 111 | example (P : ℝ → Prop) : (¬∀ x > 0, P x) ↔ ∃ x > 0, ¬P x := by 112 | -- sorry 113 | constructor 114 | · intro h 115 | by_contra H 116 | apply h 117 | intro x x_pos 118 | by_contra HP 119 | apply H 120 | use x 121 | · rintro ⟨x, xpos, hx⟩ h' 122 | exact hx (h' x xpos) 123 | -- sorry 124 | 125 | end NegationQuantifiers 126 | -------------------------------------------------------------------------------- /Tutorials/Solutions/08LimitsNegation.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | 4 | section 5 | 6 | /- 7 | The first part of this file makes sure you can negate quantified statements 8 | in your head without the help of `push_neg`. 9 | 10 | You need to complete the statement and then use the `check_me` tactic 11 | to check your answer. This tactic exists only for those exercises, 12 | it mostly calls `push_neg` and then cleans up a bit. 13 | 14 | `def SeqLimit (u : ℕ → ℝ) (l : ℝ) : Prop := ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε` 15 | -/ 16 | -- In this section, u denotes a sequence of real numbers 17 | -- f is a function from ℝ to ℝ 18 | -- x₀ and l are real numbers 19 | variable (u : ℕ → ℝ) (f : ℝ → ℝ) (x₀ l : ℝ) 20 | 21 | 22 | -- Negation of "u tends to l" 23 | -- 0062 24 | example : 25 | (¬∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε) ↔-- sorry 26 | ∃ ε > 0, ∀ N, ∃ n ≥ N, |u n - l| > ε := by 27 | -- sorry 28 | check_me 29 | -- sorry 30 | 31 | -- Negation of "f is continuous at x₀" 32 | -- 0063 33 | example : 34 | (¬∀ ε > 0, ∃ δ > 0, ∀ x, |x - x₀| ≤ δ → |f x - f x₀| ≤ ε) ↔-- sorry 35 | ∃ ε > 0, ∀ δ > 0, ∃ x, |x - x₀| ≤ δ ∧ |f x - f x₀| > ε := by 36 | -- sorry 37 | check_me 38 | -- sorry 39 | 40 | /- 41 | In the next exercise, we need to keep in mind that 42 | `∀ x x', ...` is the abbreviation of 43 | `∀ x, ∀ x', ... `. 44 | 45 | Also, `∃ x x', ...` is the abbreviation of `∃ x, ∃ x', ...`. 46 | -/ 47 | -- Negation of "f is uniformly continuous on ℝ" 48 | -- 0064 49 | example : 50 | (¬∀ ε > 0, ∃ δ > 0, ∀ x x', |x' - x| ≤ δ → |f x' - f x| ≤ ε) ↔-- sorry 51 | ∃ ε > 0, ∀ δ > 0, ∃ x x', |x' - x| ≤ δ ∧ |f x' - f x| > ε := by 52 | -- sorry 53 | check_me 54 | -- sorry 55 | 56 | -- Negation of "f is sequentially continuous at x₀" 57 | -- 0065 58 | example : 59 | (¬∀ u : ℕ → ℝ, 60 | (∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - x₀| ≤ ε) → 61 | ∀ ε > 0, ∃ N, ∀ n ≥ N, |(f ∘ u) n - f x₀| ≤ ε) ↔-- sorry 62 | ∃ u : ℕ → ℝ, 63 | (∀ δ > 0, ∃ N, ∀ n ≥ N, |u n - x₀| ≤ δ) ∧ ∃ ε > 0, ∀ N, ∃ n ≥ N, |f (u n) - f x₀| > ε := by 64 | -- sorry 65 | check_me 66 | -- sorry 67 | 68 | end 69 | 70 | /- 71 | We now turn to elementary applications of negations to limits of sequences. 72 | Remember that `linarith` can find easy numerical contradictions. 73 | 74 | Also recall the following lemmas: 75 | 76 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 77 | 78 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 79 | 80 | `le_max_left p q : p ≤ max p q` 81 | 82 | `le_max_right p q : q ≤ max p q` 83 | 84 | /-- The sequence `u` tends to `+∞`. -/ 85 | `def TendstoInfinity (u : ℕ → ℝ) := ∀ A, ∃ N, ∀ n ≥ N, u n ≥ A` 86 | -/ 87 | 88 | -- 0066 89 | example {u : ℕ → ℝ} : TendstoInfinity u → ∀ l, ¬SeqLimit u l := by 90 | -- sorry 91 | intro lim_infinie l lim_l 92 | rcases lim_l 1 (by linarith) with ⟨N, hN⟩ 93 | rcases lim_infinie (l + 2) with ⟨N', hN'⟩ 94 | let N₀ := max N N' 95 | specialize hN N₀ (le_max_left _ _) 96 | specialize hN' N₀ (le_max_right _ _) 97 | rw [abs_le] at hN 98 | linarith 99 | -- sorry 100 | 101 | def NondecreasingSeq (u : ℕ → ℝ) := 102 | ∀ n m, n ≤ m → u n ≤ u m 103 | 104 | -- 0067 105 | example (u : ℕ → ℝ) (l : ℝ) (h : SeqLimit u l) (h' : NondecreasingSeq u) : ∀ n, u n ≤ l := by 106 | -- sorry 107 | intro n 108 | by_contra H 109 | push_neg at H 110 | rcases h ((u n - l) / 2) (by linarith) with ⟨N, hN⟩ 111 | specialize hN (max n N) (le_max_right _ _) 112 | specialize h' n (max n N) (le_max_left _ _) 113 | rw [abs_le] at hN 114 | linarith 115 | -- sorry 116 | 117 | /- 118 | In the following exercises, `A : set ℝ` means that A is a set of real numbers. 119 | We can use the usual notation `x ∈ A`. 120 | 121 | The notation `∀ x ∈ A, ...` is the abbreviation of `∀ x, x ∈ A → ... ` 122 | 123 | The notation `∃ x ∈ A, ...` is the abbreviation of `∃ x, x ∈ A ∧ ... `. 124 | 125 | We'll work with upper bounds and supremums. 126 | Again we'll introduce specialized definitions for the sake of exercises, but mathlib 127 | has more general versions. 128 | 129 | `def UpperBound (A : set ℝ) (x : ℝ) := ∀ a ∈ A, a ≤ x` 130 | 131 | `def IsSup (A : set ℝ) (x : ℝ) := UpperBound A x ∧ ∀ y, UpperBound A y → x ≤ y` 132 | 133 | 134 | Remark: one can easily show that a set of real numbers has at most one sup, 135 | but we won't need this. 136 | -/ 137 | -- 0068 138 | example {A : Set ℝ} {x : ℝ} (hx : IsSup A x) : ∀ y, y < x → ∃ a ∈ A, y < a := by 139 | -- sorry 140 | intro y 141 | contrapose! 142 | exact hx.right y 143 | -- sorry 144 | 145 | /- 146 | Let's do a variation on an example from file 07 that will be useful in the last 147 | exercise below. 148 | -/ 149 | -- 0069 150 | theorem le_of_le_add_all' {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 151 | -- sorry 152 | contrapose! 153 | intro h 154 | use (y - x) / 2 155 | constructor <;> linarith 156 | -- sorry 157 | 158 | -- 0070 159 | example {x y : ℝ} {u : ℕ → ℝ} (hu : SeqLimit u x) (ineg : ∀ n, u n ≤ y) : x ≤ y := by 160 | -- sorry 161 | apply le_of_le_add_all' 162 | intro ε ε_pos 163 | rcases hu ε ε_pos with ⟨N, hN⟩ 164 | specialize hN N (by linarith) 165 | rw [abs_le] at hN 166 | linarith [ineg N] 167 | -- sorry 168 | -------------------------------------------------------------------------------- /Tutorials/Solutions/09LimitsFinal.lean: -------------------------------------------------------------------------------- 1 | import Tutorials.TutoLib 2 | 3 | /- 4 | This is the final file in the series. Here we use everything covered 5 | in previous files to prove a couple of famous theorems from 6 | elementary real analysis. Of course they all have more general versions 7 | in mathlib. 8 | 9 | As usual, keep in mind the following: 10 | 11 | `abs_le {x y : ℝ} : |x| ≤ y ↔ -y ≤ x ∧ x ≤ y` 12 | 13 | `ge_max_iff (p q r) : r ≥ max p q ↔ r ≥ p ∧ r ≥ q` 14 | 15 | `le_max_left p q : p ≤ max p q` 16 | 17 | `le_max_right p q : q ≤ max p q` 18 | 19 | as well as a lemma from the previous file: 20 | 21 | `le_of_le_add_all : (∀ ε > 0, y ≤ x + ε) → y ≤ x` 22 | 23 | Let's start with a variation on a known exercise. 24 | -/ 25 | -- 0071 26 | theorem le_lim' {x y : ℝ} {u : ℕ → ℝ} (hu : SeqLimit u x) (ineg : ∃ N, ∀ n ≥ N, y ≤ u n) : y ≤ x := by 27 | -- sorry 28 | apply le_of_le_add_all 29 | intro ε ε_pos 30 | rcases hu ε ε_pos with ⟨N, hN⟩ 31 | rcases ineg with ⟨N', hN'⟩ 32 | let N₀ := max N N' 33 | specialize hN N₀ (le_max_left N N') 34 | specialize hN' N₀ (le_max_right N N') 35 | rw [abs_le] at hN 36 | linarith 37 | -- sorry 38 | 39 | /- 40 | Let's now return to the result proved in the `00_` file of this series, 41 | and prove again the sequential characterization of upper bounds (with a slighly 42 | different proof). 43 | 44 | For this, and other exercises below, we'll need many things that we proved in previous files, 45 | and a couple of extras. 46 | 47 | From the 5th file: 48 | 49 | `limit_const (x : ℝ) : SeqLimit (λ n, x) x` 50 | 51 | 52 | `squeeze (lim_u : SeqLimit u l) (lim_w : SeqLimit w l) (hu : ∀ n, u n ≤ v n) (hw : ∀ n, v n ≤ w n) : SeqLimit v l` 53 | 54 | From the 8th: 55 | 56 | `def UpperBound (A : set ℝ) (x : ℝ) := ∀ a ∈ A, a ≤ x` 57 | 58 | `def IsSup (A : set ℝ) (x : ℝ) := UpperBound A x ∧ ∀ y, UpperBound A y → x ≤ y` 59 | 60 | `lt_sup (hx : is_sup A x) : ∀ y, y < x → ∃ a ∈ A, y < a ` 61 | 62 | You can also use: 63 | 64 | `Nat.one_div_pos_of_nat {n : ℕ} : 0 < 1 / (n + 1 : ℝ)` 65 | 66 | `inv_succ_le_all : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1/(n + 1 : ℝ) ≤ ε` 67 | 68 | and their easy consequences: 69 | 70 | `limit_of_sub_le_inv_succ (h : ∀ n, |u n - x| ≤ 1/(n+1)) : SeqLimit u x` 71 | 72 | `limit_const_add_inv_succ (x : ℝ) : SeqLimit (λ n, x + 1/(n+1)) x` 73 | 74 | `limit_const_sub_inv_succ (x : ℝ) : SeqLimit (λ n, x - 1/(n+1)) x` 75 | 76 | as well as: 77 | 78 | `lim_le (hu : SeqLimit u x) (ineg : ∀ n, u n ≤ y) : x ≤ y` 79 | 80 | The structure of the proof is offered. It features a new tactic: 81 | `choose` which invokes the axiom of choice (observing the tactic state before and 82 | after using it should be enough to understand everything). 83 | -/ 84 | -- 0072 85 | theorem isSup_iff (A : Set ℝ) (x : ℝ) : 86 | IsSup A x ↔ UpperBound A x ∧ ∃ u : ℕ → ℝ, SeqLimit u x ∧ ∀ n, u n ∈ A := by 87 | constructor 88 | · intro h 89 | constructor 90 | ·-- sorry 91 | exact h.left 92 | -- sorry 93 | · have : ∀ n : ℕ, ∃ a ∈ A, x - 1 / (n + 1) < a := by 94 | intro n 95 | have : 1 / (n + 1 : ℝ) > 0 := Nat.one_div_pos_of_nat 96 | -- sorry 97 | exact lt_sup h _ (by linarith) 98 | -- sorry 99 | choose u hu using this 100 | -- sorry 101 | use u 102 | constructor 103 | · apply squeeze (limit_const_sub_inv_succ x) (limit_const x) 104 | · intro n 105 | exact le_of_lt (hu n).2 106 | · intro n 107 | exact h.1 _ (hu n).left 108 | · intro n 109 | exact (hu n).left 110 | -- sorry 111 | · rintro ⟨maj, u, limu, u_in⟩ 112 | -- sorry 113 | constructor 114 | · exact maj 115 | · intro y ymaj 116 | apply lim_le limu 117 | intro n 118 | apply ymaj 119 | apply u_in 120 | -- sorry 121 | 122 | /-- Continuity of a function at a point -/ 123 | def ContinuousAtPt (f : ℝ → ℝ) (x₀ : ℝ) : Prop := 124 | ∀ ε > 0, ∃ δ > 0, ∀ x, |x - x₀| ≤ δ → |f x - f x₀| ≤ ε 125 | 126 | variable {f : ℝ → ℝ} {x₀ : ℝ} {u : ℕ → ℝ} 127 | 128 | -- 0073 129 | theorem seq_continuous_of_continuous (hf : ContinuousAtPt f x₀) (hu : SeqLimit u x₀) : 130 | SeqLimit (f ∘ u) (f x₀) := by 131 | -- sorry 132 | intro ε ε_pos 133 | rcases hf ε ε_pos with ⟨δ, δ_pos, hδ⟩ 134 | rcases hu δ δ_pos with ⟨N, hN⟩ 135 | use N 136 | intro n hn 137 | apply hδ 138 | exact hN n hn 139 | -- sorry 140 | 141 | -- 0074 142 | example : (∀ u : ℕ → ℝ, SeqLimit u x₀ → SeqLimit (f ∘ u) (f x₀)) → ContinuousAtPt f x₀ := by 143 | -- sorry 144 | contrapose! 145 | intro hf 146 | unfold ContinuousAtPt at hf 147 | push_neg at hf 148 | rcases hf with ⟨ε, h⟩ 149 | rcases h with ⟨ε_pos, hf⟩ 150 | have H : ∀ n : ℕ, ∃ x, |x - x₀| ≤ 1 / (n + 1) ∧ ε < |f x - f x₀| 151 | intro n 152 | apply hf 153 | exact Nat.one_div_pos_of_nat 154 | clear hf 155 | choose u hu using H 156 | use u 157 | constructor 158 | intro η η_pos 159 | have fait : ∃ N : ℕ, ∀ n : ℕ, n ≥ N → 1 / (↑n + 1) ≤ η := inv_succ_le_all η η_pos 160 | rcases fait with ⟨N, hN⟩ 161 | use N 162 | intro n hn 163 | calc 164 | |u n - x₀| ≤ 1 / (n + 1) := (hu n).left 165 | _ ≤ η := hN n hn 166 | 167 | unfold SeqLimit 168 | push_neg 169 | use ε, ε_pos 170 | intro N 171 | use N 172 | constructor 173 | linarith 174 | exact (hu N).right 175 | -- sorry 176 | 177 | /- 178 | Recall from the 6th file: 179 | 180 | 181 | `def extraction (φ : ℕ → ℕ) := ∀ n m, n < m → φ n < φ m` 182 | 183 | `def ClusterPoint (u : ℕ → ℝ) (a : ℝ) := ∃ φ, extraction φ ∧ SeqLimit (u ∘ φ) a` 184 | 185 | 186 | `id_le_extraction : extraction φ → ∀ n, n ≤ φ n` 187 | 188 | and from the 8th file: 189 | 190 | `def TendstoInfinity (u : ℕ → ℝ) := ∀ A, ∃ N, ∀ n ≥ N, u n ≥ A` 191 | 192 | `not_seqLimit_of_tendstoInfinity : TendstoInfinity u → ∀ l, ¬ SeqLimit u l` 193 | -/ 194 | variable {φ : ℕ → ℕ} 195 | 196 | -- 0075 197 | theorem subseq_tendstoInfinity (h : TendstoInfinity u) (hφ : Extraction φ) : 198 | TendstoInfinity (u ∘ φ) := by 199 | -- sorry 200 | intro A 201 | rcases h A with ⟨N, hN⟩ 202 | use N 203 | intro n hn 204 | apply hN 205 | calc 206 | N ≤ n := hn 207 | _ ≤ φ n := id_le_extraction hφ n 208 | -- sorry 209 | 210 | -- 0076 211 | theorem squeeze_infinity {u v : ℕ → ℝ} (hu : TendstoInfinity u) (huv : ∀ n, u n ≤ v n) : 212 | TendstoInfinity v := by 213 | -- sorry 214 | intro A 215 | rcases hu A with ⟨N, hN⟩ 216 | use N 217 | intro n hn 218 | specialize hN n hn 219 | specialize huv n 220 | linarith 221 | -- sorry 222 | 223 | /- 224 | We will use segments: `Icc a b := { x | a ≤ x ∧ x ≤ b }` 225 | The notation stands for Interval-closed-closed. Variations exist with 226 | o or i instead of c, where o stands for open and i for infinity. 227 | 228 | We will use the following version of Bolzano-Weierstrass 229 | 230 | `bolzano_weierstrass (h : ∀ n, u n ∈ Icc a b) : ∃ c ∈ Icc a b, ClusterPoint u c` 231 | 232 | as well as the obvious 233 | 234 | `seqLimit_id : TendstoInfinity (λ n, n)` 235 | -/ 236 | open Set 237 | 238 | -- 0077 239 | theorem bdd_above_segment {f : ℝ → ℝ} {a b : ℝ} (hf : ∀ x ∈ Icc a b, ContinuousAtPt f x) : 240 | ∃ M, ∀ x ∈ Icc a b, f x ≤ M := by 241 | -- sorry 242 | by_contra H 243 | push_neg at H 244 | have clef : ∀ n : ℕ, ∃ x, x ∈ Icc a b ∧ f x > n 245 | intro n 246 | apply H 247 | clear H 248 | choose u hu using clef 249 | have lim_infinie : TendstoInfinity (f ∘ u) 250 | apply squeeze_infinity seqLimit_id 251 | intro n 252 | specialize hu n 253 | dsimp 254 | linarith 255 | have bornes : ∀ n, u n ∈ Icc a b 256 | intro n 257 | exact (hu n).left 258 | rcases bolzano_weierstrass bornes with ⟨c, c_dans, φ, φ_extr, lim⟩ 259 | have lim_infinie_extr : TendstoInfinity (f ∘ u ∘ φ) := subseq_tendstoInfinity lim_infinie φ_extr 260 | have lim_extr : SeqLimit (f ∘ u ∘ φ) (f c) := seq_continuous_of_continuous (hf c c_dans) lim 261 | exact not_seqLimit_of_tendstoInfinity lim_infinie_extr (f c) lim_extr 262 | -- sorry 263 | 264 | /- 265 | In the next exercise, we can use: 266 | 267 | `abs_neg x : |-x| = |x|` 268 | -/ 269 | -- 0078 270 | theorem continuous_opposite {f : ℝ → ℝ} {x₀ : ℝ} (h : ContinuousAtPt f x₀) : 271 | ContinuousAtPt (fun x => -f x) x₀ := by 272 | -- sorry 273 | intro ε ε_pos 274 | rcases h ε ε_pos with ⟨δ, h⟩ 275 | rcases h with ⟨δ_pos, h⟩ 276 | use δ, δ_pos 277 | intro y hy 278 | have : -f y - -f x₀ = -(f y - f x₀); ring 279 | rw [this, abs_neg] 280 | exact h y hy 281 | -- sorry 282 | 283 | /- 284 | Now let's combine the two exercises above 285 | -/ 286 | -- 0079 287 | theorem bdd_below_segment {f : ℝ → ℝ} {a b : ℝ} (hf : ∀ x ∈ Icc a b, ContinuousAtPt f x) : 288 | ∃ m, ∀ x ∈ Icc a b, m ≤ f x := by 289 | -- sorry 290 | have : ∃ M, ∀ x ∈ Icc a b, -f x ≤ M := by 291 | apply bdd_above_segment 292 | intro x x_dans 293 | exact continuous_opposite (hf x x_dans) 294 | rcases this with ⟨M, hM⟩ 295 | use -M 296 | intro x x_dans 297 | specialize hM x x_dans 298 | linarith 299 | -- sorry 300 | 301 | /- 302 | Remember from the 5th file: 303 | 304 | `unique_limit : SeqLimit u l → SeqLimit u l' → l = l'` 305 | 306 | and from the 6th one: 307 | 308 | `subseq_tendsto_of_tendsto (h : SeqLimit u l) (hφ : extraction φ) : SeqLimit (u ∘ φ) l` 309 | 310 | We now admit the following version of the least upper bound theorem 311 | (that cannot be proved without discussing the construction of real numbers 312 | or admitting another strong theorem). 313 | 314 | `sup_segment {a b : ℝ} {A : set ℝ} (hnonvide : ∃ x, x ∈ A) (h : A ⊆ Icc a b) :` 315 | `∃ x ∈ Icc a b, is_sup A x` 316 | 317 | In the next exercise, it can be useful to prove inclusions of sets of real number. 318 | By definition, `A ⊆ B` means : `∀ x, x ∈ A → x ∈ B`. 319 | Hence one can start a proof of `A ⊆ B` by `intro x x_in`, 320 | which brings `x : ℝ` and `x_in : x ∈ A` in the local context, 321 | and then prove `x ∈ B`. 322 | 323 | Note also the use of 324 | `{x | P x}` 325 | which denotes the set of `x` satisfying predicate `P`. 326 | 327 | Hence `x' ∈ { x | P x} ↔ P x'`, by definition. 328 | -/ 329 | -- 0080 330 | example {a b : ℝ} (hab : a ≤ b) (hf : ∀ x ∈ Icc a b, ContinuousAtPt f x) : 331 | ∃ x₀ ∈ Icc a b, ∀ x ∈ Icc a b, f x ≤ f x₀ := by 332 | -- sorry 333 | rcases bdd_below_segment hf with ⟨m, hm⟩ 334 | rcases bdd_above_segment hf with ⟨M, hM⟩ 335 | let A := { y | ∃ x ∈ Icc a b, y = f x } 336 | obtain ⟨y₀, -, y_sup⟩ : ∃ y₀ ∈ Icc m M, IsSup A y₀ := by 337 | apply sup_segment 338 | · exact ⟨f a, a, ⟨by linarith, hab⟩, by ring⟩ 339 | · rintro y ⟨x, x_in, rfl⟩ 340 | exact ⟨hm x x_in, hM x x_in⟩ 341 | rw [isSup_iff] at y_sup 342 | rcases y_sup with ⟨y_maj, u, lim_u, u_in⟩ 343 | choose v hv using u_in 344 | rcases forall_and.mp hv with ⟨v_dans, hufv⟩ 345 | replace hufv : u = f ∘ v := funext hufv 346 | rcases bolzano_weierstrass v_dans with ⟨x₀, x₀_in, φ, φ_extr, lim_vφ⟩ 347 | use x₀, x₀_in 348 | intro x x_dans 349 | have lim : SeqLimit (f ∘ v ∘ φ) (f x₀) := by 350 | apply seq_continuous_of_continuous 351 | exact hf x₀ x₀_in 352 | exact lim_vφ 353 | have unique : f x₀ = y₀ := by 354 | apply unique_limit lim 355 | rw [hufv] at lim_u 356 | exact subseq_tendsto_of_tendsto lim_u φ_extr 357 | rw [unique] 358 | apply y_maj 359 | use x, x_dans 360 | -- sorry 361 | 362 | -- The following stupid lemma can be used below. 363 | lemma stupid {a b x : ℝ} (h : x ∈ Icc a b) (h' : x ≠ b) : x < b := 364 | lt_of_le_of_ne h.right h' 365 | 366 | /- 367 | And now the final boss... 368 | -/ 369 | def I := (Icc 0 1 : Set ℝ) -- the type ascription makes sure 0 and 1 are real numbers here 370 | 371 | -- 0081 372 | example (f : ℝ → ℝ) (hf : ∀ x, ContinuousAtPt f x) (h₀ : f 0 < 0) (h₁ : f 1 > 0) : 373 | ∃ x₀ ∈ I, f x₀ = 0 := by 374 | let A := { x | x ∈ I ∧ f x < 0 } 375 | have ex_x₀ : ∃ x₀ ∈ I, IsSup A x₀ := by 376 | -- sorry 377 | apply sup_segment 378 | use 0 379 | constructor 380 | constructor 381 | linarith 382 | linarith 383 | exact h₀ 384 | intro x hx 385 | exact hx.left 386 | -- sorry 387 | rcases ex_x₀ with ⟨x₀, x₀_in, x₀_sup⟩ 388 | use x₀, x₀_in 389 | have : f x₀ ≤ 0 := by 390 | -- sorry 391 | rw [isSup_iff] at x₀_sup 392 | rcases x₀_sup with ⟨_maj_x₀, u, lim_u, u_in⟩ 393 | have : SeqLimit (f ∘ u) (f x₀) := seq_continuous_of_continuous (hf x₀) lim_u 394 | apply lim_le this 395 | intro n 396 | have : f (u n) < 0 := (u_in n).right 397 | dsimp 398 | linarith 399 | -- sorry 400 | have x₀_1 : x₀ < 1 := by 401 | -- sorry 402 | apply stupid x₀_in 403 | intro h 404 | rw [← h] at h₁ 405 | linarith 406 | -- sorry 407 | have : f x₀ ≥ 0 := by 408 | have in_I : ∃ N : ℕ, ∀ n ≥ N, x₀ + 1 / (n + 1) ∈ I := by 409 | have : ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) ≤ 1 - x₀ := by 410 | -- sorry 411 | apply inv_succ_le_all 412 | linarith 413 | -- sorry 414 | -- sorry 415 | rcases this with ⟨N, hN⟩ 416 | use N 417 | intro n hn 418 | specialize hN n hn 419 | have : 1 / (n + 1 : ℝ) > 0 := Nat.one_div_pos_of_nat 420 | change 0 ≤ x₀ ∧ x₀ ≤ 1 at x₀_in 421 | constructor <;> linarith 422 | -- sorry 423 | have not_in : ∀ n : ℕ, x₀ + 1 / (n + 1) ∉ A := by 424 | -- By definition, x ∉ A means ¬ (x ∈ A). 425 | -- sorry 426 | intro n hn 427 | rcases x₀_sup with ⟨x₀_maj, _⟩ 428 | specialize x₀_maj _ hn 429 | have : 1 / (n + 1 : ℝ) > 0 := Nat.one_div_pos_of_nat 430 | linarith 431 | -- sorry 432 | dsimp at not_in 433 | -- sorry 434 | push_neg at not_in 435 | have lim : SeqLimit (fun n => f (x₀ + 1 / (n + 1))) (f x₀) := by 436 | apply seq_continuous_of_continuous (hf x₀) 437 | apply limit_const_add_inv_succ 438 | apply le_lim' lim 439 | rcases in_I with ⟨N, hN⟩ 440 | use N 441 | intro n hn 442 | exact not_in n (hN n hn) 443 | -- sorry 444 | linarith 445 | -------------------------------------------------------------------------------- /Tutorials/TutoLib.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Analysis.SpecificLimits.Basic 2 | import Mathlib.Data.Int.Parity 3 | import Mathlib.Topology.Sequences 4 | 5 | attribute [instance] Classical.propDecidable 6 | 7 | def NonDecreasing (f : ℝ → ℝ) := 8 | ∀ x₁ x₂, x₁ ≤ x₂ → f x₁ ≤ f x₂ 9 | 10 | def NonIncreasing (f : ℝ → ℝ) := 11 | ∀ x₁ x₂, x₁ ≤ x₂ → f x₁ ≥ f x₂ 12 | 13 | def EvenFun (f : ℝ → ℝ) := ∀ x, f (-x) = f x 14 | 15 | def OddFun (f : ℝ → ℝ) := ∀ x, f (-x) = -f x 16 | 17 | /- 18 | Lemmas from that file were hidden in my course, or restating things which 19 | were proved without name in previous files. 20 | -/ 21 | -- The mathlib version is unusable because it is stated in terms of ≤ 22 | theorem ge_max_iff {α : Type _} [LinearOrder α] {p q r : α} : r ≥ max p q ↔ r ≥ p ∧ r ≥ q := 23 | max_le_iff 24 | 25 | -- No idea why this is not in mathlib 26 | theorem eq_of_abs_sub_le_all (x y : ℝ) : (∀ ε > 0, |x - y| ≤ ε) → x = y := by 27 | intro h 28 | apply eq_of_abs_sub_nonpos 29 | by_contra H 30 | push_neg at H 31 | specialize h (|x - y| / 2) (by linarith) 32 | linarith 33 | 34 | def SeqLimit (u : ℕ → ℝ) (l : ℝ) : Prop := 35 | ∀ ε > 0, ∃ N, ∀ n ≥ N, |u n - l| ≤ ε 36 | 37 | theorem unique_limit {u l l'} : SeqLimit u l → SeqLimit u l' → l = l' := by 38 | intro hl hl' 39 | apply eq_of_abs_sub_le_all 40 | intro ε ε_pos 41 | specialize hl (ε / 2) (by linarith) 42 | rcases hl with ⟨N, hN⟩ 43 | specialize hl' (ε / 2) (by linarith) 44 | rcases hl' with ⟨N', hN'⟩ 45 | specialize hN (max N N') (le_max_left _ _) 46 | specialize hN' (max N N') (le_max_right _ _) 47 | calc 48 | |l - l'| = |l - u (max N N') + (u (max N N') - l')| := by ring_nf 49 | _ ≤ |l - u (max N N')| + |u (max N N') - l'| := by apply abs_add 50 | _ = |u (max N N') - l| + |u (max N N') - l'| := by rw [abs_sub_comm] 51 | _ ≤ ε / 2 + ε / 2 := by linarith 52 | _ = ε := by ring 53 | 54 | 55 | theorem le_of_le_add_all {x y : ℝ} : (∀ ε > 0, y ≤ x + ε) → y ≤ x := by 56 | contrapose! 57 | intro h 58 | use (y - x) / 2 59 | constructor <;> linarith 60 | 61 | def UpperBound (A : Set ℝ) (x : ℝ) := 62 | ∀ a ∈ A, a ≤ x 63 | 64 | def IsSup (A : Set ℝ) (x : ℝ) := 65 | UpperBound A x ∧ ∀ y, UpperBound A y → x ≤ y 66 | 67 | theorem lt_sup {A : Set ℝ} {x : ℝ} (hx : IsSup A x) : ∀ y, y < x → ∃ a ∈ A, y < a := by 68 | intro y 69 | contrapose! 70 | exact hx.right y 71 | 72 | theorem squeeze {u v w : ℕ → ℝ} {l} (hu : SeqLimit u l) (hw : SeqLimit w l) (h : ∀ n, u n ≤ v n) 73 | (h' : ∀ n, v n ≤ w n) : SeqLimit v l := by 74 | intro ε ε_pos 75 | rcases hu ε ε_pos with ⟨N, hN⟩ 76 | rcases hw ε ε_pos with ⟨N', hN'⟩ 77 | use max N N' 78 | intro n hn 79 | rw [ge_max_iff] at hn 80 | specialize hN n (by linarith) 81 | specialize hN' n (by linarith) 82 | specialize h n 83 | specialize h' n 84 | rw [abs_le] at * 85 | constructor <;> linarith 86 | 87 | def Extraction (φ : ℕ → ℕ) := 88 | ∀ n m, n < m → φ n < φ m 89 | 90 | def TendstoInfinity (u : ℕ → ℝ) := 91 | ∀ A, ∃ N, ∀ n ≥ N, u n ≥ A 92 | 93 | theorem lim_le {x y : ℝ} {u : ℕ → ℝ} (hu : SeqLimit u x) (ineg : ∀ n, u n ≤ y) : x ≤ y := by 94 | apply le_of_le_add_all 95 | intro ε ε_pos 96 | rcases hu ε ε_pos with ⟨N, hN⟩ 97 | specialize hN N (by linarith) 98 | specialize ineg N 99 | rw [abs_le] at hN 100 | linarith 101 | 102 | theorem inv_succ_le_all : ∀ ε > 0, ∃ N : ℕ, ∀ n ≥ N, 1 / (n + 1 : ℝ) ≤ ε := by 103 | convert Metric.tendsto_atTop.mp tendsto_one_div_add_atTop_nhds_0_nat using 0 104 | simp only [Real.dist_eq, sub_zero] 105 | constructor 106 | intro h ε ε_pos 107 | rcases h (ε / 2) (by linarith) with ⟨N, hN⟩ 108 | use N 109 | intro n hn 110 | rw [abs_of_pos (Nat.one_div_pos_of_nat : 1 / (n + 1 : ℝ) > 0)] 111 | specialize hN n hn 112 | linarith 113 | intro h ε ε_pos 114 | rcases h ε (by linarith) with ⟨N, hN⟩ 115 | use N 116 | intro n hn 117 | specialize hN n hn 118 | rw [abs_of_pos (@Nat.one_div_pos_of_nat ℝ _ n)] at hN 119 | linarith 120 | 121 | theorem limit_const (x : ℝ) : SeqLimit (fun _ => x) x := fun ε ε_pos => 122 | ⟨0, fun _ _ => by simp [le_of_lt ε_pos]⟩ 123 | 124 | theorem limit_of_sub_le_inv_succ {u : ℕ → ℝ} {x : ℝ} (h : ∀ n, |u n - x| ≤ 1 / (n + 1)) : 125 | SeqLimit u x := by 126 | intro ε ε_pos 127 | rcases inv_succ_le_all ε ε_pos with ⟨N, hN⟩ 128 | use N 129 | intro n hn 130 | specialize h n 131 | specialize hN n hn 132 | linarith 133 | 134 | theorem limit_const_add_inv_succ (x : ℝ) : SeqLimit (fun n => x + 1 / (n + 1)) x := 135 | limit_of_sub_le_inv_succ fun n => by rw [abs_of_pos] <;> linarith [@Nat.one_div_pos_of_nat ℝ _ n] 136 | 137 | theorem limit_const_sub_inv_succ (x : ℝ) : SeqLimit (fun n => x - 1 / (n + 1)) x := by 138 | refine' limit_of_sub_le_inv_succ fun n => _ 139 | rw [show x - 1 / (n + 1) - x = -(1 / (n + 1)) by ring, abs_neg, abs_of_pos] 140 | linarith [@Nat.one_div_pos_of_nat ℝ _ n] 141 | 142 | theorem id_le_extraction {φ} : Extraction φ → ∀ n, n ≤ φ n := by 143 | intro hyp n 144 | induction' n with n hn 145 | · exact Nat.zero_le _ 146 | · exact Nat.succ_le_of_lt (by linarith [hyp n (n + 1) (by linarith)]) 147 | 148 | theorem seqLimit_id : TendstoInfinity fun n => n := by 149 | intro A 150 | rcases exists_nat_gt A with ⟨N, hN⟩ 151 | use N 152 | intro n hn 153 | have : (n : ℝ) ≥ N; exact_mod_cast hn 154 | linarith 155 | 156 | variable {u : ℕ → ℝ} {l : ℝ} {φ : ℕ → ℕ} 157 | 158 | open Set Filter 159 | 160 | def ClusterPoint (u : ℕ → ℝ) (a : ℝ) := 161 | ∃ φ, Extraction φ ∧ SeqLimit (u ∘ φ) a 162 | 163 | theorem bolzano_weierstrass {a b : ℝ} {u : ℕ → ℝ} (h : ∀ n, u n ∈ Icc a b) : 164 | ∃ c ∈ Icc a b, ClusterPoint u c := by 165 | rcases(isCompact_Icc : IsCompact (Icc a b)).tendsto_subseq h with ⟨c, c_in, φ, hφ, lim⟩ 166 | use c, c_in, φ, hφ 167 | simp_rw [Metric.tendsto_nhds, eventually_atTop, Real.dist_eq] at lim 168 | intro ε ε_pos 169 | rcases lim ε ε_pos with ⟨N, hN⟩ 170 | use N 171 | intro n hn 172 | exact le_of_lt (hN n hn) 173 | 174 | theorem not_seqLimit_of_tendstoInfinity {u : ℕ → ℝ} : TendstoInfinity u → ∀ x, ¬SeqLimit u x := by 175 | intro lim_infinie x lim_x 176 | rcases lim_x 1 (by linarith) with ⟨N, hN⟩ 177 | rcases lim_infinie (x + 2) with ⟨N', hN'⟩ 178 | let N₀ := max N N' 179 | specialize hN N₀ (le_max_left _ _) 180 | specialize hN' N₀ (le_max_right _ _) 181 | rw [abs_le] at hN 182 | linarith 183 | 184 | open Real 185 | 186 | theorem sup_segment {a b : ℝ} {A : Set ℝ} (hnonvide : ∃ x, x ∈ A) (h : A ⊆ Icc a b) : 187 | ∃ x ∈ Icc a b, IsSup A x := by 188 | have b_maj : ∀ y : ℝ, y ∈ A → y ≤ b := fun y y_in => (h y_in).2 189 | have Sup_maj : UpperBound A (sSup A) := by 190 | intro x 191 | apply le_csSup 192 | use b 193 | exact b_maj 194 | refine' ⟨sSup A, _, _⟩ 195 | · constructor 196 | · rcases hnonvide with ⟨x, x_in⟩ 197 | exact le_trans (h x_in).1 (Sup_maj _ x_in) 198 | · apply csSup_le hnonvide b_maj 199 | · exact ⟨Sup_maj, fun y => csSup_le hnonvide⟩ 200 | 201 | theorem subseq_tendsto_of_tendsto (h : SeqLimit u l) (hφ : Extraction φ) : SeqLimit (u ∘ φ) l := by 202 | intro ε ε_pos 203 | rcases h ε ε_pos with ⟨N, hN⟩ 204 | use N 205 | intro n hn 206 | apply hN 207 | calc 208 | N ≤ n := hn 209 | _ ≤ φ n := id_le_extraction hφ n 210 | 211 | 212 | open Lean Elab Tactic 213 | 214 | macro "check_me" : tactic => `(tactic| ( 215 | repeat unfold SeqLimit 216 | repeat unfold continue_en 217 | push_neg 218 | try simp only [exists_prop] 219 | try exact Iff.rfl 220 | first | done | fail "That's not quite right. Please try again.")) 221 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": 5, 2 | "packagesDir": "lake-packages", 3 | "packages": 4 | [{"git": 5 | {"url": "https://github.com/leanprover-community/mathlib4.git", 6 | "subDir?": null, 7 | "rev": "0fd376a2a6771146b41dfd81b6726431ebc5f2c4", 8 | "opts": {}, 9 | "name": "mathlib", 10 | "inputRev?": null, 11 | "inherited": false}}, 12 | {"git": 13 | {"url": "https://github.com/mhuisi/lean4-cli.git", 14 | "subDir?": null, 15 | "rev": "21dac2e9cc7e3cf7da5800814787b833e680b2fd", 16 | "opts": {}, 17 | "name": "Cli", 18 | "inputRev?": "nightly", 19 | "inherited": true}}, 20 | {"git": 21 | {"url": "https://github.com/gebner/quote4", 22 | "subDir?": null, 23 | "rev": "e75daed95ad1c92af4e577fea95e234d7a8401c1", 24 | "opts": {}, 25 | "name": "Qq", 26 | "inputRev?": "master", 27 | "inherited": true}}, 28 | {"git": 29 | {"url": "https://github.com/JLimperg/aesop", 30 | "subDir?": null, 31 | "rev": "1a0cded2be292b5496e659b730d2accc742de098", 32 | "opts": {}, 33 | "name": "aesop", 34 | "inputRev?": "master", 35 | "inherited": true}}, 36 | {"git": 37 | {"url": "https://github.com/leanprover/std4", 38 | "subDir?": null, 39 | "rev": "67855403d60daf181775fa1ec63b04e70bcc3921", 40 | "opts": {}, 41 | "name": "std", 42 | "inputRev?": "main", 43 | "inherited": true}}, 44 | {"git": 45 | {"url": "https://github.com/EdAyers/ProofWidgets4", 46 | "subDir?": null, 47 | "rev": "65bba7286e2395f3163fd0277110578f19b8170f", 48 | "opts": {}, 49 | "name": "proofwidgets", 50 | "inputRev?": "v0.0.16", 51 | "inherited": true}}]} 52 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | open Lake DSL 3 | 4 | def moreLeanArgs := #[ 5 | "-Dpp.unicode.fun=true" -- pretty-prints `fun a ↦ b` 6 | ] 7 | 8 | def moreServerArgs := moreLeanArgs 9 | 10 | package tutorials where 11 | moreLeanArgs := moreLeanArgs 12 | moreServerArgs := moreServerArgs 13 | 14 | @[default_target] 15 | lean_lib Tutorials where 16 | moreLeanArgs := moreLeanArgs 17 | 18 | require mathlib from git 19 | "https://github.com/leanprover-community/mathlib4.git" 20 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.1.0-rc1 2 | -------------------------------------------------------------------------------- /mk_exercises.py: -------------------------------------------------------------------------------- 1 | #!/usr/bin/env python3 2 | 3 | import re 4 | from pathlib import Path 5 | 6 | sorry_regex = re.compile(r'(.*)-- sorry.*') 7 | root = Path(__file__).parent/'Tutorials' 8 | 9 | if __name__ == '__main__': 10 | for path in (root/'Solutions').glob('**/*.lean'): 11 | if path.name == 'TutoLib.lean': 12 | continue 13 | print(path) 14 | out = root/'Exercises'/path.relative_to(root/'Solutions') 15 | with out.open('w') as outp: 16 | with path.open() as inp: 17 | state = 'normal' 18 | for line in inp: 19 | m = sorry_regex.match(line) 20 | if state == 'normal': 21 | if m: 22 | state = 'sorry' 23 | else: 24 | outp.write(line) 25 | else: 26 | if m: 27 | state = 'normal' 28 | outp.write(m.group(1)+ 'sorry\n') 29 | 30 | outp.write('\n') 31 | 32 | 33 | --------------------------------------------------------------------------------